diff --git a/stage0/src/CMakeLists.txt b/stage0/src/CMakeLists.txt index 95eece17470a2..2f03e98dc4dbf 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 57e0a0358ec42..1e8514eff38fc 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 d0d61a8eed9b3..88ff1496995a3 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 793cf6403277a..243b53635fda1 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 9cec464e5fd69..c556540612b9f 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 feac201c88e7b..e42f20f90d137 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 b0e6844dca24e..dc4034e9ffc68 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 3ab2ce643b72d..ac387564aa753 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 7fbbb41851eb9..30c2c8924af4f 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 724e8c2b10444..d4134090fb75c 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 49556be1f8622..1d66ee423ff7e 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 f67e89f2cb59b..e37bdaf7b3e21 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 ed8768cfc8455..2dd2c21c7d130 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 2dbdfa40dd59e..8995a95d84e9e 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 a1397a9c64496..e4b2b506b8a1b 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 eb75c93a2737d..bed22bb5a5994 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 16c1fe8b3b6e1..d620c53aa6443 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 057b30356796b..44db877b0c8cc 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 82bd830a75884..4488f16538fe1 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 35bb6c375446e..1bc325dac84d0 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 db3746a770d49..a4ef44b7bd0c6 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 2f838dad1f2a5..3e7d1da698dc4 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 04ab99872daab..bccc9ba42b0f5 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 42057654f34aa..727ec117cc5c2 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 5ced492633216..15f3ffd7796d1 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 5ab58f6cdfc62..6388a77d97fd3 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 81744fdd800a0..06026945c0174 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 447f8929209d1..5c00d3a89fe9a 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 84c6d3194bfa9..c254121630e29 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 929b49e6a3cd3..234ce10ecfbd5 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 1c8b39d4585fd..5c15fdb910b3d 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 343128c0c5195..a6279e953ef04 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 2b611a1b59f94..a71e2fcc55c5b 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 267c60c4ccb27..8a6afe4ef032e 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 375cb3ba37baa..30b68954d8c65 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 eac9ad41f13dd..1b5a09986ab89 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 9bfaa62a3d1d7..7a0f5707546bc 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 2c499224f3cd1..528b28bf3b42e 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 8a9a092269a27..df91f21ee5b47 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 89d5fde795fd2..b64f23a20a85d 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())); } @@ -217,12 +217,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())); } @@ -359,7 +359,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; @@ -397,7 +397,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 { @@ -812,7 +812,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) { @@ -935,7 +935,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); @@ -983,7 +983,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); } @@ -1094,24 +1094,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(); } @@ -1138,7 +1148,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 00a65a34fe623..c3a0022656df0 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 176db6baae964..e7138092824ae 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 c4f921467c533..ded6caee37721 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 9bcd88d4cf593..95949e744e887 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 58d3c9c03c602..79fbb4967d70f 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 9167aeb8613bc..671c9183c5523 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 779f5e9583578..35bd795b77fe5 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 25a4b4f11b4c1..d5af9bbb19f7f 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 87c61c6ffee58..3b6bebadc1dfc 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 22e32381867b1..cd5135e5d2b7a 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 f16a00bb58526..9c131c75dc2de 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 84f34ad9c82c5..7c18d2d179f8f 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 af9bb6e1d7420..35dab66e24582 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 6874412c894e3..6340fabf7b2db 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 a1651e39c967d..63061244d1a6e 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 50d26c8749939..9f28c17cc2f06 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 5b534d95e34b3..de18defb53ece 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 96f8ed4c4da7d..5d6fa1e37476a 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 bb7e45788403a..49a3844e38f09 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 6710b0b0dd79c..8a3470c4d4db1 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 b15317c90a24f..951959c7fb026 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 a234c172fdb36..0e40d4b6ce050 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 0000000000000..537169f4ec512 --- /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 0000000000000..db21ccf13b493 --- /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 dad1df4649065..91d65d22db0c7 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 ae6e508c0d827..cbf1618347257 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 62a2319022dde..909b385ad0c2f 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 3ade6827e63c2..89e88d5004476 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 16a0aad85f1a3..a2f009f6406cc 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 c58e210426cec..f1fa9e44d2a79 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 a3d118fd9f112..b647d85f376c7 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 02027d0265f93..ebf192d59d6ad 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" @@ -332,7 +332,7 @@ extern "C" object * lean_run_frontend( uint8_t json_output, 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, @@ -344,7 +344,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), @@ -401,7 +401,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())); } } @@ -633,7 +633,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; @@ -705,8 +704,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); - env = r.fst(); + pair_ref r = run_new_frontend(contents, opts, mod_fn, *main_module_name, trust_lvl, ilean_fn, json_output); + elab_environment env = r.fst(); bool ok = unbox(r.snd().raw()); if (stats) { diff --git a/stage0/stdlib/Lake/DSL/AttributesCore.c b/stage0/stdlib/Lake/DSL/AttributesCore.c index eef75dd5a0d0a..6706e8d06968b 100644 --- a/stage0/stdlib/Lake/DSL/AttributesCore.c +++ b/stage0/stdlib/Lake/DSL/AttributesCore.c @@ -729,7 +729,6 @@ LEAN_EXPORT lean_object* l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_108_ { uint8_t x_3; lean_object* x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_108____lambda__2(x_1, x_2); -lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; @@ -996,16 +995,19 @@ LEAN_EXPORT uint8_t l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____la { lean_object* x_3; uint8_t x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____lambda__1___closed__1; +lean_inc(x_2); x_4 = l_Lake_OrderedTagAttribute_hasTag(x_3, x_2, x_1); if (x_4 == 0) { lean_object* x_5; uint8_t x_6; x_5 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____lambda__1___closed__2; +lean_inc(x_2); x_6 = l_Lake_OrderedTagAttribute_hasTag(x_5, x_2, x_1); if (x_6 == 0) { lean_object* x_7; uint8_t x_8; x_7 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____lambda__1___closed__3; +lean_inc(x_2); x_8 = l_Lake_OrderedTagAttribute_hasTag(x_7, x_2, x_1); if (x_8 == 0) { @@ -1017,6 +1019,7 @@ return x_10; else { uint8_t x_11; +lean_dec(x_2); x_11 = 1; return x_11; } @@ -1024,6 +1027,7 @@ return x_11; else { uint8_t x_12; +lean_dec(x_2); x_12 = 1; return x_12; } @@ -1031,6 +1035,7 @@ return x_12; else { uint8_t x_13; +lean_dec(x_2); x_13 = 1; return x_13; } @@ -1206,7 +1211,6 @@ LEAN_EXPORT lean_object* l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302_ { uint8_t x_3; lean_object* x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____lambda__1(x_1, x_2); -lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; @@ -1217,11 +1221,13 @@ LEAN_EXPORT uint8_t l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_403____la { lean_object* x_3; uint8_t x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_108____lambda__2___closed__1; +lean_inc(x_2); x_4 = l_Lake_OrderedTagAttribute_hasTag(x_3, x_2, x_1); if (x_4 == 0) { lean_object* x_5; uint8_t x_6; x_5 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_302____lambda__1___closed__2; +lean_inc(x_2); x_6 = l_Lake_OrderedTagAttribute_hasTag(x_5, x_2, x_1); if (x_6 == 0) { @@ -1233,6 +1239,7 @@ return x_8; else { uint8_t x_9; +lean_dec(x_2); x_9 = 1; return x_9; } @@ -1240,6 +1247,7 @@ return x_9; else { uint8_t x_10; +lean_dec(x_2); x_10 = 1; return x_10; } @@ -1415,7 +1423,6 @@ LEAN_EXPORT lean_object* l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_403_ { uint8_t x_3; lean_object* x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_403____lambda__1(x_1, x_2); -lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; @@ -1426,6 +1433,7 @@ LEAN_EXPORT uint8_t l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_498____la { lean_object* x_3; uint8_t x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_108____lambda__2___closed__1; +lean_inc(x_2); x_4 = l_Lake_OrderedTagAttribute_hasTag(x_3, x_2, x_1); if (x_4 == 0) { @@ -1437,6 +1445,7 @@ return x_6; else { uint8_t x_7; +lean_dec(x_2); x_7 = 1; return x_7; } @@ -1612,7 +1621,6 @@ LEAN_EXPORT lean_object* l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_498_ { uint8_t x_3; lean_object* x_4; x_3 = l_Lake_initFn____x40_Lake_DSL_AttributesCore___hyg_498____lambda__1(x_1, x_2); -lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; diff --git a/stage0/stdlib/Lake/DSL/Config.c b/stage0/stdlib/Lake/DSL/Config.c index 083f1fdf4eea1..15ee7ff59c9eb 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; @@ -333,7 +333,6 @@ lean_dec(x_12); x_15 = lean_box(0); x_16 = l_Lake_DSL_elabDirConst___closed__1; x_17 = l_Lean_EnvExtension_getState___rarg(x_15, x_16, x_14); -lean_dec(x_14); if (lean_obj_tag(x_17) == 0) { 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; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -405,7 +404,6 @@ lean_dec(x_45); x_48 = lean_box(0); x_49 = l_Lake_DSL_elabDirConst___closed__1; x_50 = l_Lean_EnvExtension_getState___rarg(x_48, x_49, x_47); -lean_dec(x_47); if (lean_obj_tag(x_50) == 0) { 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; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; @@ -1104,7 +1102,6 @@ lean_dec(x_19); x_22 = lean_box(0); x_23 = l_Lake_DSL_elabGetConfig___closed__1; x_24 = l_Lean_EnvExtension_getState___rarg(x_22, x_23, x_21); -lean_dec(x_21); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -1209,7 +1206,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 +1290,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 +1387,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; @@ -1435,7 +1435,6 @@ lean_dec(x_162); x_165 = lean_box(0); x_166 = l_Lake_DSL_elabGetConfig___closed__1; x_167 = l_Lean_EnvExtension_getState___rarg(x_165, x_166, x_164); -lean_dec(x_164); if (lean_obj_tag(x_167) == 0) { lean_object* x_168; uint8_t x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; @@ -1546,7 +1545,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 +1645,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/DSL/DeclUtil.c b/stage0/stdlib/Lake/DSL/DeclUtil.c index c438367ca12d2..9e470ffea05b9 100644 --- a/stage0/stdlib/Lake/DSL/DeclUtil.c +++ b/stage0/stdlib/Lake/DSL/DeclUtil.c @@ -2623,7 +2623,6 @@ lean_inc(x_45); lean_dec(x_43); lean_inc(x_1); x_46 = l_Lean_findField_x3f(x_45, x_1, x_33); -lean_dec(x_45); if (lean_obj_tag(x_46) == 0) { uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; @@ -2706,7 +2705,6 @@ lean_inc(x_66); lean_dec(x_64); lean_inc(x_1); x_67 = l_Lean_findField_x3f(x_66, x_1, x_33); -lean_dec(x_66); if (lean_obj_tag(x_67) == 0) { uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; @@ -2794,7 +2792,6 @@ lean_inc(x_90); lean_dec(x_87); lean_inc(x_1); x_91 = l_Lean_findField_x3f(x_90, x_1, x_33); -lean_dec(x_90); if (lean_obj_tag(x_91) == 0) { 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; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; @@ -3394,7 +3391,6 @@ lean_inc(x_47); lean_dec(x_45); lean_inc(x_1); x_48 = l_Lean_findField_x3f(x_47, x_1, x_35); -lean_dec(x_47); if (lean_obj_tag(x_48) == 0) { 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; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; @@ -3477,7 +3473,6 @@ lean_inc(x_68); lean_dec(x_66); lean_inc(x_1); x_69 = l_Lean_findField_x3f(x_68, x_1, x_35); -lean_dec(x_68); if (lean_obj_tag(x_69) == 0) { uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; @@ -3565,7 +3560,6 @@ lean_inc(x_92); lean_dec(x_89); lean_inc(x_1); x_93 = l_Lean_findField_x3f(x_92, x_1, x_35); -lean_dec(x_92); if (lean_obj_tag(x_93) == 0) { uint8_t 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; uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; @@ -4143,7 +4137,6 @@ lean_inc(x_46); lean_dec(x_44); lean_inc(x_1); x_47 = l_Lean_findField_x3f(x_46, x_1, x_34); -lean_dec(x_46); if (lean_obj_tag(x_47) == 0) { 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; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; @@ -4226,7 +4219,6 @@ lean_inc(x_67); lean_dec(x_65); lean_inc(x_1); x_68 = l_Lean_findField_x3f(x_67, x_1, x_34); -lean_dec(x_67); if (lean_obj_tag(x_68) == 0) { uint8_t 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; @@ -4314,7 +4306,6 @@ lean_inc(x_91); lean_dec(x_88); lean_inc(x_1); x_92 = l_Lean_findField_x3f(x_91, x_1, x_34); -lean_dec(x_91); if (lean_obj_tag(x_92) == 0) { uint8_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; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; diff --git a/stage0/stdlib/Lake/Load/Lean/Elab.c b/stage0/stdlib/Lake/Load/Lean/Elab.c index 9badd66468f8a..41bbb3f6285d8 100644 --- a/stage0/stdlib/Lake/Load/Lean/Elab.c +++ b/stage0/stdlib/Lake/Load/Lean/Elab.c @@ -142,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,7 +282,7 @@ 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_EXPORT lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_addToEnv___boxed(lean_object*, lean_object*); @@ -2881,7 +2881,7 @@ x_97 = lean_ctor_get(x_95, 1); lean_inc(x_97); lean_dec(x_95); x_98 = l_Lake_configModuleName; -x_99 = lean_environment_set_main_module(x_96, x_98); +x_99 = l_Lean_Environment_setMainModule(x_96, x_98); x_100 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_100, 0, x_1); x_101 = l_Lake_elabConfigFile___closed__2; @@ -2916,7 +2916,7 @@ x_113 = lean_ctor_get(x_110, 1); lean_inc(x_113); lean_dec(x_110); x_114 = l_Lake_configModuleName; -x_115 = lean_environment_set_main_module(x_112, x_114); +x_115 = l_Lean_Environment_setMainModule(x_112, x_114); x_116 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_116, 0, x_1); x_117 = l_Lake_elabConfigFile___closed__2; @@ -3071,7 +3071,7 @@ x_155 = lean_ctor_get(x_151, 1); lean_inc(x_155); lean_dec(x_151); x_156 = l_Lake_configModuleName; -x_157 = lean_environment_set_main_module(x_154, x_156); +x_157 = l_Lean_Environment_setMainModule(x_154, x_156); x_158 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_158, 0, x_1); x_159 = l_Lake_elabConfigFile___closed__2; @@ -3332,7 +3332,7 @@ x_224 = lean_ctor_get(x_220, 1); lean_inc(x_224); lean_dec(x_220); x_225 = l_Lake_configModuleName; -x_226 = lean_environment_set_main_module(x_223, x_225); +x_226 = l_Lean_Environment_setMainModule(x_223, x_225); x_227 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_227, 0, x_1); x_228 = l_Lake_elabConfigFile___closed__2; @@ -3484,7 +3484,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; } } @@ -4229,7 +4229,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; diff --git a/stage0/stdlib/Lake/Load/Lean/Eval.c b/stage0/stdlib/Lake/Load/Lean/Eval.c index 443b1dc53ac78..ec025cd7bc1cc 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; @@ -340,7 +340,7 @@ LEAN_EXPORT lean_object* l_Lake_unsafeEvalConstCheck(lean_object* x_1, lean_obje 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; @@ -5385,7 +5385,6 @@ LEAN_EXPORT lean_object* l_Lake_mkTagMap___rarg___boxed(lean_object* x_1, lean_o lean_object* x_5; x_5 = l_Lake_mkTagMap___rarg(x_1, x_2, x_3, x_4); lean_dec(x_2); -lean_dec(x_1); return x_5; } } @@ -5573,7 +5572,6 @@ LEAN_EXPORT lean_object* l_Lake_mkOrdTagMap___rarg___boxed(lean_object* x_1, lea lean_object* x_5; x_5 = l_Lake_mkOrdTagMap___rarg(x_1, x_2, x_3, x_4); lean_dec(x_2); -lean_dec(x_1); return x_5; } } @@ -5661,6 +5659,7 @@ LEAN_EXPORT lean_object* l_Lake_PackageConfig_loadFromEnv(lean_object* x_1, lean { lean_object* x_3; lean_object* x_4; lean_object* x_5; x_3 = l_Lake_PackageConfig_loadFromEnv___closed__1; +lean_inc(x_1); x_4 = l_Lake_OrderedTagAttribute_getAllEntries(x_3, x_1); x_5 = lean_array_to_list(x_4); if (lean_obj_tag(x_5) == 0) @@ -19214,7 +19213,6 @@ lean_inc(x_17); lean_dec(x_15); x_18 = 1; x_19 = l_Lean_findDocString_x3f(x_3, x_5, x_18, x_17); -lean_dec(x_3); x_20 = !lean_is_exclusive(x_19); if (x_20 == 0) { @@ -20078,6 +20076,7 @@ lean_closure_set(x_17, 1, x_15); lean_closure_set(x_17, 2, x_2); lean_closure_set(x_17, 3, x_3); x_18 = l_Lake_Package_loadFromEnv___closed__1; +lean_inc(x_2); x_19 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__2(x_2, x_18, x_17, x_4, x_5); if (lean_obj_tag(x_19) == 0) { @@ -20096,6 +20095,7 @@ x_23 = lean_ctor_get(x_20, 1); lean_inc(x_23); lean_dec(x_20); x_24 = l_Lake_Package_loadFromEnv___closed__2; +lean_inc(x_2); x_25 = l_Lake_OrderedTagAttribute_getAllEntries(x_24, x_2); x_26 = lean_array_size(x_25); x_27 = 0; @@ -20127,6 +20127,7 @@ x_465 = lean_alloc_closure((void*)(l_Lake_Package_loadFromEnv___lambda__9___boxe lean_closure_set(x_465, 0, x_2); lean_closure_set(x_465, 1, x_3); x_466 = l_Lake_Package_loadFromEnv___closed__12; +lean_inc(x_2); x_467 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__26(x_2, x_466, x_465); x_468 = l_IO_ofExcept___at_Lake_Package_loadFromEnv___spec__28(x_467, x_30); if (lean_obj_tag(x_468) == 0) @@ -20181,6 +20182,7 @@ x_296 = lean_alloc_closure((void*)(l_Lake_Package_loadFromEnv___lambda__8___boxe lean_closure_set(x_296, 0, x_2); lean_closure_set(x_296, 1, x_3); x_297 = l_Lake_Package_loadFromEnv___closed__11; +lean_inc(x_2); x_298 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__23(x_2, x_297, x_296); x_299 = l_IO_ofExcept___at_Lake_Package_loadFromEnv___spec__25(x_298, x_36); if (lean_obj_tag(x_299) == 0) @@ -20238,6 +20240,7 @@ lean_closure_set(x_44, 1, x_3); lean_closure_set(x_44, 2, x_15); lean_closure_set(x_44, 3, x_13); x_45 = l_Lake_Package_loadFromEnv___closed__3; +lean_inc(x_2); x_46 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__9(x_1, x_2, x_45, x_44, x_43, x_41); if (lean_obj_tag(x_46) == 0) { @@ -20264,6 +20267,7 @@ lean_closure_set(x_51, 1, x_3); lean_closure_set(x_51, 2, x_15); lean_closure_set(x_51, 3, x_13); x_52 = l_Lake_Package_loadFromEnv___closed__4; +lean_inc(x_2); x_53 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__14(x_1, x_2, x_52, x_51, x_50, x_48); if (lean_obj_tag(x_53) == 0) { @@ -20282,8 +20286,10 @@ x_57 = lean_ctor_get(x_54, 1); lean_inc(x_57); lean_dec(x_54); x_58 = l_Lake_Package_loadFromEnv___closed__5; +lean_inc(x_2); x_59 = l_Lake_OrderedTagAttribute_getAllEntries(x_58, x_2); x_60 = l_Lake_Package_loadFromEnv___closed__6; +lean_inc(x_2); x_61 = l_Lake_OrderedTagAttribute_getAllEntries(x_60, x_2); x_62 = lean_array_size(x_61); lean_inc(x_13); @@ -20312,6 +20318,7 @@ lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean x_68 = lean_ctor_get(x_64, 0); x_69 = lean_ctor_get(x_64, 1); x_162 = l_Lake_Package_loadFromEnv___closed__10; +lean_inc(x_2); x_163 = l_Lake_OrderedTagAttribute_getAllEntries(x_162, x_2); x_164 = lean_array_size(x_163); lean_inc(x_2); @@ -20365,6 +20372,7 @@ lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean x_73 = lean_ctor_get(x_70, 0); x_74 = lean_ctor_get(x_70, 1); x_75 = l_Lake_Package_loadFromEnv___closed__7; +lean_inc(x_2); x_76 = l_Lake_OrderedTagAttribute_getAllEntries(x_75, x_2); x_77 = lean_array_get_size(x_76); x_78 = lean_unsigned_to_nat(1u); @@ -20384,7 +20392,6 @@ lean_dec(x_66); x_82 = lean_ctor_get(x_8, 17); lean_inc(x_82); x_83 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_73, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_68, x_13, x_15, x_82, x_74, x_71); -lean_dec(x_2); return x_83; } else @@ -20452,7 +20459,6 @@ lean_dec(x_76); x_99 = 1; x_100 = l_Lean_Name_toString(x_98, x_99, x_15); x_101 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_73, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_68, x_13, x_15, x_100, x_74, x_71); -lean_dec(x_2); return x_101; } } @@ -20514,6 +20520,7 @@ lean_inc(x_114); lean_inc(x_113); lean_dec(x_70); x_115 = l_Lake_Package_loadFromEnv___closed__7; +lean_inc(x_2); x_116 = l_Lake_OrderedTagAttribute_getAllEntries(x_115, x_2); x_117 = lean_array_get_size(x_116); x_118 = lean_unsigned_to_nat(1u); @@ -20532,7 +20539,6 @@ lean_dec(x_66); x_122 = lean_ctor_get(x_8, 17); lean_inc(x_122); x_123 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_113, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_68, x_13, x_15, x_122, x_114, x_71); -lean_dec(x_2); return x_123; } else @@ -20599,7 +20605,6 @@ lean_dec(x_116); x_140 = 1; x_141 = l_Lean_Name_toString(x_139, x_140, x_15); x_142 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_113, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_68, x_13, x_15, x_141, x_114, x_71); -lean_dec(x_2); return x_142; } } @@ -20718,6 +20723,7 @@ lean_inc(x_177); lean_inc(x_176); lean_dec(x_64); x_229 = l_Lake_Package_loadFromEnv___closed__10; +lean_inc(x_2); x_230 = l_Lake_OrderedTagAttribute_getAllEntries(x_229, x_2); x_231 = lean_array_size(x_230); lean_inc(x_2); @@ -20779,6 +20785,7 @@ if (lean_is_exclusive(x_178)) { x_182 = lean_box(0); } x_183 = l_Lake_Package_loadFromEnv___closed__7; +lean_inc(x_2); x_184 = l_Lake_OrderedTagAttribute_getAllEntries(x_183, x_2); x_185 = lean_array_get_size(x_184); x_186 = lean_unsigned_to_nat(1u); @@ -20798,7 +20805,6 @@ lean_dec(x_66); x_190 = lean_ctor_get(x_8, 17); lean_inc(x_190); x_191 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_180, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_176, x_13, x_15, x_190, x_181, x_179); -lean_dec(x_2); return x_191; } else @@ -20871,7 +20877,6 @@ lean_dec(x_184); x_208 = 1; x_209 = l_Lean_Name_toString(x_207, x_208, x_15); x_210 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_180, x_38, x_42, x_49, x_56, x_59, x_22, x_33, x_176, x_13, x_15, x_209, x_181, x_179); -lean_dec(x_2); return x_210; } } @@ -21355,6 +21360,7 @@ x_443 = lean_alloc_closure((void*)(l_Lake_Package_loadFromEnv___lambda__8___boxe lean_closure_set(x_443, 0, x_2); lean_closure_set(x_443, 1, x_3); x_444 = l_Lake_Package_loadFromEnv___closed__11; +lean_inc(x_2); x_445 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__23(x_2, x_444, x_443); x_446 = l_IO_ofExcept___at_Lake_Package_loadFromEnv___spec__25(x_445, x_36); if (lean_obj_tag(x_446) == 0) @@ -21414,6 +21420,7 @@ lean_closure_set(x_315, 1, x_3); lean_closure_set(x_315, 2, x_15); lean_closure_set(x_315, 3, x_13); x_316 = l_Lake_Package_loadFromEnv___closed__3; +lean_inc(x_2); x_317 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__9(x_1, x_2, x_316, x_315, x_314, x_312); if (lean_obj_tag(x_317) == 0) { @@ -21440,6 +21447,7 @@ lean_closure_set(x_322, 1, x_3); lean_closure_set(x_322, 2, x_15); lean_closure_set(x_322, 3, x_13); x_323 = l_Lake_Package_loadFromEnv___closed__4; +lean_inc(x_2); x_324 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__14(x_1, x_2, x_323, x_322, x_321, x_319); if (lean_obj_tag(x_324) == 0) { @@ -21458,8 +21466,10 @@ x_328 = lean_ctor_get(x_325, 1); lean_inc(x_328); lean_dec(x_325); x_329 = l_Lake_Package_loadFromEnv___closed__5; +lean_inc(x_2); x_330 = l_Lake_OrderedTagAttribute_getAllEntries(x_329, x_2); x_331 = l_Lake_Package_loadFromEnv___closed__6; +lean_inc(x_2); x_332 = l_Lake_OrderedTagAttribute_getAllEntries(x_331, x_2); x_333 = lean_array_size(x_332); lean_inc(x_13); @@ -21494,6 +21504,7 @@ if (lean_is_exclusive(x_335)) { x_340 = lean_box(0); } x_392 = l_Lake_Package_loadFromEnv___closed__10; +lean_inc(x_2); x_393 = l_Lake_OrderedTagAttribute_getAllEntries(x_392, x_2); x_394 = lean_array_size(x_393); lean_inc(x_2); @@ -21564,6 +21575,7 @@ if (lean_is_exclusive(x_341)) { x_345 = lean_box(0); } x_346 = l_Lake_Package_loadFromEnv___closed__7; +lean_inc(x_2); x_347 = l_Lake_OrderedTagAttribute_getAllEntries(x_346, x_2); x_348 = lean_array_get_size(x_347); x_349 = lean_unsigned_to_nat(1u); @@ -21583,7 +21595,6 @@ lean_dec(x_337); x_353 = lean_ctor_get(x_8, 17); lean_inc(x_353); x_354 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_343, x_309, x_313, x_320, x_327, x_330, x_22, x_33, x_338, x_13, x_15, x_353, x_344, x_342); -lean_dec(x_2); return x_354; } else @@ -21656,7 +21667,6 @@ lean_dec(x_347); x_371 = 1; x_372 = l_Lean_Name_toString(x_370, x_371, x_15); x_373 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_343, x_309, x_313, x_320, x_327, x_330, x_22, x_33, x_338, x_13, x_15, x_372, x_344, x_342); -lean_dec(x_2); return x_373; } } @@ -22133,6 +22143,7 @@ x_638 = lean_alloc_closure((void*)(l_Lake_Package_loadFromEnv___lambda__9___boxe lean_closure_set(x_638, 0, x_2); lean_closure_set(x_638, 1, x_3); x_639 = l_Lake_Package_loadFromEnv___closed__12; +lean_inc(x_2); x_640 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__26(x_2, x_639, x_638); x_641 = l_IO_ofExcept___at_Lake_Package_loadFromEnv___spec__28(x_640, x_30); if (lean_obj_tag(x_641) == 0) @@ -22195,6 +22206,7 @@ x_617 = lean_alloc_closure((void*)(l_Lake_Package_loadFromEnv___lambda__8___boxe lean_closure_set(x_617, 0, x_2); lean_closure_set(x_617, 1, x_3); x_618 = l_Lake_Package_loadFromEnv___closed__11; +lean_inc(x_2); x_619 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__23(x_2, x_618, x_617); x_620 = l_IO_ofExcept___at_Lake_Package_loadFromEnv___spec__25(x_619, x_481); if (lean_obj_tag(x_620) == 0) @@ -22263,6 +22275,7 @@ lean_closure_set(x_489, 1, x_3); lean_closure_set(x_489, 2, x_15); lean_closure_set(x_489, 3, x_13); x_490 = l_Lake_Package_loadFromEnv___closed__3; +lean_inc(x_2); x_491 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__9(x_1, x_2, x_490, x_489, x_488, x_486); if (lean_obj_tag(x_491) == 0) { @@ -22289,6 +22302,7 @@ lean_closure_set(x_496, 1, x_3); lean_closure_set(x_496, 2, x_15); lean_closure_set(x_496, 3, x_13); x_497 = l_Lake_Package_loadFromEnv___closed__4; +lean_inc(x_2); x_498 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__14(x_1, x_2, x_497, x_496, x_495, x_493); if (lean_obj_tag(x_498) == 0) { @@ -22307,8 +22321,10 @@ x_502 = lean_ctor_get(x_499, 1); lean_inc(x_502); lean_dec(x_499); x_503 = l_Lake_Package_loadFromEnv___closed__5; +lean_inc(x_2); x_504 = l_Lake_OrderedTagAttribute_getAllEntries(x_503, x_2); x_505 = l_Lake_Package_loadFromEnv___closed__6; +lean_inc(x_2); x_506 = l_Lake_OrderedTagAttribute_getAllEntries(x_505, x_2); x_507 = lean_array_size(x_506); lean_inc(x_13); @@ -22343,6 +22359,7 @@ if (lean_is_exclusive(x_509)) { x_514 = lean_box(0); } x_566 = l_Lake_Package_loadFromEnv___closed__10; +lean_inc(x_2); x_567 = l_Lake_OrderedTagAttribute_getAllEntries(x_566, x_2); x_568 = lean_array_size(x_567); lean_inc(x_2); @@ -22413,6 +22430,7 @@ if (lean_is_exclusive(x_515)) { x_519 = lean_box(0); } x_520 = l_Lake_Package_loadFromEnv___closed__7; +lean_inc(x_2); x_521 = l_Lake_OrderedTagAttribute_getAllEntries(x_520, x_2); x_522 = lean_array_get_size(x_521); x_523 = lean_unsigned_to_nat(1u); @@ -22432,7 +22450,6 @@ lean_dec(x_511); x_527 = lean_ctor_get(x_8, 17); lean_inc(x_527); x_528 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_517, x_482, x_487, x_494, x_501, x_504, x_22, x_478, x_512, x_13, x_15, x_527, x_518, x_516); -lean_dec(x_2); return x_528; } else @@ -22505,7 +22522,6 @@ lean_dec(x_521); x_545 = 1; x_546 = l_Lean_Name_toString(x_544, x_545, x_15); x_547 = l_Lake_Package_loadFromEnv___lambda__7(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_517, x_482, x_487, x_494, x_501, x_504, x_22, x_478, x_512, x_13, x_15, x_546, x_518, x_516); -lean_dec(x_2); return x_547; } } @@ -23161,7 +23177,6 @@ LEAN_EXPORT lean_object* l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__2 lean_object* x_6; x_6 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); -lean_dec(x_1); return x_6; } } @@ -23248,7 +23263,6 @@ LEAN_EXPORT lean_object* l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__9 lean_object* x_7; x_7 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__9(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); return x_7; } @@ -23300,7 +23314,6 @@ LEAN_EXPORT lean_object* l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__1 lean_object* x_7; x_7 = l_Lake_mkTagMap___at_Lake_Package_loadFromEnv___spec__14(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); return x_7; } @@ -23365,7 +23378,6 @@ LEAN_EXPORT lean_object* l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec lean_object* x_4; x_4 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__23(x_1, x_2, x_3); lean_dec(x_2); -lean_dec(x_1); return x_4; } } @@ -23388,7 +23400,6 @@ LEAN_EXPORT lean_object* l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec lean_object* x_4; x_4 = l_Lake_mkOrdTagMap___at_Lake_Package_loadFromEnv___spec__26(x_1, x_2, x_3); lean_dec(x_2); -lean_dec(x_1); return x_4; } } @@ -23538,7 +23549,6 @@ lean_object* x_22 = _args[21]; { lean_object* x_23; x_23 = l_Lake_Package_loadFromEnv___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, x_18, x_19, x_20, x_21, x_22); -lean_dec(x_1); return x_23; } } @@ -23829,6 +23839,7 @@ LEAN_EXPORT lean_object* l_Lake_Workspace_addFacetsFromEnv(lean_object* x_1, lea { lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; x_4 = l_Lake_Workspace_addFacetsFromEnv___closed__1; +lean_inc(x_1); x_5 = l_Lake_OrderedTagAttribute_getAllEntries(x_4, x_1); x_6 = lean_box(0); x_7 = lean_array_size(x_5); @@ -23863,6 +23874,7 @@ x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); x_14 = l_Lake_Workspace_addFacetsFromEnv___closed__2; +lean_inc(x_1); x_15 = l_Lake_OrderedTagAttribute_getAllEntries(x_14, x_1); x_16 = lean_array_size(x_15); lean_inc(x_1); @@ -23895,6 +23907,7 @@ x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); x_22 = l_Lake_Workspace_addFacetsFromEnv___closed__3; +lean_inc(x_1); x_23 = l_Lake_OrderedTagAttribute_getAllEntries(x_22, x_1); x_24 = lean_array_size(x_23); x_25 = l_Array_forIn_x27Unsafe_loop___at_Lake_Workspace_addFacetsFromEnv___spec__3(x_1, x_2, x_6, x_23, x_23, x_24, x_8, x_21); diff --git a/stage0/stdlib/Lake/Util/OrderedTagAttribute.c b/stage0/stdlib/Lake/Util/OrderedTagAttribute.c index 7e0988ee65b3c..3bc83ba378e5d 100644 --- a/stage0/stdlib/Lake/Util/OrderedTagAttribute.c +++ b/stage0/stdlib/Lake/Util/OrderedTagAttribute.c @@ -1641,7 +1641,6 @@ LEAN_EXPORT lean_object* l_Lake_OrderedTagAttribute_hasTag___boxed(lean_object* uint8_t x_4; lean_object* x_5; x_4 = l_Lake_OrderedTagAttribute_hasTag(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; @@ -1756,7 +1755,6 @@ LEAN_EXPORT lean_object* l_Lake_OrderedTagAttribute_getAllEntries___boxed(lean_o { lean_object* x_3; x_3 = l_Lake_OrderedTagAttribute_getAllEntries(x_1, x_2); -lean_dec(x_2); lean_dec(x_1); return x_3; } diff --git a/stage0/stdlib/Lake/Util/Version.c b/stage0/stdlib/Lake/Util/Version.c index db668e232cd0f..5bfb76386e774 100644 --- a/stage0/stdlib/Lake/Util/Version.c +++ b/stage0/stdlib/Lake/Util/Version.c @@ -246,7 +246,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 8940990341250..246a8d9893813 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 d08486d9b0c9e..d4ce2b6d56001 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 c479d0196197d..2e2bc4bdf56b0 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 c5f1302c6c0db..2916c68add5ae 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 dd5e6f74848be..0b9a5759fb194 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; @@ -124,7 +124,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*) uint64_t l_Lean_Name_hash___override(lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137____closed__8; uint64_t lean_uint64_xor(uint64_t, uint64_t); -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__4(lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137____spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_471____closed__12; @@ -178,6 +177,7 @@ static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Compiler_CSimp_replaceConstants___spec__2(lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Compiler_CSimp_replaceConstants___spec__3(lean_object*, size_t, lean_object*); +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__4(lean_object*); static lean_object* _init_l_Lean_Compiler_CSimp_instInhabitedEntry___closed__1() { _start: { @@ -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_6597____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7020____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_6597____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7020____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 79009f99e864a..42f690adc0c5d 100644 --- a/stage0/stdlib/Lean/Compiler/ExternAttr.c +++ b/stage0/stdlib/Lean/Compiler/ExternAttr.c @@ -42,7 +42,7 @@ LEAN_EXPORT lean_object* l_Lean_getExternEntryFor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExternEntryForAux(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_object* l_List_getD___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_755____spec__1(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*); @@ -58,9 +58,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternC LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_755____lambda__4(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); -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_755____closed__4; @@ -83,6 +81,7 @@ static lean_object* l_Lean_getExternConstArityExport___closed__8; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__12; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_755____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -169,6 +168,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_755____c 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_initFn____x40_Lean_Compiler_ExternAttr___hyg_755____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isExtern___boxed(lean_object*, lean_object*); @@ -933,7 +933,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) @@ -948,9 +947,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; @@ -990,9 +988,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; @@ -1047,7 +1044,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) @@ -1064,9 +1060,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; @@ -1106,9 +1101,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; @@ -3013,15 +3007,14 @@ return x_2; static lean_object* _init_l_Lean_getExternConstArityExport___closed__11() { _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_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__7; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__7; +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_getExternConstArityExport___closed__12() { @@ -3131,7 +3124,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) { @@ -3358,7 +3351,7 @@ x_86 = lean_ctor_get(x_83, 0); x_87 = lean_ctor_get(x_83, 4); lean_dec(x_87); x_88 = l_Lean_getExternConstArityExport___closed__15; -x_89 = l_Lean_Kernel_enableDiag(x_86, x_88); +x_89 = l_Lean_Kernel_Environment_enableDiag(x_86, x_88); lean_ctor_set(x_83, 4, x_55); lean_ctor_set(x_83, 0, x_89); x_90 = lean_st_ref_set(x_74, x_83, x_84); @@ -3473,7 +3466,7 @@ lean_inc(x_113); lean_inc(x_112); lean_dec(x_83); x_119 = l_Lean_getExternConstArityExport___closed__15; -x_120 = l_Lean_Kernel_enableDiag(x_112, x_119); +x_120 = l_Lean_Kernel_Environment_enableDiag(x_112, x_119); x_121 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_121, 0, x_120); lean_ctor_set(x_121, 1, x_113); diff --git a/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c b/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c index fc5c9d4851873..99c09d448a8d2 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 3e0f4edd80b22..23bd97f2ba9e3 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 4986c361a333a..ffa36bbb59eb1 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*); @@ -58538,7 +58538,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) @@ -58597,7 +58596,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) { @@ -58784,7 +58783,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) @@ -58951,7 +58949,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) @@ -59041,7 +59038,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) @@ -59239,7 +59235,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); @@ -59313,7 +59308,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); @@ -59516,7 +59510,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); @@ -59590,7 +59583,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); @@ -59731,7 +59723,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); @@ -60019,6 +60010,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; } } @@ -60034,6 +60026,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; } } @@ -60048,6 +60041,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 8785c2909aa5c..ec87161ef74bc 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 ddd4d15ba90af..715312b6dcab5 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 410aa8ba456ac..0f138c06af5d2 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); @@ -3828,7 +3826,7 @@ if (x_27 == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Lean_KernelException_toMessageData(x_28, x_13); +x_29 = l_Lean_Kernel_Exception_toMessageData(x_28, x_13); x_30 = lean_ctor_get(x_3, 5); x_31 = l_Lean_MessageData_toString(x_29, x_23); if (lean_obj_tag(x_31) == 0) @@ -3908,7 +3906,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_26, 0); lean_inc(x_52); lean_dec(x_26); -x_53 = l_Lean_KernelException_toMessageData(x_52, x_13); +x_53 = l_Lean_Kernel_Exception_toMessageData(x_52, x_13); x_54 = lean_ctor_get(x_3, 5); x_55 = l_Lean_MessageData_toString(x_53, x_23); if (lean_obj_tag(x_55) == 0) @@ -4124,7 +4122,7 @@ if (lean_is_exclusive(x_113)) { lean_dec_ref(x_113); x_115 = lean_box(0); } -x_116 = l_Lean_KernelException_toMessageData(x_114, x_13); +x_116 = l_Lean_Kernel_Exception_toMessageData(x_114, x_13); x_117 = lean_ctor_get(x_3, 5); x_118 = l_Lean_MessageData_toString(x_116, x_110); if (lean_obj_tag(x_118) == 0) @@ -4326,7 +4324,7 @@ if (lean_is_exclusive(x_173)) { lean_dec_ref(x_173); x_175 = lean_box(0); } -x_176 = l_Lean_KernelException_toMessageData(x_174, x_159); +x_176 = l_Lean_Kernel_Exception_toMessageData(x_174, x_159); x_177 = lean_ctor_get(x_3, 5); x_178 = l_Lean_MessageData_toString(x_176, x_169); if (lean_obj_tag(x_178) == 0) diff --git a/stage0/stdlib/Lean/Compiler/InlineAttrs.c b/stage0/stdlib/Lean/Compiler/InlineAttrs.c index 8bc0c3f308aa4..131c2731bf40f 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 cd72495046400..b05da67d06289 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*); @@ -8390,8 +8390,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; @@ -8431,8 +8430,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 af48d6dd5bb42..01d83cbf88d98 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 20320b0aa9002..5d5a013cedbeb 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c @@ -96,7 +96,7 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ElimDe 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*); @@ -2542,9 +2542,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; @@ -2709,9 +2708,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; @@ -10729,8 +10727,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; @@ -10865,8 +10862,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 ccf424f88f703..aeca1ddbbdaeb 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*); @@ -8452,8 +8452,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; @@ -8493,8 +8492,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; @@ -9048,7 +9046,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; @@ -9878,7 +9877,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; @@ -10436,7 +10436,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 d246b12615370..3f4c4a922808f 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -94,11 +94,9 @@ 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; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1412____closed__20; 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; @@ -133,6 +131,7 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main__ 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; @@ -263,6 +262,7 @@ 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*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1412____closed__3; +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*); @@ -1693,7 +1693,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) { @@ -1743,7 +1743,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); @@ -1820,7 +1820,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 e2e5a785a33fd..68fd491047d02 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*); @@ -180,6 +179,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; @@ -4810,7 +4810,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) { @@ -4860,7 +4860,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); @@ -4890,7 +4890,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 c220d2264e817..3346d29c08e89 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 c9cf9dd134b94..f32a13ef133e6 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_Decl_getArity(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 09f0dd7f2d355..7e673a1b48a17 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 4aeb9c78f0a60..12f5da87089d3 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 0d36d26c820dc..079261935de85 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 9f9b88f044509..31ca76d0f5eb7 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c @@ -44,7 +44,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*); @@ -2428,11 +2428,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; } @@ -2475,11 +2477,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 541e6165c8dc8..61fad5c582be3 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*); @@ -7991,7 +7991,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; @@ -11572,9 +11573,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; @@ -17053,8 +17053,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; @@ -17094,8 +17093,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 ce3f2aa649053..62b72df81aecd 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; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono___hyg_2080____closed__16; lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); 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_LetValue_toMono___closed__1; @@ -286,7 +286,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; @@ -551,7 +551,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; @@ -739,6 +739,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_9; } } @@ -1667,7 +1668,6 @@ if (x_75 == 0) lean_object* x_76; lean_inc(x_6); lean_inc(x_5); -lean_inc(x_69); x_76 = 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_76) == 0) { @@ -1689,8 +1689,7 @@ lean_dec(x_79); x_82 = lean_ctor_get(x_80, 0); lean_inc(x_82); lean_dec(x_80); -lean_inc(x_69); -x_83 = lean_environment_find(x_82, x_69); +x_83 = l_Lean_Environment_find_x3f(x_82, x_69); if (lean_obj_tag(x_83) == 0) { size_t x_84; size_t x_85; lean_object* x_86; uint8_t x_87; @@ -1878,7 +1877,6 @@ if (x_123 == 0) lean_object* x_124; lean_inc(x_6); lean_inc(x_5); -lean_inc(x_118); x_124 = l_Lean_Compiler_LCNF_isTrivialConstructorApp_x3f(x_118, x_119, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_124) == 0) { @@ -1900,8 +1898,7 @@ lean_dec(x_127); x_130 = lean_ctor_get(x_128, 0); lean_inc(x_130); lean_dec(x_128); -lean_inc(x_118); -x_131 = lean_environment_find(x_130, x_118); +x_131 = l_Lean_Environment_find_x3f(x_130, x_118); if (lean_obj_tag(x_131) == 0) { size_t x_132; size_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Types.c b/stage0/stdlib/Lean/Compiler/LCNF/Types.c index 953d81b72b75e..ef89b58607e1c 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 ac553e1c17747..b35eeb965886f 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 8159c99833fe9..14bbc560e0a2c 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 21b152902f34f..c4d77376a00a0 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -25,7 +25,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_insta static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_185_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_reportMessageKind(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__1; static lean_object* l_Lean_Core_instMonadCoreM___closed__5; @@ -208,6 +207,7 @@ lean_object* l_Lean_Compiler_getDeclNamesForCodeGen(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_wrapAsyncAsSnapshot___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__3(lean_object*, lean_object*, lean_object*); lean_object* lean_io_get_num_heartbeats(lean_object*); +LEAN_EXPORT lean_object* l_Lean_logMessageKind(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Core_wrapAsyncAsSnapshot___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; size_t lean_usize_of_nat(lean_object*); @@ -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; @@ -575,15 +574,15 @@ static lean_object* l_Lean_Core_resetMessageLog___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM; static lean_object* l_Lean_compileDecl___closed__1; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_185____closed__7; -LEAN_EXPORT lean_object* l_Lean_reportMessageKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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; @@ -593,6 +592,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMonadLiftIOCoreM___rarg___boxed(lean_ob LEAN_EXPORT lean_object* l_Lean_Core_wrapAsync___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_CoreM___hyg_4099____closed__7; LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logMessageKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___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*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalIds___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -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); @@ -6617,7 +6619,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; @@ -6633,7 +6636,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); @@ -6692,7 +6696,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); @@ -6763,7 +6768,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) { @@ -6818,7 +6823,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); @@ -6982,7 +6987,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); @@ -7184,7 +7189,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) { @@ -7238,7 +7243,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); @@ -7322,7 +7327,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); @@ -7512,7 +7517,7 @@ lean_dec(x_45); x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); lean_dec(x_46); -x_49 = l_Lean_Kernel_isDiagnosticsEnabled(x_48); +x_49 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_48); lean_dec(x_48); if (x_49 == 0) { @@ -7566,7 +7571,7 @@ lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean x_55 = lean_ctor_get(x_52, 0); x_56 = lean_ctor_get(x_52, 4); lean_dec(x_56); -x_57 = l_Lean_Kernel_enableDiag(x_55, x_44); +x_57 = l_Lean_Kernel_Environment_enableDiag(x_55, x_44); x_58 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_52, 4, x_58); lean_ctor_set(x_52, 0, x_57); @@ -7678,7 +7683,7 @@ lean_inc(x_83); lean_inc(x_82); lean_inc(x_81); lean_dec(x_52); -x_88 = l_Lean_Kernel_enableDiag(x_81, x_44); +x_88 = l_Lean_Kernel_Environment_enableDiag(x_81, x_44); x_89 = l_Lean_Core_instInhabitedCache___closed__3; x_90 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_90, 0, x_88); @@ -7908,7 +7913,7 @@ lean_dec(x_150); x_153 = lean_ctor_get(x_151, 0); lean_inc(x_153); lean_dec(x_151); -x_154 = l_Lean_Kernel_isDiagnosticsEnabled(x_153); +x_154 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_153); lean_dec(x_153); if (x_154 == 0) { @@ -7983,7 +7988,7 @@ if (lean_is_exclusive(x_157)) { lean_dec_ref(x_157); x_166 = lean_box(0); } -x_167 = l_Lean_Kernel_enableDiag(x_159, x_149); +x_167 = l_Lean_Kernel_Environment_enableDiag(x_159, x_149); x_168 = l_Lean_Core_instInhabitedCache___closed__3; if (lean_is_scalar(x_166)) { x_169 = lean_alloc_ctor(0, 8, 0); @@ -9360,15 +9365,14 @@ return x_5; static lean_object* _init_l_Lean_Core_resetMessageLog___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__4; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__4; +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; } } LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -10251,7 +10255,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) { @@ -10305,7 +10309,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); @@ -10389,7 +10393,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); @@ -17235,15 +17239,14 @@ return x_5; static lean_object* _init_l_Lean_Core_wrapAsyncAsSnapshot___lambda__5___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Core_wrapAsyncAsSnapshot___lambda__5___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_wrapAsyncAsSnapshot___lambda__5___closed__3; +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_Core_wrapAsyncAsSnapshot___lambda__5___closed__5() { @@ -18788,7 +18791,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) { @@ -18847,7 +18849,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) { @@ -21537,7 +21538,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; @@ -23528,7 +23529,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) { @@ -23582,7 +23583,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); @@ -23611,7 +23612,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); @@ -23693,7 +23694,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) { @@ -23768,7 +23769,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 { @@ -23979,7 +23980,7 @@ lean_dec(x_74); x_77 = lean_ctor_get(x_75, 0); lean_inc(x_77); lean_dec(x_75); -x_78 = l_Lean_Kernel_isDiagnosticsEnabled(x_77); +x_78 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_77); lean_dec(x_77); if (x_78 == 0) { @@ -24197,7 +24198,7 @@ x_84 = lean_ctor_get(x_81, 0); x_85 = lean_ctor_get(x_81, 4); lean_dec(x_85); x_86 = l_Lean_ImportM_runCoreM___rarg___closed__10; -x_87 = l_Lean_Kernel_enableDiag(x_84, x_86); +x_87 = l_Lean_Kernel_Environment_enableDiag(x_84, x_86); lean_ctor_set(x_81, 4, x_53); lean_ctor_set(x_81, 0, x_87); x_88 = lean_st_ref_set(x_72, x_81, x_82); @@ -24312,7 +24313,7 @@ lean_inc(x_111); lean_inc(x_110); lean_dec(x_81); x_117 = l_Lean_ImportM_runCoreM___rarg___closed__10; -x_118 = l_Lean_Kernel_enableDiag(x_110, x_117); +x_118 = l_Lean_Kernel_Environment_enableDiag(x_110, x_117); x_119 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_119, 0, x_118); lean_ctor_set(x_119, 1, x_111); @@ -25080,7 +25081,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_mapCoreM___rarg), 5, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_reportMessageKind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_logMessageKind(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; @@ -25094,7 +25095,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 5); lean_inc(x_9); lean_dec(x_7); -x_10 = lean_ctor_get(x_9, 1); +x_10 = lean_ctor_get(x_9, 2); lean_inc(x_10); lean_dec(x_9); x_11 = l_Lean_NameSet_contains(x_10, x_1); @@ -25121,10 +25122,10 @@ 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; uint8_t x_23; -x_19 = lean_ctor_get(x_14, 1); +x_19 = lean_ctor_get(x_14, 2); x_20 = lean_box(0); x_21 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_19, x_1, x_20); -lean_ctor_set(x_14, 1, x_21); +lean_ctor_set(x_14, 2, x_21); x_22 = lean_st_ref_set(x_3, x_13, x_15); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) @@ -25153,19 +25154,20 @@ return x_30; } else { -uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; -x_31 = lean_ctor_get_uint8(x_14, sizeof(void*)*2); -x_32 = lean_ctor_get(x_14, 0); -x_33 = lean_ctor_get(x_14, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_31 = lean_ctor_get(x_14, 0); +x_32 = lean_ctor_get(x_14, 1); +x_33 = lean_ctor_get(x_14, 2); lean_inc(x_33); lean_inc(x_32); +lean_inc(x_31); lean_dec(x_14); x_34 = lean_box(0); x_35 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_33, x_1, x_34); -x_36 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_36, 0, x_32); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*2, x_31); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_31); +lean_ctor_set(x_36, 1, x_32); +lean_ctor_set(x_36, 2, x_35); lean_ctor_set(x_13, 5, x_36); x_37 = lean_st_ref_set(x_3, x_13, x_15); x_38 = lean_ctor_get(x_37, 1); @@ -25192,7 +25194,7 @@ 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; uint8_t 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_object* x_63; +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; uint8_t x_61; lean_object* x_62; lean_object* x_63; x_43 = lean_ctor_get(x_13, 0); x_44 = lean_ctor_get(x_13, 1); x_45 = lean_ctor_get(x_13, 2); @@ -25208,14 +25210,16 @@ lean_inc(x_45); lean_inc(x_44); lean_inc(x_43); lean_dec(x_13); -x_50 = lean_ctor_get_uint8(x_14, sizeof(void*)*2); -x_51 = lean_ctor_get(x_14, 0); +x_50 = lean_ctor_get(x_14, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_14, 1); lean_inc(x_51); -x_52 = lean_ctor_get(x_14, 1); +x_52 = lean_ctor_get(x_14, 2); lean_inc(x_52); 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_53 = x_14; } else { lean_dec_ref(x_14); @@ -25224,13 +25228,13 @@ if (lean_is_exclusive(x_14)) { x_54 = lean_box(0); x_55 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_52, x_1, x_54); if (lean_is_scalar(x_53)) { - x_56 = lean_alloc_ctor(0, 2, 1); + x_56 = lean_alloc_ctor(0, 3, 0); } else { x_56 = x_53; } -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set_uint8(x_56, sizeof(void*)*2, x_50); +lean_ctor_set(x_56, 0, x_50); +lean_ctor_set(x_56, 1, x_51); +lean_ctor_set(x_56, 2, x_55); x_57 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_57, 0, x_43); lean_ctor_set(x_57, 1, x_44); @@ -25284,14 +25288,14 @@ lean_dec(x_5); x_68 = lean_ctor_get(x_66, 5); lean_inc(x_68); lean_dec(x_66); -x_69 = lean_ctor_get(x_68, 1); +x_69 = lean_ctor_get(x_68, 2); lean_inc(x_69); lean_dec(x_68); x_70 = l_Lean_NameSet_contains(x_69, x_1); lean_dec(x_69); if (x_70 == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; 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; uint8_t x_94; lean_object* x_95; lean_object* x_96; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; x_71 = lean_st_ref_take(x_3, x_67); x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); @@ -25328,14 +25332,16 @@ if (lean_is_exclusive(x_72)) { lean_dec_ref(x_72); x_82 = lean_box(0); } -x_83 = lean_ctor_get_uint8(x_73, sizeof(void*)*2); -x_84 = lean_ctor_get(x_73, 0); +x_83 = lean_ctor_get(x_73, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_73, 1); lean_inc(x_84); -x_85 = lean_ctor_get(x_73, 1); +x_85 = lean_ctor_get(x_73, 2); lean_inc(x_85); if (lean_is_exclusive(x_73)) { lean_ctor_release(x_73, 0); lean_ctor_release(x_73, 1); + lean_ctor_release(x_73, 2); x_86 = x_73; } else { lean_dec_ref(x_73); @@ -25344,13 +25350,13 @@ if (lean_is_exclusive(x_73)) { x_87 = lean_box(0); x_88 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_85, x_1, x_87); if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(0, 2, 1); + x_89 = lean_alloc_ctor(0, 3, 0); } else { x_89 = x_86; } -lean_ctor_set(x_89, 0, x_84); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set_uint8(x_89, sizeof(void*)*2, x_83); +lean_ctor_set(x_89, 0, x_83); +lean_ctor_set(x_89, 1, x_84); +lean_ctor_set(x_89, 2, x_88); if (lean_is_scalar(x_82)) { x_90 = lean_alloc_ctor(0, 8, 0); } else { @@ -25400,11 +25406,11 @@ return x_99; } } } -LEAN_EXPORT lean_object* l_Lean_reportMessageKind___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_logMessageKind___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_reportMessageKind(x_1, x_2, x_3, x_4); +x_5 = l_Lean_logMessageKind(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; diff --git a/stage0/stdlib/Lean/Declaration.c b/stage0/stdlib/Lean/Declaration.c index 8f10d31518d2b..fda9a2077309a 100644 --- a/stage0/stdlib/Lean/Declaration.c +++ b/stage0/stdlib/Lean/Declaration.c @@ -23,6 +23,8 @@ LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___boxed(lean LEAN_EXPORT lean_object* l_Lean_mkQuotValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RecursorVal_isUnsafeEx___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687_(uint8_t, lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__9; +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*); static lean_object* l_Lean_ConstantInfo_value_x21___closed__2; LEAN_EXPORT lean_object* l_Lean_instBEqConstructor; @@ -30,9 +32,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*); @@ -50,7 +54,7 @@ LEAN_EXPORT lean_object* l_Lean_mkRecursorValEx___boxed(lean_object*, lean_objec LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqDefinitionSafety____x40_Lean_Declaration___hyg_669_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_isTheorem___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTheoremVal; -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117_(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); @@ -68,18 +72,19 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedDeclaration; LEAN_EXPORT lean_object* l_Lean_instInhabitedQuotVal; LEAN_EXPORT lean_object* l_Lean_RecursorVal_getFirstMinorIdx(lean_object*); static lean_object* l_Lean_instBEqDefinitionSafety___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstructorVal_isUnsafeEx___boxed(lean_object*); LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_434____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkConstructorValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedAxiomVal___closed__1; -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_compare___boxed(lean_object*, lean_object*); 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*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2991____boxed(lean_object*, 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*); @@ -92,7 +97,6 @@ LEAN_EXPORT uint8_t l_Lean_ReducibilityHints_compare(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_toConstantVal___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x3f(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177____boxed(lean_object*, lean_object*); 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 uint8_t lean_quot_val_kind(lean_object*); @@ -104,9 +108,7 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__2; LEAN_EXPORT uint8_t lean_definition_val_get_safety(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__16; -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2991_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqInductiveType____x40_Lean_Declaration___hyg_1593____spec__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqDefinitionSafety; LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_toCtorIdx___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isTheorem(lean_object*); @@ -135,8 +137,9 @@ static lean_object* l_Lean_instInhabitedConstantVal___closed__2; LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqDeclaration____x40_Lean_Declaration___hyg_1777____spec__2___boxed(lean_object*, lean_object*); 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 lean_object* l_Lean_DefinitionSafety_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2780____boxed(lean_object*, lean_object*); 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*); @@ -151,6 +154,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructor____ LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqReducibilityHints____x40_Lean_Declaration___hyg_28____boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTheoremVal___closed__1; +LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqTheoremVal____x40_Lean_Declaration___hyg_1098____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__5; LEAN_EXPORT uint8_t lean_axiom_val_is_unsafe(lean_object*); @@ -162,6 +166,7 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_numLevelParams___boxed(lean_object* lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__18; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Declaration_getNames(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40_Lean_Declaration___hyg_555_(lean_object*, lean_object*); @@ -178,13 +183,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; LEAN_EXPORT lean_object* l_Lean_instBEqDeclaration; +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(lean_object*, lean_object*); 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_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqDeclaration____x40_Lean_Declaration___hyg_1777____spec__1___boxed(lean_object*, lean_object*); @@ -209,6 +218,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40 lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); LEAN_EXPORT uint8_t lean_inductive_val_is_reflexive(lean_object*); LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqDeclaration____x40_Lean_Declaration___hyg_1777____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__7; static lean_object* l_Lean_ConstantInfo_value_x21___closed__1; @@ -218,6 +228,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedInductiveVal___closed__1; static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____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; static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__13; @@ -231,9 +242,11 @@ uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); 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*); 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_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__20; LEAN_EXPORT lean_object* l_Lean_instBEqConstructorVal; LEAN_EXPORT lean_object* l_Lean_instReprDefinitionSafety; @@ -251,6 +264,8 @@ LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__4(lean_ob 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*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__17; +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; @@ -273,10 +288,13 @@ 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_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____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 uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(lean_object*, 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); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_687____closed__10; LEAN_EXPORT uint8_t l_Lean_instInhabitedQuotKind; @@ -292,7 +310,6 @@ LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstantVal____x40_L 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*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqDeclaration____x40_Lean_Declaration___hyg_1777_(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2780_(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; LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_getHeightEx___boxed(lean_object*); @@ -2506,6 +2523,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: { @@ -3228,7 +3524,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_2780_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(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; @@ -3319,11 +3615,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2780____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906____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_2780_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3334,7 +3630,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_2780____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906____boxed), 2, 0); return x_1; } } @@ -3414,7 +3710,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_2991_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117_(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; @@ -3450,11 +3746,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2991____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117____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_2991_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3465,7 +3761,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_2991____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117____boxed), 2, 0); return x_1; } } @@ -3506,7 +3802,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_3177____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_3303____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3539,7 +3835,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_2991_(x_6, x_8); +x_10 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3117_(x_6, x_8); if (x_10 == 0) { uint8_t x_11; @@ -3556,7 +3852,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(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; @@ -3638,7 +3934,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_3177____spec__1(x_9, x_18); +x_38 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____spec__1(x_9, x_18); if (x_38 == 0) { uint8_t x_39; @@ -3721,22 +4017,22 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177____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_3303____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_3177____spec__1(x_1, x_2); +x_3 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____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_3177____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____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_3177_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3747,7 +4043,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_3177____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303____boxed), 2, 0); return x_1; } } @@ -4422,7 +4718,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(461u); +x_3 = lean_unsigned_to_nat(474u); x_4 = lean_unsigned_to_nat(33u); 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); @@ -4608,7 +4904,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(481u); +x_3 = lean_unsigned_to_nat(494u); 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); @@ -4862,6 +5158,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 495d6483a7e10..624d4ec710c70 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -298,7 +298,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); @@ -1092,7 +1092,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___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; 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*); @@ -21719,7 +21719,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); @@ -21738,7 +21739,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); @@ -33160,7 +33162,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); @@ -33214,6 +33215,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; } } @@ -44526,8 +44528,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); @@ -44541,7 +44541,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; @@ -45193,7 +45193,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; @@ -45472,7 +45472,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); @@ -45815,6 +45814,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 @@ -46320,6 +46320,7 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_4); +lean_dec(x_1); return x_14; } } @@ -52572,6 +52573,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; @@ -52652,6 +52654,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; @@ -52776,6 +52779,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; @@ -52977,6 +52981,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 eef27b8678f9c..8a64bd28fe116 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 c70d7a2bb1476..b3c35e6d7b79f 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 0cea934c96528..bf3fcce7c4fb4 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -244,7 +244,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*); @@ -876,6 +876,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*); @@ -1140,7 +1141,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*); @@ -5924,7 +5924,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; } @@ -8663,8 +8663,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; @@ -8705,8 +8704,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; @@ -11027,8 +11025,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; @@ -11068,8 +11065,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; @@ -24185,15 +24181,14 @@ return x_8; static lean_object* _init_l_Lean_Elab_Command_failIfSucceeds___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_elabOpen___spec__15___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_elabOpen___spec__15___closed__3; +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_Elab_Command_failIfSucceeds___closed__2() { diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index 2019dcafa77ec..3aa0b60e8a673 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*); @@ -862,7 +862,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; @@ -3665,7 +3665,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; @@ -3756,6 +3757,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; @@ -3828,6 +3830,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; @@ -3913,6 +3916,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; @@ -4115,7 +4119,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; @@ -4214,6 +4219,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; @@ -9385,7 +9391,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); @@ -9493,7 +9500,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); @@ -9617,7 +9625,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; @@ -9652,7 +9661,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; @@ -10408,7 +10418,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; @@ -10438,7 +10449,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; @@ -10551,7 +10563,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; @@ -11270,7 +11283,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; @@ -11281,7 +11295,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); @@ -11330,7 +11345,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; @@ -11341,7 +11357,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); @@ -12944,7 +12961,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_Term_elabSorry___closed__3; lean_inc(x_12); lean_ctor_set_tag(x_14, 2); @@ -13011,7 +13029,8 @@ lean_dec(x_14); 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_Term_elabSorry___closed__3; lean_inc(x_12); x_51 = lean_alloc_ctor(2, 2, 0); @@ -16268,7 +16287,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); @@ -16450,7 +16470,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 9856730051524..1cd32d2b9dbf9 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; @@ -983,6 +982,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 15068446f903b..433ff8e50c157 100644 --- a/stage0/stdlib/Lean/Elab/Calc.c +++ b/stage0/stdlib/Lean/Elab/Calc.c @@ -191,7 +191,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*); @@ -16048,7 +16048,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; @@ -16076,7 +16077,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; @@ -16134,7 +16136,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 d2d8402f591a2..8bfeea03d74a4 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -45,7 +45,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*); @@ -300,7 +298,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; @@ -314,7 +311,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; @@ -325,7 +321,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*); @@ -443,6 +438,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_ 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; @@ -480,7 +477,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*); @@ -746,7 +742,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*); @@ -849,7 +845,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*); @@ -953,6 +948,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,21 +3812,20 @@ 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: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_Command_mkState___closed__11; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_mkState___closed__11; +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___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; @@ -3857,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); @@ -3887,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; } } } @@ -4549,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; } } } @@ -20315,15 +20292,14 @@ return x_5; static lean_object* _init_l_Lean_Elab_Command_wrapAsyncAsSnapshot___lambda__5___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_Command_wrapAsyncAsSnapshot___lambda__5___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_wrapAsyncAsSnapshot___lambda__5___closed__3; +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_Elab_Command_wrapAsyncAsSnapshot___lambda__5___closed__5() { @@ -21397,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; } @@ -21412,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); @@ -26342,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); @@ -26523,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); @@ -26946,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); @@ -27127,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); @@ -33722,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); @@ -34075,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); @@ -40774,7 +40756,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); @@ -41327,7 +41309,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); @@ -41678,7 +41660,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); @@ -43949,10 +43931,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 a4caad04db0f8..fb812a9202bc2 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*); @@ -4169,8 +4169,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; @@ -4210,8 +4209,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 1f95560e20470..78a45c2ba55c4 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 4b8c4164d13f1..a17fbdc4c44ce 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; } } @@ -6025,7 +6026,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); @@ -6270,7 +6272,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); @@ -7027,7 +7030,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); @@ -7209,7 +7213,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 c7e383470601a..701f004918e9f 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 86fc53cbaf5d9..a2190cac2d28b 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*); @@ -605,15 +605,14 @@ return x_5; static lean_object* _init_l_Lean_Elab_instToSnapshotTreeHeaderProcessedSnapshot___lambda__1___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_instToSnapshotTreeHeaderProcessedSnapshot___lambda__1___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_instToSnapshotTreeHeaderProcessedSnapshot___lambda__1___closed__3; +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_Elab_instToSnapshotTreeHeaderProcessedSnapshot___lambda__1___closed__5() { @@ -1956,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); @@ -2137,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); @@ -2560,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); @@ -2741,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 ed12ae27b896d..affdc93b9ebee 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -70,7 +70,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_3922____closed__3; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; 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*); @@ -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*); @@ -740,7 +740,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; @@ -842,7 +843,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; @@ -929,7 +931,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; @@ -1042,7 +1045,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; @@ -2440,7 +2444,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); @@ -2613,7 +2618,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); @@ -4319,7 +4325,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; @@ -5033,7 +5040,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; @@ -5322,7 +5330,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; @@ -6740,7 +6749,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); @@ -6833,7 +6843,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); @@ -6939,7 +6950,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); @@ -7042,7 +7054,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); @@ -7674,7 +7687,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); @@ -7739,7 +7753,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); @@ -8426,7 +8441,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); @@ -8588,7 +8604,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); @@ -9226,8 +9243,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; @@ -9268,8 +9284,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; @@ -10225,7 +10240,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; @@ -10271,7 +10286,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; @@ -10324,6 +10339,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); @@ -10529,6 +10545,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 c43b25505a079..491da4f38ea00 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 45cd2c0108a8a..0256e70d38a1c 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -87,7 +87,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*); @@ -326,7 +326,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_ 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_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4877____closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4877____closed__11; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__1; @@ -750,7 +750,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); @@ -823,7 +824,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); @@ -1403,7 +1405,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); @@ -1619,7 +1622,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); @@ -2045,7 +2049,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); @@ -2180,7 +2185,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); @@ -2427,7 +2433,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); @@ -2469,7 +2476,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); @@ -6081,7 +6089,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); @@ -6260,7 +6269,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); @@ -6431,7 +6441,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); @@ -6640,7 +6651,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); @@ -7032,7 +7044,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); @@ -7468,7 +7481,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); @@ -8725,7 +8739,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; @@ -8922,7 +8937,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; @@ -10514,8 +10530,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; @@ -10556,8 +10571,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 7b56728d9eec1..5d636a720fce6 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; @@ -384,7 +384,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); @@ -4725,7 +4734,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; @@ -5000,7 +5010,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; @@ -5280,7 +5291,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; @@ -5534,7 +5546,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); @@ -5642,7 +5655,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); @@ -5774,7 +5788,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); @@ -5909,7 +5924,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; @@ -6002,7 +6018,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; @@ -6119,7 +6136,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; @@ -6271,7 +6289,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; @@ -6332,7 +6351,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; @@ -6407,7 +6427,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; @@ -6513,7 +6534,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; @@ -6606,7 +6628,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; @@ -6723,7 +6746,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; @@ -6837,7 +6861,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; @@ -6870,7 +6895,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; @@ -7812,7 +7838,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); @@ -7864,7 +7891,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); @@ -8128,7 +8156,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); @@ -8393,7 +8422,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); @@ -8451,7 +8481,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); @@ -10095,7 +10126,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); @@ -10177,7 +10209,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); @@ -10269,7 +10302,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); @@ -10340,7 +10374,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); @@ -10438,7 +10473,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); @@ -10540,7 +10576,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); @@ -10950,7 +10987,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); @@ -11161,7 +11199,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); @@ -11615,7 +11654,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; @@ -11650,7 +11690,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; @@ -11711,7 +11752,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; @@ -11795,7 +11837,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; @@ -11836,7 +11879,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; @@ -12673,7 +12717,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); @@ -12759,7 +12804,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); @@ -12878,7 +12924,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); @@ -12964,7 +13011,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); @@ -13202,7 +13250,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); @@ -13349,7 +13398,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); @@ -13946,7 +13996,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); @@ -14039,7 +14090,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); @@ -14181,7 +14233,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); @@ -14284,7 +14337,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); @@ -14412,7 +14466,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); @@ -15007,7 +15062,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); @@ -15117,7 +15173,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); @@ -15252,7 +15309,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); @@ -15435,7 +15493,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); @@ -15555,7 +15614,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); @@ -15700,7 +15760,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); @@ -15866,7 +15927,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); @@ -17682,7 +17744,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; @@ -17728,7 +17790,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; @@ -17781,6 +17843,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); @@ -18017,6 +18080,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 b8f6eef248d01..eff745b58f9fa 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -57,7 +57,7 @@ static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_ static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__23; 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; -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_2336____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2336_(lean_object*); @@ -166,7 +166,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); 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*); @@ -1148,7 +1148,8 @@ x_67 = lean_ctor_get(x_65, 0); x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); lean_dec(x_67); -x_69 = lean_environment_main_module(x_68); +x_69 = l_Lean_Environment_mainModule(x_68); +lean_dec(x_68); x_70 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_64); lean_inc(x_69); @@ -1213,7 +1214,8 @@ lean_dec(x_65); x_96 = lean_ctor_get(x_94, 0); lean_inc(x_96); lean_dec(x_94); -x_97 = lean_environment_main_module(x_96); +x_97 = l_Lean_Environment_mainModule(x_96); +lean_dec(x_96); x_98 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_64); lean_inc(x_97); @@ -1294,7 +1296,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_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_135 = l_Lean_addMacroScope(x_133, x_134, x_128); x_136 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1381,7 +1384,8 @@ lean_dec(x_129); x_169 = lean_ctor_get(x_167, 0); lean_inc(x_169); lean_dec(x_167); -x_170 = lean_environment_main_module(x_169); +x_170 = l_Lean_Environment_mainModule(x_169); +lean_dec(x_169); x_171 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_172 = l_Lean_addMacroScope(x_170, x_171, x_128); x_173 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1491,7 +1495,8 @@ if (lean_is_exclusive(x_211)) { x_215 = lean_ctor_get(x_212, 0); lean_inc(x_215); lean_dec(x_212); -x_216 = lean_environment_main_module(x_215); +x_216 = l_Lean_Environment_mainModule(x_215); +lean_dec(x_215); x_217 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_218 = l_Lean_addMacroScope(x_216, x_217, x_210); x_219 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1603,7 +1608,8 @@ x_260 = lean_ctor_get(x_258, 0); x_261 = lean_ctor_get(x_260, 0); lean_inc(x_261); lean_dec(x_260); -x_262 = lean_environment_main_module(x_261); +x_262 = l_Lean_Environment_mainModule(x_261); +lean_dec(x_261); x_263 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_257); lean_inc(x_262); @@ -1668,7 +1674,8 @@ lean_dec(x_258); x_289 = lean_ctor_get(x_287, 0); lean_inc(x_289); lean_dec(x_287); -x_290 = lean_environment_main_module(x_289); +x_290 = l_Lean_Environment_mainModule(x_289); +lean_dec(x_289); x_291 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_257); lean_inc(x_290); @@ -3954,7 +3961,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); @@ -4047,7 +4055,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); @@ -4153,7 +4162,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); @@ -4256,7 +4266,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); @@ -5503,7 +5514,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; @@ -5549,7 +5560,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; @@ -5602,6 +5613,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); @@ -5820,6 +5832,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 88144fef7e0f4..53cc80aa35343 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c @@ -74,7 +74,7 @@ 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*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2632____closed__9; 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*); @@ -235,7 +235,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*); @@ -6798,7 +6798,8 @@ 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 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_34); lean_ctor_set_tag(x_53, 2); @@ -6848,7 +6849,8 @@ lean_dec(x_53); 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_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_34); x_80 = lean_alloc_ctor(2, 2, 0); @@ -6961,7 +6963,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_main_module(x_116); +x_117 = l_Lean_Environment_mainModule(x_116); +lean_dec(x_116); x_118 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_34); if (lean_is_scalar(x_115)) { @@ -7105,7 +7108,8 @@ if (lean_is_exclusive(x_161)) { x_165 = lean_ctor_get(x_162, 0); lean_inc(x_165); lean_dec(x_162); -x_166 = lean_environment_main_module(x_165); +x_166 = l_Lean_Environment_mainModule(x_165); +lean_dec(x_165); x_167 = 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_164)) { @@ -7365,7 +7369,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_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__4; x_36 = l_Lean_addMacroScope(x_34, x_35, x_28); lean_inc(x_2); @@ -7403,7 +7408,8 @@ lean_dec(x_29); 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_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__4; x_48 = l_Lean_addMacroScope(x_46, x_47, x_28); lean_inc(x_2); @@ -7946,7 +7952,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; @@ -8625,7 +8632,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; @@ -8890,7 +8898,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; @@ -9171,7 +9180,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; @@ -12112,8 +12122,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; @@ -12154,8 +12163,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; @@ -13143,7 +13151,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; @@ -13189,7 +13197,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; @@ -13242,6 +13250,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); @@ -13481,6 +13490,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 7795a0790c2d7..711360202f25a 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 d4f08def4ed07..5660983bc4f1e 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -64,7 +64,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*, 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_mkMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,7 +202,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2785____closed__8; 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; @@ -1056,7 +1056,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); @@ -1115,7 +1116,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); @@ -2305,7 +2307,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); @@ -2375,7 +2378,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); @@ -2414,7 +2418,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); @@ -2463,7 +2468,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); @@ -2515,7 +2521,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); @@ -2561,7 +2568,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); @@ -2658,7 +2666,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); @@ -2704,7 +2713,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); @@ -2795,7 +2805,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); @@ -2879,7 +2890,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); @@ -2925,7 +2937,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); @@ -3036,7 +3049,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); @@ -3125,7 +3139,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); @@ -3171,7 +3186,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); @@ -3318,7 +3334,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); @@ -3407,7 +3424,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); @@ -3453,7 +3471,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); @@ -3637,7 +3656,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); @@ -3726,7 +3746,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); @@ -3772,7 +3793,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); @@ -3994,7 +4016,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); @@ -4083,7 +4106,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); @@ -4129,7 +4153,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); @@ -4374,7 +4399,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); @@ -4463,7 +4489,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); @@ -4509,7 +4536,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); @@ -4769,7 +4797,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); @@ -4858,7 +4887,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); @@ -4904,7 +4934,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); @@ -5248,7 +5279,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); @@ -5337,7 +5369,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); @@ -5383,7 +5416,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); @@ -5745,7 +5779,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); @@ -5834,7 +5869,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); @@ -5880,7 +5916,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); @@ -6267,7 +6304,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); @@ -6356,7 +6394,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); @@ -6402,7 +6441,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); @@ -7565,7 +7605,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); @@ -7658,7 +7699,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); @@ -7779,7 +7821,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); @@ -7882,7 +7925,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); @@ -8609,7 +8653,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); @@ -8674,7 +8719,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); @@ -9377,7 +9423,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; @@ -9423,7 +9469,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; @@ -9476,6 +9522,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); @@ -9694,6 +9741,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 55f2e31935669..3b1bf5c1e8e7a 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -62,7 +62,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 @@ LEAN_EXPORT uint8_t l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBo 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*); static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3579____closed__13; -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; @@ -616,7 +616,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); @@ -710,7 +711,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); @@ -1486,7 +1488,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); @@ -1635,7 +1638,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); @@ -2031,7 +2035,8 @@ x_49 = lean_ctor_get(x_46, 1); x_50 = lean_ctor_get(x_48, 0); lean_inc(x_50); lean_dec(x_48); -x_51 = lean_environment_main_module(x_50); +x_51 = l_Lean_Environment_mainModule(x_50); +lean_dec(x_50); x_52 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_44); lean_ctor_set_tag(x_46, 2); @@ -2188,7 +2193,8 @@ lean_dec(x_46); 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_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_44); x_91 = lean_alloc_ctor(2, 2, 0); @@ -2569,7 +2575,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; @@ -2621,7 +2628,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; @@ -2851,7 +2859,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); @@ -3547,7 +3556,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); @@ -3620,7 +3630,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); @@ -3709,7 +3720,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); @@ -3767,7 +3779,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); @@ -3843,7 +3856,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); @@ -3906,7 +3920,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); @@ -4211,7 +4226,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); @@ -4310,7 +4326,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); @@ -4396,7 +4413,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); @@ -5513,7 +5531,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); @@ -5598,7 +5617,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); @@ -5772,7 +5792,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); @@ -5981,7 +6002,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); @@ -6212,7 +6234,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); @@ -6422,7 +6445,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); @@ -6529,7 +6553,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)) { @@ -6775,7 +6800,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); @@ -6886,7 +6912,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)) { @@ -8157,7 +8184,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); @@ -8250,7 +8278,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); @@ -8356,7 +8385,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); @@ -8459,7 +8489,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); @@ -9691,7 +9722,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; @@ -9737,7 +9768,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; @@ -9790,6 +9821,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); @@ -10008,6 +10040,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 f46b42bbdd087..e6f203a6f1954 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/TypeName.c b/stage0/stdlib/Lean/Elab/Deriving/TypeName.c index 2c9f03eaf4f16..c478a486db1da 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 510c4893c1fea..d32abf4d4d44a 100644 --- a/stage0/stdlib/Lean/Elab/Do.c +++ b/stage0/stdlib/Lean/Elab/Do.c @@ -1091,7 +1091,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*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_37390____closed__9; 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*); @@ -7509,7 +7509,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); @@ -8757,7 +8758,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); @@ -9318,7 +9320,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); @@ -9460,7 +9463,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); @@ -9717,7 +9721,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); @@ -9738,7 +9743,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); @@ -13983,7 +13989,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); @@ -14269,7 +14276,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); @@ -30483,7 +30491,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); @@ -33271,7 +33280,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); @@ -34364,7 +34374,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); @@ -34646,7 +34657,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); @@ -36461,7 +36473,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); @@ -36482,7 +36495,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); @@ -36975,7 +36989,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); @@ -37310,7 +37325,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; @@ -37400,7 +37416,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; @@ -37689,7 +37706,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; @@ -38906,7 +38924,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); @@ -39593,7 +39612,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); @@ -39694,7 +39714,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); @@ -39877,7 +39898,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); @@ -40489,7 +40511,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); @@ -40535,7 +40558,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); @@ -41366,7 +41390,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); @@ -41403,7 +41428,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); @@ -41473,7 +41499,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); @@ -41568,7 +41595,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); @@ -41603,7 +41631,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); @@ -41635,7 +41664,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); @@ -41827,7 +41857,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); @@ -42515,7 +42546,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); @@ -42724,7 +42756,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); @@ -43079,7 +43112,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); @@ -43237,7 +43271,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); @@ -43339,7 +43374,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 4071e38e6b896..915a6220b2d73 100644 --- a/stage0/stdlib/Lean/Elab/Extra.c +++ b/stage0/stdlib/Lean/Elab/Extra.c @@ -372,7 +372,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*); @@ -1482,7 +1482,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); @@ -1547,7 +1548,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); @@ -1631,7 +1633,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)) { @@ -4074,7 +4077,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); @@ -4139,7 +4143,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); @@ -4223,7 +4228,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/Frontend.c b/stage0/stdlib/Lean/Elab/Frontend.c index dc6123f0fa0e9..253320e6a869e 100644 --- a/stage0/stdlib/Lean/Elab/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Frontend.c @@ -597,15 +597,14 @@ return x_5; static lean_object* _init_l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__3; +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; } } LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { diff --git a/stage0/stdlib/Lean/Elab/GuardMsgs.c b/stage0/stdlib/Lean/Elab/GuardMsgs.c index 95ef7ecf7e774..1d173b11bdd30 100644 --- a/stage0/stdlib/Lean/Elab/GuardMsgs.c +++ b/stage0/stdlib/Lean/Elab/GuardMsgs.c @@ -9066,15 +9066,14 @@ return x_5; static lean_object* _init_l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__3; +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_Elab_Tactic_GuardMsgs_elabGuardMsgs___lambda__1___closed__5() { diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 261e7304033a1..d3c33153e9b4f 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 0b9b207824656..08080c93725fe 100644 --- a/stage0/stdlib/Lean/Elab/InfoTree/Main.c +++ b/stage0/stdlib/Lean/Elab/InfoTree/Main.c @@ -109,7 +109,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; LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatElabInfo(lean_object*, lean_object*); @@ -150,13 +150,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; @@ -210,6 +208,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___rarg___lambda__6___bo lean_object* l_Lean_MessageData_ofFormat(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*); @@ -451,6 +450,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*); @@ -2404,7 +2404,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) { @@ -2458,7 +2458,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); @@ -2487,7 +2487,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); @@ -2569,7 +2569,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) { @@ -2644,7 +2644,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 { @@ -2785,15 +2785,14 @@ return x_2; static lean_object* _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__11() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; +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_Elab_ContextInfo_runCoreM___rarg___closed__12() { @@ -2916,7 +2915,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) { @@ -3134,7 +3133,7 @@ x_86 = lean_ctor_get(x_83, 0); x_87 = lean_ctor_get(x_83, 4); lean_dec(x_87); x_88 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__15; -x_89 = l_Lean_Kernel_enableDiag(x_86, x_88); +x_89 = l_Lean_Kernel_Environment_enableDiag(x_86, x_88); lean_ctor_set(x_83, 4, x_56); lean_ctor_set(x_83, 0, x_89); x_90 = lean_st_ref_set(x_74, x_83, x_84); @@ -3249,7 +3248,7 @@ lean_inc(x_113); lean_inc(x_112); lean_dec(x_83); x_119 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__15; -x_120 = l_Lean_Kernel_enableDiag(x_112, x_119); +x_120 = l_Lean_Kernel_Environment_enableDiag(x_112, x_119); x_121 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_121, 0, x_120); lean_ctor_set(x_121, 1, x_113); @@ -9398,8 +9397,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; @@ -9439,8 +9437,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 0c8deb1316a6d..139047cffd5d8 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 550eafd03602e..c9775fd0f40f8 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 67bb2b0b937f2..aae5367706c31 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -221,7 +221,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; @@ -325,12 +325,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_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*); @@ -879,7 +878,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*); @@ -1126,6 +1125,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); @@ -11658,7 +11660,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; @@ -11977,7 +11980,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; @@ -12313,7 +12317,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; @@ -12782,7 +12787,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; @@ -14967,7 +14973,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; @@ -15203,7 +15210,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; @@ -15459,7 +15467,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; @@ -15991,7 +16000,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); @@ -16430,7 +16440,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; @@ -16709,7 +16720,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; @@ -23350,7 +23362,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; @@ -24755,7 +24768,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; @@ -24945,7 +24959,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; @@ -30013,7 +30028,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) { @@ -30063,7 +30078,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); @@ -30093,7 +30108,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); @@ -45419,7 +45434,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); @@ -45601,7 +45617,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); @@ -48582,7 +48599,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); @@ -48626,7 +48644,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); @@ -48665,7 +48684,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); @@ -48735,7 +48755,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); @@ -48786,7 +48807,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); @@ -50403,8 +50425,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; @@ -50517,8 +50538,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; @@ -50913,7 +50933,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); @@ -51095,7 +51116,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); @@ -52270,7 +52292,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); @@ -52312,7 +52335,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); @@ -52351,7 +52375,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); @@ -52421,7 +52446,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); @@ -52470,7 +52496,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); @@ -53871,7 +53898,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); @@ -53895,7 +53923,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/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 168e9708374a2..e4a9f14949b2b 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -4023,15 +4023,14 @@ return x_5; static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__4() { _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_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__3; +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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__5() { diff --git a/stage0/stdlib/Lean/Elab/Open.c b/stage0/stdlib/Lean/Elab/Open.c index 3defbf6367a9e..3214c1fab7c69 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 2d9c7364c5e8f..e4278531177fc 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*); @@ -359,7 +359,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*); @@ -392,7 +392,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; @@ -466,6 +465,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); @@ -8347,8 +8346,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; @@ -8388,8 +8386,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; @@ -12182,7 +12179,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; @@ -12211,7 +12209,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; @@ -12367,7 +12366,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); @@ -14144,8 +14144,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; @@ -14250,8 +14249,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; @@ -16573,7 +16571,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); @@ -16751,7 +16750,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 0b8768943578d..9a24be587c689 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c @@ -113,12 +113,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*); @@ -165,6 +163,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*); @@ -340,6 +339,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*); @@ -11096,7 +11096,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) { @@ -11142,7 +11142,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); @@ -11172,7 +11172,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 3f93c275b19fe..d6ff4d8a082b3 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); @@ -2220,7 +2220,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) { @@ -2255,7 +2254,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 5f287a283f3bd..f69661e22a58c 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c @@ -159,14 +159,12 @@ lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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_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; -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*); @@ -234,6 +232,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*); @@ -521,6 +520,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*); @@ -3670,7 +3670,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) { @@ -3720,7 +3720,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); @@ -3750,7 +3750,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); @@ -30807,7 +30807,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) { @@ -30858,7 +30858,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); @@ -30891,7 +30891,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); @@ -30949,7 +30949,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 f952e27b30866..b9aff272a7abd 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c @@ -33,7 +33,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Nonrec_mkEqns___lambda__2___boxed(lean_obje 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*); static lean_object* l_Lean_Elab_Nonrec_initFn____x40_Lean_Elab_PreDefinition_Nonrec_Eqns___hyg_1608____closed__1; -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*); @@ -44,10 +44,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*, 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; @@ -65,6 +63,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*); @@ -147,6 +146,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); @@ -701,8 +701,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; @@ -766,8 +765,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; @@ -4882,7 +4880,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) { @@ -4928,7 +4926,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); @@ -4958,7 +4956,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); @@ -5080,8 +5078,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; @@ -5521,8 +5518,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 ce2a3e9fb3969..d68f864700f34 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 ead5f8861e2a6..b87669060076f 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c @@ -60,11 +60,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___lambda__1___boxed(lean_ 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); static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1495____closed__18; 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_EXPORT lean_object* l_Lean_Elab_Structural_eqnInfoExt; @@ -86,6 +84,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_ 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_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1495____closed__14; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1240____closed__3; 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*); @@ -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*); static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1495____closed__19; @@ -3402,7 +3402,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) { @@ -3448,7 +3448,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); @@ -3478,7 +3478,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 6f77bfaca8c2f..ac7f71aff7963 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 8a92378a37dd1..7f88b2bcf216a 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; @@ -8657,8 +8657,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; @@ -8698,8 +8697,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 d10a40ec16e6d..54e2bb325b3ca 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c @@ -8828,7 +8828,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 4401391335fc6..4f834e06ebf06 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -76,12 +76,10 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ 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*); static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_1553____closed__4; -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_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_1828____closed__18; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__5; @@ -108,6 +106,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*); @@ -218,6 +217,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*); static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_1828____closed__1; 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*); @@ -4523,7 +4523,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) { @@ -4569,7 +4569,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); @@ -4600,7 +4600,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 e180b9de42b23..fa4428cea4c5b 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*); @@ -1199,8 +1199,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; @@ -1240,8 +1239,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; @@ -1333,8 +1331,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; @@ -1374,8 +1371,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 50e555d038790..7fbc3958cce38 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c @@ -194,7 +194,7 @@ LEAN_EXPORT lean_object* l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex 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*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_4200____closed__7; -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*); @@ -4165,8 +4165,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; @@ -4206,8 +4205,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; @@ -4307,8 +4305,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; @@ -4348,8 +4345,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 42b0e414df6a9..a381b25f418c6 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c @@ -102,9 +102,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; @@ -152,6 +150,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*); @@ -326,6 +325,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); @@ -8912,7 +8912,7 @@ lean_dec(x_59); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -x_63 = l_Lean_Kernel_isDiagnosticsEnabled(x_62); +x_63 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_62); lean_dec(x_62); if (x_63 == 0) { @@ -9154,7 +9154,7 @@ lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean x_69 = lean_ctor_get(x_66, 0); x_70 = lean_ctor_get(x_66, 4); lean_dec(x_70); -x_71 = l_Lean_Kernel_enableDiag(x_69, x_58); +x_71 = l_Lean_Kernel_Environment_enableDiag(x_69, x_58); x_72 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__4___closed__3; lean_ctor_set(x_66, 4, x_72); lean_ctor_set(x_66, 0, x_71); @@ -9234,7 +9234,7 @@ lean_inc(x_84); lean_inc(x_83); lean_inc(x_82); lean_dec(x_66); -x_89 = l_Lean_Kernel_enableDiag(x_82, x_58); +x_89 = l_Lean_Kernel_Environment_enableDiag(x_82, x_58); x_90 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__4___closed__3; x_91 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_91, 0, x_89); diff --git a/stage0/stdlib/Lean/Elab/Print.c b/stage0/stdlib/Lean/Elab/Print.c index 6ca4ab2d48379..1fa1241b39530 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; @@ -8341,7 +8339,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); @@ -8502,7 +8499,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); @@ -9779,9 +9775,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 50202a43b472c..60ea71ddfea09 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -915,7 +915,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 cb02223ea10df..34cd97cb134fe 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 363b1d22d4829..bc5218930f508 100644 --- a/stage0/stdlib/Lean/Elab/StructInst.c +++ b/stage0/stdlib/Lean/Elab/StructInst.c @@ -650,7 +650,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; @@ -3590,7 +3590,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); @@ -3623,7 +3624,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); @@ -3676,7 +3678,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); @@ -3760,7 +3763,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); @@ -3800,7 +3804,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); @@ -6373,7 +6378,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); @@ -6571,7 +6577,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); @@ -6843,7 +6850,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); @@ -7071,7 +7079,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); @@ -7275,7 +7284,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); @@ -7469,7 +7479,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); @@ -28699,7 +28710,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); @@ -28707,7 +28717,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) { @@ -28772,7 +28781,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); @@ -28780,7 +28788,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) { @@ -28899,7 +28906,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); @@ -28907,7 +28913,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) { @@ -38851,7 +38856,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); @@ -39033,7 +39039,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 7559801b68bc9..9bc9821d38d81 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -902,7 +902,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; @@ -910,6 +910,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; @@ -1222,7 +1223,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*); @@ -3437,7 +3437,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); @@ -3619,7 +3620,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); @@ -8107,7 +8109,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; @@ -8150,7 +8153,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; @@ -8237,7 +8241,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; @@ -8405,7 +8410,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; @@ -8683,7 +8689,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; @@ -15249,7 +15256,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); @@ -24850,7 +24856,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; @@ -31475,7 +31481,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 3512919e6a09a..747557cac87ea 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -594,7 +594,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 57a9f1ed678de..c0836a88ba8b8 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -434,7 +434,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_8550____closed__12; @@ -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); @@ -10002,7 +10004,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); @@ -10180,7 +10183,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 def957e4a16b1..45b5481a346d3 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; @@ -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*); @@ -15243,8 +15243,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; @@ -15284,8 +15283,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; @@ -22638,7 +22636,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) { @@ -22692,7 +22690,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); @@ -22722,7 +22720,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); @@ -30341,7 +30339,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); @@ -30481,7 +30480,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); @@ -30638,7 +30638,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); @@ -30778,7 +30779,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); @@ -30937,7 +30939,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); @@ -31077,7 +31080,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); @@ -48745,7 +48749,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; @@ -48829,7 +48834,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 a00fde5def871..f830f06138ce4 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 e5816b88e9ed5..98123dc801a21 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_112; @@ -10836,9 +10835,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 c45f740d17c33..518176567a969 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_elabDecideConfig___lambda__2___closed__1; @@ -329,6 +327,7 @@ lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__4___closed__8; static lean_object* l_Lean_Elab_Tactic_evalRename___closed__2; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_ElabTerm___hyg_6163____closed__1; +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*); @@ -17704,7 +17704,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) { @@ -17754,7 +17754,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); @@ -17784,7 +17784,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 f6eb1010e2f76..ff5831dfbe278 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Ext.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Ext.c @@ -174,7 +174,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/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index 81d9b5cb6fe29..056a98f666f32 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -163,7 +163,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*); @@ -21633,7 +21633,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_find(x_26, x_21); +x_27 = l_Lean_Environment_find_x3f(x_26, x_21); +lean_dec(x_21); 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; @@ -21764,7 +21765,8 @@ lean_dec(x_22); x_58 = lean_ctor_get(x_56, 0); lean_inc(x_58); lean_dec(x_56); -x_59 = lean_environment_find(x_58, x_21); +x_59 = l_Lean_Environment_find_x3f(x_58, x_21); +lean_dec(x_21); 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; @@ -22430,7 +22432,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; @@ -22476,7 +22478,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; @@ -23074,7 +23076,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); @@ -23121,7 +23122,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); @@ -23340,6 +23340,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_11; } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Match.c b/stage0/stdlib/Lean/Elab/Tactic/Match.c index c6c3384b7738b..cab281e2eeb63 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 971d3a6bf7fb0..67576d3791e01 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; diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index 36172c148dcad..49cdbed080d44 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -386,7 +386,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Tactic_hashCach static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg___closed__1; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20321____closed__7; 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*, lean_object*); @@ -536,7 +536,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*); @@ -550,7 +549,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*); @@ -791,6 +789,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*); @@ -1320,7 +1319,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); 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_initFn____x40_Lean_Elab_Term___hyg_20554____lambda__3___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*); @@ -1732,6 +1731,7 @@ lean_object* l_ReaderT_bind___at_Lean_Meta_zetaReduce___spec__14___rarg(lean_obj 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_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*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2511,15 +2511,14 @@ return x_1; static lean_object* _init_l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__6; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__6; +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_Elab_Tactic_instInhabitedTacticFinishedSnapshot___closed__2() { @@ -56509,7 +56508,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); @@ -57207,7 +57207,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); @@ -57389,7 +57390,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); @@ -57827,7 +57829,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); @@ -58009,7 +58012,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); @@ -70472,8 +70476,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; @@ -70514,8 +70517,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; @@ -74370,7 +74372,7 @@ lean_dec(x_138); x_141 = lean_ctor_get(x_139, 0); lean_inc(x_141); lean_dec(x_139); -x_142 = l_Lean_Kernel_isDiagnosticsEnabled(x_141); +x_142 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_141); lean_dec(x_141); if (x_142 == 0) { @@ -74424,7 +74426,7 @@ lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; x_148 = lean_ctor_get(x_145, 0); x_149 = lean_ctor_get(x_145, 4); lean_dec(x_149); -x_150 = l_Lean_Kernel_enableDiag(x_148, x_137); +x_150 = l_Lean_Kernel_Environment_enableDiag(x_148, x_137); x_151 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; lean_ctor_set(x_145, 4, x_151); lean_ctor_set(x_145, 0, x_150); @@ -74539,7 +74541,7 @@ lean_inc(x_176); lean_inc(x_175); lean_inc(x_174); lean_dec(x_145); -x_181 = l_Lean_Kernel_enableDiag(x_174, x_137); +x_181 = l_Lean_Kernel_Environment_enableDiag(x_174, x_137); x_182 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; x_183 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_183, 0, x_181); @@ -74773,7 +74775,7 @@ lean_dec(x_243); x_246 = lean_ctor_get(x_244, 0); lean_inc(x_246); lean_dec(x_244); -x_247 = l_Lean_Kernel_isDiagnosticsEnabled(x_246); +x_247 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_246); lean_dec(x_246); if (x_247 == 0) { @@ -74848,7 +74850,7 @@ if (lean_is_exclusive(x_250)) { lean_dec_ref(x_250); x_259 = lean_box(0); } -x_260 = l_Lean_Kernel_enableDiag(x_252, x_242); +x_260 = l_Lean_Kernel_Environment_enableDiag(x_252, x_242); x_261 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; if (lean_is_scalar(x_259)) { x_262 = lean_alloc_ctor(0, 8, 0); @@ -76355,7 +76357,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); @@ -77331,6 +77332,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) @@ -79467,7 +79469,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); @@ -79621,7 +79624,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 578ac6183f15a..8f616ea07894b 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 5b515896a6800..454e23e4a9a71 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -18,10 +18,9 @@ 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*); 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*); +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_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*); @@ -30,34 +29,39 @@ 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*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__2; +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*); 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*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__5; 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_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_header___boxed(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__3; 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_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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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* l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___boxed(lean_object*); +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; @@ -65,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*); @@ -77,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); @@ -87,34 +93,28 @@ 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*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__7(lean_object*, lean_object*, size_t, size_t, 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_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_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__7___boxed(lean_object*, lean_object*, 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; 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); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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*); @@ -132,135 +132,137 @@ 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*, 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__2___boxed(lean_object*, lean_object*, lean_object*, 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*); 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_Lean_Environment_find_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__11___boxed(lean_object*, lean_object*, lean_object*, 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__6(lean_object*, lean_object*, size_t, size_t, 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_Lean_initFn____x40_Lean_Environment___hyg_6724_(lean_object*); 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*); uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_434____spec__1(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_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*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(lean_object*); uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__6; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_mkExtNameMap___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); 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 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_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__9___boxed(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*); 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*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2(lean_object*); uint8_t l_Std_Format_isNil(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__1; +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_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3(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*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__7(lean_object*, lean_object*, size_t, size_t, 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_initFn____x40_Lean_Environment___hyg_6597____lambda__2___boxed(lean_object*); 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* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__5(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__5; 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_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__7; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020_(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; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); 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*); @@ -270,7 +272,7 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___l LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck(lean_object*); 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__10(lean_object*, lean_object*, size_t, size_t, 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*); @@ -288,8 +290,8 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(l 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_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__5___boxed(lean_object*, 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*); @@ -299,9 +301,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; @@ -310,15 +312,16 @@ 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_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*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object*, lean_object*, lean_object*, lean_object*); @@ -327,19 +330,20 @@ 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_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6724____closed__1; 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_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*); @@ -347,55 +351,65 @@ 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; +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*); 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_6597____closed__2; 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; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_isSafeDefinition___boxed(lean_object*, lean_object*); static lean_object* l_Lean_finalizeImport___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__8(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); 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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__7___boxed(lean_object*, lean_object*, lean_object*, 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__24; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object*); +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; +LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2(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*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__1; 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*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__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_object* l_Lean_Declaration_getNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__5(lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(lean_object*, lean_object*, lean_object*); @@ -407,34 +421,31 @@ LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_obj LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_instToStringNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringImport___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__11(lean_object*, lean_object*, size_t, size_t, 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; 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,22 +453,21 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4( LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2(lean_object*, lean_object*, lean_object*); 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__10___boxed(lean_object*, lean_object*, 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; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__3; 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_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__1(lean_object*, lean_object*, 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; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); 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_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__6(lean_object*, lean_object*, size_t, size_t, 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*); @@ -465,51 +475,56 @@ 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*, 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_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__1___boxed(lean_object*, lean_object*, lean_object*); +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_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_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_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_object* l_Lean_withImporting___rarg(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_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_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__1(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,37 +535,41 @@ 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*); 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; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; LEAN_EXPORT lean_object* l_IO_println___at_Lean_Environment_displayStats___spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instReprImport___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___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_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object*, lean_object*, lean_object*); 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_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_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; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____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_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; @@ -566,11 +585,9 @@ 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_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__9(lean_object*, 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; @@ -578,30 +595,29 @@ uint8_t l_Lean_NameSet_contains(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*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_mkTagDeclarationExtension___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__4(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_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__2; 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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597_(lean_object*); -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; @@ -611,39 +627,39 @@ static uint32_t l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1_ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object*, lean_object*, lean_object*); 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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__8___boxed(lean_object*, lean_object*, 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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); 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; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__18; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_imports(lean_object*); @@ -655,13 +671,12 @@ 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; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__4; LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__1; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__1; LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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,60 +691,62 @@ LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__8(lean_object*, size_t, size_t, lean_object*); size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__5; lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); 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_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; +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*); 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*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__3; +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*); 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*); 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*); @@ -747,63 +764,67 @@ 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*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__21; 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*); 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_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*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__2; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__3; +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; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__6; -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); lean_object* lean_save_module_data(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7020____spec__4(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*); 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,3623 +3220,3732 @@ 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) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_ConstantInfo_value_x21(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_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() { +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; +} +} +} +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__2), 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__5() { +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_imp___elambda__1), 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__6() { +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_mkInitialExtStates), 1, 0); -return x_1; +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize), 2, 0); -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; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; 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; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp() { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; -return x_1; +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__1___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__1___rarg(x_1, x_2, x_3); -lean_dec(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); -return x_4; +lean_dec(x_1); +return x_3; } } -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___boxed(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_object* x_3; +x_3 = l_Lean_Environment_getModuleIdxFor_x3f(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__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Environment_isConstructor(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_dec(x_1); +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_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_7; +lean_dec(x_5); +x_7 = 0; +return x_7; +} } -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_isConstructor___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); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_isConstructor(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___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_Environment_isSafeDefinition(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_3; +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; 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_4; +x_4 = 0; +return x_4; +} +else { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 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_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); +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_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; +uint8_t x_10; +lean_dec(x_8); +x_10 = 0; +return x_10; } } else { -uint8_t x_15; -lean_free_object(x_1); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_11; lean_dec(x_5); -lean_dec(x_4); -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) -{ -return x_9; +x_11 = 0; +return x_11; } -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; } } } -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_EXPORT lean_object* l_Lean_Environment_isSafeDefinition___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -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; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_isSafeDefinition(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; } -else +LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -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_3; +x_3 = lean_name_eq(x_2, x_1); +return x_3; } -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; +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, 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_Environment_getModuleIdx_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +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_instInhabited(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; +x_3 = l_Lean_Environment_getModuleIdx_x3f(x_1, x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(lean_object* x_1, lean_object* x_2) { +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) { _start: { lean_object* x_3; -x_3 = l_Lean_EnvExtension_instInhabited(x_1, x_2); -lean_dec(x_2); +x_3 = l_Lean_ConstantInfo_instantiateTypeLevelParams(x_1, x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(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: { -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; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_ConstantInfo_value_x21(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; } -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; } +LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +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_setState(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_setState___rarg___boxed), 3, 0); +lean_inc(x_2); return x_2; } } -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_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_setState___rarg(x_1, x_2, x_3); -lean_dec(x_1); +x_4 = lean_apply_1(x_2, 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_instInhabitedEnvExtensionInterface___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) +lean_inc(x_3); +return x_3; +} +} +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__1() { +_start: { -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); +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); return x_2; } -else +} +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__2() { +_start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_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_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_instInhabitedEnvExtensionInterface___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object* x_1) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__4() { _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_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__2), 3, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__5() { _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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__6() { _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_1; +x_1 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object* x_1) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__7() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_getState___rarg___boxed), 3, 0); -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; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface() { _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); -return x_4; +lean_object* x_1; +x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__7; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { +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_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +x_3 = l_Lean_instInhabitedEnvExtensionInterface___lambda__1(x_1, x_2); +lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object* x_1) { +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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_registerEnvExtension___rarg), 2, 0); -return x_2; +lean_object* x_5; +x_5 = l_Lean_instInhabitedEnvExtensionInterface___lambda__3(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___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object* x_1) { +static uint32_t _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1() { _start: { -lean_object* x_2; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_1); +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_mkEmptyEnvironment___lambda__1___closed__1() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___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; +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; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object* x_1) { _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); +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; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__3() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2() { _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_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; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__4() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; +lean_object* x_2; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; +return x_2; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__5() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1833_(lean_object* x_1) { _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_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_3; +} +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; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__6() { +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1() { _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___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef; +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_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_3); -if (lean_obj_tag(x_4) == 0) +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) { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 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_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_dec(x_7); +lean_dec(x_1); +lean_ctor_set(x_5, 0, x_2); +return x_5; } else { -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_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_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; +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; } +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; } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_4); -if (x_26 == 0) +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; +} +} +} +} +else { -return x_4; +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); +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_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; +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; +} +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_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__1() { +} +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() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("environment objects cannot be created during initialization", 59, 59); +x_1 = lean_mk_string_unchecked("invalid environment extension has been accessed", 47, 47); return x_1; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__2() { +static lean_object* _init_l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg() { _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_object* x_1; +x_1 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1; +return x_1; } } -LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_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_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1() { +_start: { -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_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Environment", 16, 16); +return x_1; } -else +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2() { +_start: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_3); -if (x_9 == 0) +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_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; +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; } -else +} +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_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); +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_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; +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_9; +x_9 = lean_array_fset(x_2, x_4, x_3); +return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(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: { -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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed), 3, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint32_t x_3; lean_object* x_4; -x_3 = lean_unbox_uint32(x_1); +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_2, x_3); lean_dec(x_1); -x_4 = lean_mk_empty_environment(x_3, x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -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); +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtensionState___rarg), 1, 0); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.modifyState", 44, 44); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1(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; 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_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_instInhabitedPersistentEnvExtension___lambda__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_inc(x_1); -return x_1; -} +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_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -return x_2; +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_instInhabitedPersistentEnvExtension___lambda__4(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_box(0); +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed), 3, 0); return x_2; } } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__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() { +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_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed), 2, 0); -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_Lean_instInhabitedPersistentEnvExtension___closed__3() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed), 1, 0); +x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.getState", 41, 41); return x_1; } } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__4() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed), 1, 0); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_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_Lean_instInhabitedPersistentEnvExtension___closed__5() { +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; 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 = 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_5; -x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__5; -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; } -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_9; lean_dec(x_1); -return x_4; -} +x_9 = lean_array_fget(x_3, x_4); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -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_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__3(x_1); -lean_dec(x_1); +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed(lean_object* x_1) { +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_2; -x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__4(x_1); -lean_dec(x_1); -return x_2; +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_instInhabitedPersistentEnvExtension___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* 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) { _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_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; 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_16; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_5); +return x_14; +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_5); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } -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: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = l_Lean_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_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); +} +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 lean_object* l_Lean_PersistentEnvExtension_getModuleEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1() { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed), 4, 0); -return x_4; +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; } } -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_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2() { _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; 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_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_ctor_get(x_1, 3); +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); -lean_dec(x_1); -x_5 = !lean_is_exclusive(x_3); +x_5 = lean_unbox(x_4); +lean_dec(x_4); 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); +uint8_t x_6; +lean_dec(x_1); +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_3, 0); +lean_dec(x_7); +x_8 = 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; } 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_object* x_9; lean_object* x_10; lean_object* x_11; 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); +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_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___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); -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_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg), 3, 0); -return x_4; -} +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_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -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_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object* x_1) { _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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg), 2, 0); +return x_2; } } -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_registerExt___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_PersistentEnvExtension_getState___rarg(x_1, x_2, x_3); -lean_dec(x_3); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(x_1, x_2, 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) { +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_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 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; -} -else +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_2, x_1); +if (x_5 == 0) { -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -lean_dec(x_2); +lean_object* x_6; 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); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); return x_6; } -} -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_setState___rarg___boxed), 3, 0); -return x_4; -} -} -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_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_4; -x_4 = l_Lean_PersistentEnvExtension_setState___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} +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; } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_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_18; +lean_dec(x_9); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 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; +return x_11; } 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_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_Lean_PersistentEnvExtension_modifyState___rarg(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; 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_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_PersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2, lean_object* 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) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___boxed), 3, 0); -return x_4; +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; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +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_4; -x_4 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); -lean_dec(x_1); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2639_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object* x_1) { _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) -{ -return x_3; -} -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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__1() { +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_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__2() { +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_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -return x_1; -} +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; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; +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__4() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__5() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(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__4; -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_setState___rarg(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__6() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__7() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg(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; -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_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__8() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("null", 4, 4); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg), 2, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__9() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(lean_object* x_1, lean_object* x_2) { _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); +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__10() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("exact", 5, 5); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed), 2, 0); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_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; -} -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__12() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2() { _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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4), 1, 0); +return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__13() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3() { _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__12; -x_3 = lean_array_push(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3), 1, 0); +return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__14() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Term", 4, 4); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2), 1, 0); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__15() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("declName", 8, 8); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1), 1, 0); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__16() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_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_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates), 1, 0); +return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__17() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("decl_name%", 10, 10); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize), 2, 0); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__18() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8() { _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; +lean_object* x_1; lean_object* x_2; 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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__19() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp() { _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 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; +return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__20() { +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_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); +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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__21() { +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_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; +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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__22() { +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_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); +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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__23() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_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_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(x_1, x_2); +lean_dec(x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__24() { +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(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; -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_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_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); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); +return x_4; +} +else +{ +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); +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; } } -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_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_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; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__26() { -_start: +} +else +{ +uint8_t x_28; +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_4); +if (x_28 == 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__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() { -_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__26; -x_3 = lean_array_push(x_1, x_2); -return x_3; +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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__28() { +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited(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; -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); -return x_4; +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; +return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684_() { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; +lean_object* x_3; +x_3 = l_Lean_EnvExtension_instInhabited(x_1, x_2); +lean_dec(x_2); +return x_3; } } -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_Lean_EnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; 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) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 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_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 { -uint8_t x_13; -x_13 = 1; -return x_13; +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_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; } } -else -{ -uint8_t x_14; -x_14 = 0; -return x_14; } +LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_setState___rarg___boxed), 3, 0); +return x_2; } } -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_EnvExtension_setState___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_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___boxed), 4, 0); +x_4 = l_Lean_EnvExtension_setState___rarg(x_1, x_2, x_3); +lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object* x_1, lean_object* x_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_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; -} -} -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1() { -_start: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1), 2, 0); -return x_1; -} +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; } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_persistentEnvExtensionsRef; -return x_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); +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_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object* x_1) { _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_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_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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_modifyState___rarg___boxed), 3, 0); +return x_2; } -else +} +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_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_4; +x_4 = l_Lean_EnvExtension_modifyState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -else -{ -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_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -return x_13; +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; } -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 lean_object* l_Lean_EnvExtension_getState(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_getState___rarg___boxed), 3, 0); +return x_2; } } -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_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_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_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_4; +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_ctor_set(x_50, 0, x_41); -lean_ctor_set(x_50, 1, x_48); -return x_50; } -else +LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { +_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); +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +return x_3; } -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_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_registerEnvExtension___rarg), 2, 0); +return x_2; } } +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_1); +return x_2; } } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1() { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid environment extension, '", 32, 32); -return x_1; +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_registerPersistentEnvExtensionUnsafe___rarg___closed__2() { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' has already been used", 23, 23); -return x_1; +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_registerPersistentEnvExtensionUnsafe___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__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_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; 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; } -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_mkEmptyEnvironment___lambda__1___closed__4() { +_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); +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_mkEmptyEnvironment___lambda__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_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; 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_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_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_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; +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; +} } 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) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_4); +if (x_28 == 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; +return x_4; } else { -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_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_registerPersistentEnvExtensionUnsafe(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_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) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__1() { _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 = lean_mk_string_unchecked("environment objects cannot be created during initialization", 59, 59); +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) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__2() { _start: { -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; +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_registerPersistentEnvExtensionUnsafe___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t x_1, lean_object* x_2) { _start: { -lean_object* x_5; -x_5 = l_Lean_registerPersistentEnvExtensionUnsafe(x_1, x_2, x_3, 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); -return 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: -{ -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) +if (x_5 == 0) { -lean_dec(x_4); +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); -lean_dec(x_1); -return x_2; +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_7; -x_7 = lean_nat_dec_le(x_4, x_4); -if (x_7 == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_3); +if (x_9 == 0) { -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -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 { -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_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_mkStateFromImportedEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +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_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; +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; } -else -{ -uint8_t x_8; -x_8 = lean_nat_dec_le(x_5, x_5); -if (x_8 == 0) -{ -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_2; } -else +LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -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; +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; } } +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(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); +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_mkStateFromImportedEntries(lean_object* x_1, lean_object* x_2) { +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_mkStateFromImportedEntries___rarg), 3, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtensionState___rarg), 1, 0); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3062_() { +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_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; +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_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_instInhabitedPersistentEnvExtension___lambda__2(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_inc(x_1); +return x_1; } } -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_instInhabitedPersistentEnvExtension___lambda__3(lean_object* x_1) { _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_2; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; 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_inc(x_11); -lean_inc(x_10); -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); -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_instInhabitedPersistentEnvExtension___lambda__4(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_box(0); +return x_2; } } -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("number of local entries: ", 25, 25); +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__2() { _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); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object* x_1) { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__3() { _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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed), 1, 0); +return x_1; } } -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__5() { _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_1; 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_registerSimplePersistentEnvExtension(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +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_5; +x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__5; +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: { lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg), 2, 0); +x_4 = l_Lean_instInhabitedPersistentEnvExtension___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +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_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _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_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_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(x_1); +x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__3(x_1); lean_dec(x_1); 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_instInhabitedPersistentEnvExtension___lambda__4___boxed(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_registerSimplePersistentEnvExtension(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; +lean_object* x_2; +x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__4(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(lean_object* x_1) { +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_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_5; +x_5 = l_Lean_instInhabitedPersistentEnvExtension(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { +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_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_instInhabited___rarg), 1, 0); -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; } } -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_Lean_PersistentEnvExtension_getModuleEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = 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_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed), 4, 0); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object* x_1, lean_object* x_2) { +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_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed), 3, 0); +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_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_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_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +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_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_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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); +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg), 3, 0); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -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_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_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(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_getState___rarg___boxed), 3, 0); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getState___rarg___boxed), 3, 0); +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_PersistentEnvExtension_getState___rarg___boxed(lean_object* 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); +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_SimplePersistentEnvExtension_setState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -7072,34 +6971,35 @@ return x_6; } } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +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; -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_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_SimplePersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(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_setState___rarg___boxed), 3, 0); -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 lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed(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_SimplePersistentEnvExtension_setState___rarg(x_1, x_2, x_3); +x_4 = l_Lean_PersistentEnvExtension_setState___rarg(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -7128,1807 +7028,1743 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_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_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_SimplePersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState(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; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___boxed), 3, 0); +return x_4; } } -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_PersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_SimplePersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); +x_4 = l_Lean_PersistentEnvExtension_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_3379_() { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_3055_(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; 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_3; +} +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; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1() { +} +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Name_quickLt___boxed), 2, 0); +x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -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) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__2() { _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; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } -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 +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__3() { +_start: { -lean_dec(x_10); -lean_dec(x_3); -return x_11; -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__4() { _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_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__5() { _start: { -lean_object* x_2; -x_2 = l_Lean_NameSet_empty; -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___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; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__3(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__6() { _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; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; } } -else +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__7() { +_start: { -lean_dec(x_5); -lean_dec(x_3); -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___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; } } -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__8() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__1), 2, 0); +x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__9() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__2___boxed), 1, 0); -return x_1; +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; } } -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__3() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__10() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__3), 1, 0); +x_1 = lean_mk_string_unchecked("exact", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__11() { _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; 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; } } -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___auto____x40_Lean_Environment___hyg_3100____closed__12() { _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); -return x_7; +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_mkTagDeclarationExtension___lambda__2___boxed(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_mkTagDeclarationExtension___lambda__2(x_1); -lean_dec(x_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; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__14() { _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_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__15() { _start: { lean_object* x_1; -x_1 = l_Lean_TagDeclarationExtension_instInhabited___closed__1; +x_1 = lean_mk_string_unchecked("declName", 8, 8); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__16() { _start: { -lean_object* x_3; -x_3 = lean_panic_fn(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_TagDeclarationExtension_tag___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__17() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +x_1 = lean_mk_string_unchecked("decl_name%", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__18() { _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_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; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__3() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__19() { _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); +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; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__4() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__20() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.TagDeclarationExtension.tag", 32, 32); -return x_1; +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; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__5() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__21() { _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_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_Lean_TagDeclarationExtension_tag(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__22() { _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; -x_5 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_3); -return x_5; +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_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; } +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__23() { +_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__22; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -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) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__24() { _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_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; +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 -{ -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_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 +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__25() { +_start: { -uint8_t x_19; -lean_dec(x_16); -lean_dec(x_3); -x_19 = 0; -return x_19; +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; } } -else +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__26() { +_start: { -uint8_t x_20; -lean_dec(x_8); -lean_dec(x_3); -x_20 = 0; -return x_20; -} +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: { -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_24; -lean_dec(x_22); -lean_dec(x_4); -x_24 = 0; -return x_24; -} -else -{ -x_3 = x_22; -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; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_isTagged___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100_() { _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_1; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__28; +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +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_4; -x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); -if (lean_obj_tag(x_4) == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 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_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); -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; 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) +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) { -uint8_t x_16; -lean_dec(x_13); -lean_dec(x_10); -x_16 = 0; -return x_16; +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_17; -x_17 = lean_nat_dec_le(x_14, x_13); -if (x_17 == 0) -{ -uint8_t x_18; -lean_dec(x_13); -lean_dec(x_10); -x_18 = 0; -return x_18; +uint8_t x_13; +x_13 = 1; +return x_13; +} } 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; -} -} +uint8_t x_14; +x_14 = 0; +return x_14; } } } -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_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1(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; +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_Lean_TagDeclarationExtension_isTagged___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _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); +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; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3542_() { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; +x_1 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2() { _start: { -if (lean_obj_tag(x_2) == 0) -{ +lean_object* x_1; +x_1 = l_Lean_persistentEnvExtensionsRef; return x_1; } -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_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_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2(lean_object* x_1) { +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_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: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -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_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1(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; 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_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_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_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_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: +else { -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_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_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object* x_1, lean_object* x_2) { -_start: +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); -lean_dec(x_2); -x_5 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_1, 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() { -_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: +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_3; lean_object* x_4; 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; -} +return x_13; } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg), 2, 0); -return x_2; -} +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_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_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_2); -lean_dec(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(x_1); +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); -return x_2; -} -} -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: +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_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); -return 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; 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; } -static lean_object* _init_l_Lean_MapDeclarationExtension_instInhabited___closed__1() { -_start: +else { -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_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_Lean_MapDeclarationExtension_instInhabited(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_MapDeclarationExtension_instInhabited___closed__1; -return x_2; +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +return x_54; } } -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_panic_fn(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__1() { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1() { _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("invalid environment extension, '", 32, 32); return x_1; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__2() { -_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; -} -} -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__3() { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.MapDeclarationExtension.insert", 35, 35); +x_1 = lean_mk_string_unchecked("' has already been used", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__4() { +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(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_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_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_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_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; 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_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_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 { -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; -} -} +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; } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_insert___rarg), 4, 0); -return x_2; +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_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_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) +} +else { -uint8_t x_13; +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_13 = l_Lean_Name_quickLt(x_11, x_10); -lean_dec(x_10); -if (x_13 == 0) +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_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_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 { -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) +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) { -x_4 = x_18; -x_5 = lean_box(0); -goto _start; +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 { -lean_object* x_21; -lean_dec(x_18); -lean_dec(x_3); -x_21 = lean_box(0); -return x_21; -} +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; } -else -{ -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_3); -x_22 = lean_box(0); -return x_22; } } } -else -{ -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_26; -lean_dec(x_24); -lean_dec(x_4); -x_26 = lean_box(0); -return x_26; } -else +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -x_3 = x_24; -x_5 = lean_box(0); -goto _start; +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_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_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object* 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) { _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_4; +x_4 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; } } -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 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_Environment_getModuleIdxFor_x3f(x_3, x_4); -if (lean_obj_tag(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); +x_5 = l_Lean_registerPersistentEnvExtensionUnsafe(x_1, x_2, x_3, x_4); lean_dec(x_4); -lean_dec(x_7); -return x_8; +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(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; 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_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_18; -lean_dec(x_15); -lean_dec(x_12); -lean_dec(x_11); -x_18 = lean_box(0); -return x_18; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_2; } else { -uint8_t x_19; -x_19 = lean_nat_dec_le(x_16, x_15); -if (x_19 == 0) +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 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; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_2; } 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) +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_Lean_mkStateFromImportedEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_22; -x_22 = lean_box(0); -return x_22; +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_23; -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) +uint8_t x_8; +x_8 = lean_nat_dec_le(x_5, x_5); +if (x_8 == 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; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_2; } 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; -} -} -} +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_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object* x_1, lean_object* x_2) { _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_3; +x_3 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg), 3, 0); +return x_3; } } -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) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3478_() { _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_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___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__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; -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_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 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) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2(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; 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_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) { -uint8_t x_14; -lean_dec(x_8); -lean_dec(x_3); -x_14 = 1; -return x_14; +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_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_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; +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_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; } -else -{ -uint8_t x_21; -lean_dec(x_18); -lean_dec(x_3); -x_21 = 0; -return x_21; } } -else +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_22; -lean_dec(x_8); -lean_dec(x_3); -x_22 = 0; -return x_22; -} -} +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; } -else -{ -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) -{ -uint8_t x_26; -lean_dec(x_24); -lean_dec(x_4); -x_26 = 0; -return x_26; } -else +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1() { +_start: { -x_3 = x_24; -x_5 = lean_box(0); -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("number of local entries: ", 25, 25); +return x_1; } } +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2() { +_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); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object* x_1) { _start: { -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_object* x_2; 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; } } -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) { +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1() { _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_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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed), 1, 0); +return x_1; } -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_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { +_start: { -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_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; 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_registerSimplePersistentEnvExtension(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); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object* x_1) { +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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_contains___rarg___boxed), 4, 0); -return x_2; +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_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) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed(lean_object* x_1) { _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_object* x_2; +x_2 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(x_1); lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +return x_2; } } -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) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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_object* x_4; +x_4 = l_Lean_registerSimplePersistentEnvExtension(x_1, x_2, x_3); lean_dec(x_3); -lean_dec(x_2); -x_6 = lean_box(x_5); -return x_6; +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(lean_object* x_1) { _start: { -lean_object* x_5; -x_5 = lean_save_module_data(x_1, x_2, x_3, x_4); +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); -lean_dec(x_1); -return x_5; +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_read_module_data(x_1, x_2); -lean_dec(x_1); +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_instInhabited___rarg), 1, 0); 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) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_2, x_3); -if (x_6 == 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_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_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; } -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) -{ -return x_9; } -else +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(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; +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_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_SimplePersistentEnvExtension_getEntries___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_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_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* l_Lean_SimplePersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed), 3, 0); +return x_3; } -LEAN_EXPORT lean_object* lean_environment_free_regions(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: { -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_object* x_4; +x_4 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_1, x_2, x_3); 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_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; +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: { -uint8_t x_10; -x_10 = lean_nat_dec_le(x_5, x_5); -if (x_10 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 1); 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_ctor_set(x_2, 1, x_1); +return x_2; } 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_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_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_SimplePersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_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_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: +{ +lean_object* x_4; +x_4 = l_Lean_SimplePersistentEnvExtension_setState___rarg(x_1, x_2, x_3); lean_dec(x_1); -return x_8; +return x_4; } } -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_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -return x_4; +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; 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_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_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_SimplePersistentEnvExtension_modifyState___rarg(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); -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_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; } -case 1: -{ -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: +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2) { +_start: { -x_3 = x_9; -goto _start; -} +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_4; +x_4 = l_Lean_SimplePersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); lean_dec(x_1); -return x_5; +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_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +static lean_object* _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1() { _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_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_quickLt___boxed), 2, 0); +return x_1; } } -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_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: { -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_7; +x_7 = lean_nat_dec_lt(x_3, x_4); +if (x_7 == 0) { -lean_dec(x_5); -lean_dec(x_1); -return x_6; +lean_dec(x_3); +return x_2; } 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; +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_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) { +} +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object* x_1, lean_object* x_2) { _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); +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_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object* x_1) { _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); +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_lt(x_6, x_5); +x_7 = lean_nat_dec_eq(x_3, x_6); if (x_7 == 0) { -lean_dec(x_5); -lean_dec(x_1); -return x_3; -} -else -{ uint8_t x_8; -x_8 = lean_nat_dec_le(x_5, x_5); +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_1); -return x_3; +lean_dec(x_3); +return x_9; } else { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = 0; -x_10 = lean_usize_of_nat(x_5); +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); -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; -} +lean_dec(x_3); +return x_10; } } 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_dec(x_5); +lean_dec(x_3); +return x_2; } } } -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) { +static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__1() { _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_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__1), 2, 0); +return x_1; } } -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_Lean_mkTagDeclarationExtension___closed__2() { _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 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__2___boxed), 1, 0); +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) { +static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__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_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__3), 1, 0); +return x_1; } } -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_mkTagDeclarationExtension(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; 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); -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_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_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object* 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) { _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_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_mkModuleData___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2___boxed(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = lean_array_push(x_1, x_2); -return x_4; +lean_object* x_2; +x_2 = l_Lean_mkTagDeclarationExtension___lambda__2(x_1); +lean_dec(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) { +static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited___closed__1() { _start: { -lean_object* x_4; -x_4 = lean_array_push(x_1, x_3); -return x_4; +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; } } -static lean_object* _init_l_Lean_mkModuleData___closed__1() { +static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__1___boxed), 3, 0); +x_1 = l_Lean_TagDeclarationExtension_instInhabited___closed__1; return x_1; } } -static lean_object* _init_l_Lean_mkModuleData___closed__2() { +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_alloc_closure((void*)(l_Lean_mkModuleData___lambda__2___boxed), 3, 0); +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData(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; 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, 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; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `TagDeclarationExtension`\n ", 86, 86); +return x_1; } -else -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; 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); -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; } +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__3() { +_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_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) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___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_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.TagDeclarationExtension.tag", 32, 32); +return x_1; } } -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) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__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_mkModuleData___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_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_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_TagDeclarationExtension_tag(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _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_4; +x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +x_5 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_3); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); +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_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +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: { -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_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_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_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_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; } -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: +else { -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; +uint8_t x_19; +lean_dec(x_16); +lean_dec(x_3); +x_19 = 0; +return x_19; } } -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: +else { -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; +uint8_t x_20; +lean_dec(x_8); +lean_dec(x_3); +x_20 = 0; +return x_20; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___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_mkModuleData___lambda__1(x_1, x_2, x_3); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_9); lean_dec(x_3); -return x_4; -} +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_24; +lean_dec(x_22); +lean_dec(x_4); +x_24 = 0; +return x_24; } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_mkModuleData___lambda__2(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +x_3 = x_22; +x_5 = lean_box(0); +goto _start; } } -LEAN_EXPORT lean_object* lean_write_module(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +static lean_object* _init_l_Lean_TagDeclarationExtension_isTagged___closed__1() { _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_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 uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(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_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); +if (lean_obj_tag(x_4) == 0) { -uint8_t x_3; -x_3 = 0; -return x_3; +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_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_8; 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) { -x_2 = x_5; -goto _start; +uint8_t x_16; +lean_dec(x_13); +lean_dec(x_10); +x_16 = 0; +return x_16; } else { -uint8_t x_8; -x_8 = 1; -return x_8; +uint8_t x_17; +x_17 = lean_nat_dec_le(x_14, x_13); +if (x_17 == 0) +{ +uint8_t x_18; +lean_dec(x_13); +lean_dec(x_10); +x_18 = 0; +return x_18; } +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_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_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: +{ +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_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___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_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_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8937,6830 +8773,7557 @@ 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; +lean_object* x_3; 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; } -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_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2(lean_object* x_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_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_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object* x_1) { _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); +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; } -else +} +LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1(lean_object* x_1) { +_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_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _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_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_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(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_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 = 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; +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_replace___at_Lean_mkExtNameMap___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__1() { _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; +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; } -else +} +static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__2() { +_start: { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +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_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_3; lean_object* x_4; 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_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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg), 2, 0); +return x_2; } -else +} +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_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); +lean_object* x_3; +x_3 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_2); +lean_dec(x_2); return x_3; } } -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_3); -x_14 = lean_name_eq(x_11, x_1); -if (x_14 == 0) +LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object* x_1) { +_start: { -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_2; +x_2 = l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(x_1); +lean_dec(x_1); +return x_2; } -else -{ -lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 2, x_13); -return x_17; } +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___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_mkMapDeclarationExtension___rarg___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; } } +static lean_object* _init_l_Lean_MapDeclarationExtension_instInhabited___closed__1() { +_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_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, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_instInhabited(lean_object* x_1) { _start: { -uint8_t x_12; -x_12 = lean_nat_dec_lt(x_8, x_5); -if (x_12 == 0) +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_13; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_10); -lean_ctor_set(x_13, 1, x_11); -return x_13; +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `MapDeclarationExtension`\n ", 86, 86); +return x_1; +} +} +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__2() { +_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; +} +} +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.MapDeclarationExtension.insert", 35, 35); +return x_1; +} +} +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___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_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_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; } else { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_7, x_14); -if (x_15 == 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_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_7, x_16); -lean_dec(x_7); -x_18 = lean_array_fget(x_1, x_8); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 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_21; lean_object* x_22; lean_object* 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; uint8_t x_37; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_ctor_get(x_10, 1); -x_23 = lean_array_get_size(x_22); -x_24 = l_Lean_Name_hash___override(x_19); -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_23); -lean_dec(x_23); -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_22, x_35); -x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_19, x_36); -if (x_37 == 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) { -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_38 = lean_nat_add(x_21, x_16); -lean_dec(x_21); -lean_inc(x_8); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_19); -lean_ctor_set(x_39, 1, x_8); -lean_ctor_set(x_39, 2, x_36); -x_40 = lean_array_uset(x_22, x_35, x_39); -x_41 = lean_unsigned_to_nat(4u); -x_42 = lean_nat_mul(x_38, x_41); -x_43 = lean_unsigned_to_nat(3u); -x_44 = lean_nat_div(x_42, x_43); -lean_dec(x_42); -x_45 = lean_array_get_size(x_40); -x_46 = lean_nat_dec_le(x_44, x_45); -lean_dec(x_45); -lean_dec(x_44); -if (x_46 == 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_47; lean_object* x_48; -x_47 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_40); -lean_ctor_set(x_10, 1, x_47); -lean_ctor_set(x_10, 0, x_38); -x_48 = lean_nat_add(x_8, x_6); +lean_object* x_14; lean_dec(x_8); -x_7 = x_17; -x_8 = x_48; -x_9 = lean_box(0); -goto _start; +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_50; -lean_ctor_set(x_10, 1, x_40); -lean_ctor_set(x_10, 0, x_38); -x_50 = lean_nat_add(x_8, x_6); +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_7 = x_17; -x_8 = x_50; -x_9 = lean_box(0); +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_21; +lean_dec(x_18); +lean_dec(x_3); +x_21 = lean_box(0); +return x_21; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_inc(x_2); -x_52 = lean_array_uset(x_22, x_35, x_2); -lean_inc(x_8); -x_53 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_19, x_8, x_36); -x_54 = lean_array_uset(x_52, x_35, x_53); -lean_ctor_set(x_10, 1, x_54); -x_55 = lean_nat_add(x_8, x_6); +lean_object* x_22; lean_dec(x_8); -x_7 = x_17; -x_8 = x_55; -x_9 = lean_box(0); -goto _start; +lean_dec(x_3); +x_22 = lean_box(0); +return x_22; +} } } else { -lean_object* x_57; lean_object* 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; uint64_t x_65; uint64_t x_66; size_t x_67; size_t x_68; size_t x_69; size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; -x_57 = lean_ctor_get(x_10, 0); -x_58 = lean_ctor_get(x_10, 1); -lean_inc(x_58); -lean_inc(x_57); +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_dec(x_10); -x_59 = lean_array_get_size(x_58); -x_60 = l_Lean_Name_hash___override(x_19); -x_61 = 32; -x_62 = lean_uint64_shift_right(x_60, x_61); -x_63 = lean_uint64_xor(x_60, x_62); -x_64 = 16; -x_65 = lean_uint64_shift_right(x_63, x_64); -x_66 = lean_uint64_xor(x_63, x_65); -x_67 = lean_uint64_to_usize(x_66); -x_68 = lean_usize_of_nat(x_59); -lean_dec(x_59); -x_69 = 1; -x_70 = lean_usize_sub(x_68, x_69); -x_71 = lean_usize_land(x_67, x_70); -x_72 = lean_array_uget(x_58, x_71); -x_73 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_19, x_72); -if (x_73 == 0) +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_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_74 = lean_nat_add(x_57, x_16); -lean_dec(x_57); -lean_inc(x_8); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_19); -lean_ctor_set(x_75, 1, x_8); -lean_ctor_set(x_75, 2, x_72); -x_76 = lean_array_uset(x_58, x_71, 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; -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_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_85; -x_9 = lean_box(0); -x_10 = x_84; -goto _start; +lean_object* x_26; +lean_dec(x_24); +lean_dec(x_4); +x_26 = lean_box(0); +return x_26; } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_74); -lean_ctor_set(x_87, 1, x_76); -x_88 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_88; -x_9 = lean_box(0); -x_10 = x_87; +x_3 = x_24; +x_5 = lean_box(0); goto _start; } } -else -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_inc(x_2); -x_90 = lean_array_uset(x_58, x_71, x_2); -lean_inc(x_8); -x_91 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_19, x_8, x_72); -x_92 = lean_array_uset(x_90, x_71, x_91); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_57); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_94; -x_9 = lean_box(0); -x_10 = x_93; -goto _start; -} } } -else +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object* x_1) { +_start: { -lean_object* x_96; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_2); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_10); -lean_ctor_set(x_96, 1, x_11); -return x_96; -} -} +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_EXPORT lean_object* l_Lean_mkExtNameMap(lean_object* x_1, lean_object* x_2) { +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: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; 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_8); -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; -lean_inc(x_8); -lean_inc(x_1); -x_12 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_5, x_7, x_10, x_1, x_8, x_9, x_8, x_1, lean_box(0), x_11, x_6); -lean_dec(x_8); +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; lean_object* x_8; lean_dec(x_1); -lean_dec(x_10); +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_9; lean_object* x_10; 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_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +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) { -return x_12; +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 { -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); +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); -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_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; } +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; } -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: +else { -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_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; } } -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, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = 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, x_10, x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_12; } } -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) { +} +} +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object* x_1) { _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); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed), 4, 0); return x_2; } -else +} +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_8; lean_object* x_9; -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); +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); -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_dec(x_1); +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_MapDeclarationExtension_find_x3f___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_4, x_3); -if (x_7 == 0) -{ -lean_object* x_8; -lean_dec(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; 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_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 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_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: { -if (lean_obj_tag(x_2) == 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) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +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; } 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_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) { -x_2 = x_6; +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_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_21; +lean_dec(x_18); +lean_dec(x_3); +x_21 = 0; +return x_21; } } +else +{ +uint8_t x_22; +lean_dec(x_8); +lean_dec(x_3); +x_22 = 0; +return x_22; +} } } -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_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_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_26; +lean_dec(x_24); +lean_dec(x_4); +x_26 = 0; +return x_26; } 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_24; +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() { +} +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object* x_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); +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 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_MapDeclarationExtension_contains___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_8, x_7); -if (x_11 == 0) +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_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_6; lean_object* x_7; uint8_t 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_NameMap_contains___rarg(x_7, x_4); +lean_dec(x_4); +lean_dec(x_7); +return x_8; } 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_object* x_9; lean_object* x_10; 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); -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) +if (x_17 == 0) { -size_t x_32; +uint8_t x_18; lean_dec(x_15); -x_32 = lean_usize_add(x_8, x_27); -x_8 = x_32; -goto _start; +lean_dec(x_12); +lean_dec(x_11); +x_18 = 0; +return x_18; } 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_19; +x_19 = lean_nat_dec_le(x_16, x_15); +if (x_19 == 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_20; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +x_20 = 0; +return x_20; } 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_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_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_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object* x_1) { _start: { -uint8_t x_13; -x_13 = lean_nat_dec_lt(x_9, x_6); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_9); -lean_dec(x_8); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_11); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else -{ -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_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; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_8, x_17); -lean_dec(x_8); -x_19 = lean_array_fget(x_1, x_9); -x_20 = lean_ctor_get(x_19, 4); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_box(0); -x_22 = lean_array_size(x_20); -x_23 = 0; -lean_inc(x_9); -x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_2, x_3, x_9, x_20, x_21, x_20, x_22, x_23, x_11, x_12); -lean_dec(x_20); -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_nat_add(x_9, x_7); -lean_dec(x_9); -x_8 = x_18; -x_9 = x_27; -x_10 = lean_box(0); -x_11 = x_25; -x_12 = x_26; -goto _start; -} -else -{ -lean_object* x_29; -lean_dec(x_9); -lean_dec(x_8); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_11); -lean_ctor_set(x_29, 1, x_12); -return x_29; -} -} +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___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_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: { -lean_object* x_5; 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); -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_inc(x_21); -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_22, x_21, x_23, x_21, x_22, lean_box(0), x_16, x_20); -lean_dec(x_21); -lean_dec(x_24); -lean_dec(x_19); -lean_dec(x_7); +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); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) -{ -return x_25; -} -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_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_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; +x_7 = lean_box(x_6); +return x_7; } } -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_MapDeclarationExtension_contains___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_3); +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); -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; +x_6 = lean_box(x_5); +return x_6; } } -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) { +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_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(x_1, x_2); +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_3; +return x_5; } } -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_readModuleData___boxed(lean_object* x_1, lean_object* x_2) { _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_3; +x_3 = lean_read_module_data(x_1, x_2); lean_dec(x_1); -return x_4; +return x_3; } } -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_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: { -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, lean_object* x_11, lean_object* x_12) { -_start: +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) { -lean_object* x_13; -x_13 = 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, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_7; size_t x_8; lean_object* x_9; lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object* x_1, lean_object* x_2) { -_start: +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_3; -x_3 = lean_update_env_attributes(x_1, x_2); -return x_3; -} +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_EXPORT lean_object* l_Lean_getNumBuiltinAttributes___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_get_num_attributes(x_1); -return x_2; -} +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; } -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: +else { -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_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; } } -static lean_object* _init_l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1() { -_start: +} +else { -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_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* 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) { +} +LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object* x_1, lean_object* x_2) { _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_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_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); -lean_ctor_set(x_7, 0, x_4); -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_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_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); +x_4 = lean_ctor_get(x_3, 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_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_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_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_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) +uint8_t x_10; +x_10 = lean_nat_dec_le(x_5, x_5); +if (x_10 == 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_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; } else { -uint8_t x_55; +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_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: +{ +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_55 = !lean_is_exclusive(x_50); -if (x_55 == 0) +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) { +_start: { -return x_50; +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_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_6; lean_object* x_7; lean_object* x_8; 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; } } } -else +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: { -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) +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 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_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: { -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_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: { -return x_62; +x_3 = x_9; +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_dec(x_1); +return x_5; } } } +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: { -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_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) { -return x_39; +lean_dec(x_5); +lean_dec(x_1); +return x_6; } 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; +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; } } } -else +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: { -uint8_t x_75; -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_2); +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_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); -x_75 = !lean_is_exclusive(x_33); -if (x_75 == 0) +return x_3; +} +else { -return x_33; +uint8_t x_8; +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; } 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; +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 { -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) +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: { -return x_29; +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_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_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_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_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_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); -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) +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object* x_1) { +_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_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_EXPORT lean_object* l_Lean_mkModuleData___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -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_4; +x_4 = lean_array_push(x_1, x_2); +return x_4; +} +} +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_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_4; +x_4 = lean_array_push(x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_mkModuleData___closed__1() { +_start: { -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) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__1___boxed), 3, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_mkModuleData___closed__2() { +_start: { -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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__2___boxed), 3, 0); +return x_1; } -else +} +LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object* x_1, lean_object* x_2) { +_start: { -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_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_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_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_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_2); +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); +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_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); +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; } -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; } +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: +{ +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; } } -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_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_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; +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_mkModuleData___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } -else +} +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_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +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); -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; -} +return x_7; } } -else +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_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_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); -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); +return x_4; } -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_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_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; } } -else +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_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_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(x_1, x_2, x_3); 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); +return x_4; } -if (lean_is_scalar(x_133)) { - x_134 = lean_alloc_ctor(1, 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; +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_4; +x_4 = l_Lean_mkModuleData___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_object* x_4; +x_4 = l_Lean_mkModuleData___lambda__2(x_1, x_2, x_3); 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); +return x_4; } -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* lean_write_module(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_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 { -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) +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) { -return x_22; +x_2 = x_5; +goto _start; } else { -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_8; +x_8 = 1; +return x_8; } } } -else +} +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_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) +if (lean_obj_tag(x_2) == 0) { -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); +return x_1; } -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) +else { -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) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 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; -} -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) -{ -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; +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_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +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); -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); -} -if (lean_is_scalar(x_182)) { - x_183 = lean_alloc_ctor(1, 2, 0); -} else { - x_183 = x_182; +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_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_181); -return x_183; } } } -else +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_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_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_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_dec(x_2); +lean_dec(x_1); +return x_3; } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_2); +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_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); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } -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_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___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_mkExtNameMap___spec__3(x_7, x_1, x_6); +return x_8; } } -else +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_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); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_dec(x_2); 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; -} -lean_ctor_set(x_199, 0, x_196); -lean_ctor_set(x_199, 1, x_197); -return x_199; +x_4 = lean_box(0); +return x_4; } +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_name_eq(x_6, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; } 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); -} -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_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_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; -} +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +x_14 = lean_name_eq(x_11, x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_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_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); +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; } -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; } } } +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, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_lt(x_8, x_5); +if (x_12 == 0) +{ +lean_object* x_13; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_10); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -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_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_7, 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; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_7, x_16); 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_18 = lean_array_fget(x_1, x_8); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) { -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; +lean_object* x_21; lean_object* x_22; lean_object* 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; uint8_t x_37; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_ctor_get(x_10, 1); +x_23 = lean_array_get_size(x_22); +x_24 = l_Lean_Name_hash___override(x_19); +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_23); +lean_dec(x_23); +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_22, x_35); +x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_19, 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; uint8_t x_46; +x_38 = lean_nat_add(x_21, x_16); +lean_dec(x_21); +lean_inc(x_8); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_19); +lean_ctor_set(x_39, 1, x_8); +lean_ctor_set(x_39, 2, x_36); +x_40 = lean_array_uset(x_22, x_35, x_39); +x_41 = lean_unsigned_to_nat(4u); +x_42 = lean_nat_mul(x_38, x_41); +x_43 = lean_unsigned_to_nat(3u); +x_44 = lean_nat_div(x_42, x_43); +lean_dec(x_42); +x_45 = lean_array_get_size(x_40); +x_46 = lean_nat_dec_le(x_44, x_45); +lean_dec(x_45); +lean_dec(x_44); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; +x_47 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_40); +lean_ctor_set(x_10, 1, x_47); +lean_ctor_set(x_10, 0, x_38); +x_48 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_48; +x_9 = lean_box(0); +goto _start; } 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_50; +lean_ctor_set(x_10, 1, x_40); +lean_ctor_set(x_10, 0, x_38); +x_50 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_50; +x_9 = lean_box(0); +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) +} +else { -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_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; 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; +x_52 = lean_array_uset(x_22, x_35, x_2); +lean_inc(x_8); +x_53 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_19, x_8, x_36); +x_54 = lean_array_uset(x_52, x_35, x_53); +lean_ctor_set(x_10, 1, x_54); +x_55 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_55; +x_9 = lean_box(0); +goto _start; } -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_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_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) +else { -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) +lean_object* x_57; lean_object* 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; uint64_t x_65; uint64_t x_66; size_t x_67; size_t x_68; size_t x_69; size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; +x_57 = lean_ctor_get(x_10, 0); +x_58 = lean_ctor_get(x_10, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_10); +x_59 = lean_array_get_size(x_58); +x_60 = l_Lean_Name_hash___override(x_19); +x_61 = 32; +x_62 = lean_uint64_shift_right(x_60, x_61); +x_63 = lean_uint64_xor(x_60, x_62); +x_64 = 16; +x_65 = lean_uint64_shift_right(x_63, x_64); +x_66 = lean_uint64_xor(x_63, x_65); +x_67 = lean_uint64_to_usize(x_66); +x_68 = lean_usize_of_nat(x_59); +lean_dec(x_59); +x_69 = 1; +x_70 = lean_usize_sub(x_68, x_69); +x_71 = lean_usize_land(x_67, x_70); +x_72 = lean_array_uget(x_58, x_71); +x_73 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_19, x_72); +if (x_73 == 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_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_74 = lean_nat_add(x_57, x_16); +lean_dec(x_57); +lean_inc(x_8); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_19); +lean_ctor_set(x_75, 1, x_8); +lean_ctor_set(x_75, 2, x_72); +x_76 = lean_array_uset(x_58, x_71, 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_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_object* x_83; lean_object* x_84; lean_object* x_85; +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_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_85; +x_9 = lean_box(0); +x_10 = x_84; +goto _start; } 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) +lean_object* x_87; lean_object* x_88; +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_74); +lean_ctor_set(x_87, 1, x_76); +x_88 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_88; +x_9 = lean_box(0); +x_10 = x_87; +goto _start; +} +} +else { -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_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_inc(x_2); +x_90 = lean_array_uset(x_58, x_71, x_2); +lean_inc(x_8); +x_91 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_19, x_8, x_72); +x_92 = lean_array_uset(x_90, x_71, x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_57); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_94; +x_9 = lean_box(0); +x_10 = x_93; +goto _start; +} +} } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_object* x_96; +lean_dec(x_8); +lean_dec(x_7); 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; +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_10); +lean_ctor_set(x_96, 1, x_11); +return x_96; } -lean_ctor_set(x_263, 0, x_260); -lean_ctor_set(x_263, 1, x_261); -return x_263; } } } -else +LEAN_EXPORT lean_object* l_Lean_mkExtNameMap(lean_object* x_1, lean_object* x_2) { +_start: { -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_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; 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_8); 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) +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; +lean_inc(x_8); +lean_inc(x_1); +x_12 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_5, x_7, x_10, x_1, x_8, x_9, x_8, x_1, lean_box(0), x_11, x_6); +lean_dec(x_8); +lean_dec(x_1); +lean_dec(x_10); +lean_dec(x_5); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -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; +return x_12; } else { -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); -} -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_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; } } } -else +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_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); +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_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; +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, lean_object* x_10, lean_object* x_11) { +_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_object* x_12; +x_12 = 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, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); 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); -} -if (lean_is_scalar(x_282)) { - x_283 = lean_alloc_ctor(1, 2, 0); -} else { - x_283 = x_282; +return x_12; } -lean_ctor_set(x_283, 0, x_280); -lean_ctor_set(x_283, 1, x_281); -return x_283; } +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_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_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_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); 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); -} -if (lean_is_scalar(x_286)) { - x_287 = lean_alloc_ctor(1, 2, 0); -} else { - x_287 = x_286; +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_287, 0, x_284); -lean_ctor_set(x_287, 1, x_285); -return x_287; } } -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: { -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); +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_4, x_3); +if (x_7 == 0) +{ +lean_object* x_8; lean_dec(x_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_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +return x_8; } +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; 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___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_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(lean_object* x_1, lean_object* x_2) { _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); -lean_dec(x_1); -return x_7; +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; } -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: +else { -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; +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___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_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: { +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_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; -} +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_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); -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_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2() { -_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() { +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("key is not present in hash table", 32, 32); -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; } } -static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__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_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; -} +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_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: +else { -if (lean_obj_tag(x_3) == 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); +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_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; +size_t x_32; +lean_dec(x_15); +x_32 = lean_usize_add(x_8, x_27); +x_8 = x_32; +goto _start; } 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_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) { -x_3 = x_8; +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; } else { -lean_dec(x_1); -lean_inc(x_7); -return x_7; -} +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; } } } -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("import ", 7, 7); -return x_1; } } -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__2() { +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_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" failed, environment already contains '", 39, 39); -return x_1; -} -} -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__3() { -_start: +uint8_t x_13; +x_13 = lean_nat_dec_lt(x_9, x_6); +if (x_13 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' from ", 7, 7); -return x_1; -} +lean_object* x_14; +lean_dec(x_9); +lean_dec(x_8); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_12); +return x_14; } -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: +else { -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_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_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_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; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_8, x_17); +lean_dec(x_8); +x_19 = lean_array_fget(x_1, x_9); +x_20 = lean_ctor_get(x_19, 4); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_box(0); +x_22 = lean_array_size(x_20); +x_23 = 0; +lean_inc(x_9); +x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_2, x_3, x_9, x_20, x_21, x_20, x_22, x_23, x_11, x_12); +lean_dec(x_20); +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_nat_add(x_9, x_7); +lean_dec(x_9); +x_8 = x_18; +x_9 = x_27; +x_10 = lean_box(0); +x_11 = x_25; +x_12 = x_26; +goto _start; } else { -lean_object* x_90; -x_90 = lean_array_fget(x_6, x_3); -x_12 = x_90; -goto block_87; +lean_object* x_29; +lean_dec(x_9); +lean_dec(x_8); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_11); +lean_ctor_set(x_29, 1, x_12); +return x_29; } -block_87: +} +} +} +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_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_object* x_5; 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); +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_inc(x_21); +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_22, x_21, x_23, x_21, x_22, lean_box(0), x_16, x_20); +lean_dec(x_21); 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_19); lean_dec(x_7); -if (x_40 == 0) +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 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; +return x_25; } 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_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; } } -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_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_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_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_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; } } +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: +{ +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; } } -LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object* x_1) { +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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_throwAlreadyImported___rarg___boxed), 5, 0); -return x_2; +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_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* 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_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); +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_Lean_throwAlreadyImported___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_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: { -lean_object* x_6; -x_6 = l_Lean_throwAlreadyImported___rarg(x_1, x_2, x_3, x_4, x_5); +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, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = 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, 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_2); lean_dec(x_1); -return x_6; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_ImportStateM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object* x_1, lean_object* x_2) { _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_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_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_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; -} -} -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; -} -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; -} -} -} -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_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); -} -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_object* x_2; +x_2 = lean_get_num_attributes(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object* 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) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ImportStateM_run___rarg), 3, 0); -return x_2; +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_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___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); +x_1 = l_Lean_instInhabitedEnvExtensionState; +x_2 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(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) { +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: { -lean_object* x_6; -x_6 = lean_read_module_data(x_1, x_5); -if (lean_obj_tag(x_6) == 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_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); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; 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); +x_11 = lean_array_get_size(x_9); +x_12 = lean_nat_dec_lt(x_3, x_11); 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) +if (x_12 == 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; +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_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_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_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_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_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); -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; +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_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_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_44; -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_55; lean_dec(x_2); -x_44 = !lean_is_exclusive(x_12); -if (x_44 == 0) +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_50); +if (x_55 == 0) { -return x_12; +return x_50; } 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_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 { -uint8_t x_48; +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 +{ +uint8_t x_67; lean_dec(x_2); -x_48 = !lean_is_exclusive(x_6); -if (x_48 == 0) +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_62); +if (x_67 == 0) { -return x_6; +return x_62; } 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_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; } } } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1() { -_start: +else { -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: +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_1; -x_1 = lean_mk_string_unchecked("' of module ", 12, 12); -return x_1; -} +return x_39; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" does not exist", 15, 15); -return x_1; +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__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +} +else { -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); +uint8_t x_75; +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_2); 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) +x_75 = !lean_is_exclusive(x_33); +if (x_75 == 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) +return x_33; +} +else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +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; +} +} +} +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; 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); +uint8_t x_79; +lean_free_object(x_16); +lean_dec(x_27); +lean_dec(x_23); 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; +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) +{ +return x_29; } 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; +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; +} } } 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_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_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_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_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) +{ +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) +{ +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 { -uint8_t x_49; -lean_dec(x_10); -x_49 = !lean_is_exclusive(x_14); -if (x_49 == 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) { -return x_14; +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_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; +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 { -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_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_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_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 { -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; +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_65); - x_69 = lean_box(0); + lean_dec_ref(x_118); + x_125 = 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); +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); } else { - x_81 = x_69; - lean_ctor_set_tag(x_81, 1); + x_126 = x_125; } -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_68); -return x_81; +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; } -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; +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_62); - x_87 = lean_box(0); + lean_dec_ref(x_95); + x_129 = lean_box(0); } -if (lean_is_scalar(x_87)) { - x_88 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_129)) { + x_130 = 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; -} -} + x_130 = x_129; } +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +return x_130; } -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_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; 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_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); } -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; +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(1, 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_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_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); -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_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; +} } } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_20); -if (x_33 == 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_20; +return x_22; } 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_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 { -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_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_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_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_EXPORT lean_object* l_Lean_importModulesCore(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +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_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) +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) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 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) { -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 +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_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_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; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 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) { -return x_8; +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_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_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_dec(x_2); lean_dec(x_1); -return x_11; +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_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_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; } } -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) +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) { -uint8_t x_14; -x_14 = 0; -return x_14; +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_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_434____spec__1(x_15, x_16); -if (x_17 == 0) -{ -uint8_t x_18; -x_18 = 0; -return x_18; +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); } -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_434____spec__1(x_19, x_20); -return x_21; +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 { -uint8_t x_22; -x_22 = 0; -return x_22; -} +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_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 -{ -uint8_t x_23; -x_23 = 0; -return x_23; +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(1, 2, 0); +} else { + x_199 = x_198; } +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +return x_199; } } -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_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_4 = lean_box(x_3); -return x_4; +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_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_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; } } -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_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; } -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_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; } +lean_ctor_set(x_211, 0, x_208); +lean_ctor_set(x_211, 1, x_209); +return x_211; } } -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_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) { +lean_object* x_216; +lean_dec(x_212); lean_dec(x_2); lean_dec(x_1); -return x_3; +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_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_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; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object* x_1) { -_start: +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_2; 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_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_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: +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) { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_4); -if (x_8 == 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_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_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) { -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) +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_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_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_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_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_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_24); +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_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_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; } } -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; -} +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_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_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_64); +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_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_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; } } -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_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); -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_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); } -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; +if (lean_is_scalar(x_278)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_278; } -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; +lean_ctor_set(x_279, 0, x_276); +lean_ctor_set(x_279, 1, x_277); +return x_279; } } else { -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; +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); +} +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; +} } 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_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); +} +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_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_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_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; +} +} +} +} +} +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: { -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_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_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_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_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_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_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; } } -else +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1() { +_start: { -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_105); - x_120 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); +return x_1; } -if (lean_is_scalar(x_120)) { - x_121 = lean_alloc_ctor(1, 1, 0); -} else { - x_121 = x_120; } -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; +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.DHashMap.Internal.AssocList.get!", 36, 36); +return x_1; } } -block_56: +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__3() { +_start: { -if (lean_obj_tag(x_31) == 0) +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_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_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; } -else +} +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: { -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) +if (lean_obj_tag(x_3) == 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) +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 { -return x_43; +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) +{ +x_3 = x_8; +goto _start; } 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; +lean_dec(x_1); +lean_inc(x_7); +return x_7; } } -else +} +} +static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__1() { +_start: { -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_1; +x_1 = lean_mk_string_unchecked("import ", 7, 7); +return x_1; } } +static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__2() { +_start: +{ +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_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_1; +x_1 = lean_mk_string_unchecked("' from ", 7, 7); +return x_1; +} +} +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: +{ +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_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_90; +x_90 = lean_array_fget(x_6, x_3); +x_12 = x_90; +goto block_87; +} +block_87: +{ +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; +} +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; +} +} +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_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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object* x_1) { +_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_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); +return x_4; +} +} +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_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; +} +} +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; +} +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; +} +} +} +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_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); +} +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 { - lean_dec_ref(x_18); - x_155 = lean_box(0); + 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; +} } -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_EXPORT lean_object* l_Lean_importModulesCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -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_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: { -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); -} else { - x_182 = x_155; +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; } -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_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_183; -if (lean_is_scalar(x_155)) { - x_183 = lean_alloc_ctor(0, 2, 0); -} else { - x_183 = x_155; +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_434____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_434____spec__1(x_19, x_20); +return x_21; +} } -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 { -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); +uint8_t x_22; +x_22 = 0; +return x_22; } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 1, 0); -} else { - x_186 = x_185; } -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; +else +{ +uint8_t x_23; +x_23 = 0; +return x_23; } -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_equivInfo___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -if (lean_obj_tag(x_127) == 0) +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_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; +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_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_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_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_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; +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; } } } } -} -} -else +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_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_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_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; +lean_dec(x_1); +return x_3; } 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_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; } -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_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object* x_1) { +_start: { -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) +lean_object* x_2; 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_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_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; +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_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_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_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_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; } -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: +else { -if (lean_obj_tag(x_202) == 0) +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_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_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_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_214); - x_217 = lean_box(0); -} -if (lean_is_scalar(x_217)) { - x_218 = lean_alloc_ctor(1, 2, 0); -} else { - x_218 = x_217; +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_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_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_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); -} -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_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_272; lean_object* x_273; lean_object* x_274; -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); +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); -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; +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 { -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); +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; } -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; +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; } -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); } -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) +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; +} +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); -} else { - x_335 = x_308; +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_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_336; -if (lean_is_scalar(x_308)) { - x_336 = lean_alloc_ctor(0, 2, 0); -} else { - x_336 = x_308; +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_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; } -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 +{ +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_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; +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_323); - x_338 = lean_box(0); + lean_dec_ref(x_105); + x_120 = lean_box(0); } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 1, 0); } else { - x_339 = x_338; + x_121 = x_120; } -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_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_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: +block_56: { -if (lean_obj_tag(x_280) == 0) +if (lean_obj_tag(x_31) == 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_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_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; +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_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_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_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_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_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); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +return x_43; } -if (lean_is_scalar(x_295)) { - x_296 = lean_alloc_ctor(1, 2, 0); -} else { - x_296 = x_295; +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; } -lean_ctor_set(x_296, 0, x_293); -lean_ctor_set(x_296, 1, x_294); -return x_296; } else { -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_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_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_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_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_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) +if (lean_obj_tag(x_127) == 0) { -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; +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, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_nat_dec_lt(x_8, x_5); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_8); -lean_dec(x_7); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_10); -lean_ctor_set(x_13, 1, x_11); -return x_13; } else { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_7, x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_7, x_16); -lean_dec(x_7); -x_18 = !lean_is_exclusive(x_10); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; -x_19 = lean_array_fget(x_2, x_8); -x_20 = lean_ctor_get(x_19, 2); -lean_inc(x_20); -x_21 = lean_array_get_size(x_20); -x_22 = l_Array_toSubarray___rarg(x_20, x_14, x_21); -x_23 = lean_ctor_get(x_19, 1); -lean_inc(x_23); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_10); -x_26 = lean_array_size(x_23); -x_27 = 0; -lean_inc(x_8); -x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_8, x_23, x_24, x_23, x_26, x_27, x_25, x_11); -lean_dec(x_23); -if (lean_obj_tag(x_28) == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 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; 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_33; lean_object* x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_33 = lean_ctor_get(x_30, 0); -x_34 = lean_ctor_get(x_19, 3); -lean_inc(x_34); -lean_dec(x_19); -x_35 = lean_array_size(x_34); -lean_inc(x_8); -x_36 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_24, x_34, x_34, x_35, x_27, x_33, x_31); -lean_dec(x_34); -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); -lean_ctor_set(x_30, 0, x_37); -x_39 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_39; -x_9 = lean_box(0); -x_10 = x_30; -x_11 = x_38; -goto _start; +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_41; lean_object* x_42; lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_41 = lean_ctor_get(x_30, 0); -x_42 = lean_ctor_get(x_30, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_30); -x_43 = lean_ctor_get(x_19, 3); -lean_inc(x_43); -lean_dec(x_19); -x_44 = lean_array_size(x_43); -lean_inc(x_8); -x_45 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_24, x_43, x_43, x_44, x_27, x_41, x_31); -lean_dec(x_43); -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_42); -x_49 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_49; -x_9 = lean_box(0); -x_10 = x_48; -x_11 = x_47; -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; } -else +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_Kernel_Environment_find_x3f___spec__5(x_12, x_244); +if (lean_obj_tag(x_245) == 0) { -uint8_t x_51; -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_8); -x_51 = !lean_is_exclusive(x_28); -if (x_51 == 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_28; +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_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_28, 0); -x_53 = lean_ctor_get(x_28, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_28); -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_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_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; -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); -x_57 = lean_array_fget(x_2, x_8); -x_58 = lean_ctor_get(x_57, 2); -lean_inc(x_58); -x_59 = lean_array_get_size(x_58); -x_60 = l_Array_toSubarray___rarg(x_58, x_14, x_59); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_55); -lean_ctor_set(x_63, 1, x_56); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_60); -lean_ctor_set(x_64, 1, x_63); -x_65 = lean_array_size(x_61); -x_66 = 0; -lean_inc(x_8); -x_67 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_8, x_61, x_62, x_61, x_65, x_66, x_64, x_11); -lean_dec(x_61); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; size_t 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_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = lean_ctor_get(x_67, 1); -lean_inc(x_70); -lean_dec(x_67); -x_71 = lean_ctor_get(x_69, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_69, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_73 = x_69; +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_69); - x_73 = lean_box(0); + lean_dec_ref(x_245); + x_260 = lean_box(0); } -x_74 = lean_ctor_get(x_57, 3); -lean_inc(x_74); -lean_dec(x_57); -x_75 = lean_array_size(x_74); -lean_inc(x_8); -x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_62, x_74, x_74, x_75, x_66, x_71, x_70); -lean_dec(x_74); -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); -if (lean_is_scalar(x_73)) { - x_79 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_260)) { + x_261 = lean_alloc_ctor(1, 1, 0); } else { - x_79 = x_73; + x_261 = x_260; } -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_72); -x_80 = lean_nat_add(x_8, x_6); -lean_dec(x_8); -x_7 = x_17; -x_8 = x_80; -x_9 = lean_box(0); -x_10 = x_79; -x_11 = x_78; +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_57); -lean_dec(x_17); -lean_dec(x_8); -x_82 = lean_ctor_get(x_67, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_67, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_84 = x_67; +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_67); - 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_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; +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; } } -else -{ -lean_object* x_86; -lean_dec(x_8); -lean_dec(x_7); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_10); -lean_ctor_set(x_86, 1, x_11); -return x_86; } } } } -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) { -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) { -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); -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); -lean_inc(x_8); -x_23 = l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_7, x_13, x_9, x_8, x_12, x_8, x_9, lean_box(0), x_22, x_6); -lean_dec(x_8); -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) { -return x_23; -} -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; 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_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_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_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1___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_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, lean_object* x_10, lean_object* x_11) { -_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_12; -x_12 = 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, x_10, x_11); -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_12; -} +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, lean_object* x_10, lean_object* x_11) { _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: +uint8_t x_12; +x_12 = lean_nat_dec_lt(x_8, x_5); +if (x_12 == 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_13; +lean_dec(x_8); +lean_dec(x_7); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_10); +lean_ctor_set(x_13, 1, x_11); +return x_13; } -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) +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_7, x_14); +if (x_15 == 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_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_7, x_16); +lean_dec(x_7); +x_18 = !lean_is_exclusive(x_10); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_19 = lean_array_fget(x_2, x_8); +x_20 = lean_ctor_get(x_19, 2); +lean_inc(x_20); +x_21 = lean_array_get_size(x_20); +x_22 = l_Array_toSubarray___rarg(x_20, x_14, x_21); +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_10); +x_26 = lean_array_size(x_23); +x_27 = 0; +lean_inc(x_8); +x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_8, x_23, x_24, x_23, x_26, x_27, x_25, x_11); +lean_dec(x_23); +if (lean_obj_tag(x_28) == 0) { -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_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 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; +lean_object* x_33; lean_object* x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_33 = lean_ctor_get(x_30, 0); +x_34 = lean_ctor_get(x_19, 3); +lean_inc(x_34); +lean_dec(x_19); +x_35 = lean_array_size(x_34); +lean_inc(x_8); +x_36 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_24, x_34, x_34, x_35, x_27, x_33, x_31); +lean_dec(x_34); +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); +lean_ctor_set(x_30, 0, x_37); +x_39 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_39; +x_9 = lean_box(0); +x_10 = x_30; +x_11 = x_38; +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_41; lean_object* x_42; lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_41 = lean_ctor_get(x_30, 0); +x_42 = lean_ctor_get(x_30, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_30); +x_43 = lean_ctor_get(x_19, 3); +lean_inc(x_43); +lean_dec(x_19); +x_44 = lean_array_size(x_43); +lean_inc(x_8); +x_45 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_24, x_43, x_43, x_44, x_27, x_41, x_31); +lean_dec(x_43); +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_42); +x_49 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_49; +x_9 = lean_box(0); +x_10 = x_48; +x_11 = x_47; 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_51; +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_8); +x_51 = !lean_is_exclusive(x_28); +if (x_51 == 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_28; } -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_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_28, 0); +x_53 = lean_ctor_get(x_28, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_28); +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_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_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; +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); +x_57 = lean_array_fget(x_2, x_8); +x_58 = lean_ctor_get(x_57, 2); +lean_inc(x_58); +x_59 = lean_array_get_size(x_58); +x_60 = l_Array_toSubarray___rarg(x_58, x_14, x_59); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +x_62 = lean_box(0); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_55); +lean_ctor_set(x_63, 1, x_56); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_60); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_array_size(x_61); +x_66 = 0; +lean_inc(x_8); +x_67 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_8, x_61, x_62, x_61, x_65, x_66, x_64, x_11); +lean_dec(x_61); +if (lean_obj_tag(x_67) == 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_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; size_t 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_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = lean_ctor_get(x_67, 1); +lean_inc(x_70); +lean_dec(x_67); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_69, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_73 = x_69; +} else { + lean_dec_ref(x_69); + x_73 = lean_box(0); +} +x_74 = lean_ctor_get(x_57, 3); +lean_inc(x_74); +lean_dec(x_57); +x_75 = lean_array_size(x_74); +lean_inc(x_8); +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_8, x_62, x_74, x_74, x_75, x_66, x_71, x_70); +lean_dec(x_74); +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); +if (lean_is_scalar(x_73)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_73; } +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_72); +x_80 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_17; +x_8 = x_80; +x_9 = lean_box(0); +x_10 = x_79; +x_11 = x_78; +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_57); +lean_dec(x_17); +lean_dec(x_8); +x_82 = lean_ctor_get(x_67, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_67, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_84 = x_67; +} else { + lean_dec_ref(x_67); + 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_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; +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: +} +else { -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_object* x_86; +lean_dec(x_8); +lean_dec(x_7); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_10); +lean_ctor_set(x_86, 1, x_11); +return x_86; } } -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) +{ +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) +{ +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); +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_6597____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); +lean_inc(x_8); +x_23 = l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_7, x_13, x_9, x_8, x_12, x_8, x_9, lean_box(0), x_22, x_6); +lean_dec(x_8); 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 +block_55: { -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_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_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_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_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_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_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; -} +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_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; +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_6597____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_6597____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, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = 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, x_10, x_11); +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_12; } } -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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_7020____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; @@ -15933,7 +16496,7 @@ return x_5; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6597____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_7020____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; @@ -15972,7 +16535,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_6597____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_7020____spec__2(x_1, x_7, x_16, x_17, x_5); lean_dec(x_7); x_3 = x_12; x_5 = x_18; @@ -15986,1387 +16549,1008 @@ lean_dec(x_1); return x_5; } } -} -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____lambda__2___closed__3; -x_7 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____lambda__2___closed__3; -x_14 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____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_6597____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_6597____spec__4(x_33); -return x_34; -} -} -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____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_6597____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__1), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__3; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_6597____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_7020____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_6597____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_6597____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_6597____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_7020____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_6597____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_7020____spec__4(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6597____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_6597____lambda__2___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6597____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_7020____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_6724____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_6724_(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6724____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____lambda__2___closed__3; +x_7 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____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_7020____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_7020____lambda__2___closed__3; +x_14 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__1), 2, 0); +return x_1; } } -case 1: +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____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_7020____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_7020____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_7020____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_7020____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__3; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_7020____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_7020_(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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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_7020____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() { @@ -17476,7 +17660,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) @@ -17508,42 +17692,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; } } } @@ -18188,7 +18456,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); @@ -18275,7 +18543,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); @@ -18304,7 +18572,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); @@ -18720,9 +18988,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; @@ -18808,9 +19075,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; @@ -18831,7 +19096,6 @@ return x_7; else { uint8_t x_8; -lean_dec(x_2); lean_dec(x_1); x_8 = 0; return x_8; @@ -18866,6 +19130,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; } @@ -19214,10 +19479,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(); @@ -19241,7 +19514,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); @@ -19296,10 +19569,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(); @@ -19314,69 +19583,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(); @@ -19385,16 +19654,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(); @@ -19419,8 +19688,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(); @@ -19477,46 +19746,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_6597____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____lambda__2___closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6597____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6597_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____lambda__2___closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7020____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_7020_(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_6724____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6724____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6724____closed__1); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6724_(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 03503c51efe85..9b6020f7d0eea 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/Basic.c b/stage0/stdlib/Lean/Language/Basic.c index 76abb3b46a538..f3d83896c6074 100644 --- a/stage0/stdlib/Lean/Language/Basic.c +++ b/stage0/stdlib/Lean/Language/Basic.c @@ -324,15 +324,14 @@ return x_5; static lean_object* _init_l_Lean_Language_Snapshot_instInhabitedDiagnostics___closed__5() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Language_Snapshot_instInhabitedDiagnostics___closed__4; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_Snapshot_instInhabitedDiagnostics___closed__4; +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_Language_Snapshot_instInhabitedDiagnostics___closed__6() { @@ -2779,7 +2778,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_forM___at_Lean_Language_reportMessage _start: { lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); x_5 = l_Lean_PersistentArray_forM___at_Lean_Language_reportMessages___spec__2(x_2, x_4, x_3); return x_5; } diff --git a/stage0/stdlib/Lean/Language/Lean.c b/stage0/stdlib/Lean/Language/Lean.c index da965ae8dcd49..138d863f09e30 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; @@ -3109,11 +3109,11 @@ 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; 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; x_38 = l_Lean_NameSet_empty; -lean_inc(x_7); -x_39 = lean_alloc_ctor(0, 2, 1); +lean_inc_n(x_7, 2); +x_39 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_39, 0, x_7); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set_uint8(x_39, sizeof(void*)*2, x_27); +lean_ctor_set(x_39, 1, x_7); +lean_ctor_set(x_39, 2, x_38); x_40 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_402____closed__4; x_41 = lean_box(0); lean_inc_n(x_9, 2); @@ -5838,11 +5838,11 @@ lean_ctor_set(x_24, 11, x_4); lean_ctor_set_uint8(x_24, sizeof(void*)*12, x_23); lean_ctor_set_uint8(x_24, sizeof(void*)*12 + 1, x_23); x_25 = l_Lean_NameSet_empty; -lean_inc(x_5); -x_26 = lean_alloc_ctor(0, 2, 1); +lean_inc_n(x_5, 2); +x_26 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_26, 0, x_5); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set_uint8(x_26, sizeof(void*)*2, x_23); +lean_ctor_set(x_26, 1, x_5); +lean_ctor_set(x_26, 2, x_25); x_27 = 1; x_28 = l_Lean_Language_Lean_process_doElab___closed__5; x_29 = lean_alloc_ctor(0, 2, 1); @@ -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); @@ -7049,11 +7049,11 @@ lean_ctor_set(x_24, 11, x_4); lean_ctor_set_uint8(x_24, sizeof(void*)*12, x_23); lean_ctor_set_uint8(x_24, sizeof(void*)*12 + 1, x_23); x_25 = l_Lean_NameSet_empty; -lean_inc(x_5); -x_26 = lean_alloc_ctor(0, 2, 1); +lean_inc_n(x_5, 2); +x_26 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_26, 0, x_5); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set_uint8(x_26, sizeof(void*)*2, x_23); +lean_ctor_set(x_26, 1, x_5); +lean_ctor_set(x_26, 2, x_25); x_27 = 1; x_28 = l_Lean_Language_Lean_process_doElab___closed__5; x_29 = lean_alloc_ctor(0, 2, 1); @@ -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); @@ -8948,15 +8948,14 @@ return x_5; static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__12___closed__7() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__12___closed__6; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__12___closed__6; +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_Language_Lean_process_parseCmd___lambda__12___closed__8() { @@ -12946,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 b56b6c6d885d3..9d4ced534ad25 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 83759bb8497e8..a1765d27d6375 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -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; diff --git a/stage0/stdlib/Lean/Linter/UnusedVariables.c b/stage0/stdlib/Lean/Linter/UnusedVariables.c index 8862d98debe79..da45d506c444b 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 c7a5a27814bc7..054cc841053f2 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; @@ -30,25 +31,28 @@ lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__2(lean_object*, 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_Lean_Kernel_Exception_toMessageData___closed__53; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_mkErrorStringWithPos___closed__4; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__30; LEAN_EXPORT lean_object* l_Lean_MessageData_kind(lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__1; 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_2884____rarg___closed__12; 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_List_mapTR_loop___at_Lean_MessageData_orList___spec__1___closed__3; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SerialMessage_toMessage(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744_(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__26; -static lean_object* l_Lean_KernelException_toMessageData___closed__50; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__17; -static lean_object* l_Lean_KernelException_toMessageData___closed__36; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__3; +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*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_MessageLog_errorsToWarnings___spec__2(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*); LEAN_EXPORT lean_object* l_Lean_indentD(lean_object*); @@ -66,7 +69,7 @@ static lean_object* l_Lean_termM_x21_____closed__1; lean_object* l_EStateM_instInhabited___rarg(lean_object*, lean_object*); 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_Lean_MessageData_formatAux___closed__1; @@ -74,6 +77,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___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__19; static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__3; @@ -87,10 +91,9 @@ static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_ static lean_object* l_Lean_MessageData_ofName___closed__1; 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_Lean_KernelException_toMessageData___closed__8; 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_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; @@ -102,19 +105,20 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_MessageLog_errors static lean_object* l_Lean_MessageData_hasSyntheticSorry_visit___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__23; 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_instToMessageDataOption___rarg(lean_object*, lean_object*); static lean_object* 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_MessageData_ofList___closed__7; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__29; +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*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MessageData_hasTag___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -126,7 +130,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_beqMessageSeverity____ LEAN_EXPORT lean_object* l_Lean_MessageData_formatAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____spec__1___boxed(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_3364____closed__16; LEAN_EXPORT lean_object* l_Lean_mkErrorStringWithPos(lean_object*, lean_object*, lean_object*, lean_object*); @@ -136,14 +142,15 @@ 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___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__22; 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; LEAN_EXPORT lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); +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_Lean_Kernel_Exception_toMessageData___closed__13; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeListExpr(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__17; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__16; @@ -151,7 +158,6 @@ 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_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*); @@ -159,6 +165,7 @@ 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___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__27; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__40; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeString; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__7; @@ -167,6 +174,7 @@ static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__L LEAN_EXPORT lean_object* l_Lean_MessageData_instAppend(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___rarg(lean_object*); static lean_object* l_Lean_MessageData_orList___closed__4; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__42; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeSyntax; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__2; LEAN_EXPORT lean_object* l_Lean_instToMessageDataOfToFormat___rarg(lean_object*); @@ -180,16 +188,21 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedBaseMessage___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__6; 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___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884_(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_KernelException_toMessageData___closed__44; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__27; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__28; static lean_object* l_Lean_MessageData_instCoeList___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__9; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__50; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__6; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__20; +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_Lean_mkErrorStringWithPos___closed__3; static lean_object* l_Lean_MessageData_formatAux___closed__11; @@ -199,10 +212,13 @@ 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___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__13; +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; @@ -212,8 +228,6 @@ static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_ static lean_object* l_Lean_MessageData_ofList___closed__6; 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_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_formatAux___closed__8; @@ -221,11 +235,10 @@ 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; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__7; 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; static lean_object* l_Lean_instToMessageDataOption___rarg___closed__3; static lean_object* l___private_Lean_Message_0__Lean_toJsonMessageSeverity____x40_Lean_Message___hyg_125____closed__3; @@ -243,20 +256,18 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax(lean_object*); 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; -static lean_object* l_Lean_KernelException_toMessageData___closed__3; +lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); 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_Kernel_Exception_toMessageData___closed__26; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744____rarg___closed__4; -static lean_object* l_Lean_KernelException_toMessageData___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364_(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__2___closed__2; 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 +277,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; LEAN_EXPORT lean_object* l_Lean_MessageData_ofExpr___elambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqMessageSeverity; LEAN_EXPORT lean_object* l_Lean_Message_kind(lean_object*); @@ -275,24 +284,23 @@ static lean_object* l_Lean_MessageData_ofConstName___elambda__2___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__15; 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; lean_object* lean_array_to_list(lean_object*); static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_MessageLog_getInfoMessages___spec__2___closed__1; static lean_object* l_Lean_instImpl____x40_Lean_Message___hyg_606____closed__3; 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___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__20; 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; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__9; 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_Lean_KernelException_toMessageData___closed__2; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__15; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); static lean_object* l_Lean_instToMessageDataMessageData___closed__1; @@ -309,6 +317,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MessageLog_hasError LEAN_EXPORT lean_object* l_Lean_instFromJsonBaseMessage(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__19; lean_object* l_Lean_ppConstNameWithInfos(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__12; LEAN_EXPORT lean_object* l_Lean_MessageData_orList(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_getInfoMessages___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_ofSyntax___elambda__1(lean_object*); @@ -316,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*); extern lean_object* l_IO_instInhabitedError; 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_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__10; LEAN_EXPORT lean_object* l_Lean_MessageData_ofConstName___elambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -334,7 +343,6 @@ lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_MessageData_ofLevel___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg(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*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_MessageLog_getInfoMessages___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -349,17 +357,21 @@ extern lean_object* l_Std_Format_defWidth; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__8; 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_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__11; 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; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__33; 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_Lean_Kernel_Exception_toMessageData___closed__44; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744____rarg___closed__2; -static lean_object* l_Lean_KernelException_toMessageData___closed__40; +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); @@ -380,21 +392,17 @@ 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_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__11; 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_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_Lean_KernelException_toMessageData___closed__43; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____spec__1___lambda__1___boxed(lean_object*, lean_object*); 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_Lean_KernelException_toMessageData___closed__42; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__1; static lean_object* l_Lean_MessageData_formatAux___closed__9; LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion___rarg___lambda__1___boxed(lean_object*); @@ -403,12 +411,12 @@ static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__4; 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_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_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_toJsonMessageSeverity____x40_Lean_Message___hyg_125____closed__4; static lean_object* l_Lean_MessageData_formatAux___closed__10; @@ -418,18 +426,19 @@ static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_ 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___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__18; static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__2; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744____rarg___closed__5; -static lean_object* l_Lean_KernelException_toMessageData___closed__48; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_toJsonSerialMessage____x40_Lean_Message___hyg_3228_(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__5; 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_Lean_MessageData_ofConstName___elambda__2___closed__3; @@ -437,10 +446,13 @@ 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); 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___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__21; 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*); @@ -456,25 +468,21 @@ static lean_object* l_Lean_MessageData_formatAux___closed__12; 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_2884____spec__4___boxed(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; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____spec__2(lean_object*, lean_object*); -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*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__32; 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_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____spec__3___boxed(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__31; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__14; -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*); @@ -482,9 +490,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageData_formatAux_ static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__14; 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_Lean_MessageLog_empty___closed__2; LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax___elambda__1___boxed(lean_object*); -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_Lean_instFromJsonSerialMessage___closed__1; LEAN_EXPORT lean_object* l_Lean_toMessageList(lean_object*); @@ -504,7 +513,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_Lean_MessageData_paren___closed__1; LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax(lean_object*); static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__5; @@ -515,27 +523,22 @@ 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___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__10; 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___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2884____rarg___closed__3; static lean_object* l_Lean_MessageData_ofList___closed__3; @@ -545,19 +548,16 @@ 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; -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*); 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*); @@ -570,6 +570,8 @@ 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; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__5; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeExpr; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3364____closed__12; LEAN_EXPORT lean_object* l_Lean_instToMessageDataSubarray___rarg(lean_object*, lean_object*); @@ -582,7 +584,6 @@ 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___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744____rarg___closed__9; -static lean_object* l_Lean_KernelException_toMessageData___closed__23; 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*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -607,10 +608,13 @@ 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*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -621,7 +625,7 @@ 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___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2744____rarg___closed__3; 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*); @@ -630,10 +634,10 @@ 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_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*); @@ -9393,15 +9397,14 @@ return x_5; static lean_object* _init_l_Lean_instInhabitedMessageLog___closed__5() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_instInhabitedMessageLog___closed__4; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedMessageLog___closed__4; +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_instInhabitedMessageLog() { @@ -9451,15 +9454,14 @@ return x_5; static lean_object* _init_l_Lean_MessageLog_empty___closed__4() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_MessageLog_empty___closed__3; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_MessageLog_empty___closed__3; +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_MessageLog_empty() { @@ -9474,7 +9476,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_msgs(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_ctor_get(x_1, 0); +x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); return x_2; } @@ -9488,11 +9490,25 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_MessageLog_reportedPlusUnreported(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_PersistentArray_append___rarg(x_2, x_3); +lean_dec(x_3); +return x_4; +} +} LEAN_EXPORT uint8_t l_Lean_MessageLog_hasUnreported(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = lean_ctor_get(x_1, 0); +x_2 = lean_ctor_get(x_1, 1); x_3 = l_Lean_PersistentArray_isEmpty___rarg(x_2); if (x_3 == 0) { @@ -9526,96 +9542,67 @@ 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, 0); +x_4 = lean_ctor_get(x_2, 1); x_5 = l_Lean_PersistentArray_push___rarg(x_4, x_1); -lean_ctor_set(x_2, 0, x_5); +lean_ctor_set(x_2, 1, x_5); return x_2; } else { -uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_6 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_2, 1); +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_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_2, 2); lean_inc(x_8); lean_inc(x_7); +lean_inc(x_6); lean_dec(x_2); x_9 = l_Lean_PersistentArray_push___rarg(x_7, x_1); -x_10 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_6); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_9); +lean_ctor_set(x_10, 2, x_8); return x_10; } } } +static lean_object* _init_l_Lean_MessageLog_append___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_quickCmp___boxed), 2, 0); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_MessageLog_append(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -x_4 = lean_ctor_get(x_1, 0); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -lean_dec(x_1); -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_2, 0); +x_5 = l_Lean_PersistentArray_append___rarg(x_3, x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +x_8 = l_Lean_PersistentArray_append___rarg(x_6, x_7); lean_dec(x_7); -x_8 = l_Lean_PersistentArray_append___rarg(x_4, x_6); -lean_dec(x_6); -if (x_3 == 0) -{ -lean_object* x_9; -x_9 = l_Lean_NameSet_empty; -lean_ctor_set(x_2, 1, x_9); -lean_ctor_set(x_2, 0, x_8); -return x_2; -} -else -{ -uint8_t x_10; lean_object* x_11; -x_10 = 1; -x_11 = l_Lean_NameSet_empty; -lean_ctor_set(x_2, 1, x_11); -lean_ctor_set(x_2, 0, x_8); -lean_ctor_set_uint8(x_2, sizeof(void*)*2, x_10); -return x_2; -} -} -else -{ -uint8_t x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -x_13 = lean_ctor_get(x_2, 0); -lean_inc(x_13); +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); lean_dec(x_2); -x_14 = l_Lean_PersistentArray_append___rarg(x_4, x_13); -lean_dec(x_13); -if (x_3 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = l_Lean_NameSet_empty; -x_16 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*2, x_12); -return x_16; -} -else -{ -uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_17 = 1; -x_18 = l_Lean_NameSet_empty; -x_19 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_17); -return x_19; -} -} +x_11 = l_Lean_MessageLog_append___closed__1; +x_12 = l_Lean_RBTree_union___rarg(x_11, x_9, x_10); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_8); +lean_ctor_set(x_13, 2, x_12); +return x_13; } } static lean_object* _init_l_Lean_MessageLog_instAppend___closed__1() { @@ -9796,20 +9783,21 @@ return x_3; LEAN_EXPORT uint8_t l_Lean_MessageLog_hasErrors(lean_object* x_1) { _start: { -uint8_t x_2; -x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_2 == 0) +lean_object* x_2; uint8_t x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_2); +if (x_3 == 0) { -lean_object* x_3; uint8_t x_4; -x_3 = lean_ctor_get(x_1, 0); -x_4 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_3); -return x_4; +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_4); +return x_5; } else { -uint8_t x_5; -x_5 = 1; -return x_5; +uint8_t x_6; +x_6 = 1; +return x_6; } } } @@ -9874,31 +9862,38 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_MessageLog_markAllReported(lean_object* x_1) { _start: { -uint8_t x_2; uint8_t x_3; -x_2 = l_Lean_MessageLog_hasErrors(x_1); -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = l_Lean_PersistentArray_append___rarg(x_3, x_4); lean_dec(x_4); -x_5 = l_Lean_MessageLog_empty___closed__3; +x_6 = l_Lean_MessageLog_empty___closed__3; +lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_2); return x_1; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); +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_1, 0); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); lean_dec(x_1); -x_7 = l_Lean_MessageLog_empty___closed__3; -x_8 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -lean_ctor_set_uint8(x_8, sizeof(void*)*2, x_2); -return x_8; +x_10 = l_Lean_PersistentArray_append___rarg(x_7, x_8); +lean_dec(x_8); +x_11 = l_Lean_MessageLog_empty___closed__3; +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set(x_12, 2, x_9); +return x_12; } } } @@ -10126,17 +10121,17 @@ return x_18; LEAN_EXPORT lean_object* l_Lean_MessageLog_errorsToWarnings(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; -x_2 = lean_ctor_get(x_1, 0); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); lean_dec(x_1); x_3 = l_Lean_PersistentArray_mapM___at_Lean_MessageLog_errorsToWarnings___spec__1(x_2); -x_4 = 0; +x_4 = l_Lean_MessageLog_empty___closed__3; x_5 = l_Lean_NameSet_empty; -x_6 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_5); -lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_4); +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_5); return x_6; } } @@ -10508,18 +10503,17 @@ return x_35; LEAN_EXPORT lean_object* l_Lean_MessageLog_getInfoMessages(lean_object* x_1) { _start: { -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; -x_2 = lean_ctor_get(x_1, 0); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_ctor_get(x_1, 1); x_3 = l_Lean_MessageLog_empty___closed__3; x_4 = lean_unsigned_to_nat(0u); x_5 = l_Lean_PersistentArray_foldlM___at_Lean_MessageLog_getInfoMessages___spec__1(x_2, x_3, x_4); -x_6 = 0; -x_7 = l_Lean_NameSet_empty; -x_8 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_7); -lean_ctor_set_uint8(x_8, sizeof(void*)*2, x_6); -return x_8; +x_6 = l_Lean_NameSet_empty; +x_7 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_5); +lean_ctor_set(x_7, 2, x_6); +return x_7; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MessageLog_getInfoMessages___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -10593,7 +10587,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_forM___rarg(lean_object* x_1, lean_ob _start: { lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); lean_inc(x_4); lean_dec(x_2); x_5 = l_Lean_PersistentArray_forM___rarg(x_1, x_4, x_3); @@ -10612,7 +10606,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_toList(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 0); +x_2 = lean_ctor_get(x_1, 1); x_3 = l_Lean_PersistentArray_toList___rarg(x_2); return x_3; } @@ -10630,7 +10624,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_toArray(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 0); +x_2 = lean_ctor_get(x_1, 1); x_3 = l_Lean_PersistentArray_toArray___rarg(x_2); return x_3; } @@ -11876,7 +11870,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; @@ -11892,7 +11886,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; @@ -11900,16 +11894,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; @@ -11918,7 +11912,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; @@ -11926,16 +11920,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; @@ -11943,26 +11937,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; @@ -11970,16 +11964,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; @@ -11987,16 +11981,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; @@ -12004,16 +11998,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; @@ -12022,7 +12016,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; @@ -12030,16 +12024,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; @@ -12047,16 +12041,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; @@ -12064,16 +12058,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; @@ -12081,16 +12075,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; @@ -12098,16 +12092,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; @@ -12115,16 +12109,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; @@ -12132,16 +12126,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; @@ -12149,16 +12143,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; @@ -12166,16 +12160,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; @@ -12183,16 +12177,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; @@ -12200,16 +12194,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; @@ -12217,16 +12211,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; @@ -12234,16 +12228,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; @@ -12251,26 +12245,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; @@ -12278,26 +12272,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; @@ -12305,26 +12299,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; @@ -12332,26 +12326,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)) { @@ -12365,16 +12359,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 @@ -12386,16 +12380,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; } } @@ -12410,16 +12404,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 @@ -12432,16 +12426,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; } } @@ -12471,11 +12465,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); @@ -12483,7 +12477,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); @@ -12491,12 +12485,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: @@ -12519,11 +12513,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); @@ -12531,7 +12525,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); @@ -12539,12 +12533,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: @@ -12555,8 +12549,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; } } @@ -12571,16 +12565,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: @@ -12593,16 +12587,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: @@ -12616,15 +12610,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: @@ -12638,15 +12632,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: @@ -12660,15 +12654,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: @@ -12682,15 +12676,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: @@ -12708,11 +12702,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); @@ -12720,7 +12714,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); @@ -12728,11 +12722,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: @@ -12746,15 +12740,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: @@ -12769,11 +12763,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); @@ -12781,12 +12775,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: @@ -12798,11 +12792,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); @@ -12812,28 +12806,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; } } @@ -13279,6 +13273,8 @@ l_Lean_MessageLog_empty___closed__4 = _init_l_Lean_MessageLog_empty___closed__4( lean_mark_persistent(l_Lean_MessageLog_empty___closed__4); l_Lean_MessageLog_empty = _init_l_Lean_MessageLog_empty(); lean_mark_persistent(l_Lean_MessageLog_empty); +l_Lean_MessageLog_append___closed__1 = _init_l_Lean_MessageLog_append___closed__1(); +lean_mark_persistent(l_Lean_MessageLog_append___closed__1); l_Lean_MessageLog_instAppend___closed__1 = _init_l_Lean_MessageLog_instAppend___closed__1(); lean_mark_persistent(l_Lean_MessageLog_instAppend___closed__1); l_Lean_MessageLog_instAppend = _init_l_Lean_MessageLog_instAppend(); @@ -13381,112 +13377,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 a3af27db79c3d..65fbbf0763710 100644 --- a/stage0/stdlib/Lean/Meta/AppBuilder.c +++ b/stage0/stdlib/Lean/Meta/AppBuilder.c @@ -129,7 +129,7 @@ static lean_object* l_Lean_Meta_congrArg_x3f___lambda__4___closed__9; static lean_object* l_Lean_Meta_mkImpCongr___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7390____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_Lean_Meta_mkSorry___closed__7; 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; @@ -12930,7 +12930,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; @@ -13135,7 +13136,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 788d04b387878..439e7056b299e 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -279,6 +279,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*); LEAN_EXPORT lean_object* l_Lean_Meta_withInstImplicitAsImplict(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_getKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -287,7 +288,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*); @@ -383,7 +384,6 @@ LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_B 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_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_beqSynthInstanceCacheKey____x40_Lean_Meta_Basic___hyg_1168____boxed(lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_3106____closed__7; 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*); @@ -395,12 +395,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getNumPostponed___rarg(lean_object*, lean_o 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); 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*); -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*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_3106____closed__4; uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -530,7 +528,6 @@ lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(le lean_object* l_Lean_ConstantInfo_numLevelParams(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_3106____closed__6; 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*); @@ -574,6 +571,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedPostponedEntry; LEAN_EXPORT lean_object* l_Lean_Meta_ProjReductionKind_noConfusion___rarg___lambda__1___boxed(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*); 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*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_3106____closed__9; @@ -935,7 +933,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_Lean_Meta_mapForallTelescope_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1009,6 +1007,7 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__ LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg(lean_object*, lean_object*, lean_object*, 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*); 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*); @@ -1016,6 +1015,7 @@ lean_object* l_Repr_addAppParen(lean_object*, 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*); @@ -1211,6 +1211,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassImp_x3f static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprProjReductionKind____x40_Lean_Meta_Basic___hyg_210____closed__9; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2138____closed__8; 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_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*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_liftMkBindingM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5594,7 +5595,7 @@ lean_dec(x_80); x_83 = lean_ctor_get(x_81, 0); lean_inc(x_83); lean_dec(x_81); -x_84 = l_Lean_Kernel_isDiagnosticsEnabled(x_83); +x_84 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_83); lean_dec(x_83); if (x_84 == 0) { @@ -5648,7 +5649,7 @@ lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean x_90 = lean_ctor_get(x_87, 0); x_91 = lean_ctor_get(x_87, 4); lean_dec(x_91); -x_92 = l_Lean_Kernel_enableDiag(x_90, x_79); +x_92 = l_Lean_Kernel_Environment_enableDiag(x_90, x_79); x_93 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; lean_ctor_set(x_87, 4, x_93); lean_ctor_set(x_87, 0, x_92); @@ -5763,7 +5764,7 @@ lean_inc(x_118); lean_inc(x_117); lean_inc(x_116); lean_dec(x_87); -x_123 = l_Lean_Kernel_enableDiag(x_116, x_79); +x_123 = l_Lean_Kernel_Environment_enableDiag(x_116, x_79); x_124 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; x_125 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_125, 0, x_123); @@ -5997,7 +5998,7 @@ lean_dec(x_185); x_188 = lean_ctor_get(x_186, 0); lean_inc(x_188); lean_dec(x_186); -x_189 = l_Lean_Kernel_isDiagnosticsEnabled(x_188); +x_189 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_188); lean_dec(x_188); if (x_189 == 0) { @@ -6072,7 +6073,7 @@ if (lean_is_exclusive(x_192)) { lean_dec_ref(x_192); x_201 = lean_box(0); } -x_202 = l_Lean_Kernel_enableDiag(x_194, x_184); +x_202 = l_Lean_Kernel_Environment_enableDiag(x_194, x_184); x_203 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; if (lean_is_scalar(x_201)) { x_204 = lean_alloc_ctor(0, 8, 0); @@ -8239,12 +8240,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); @@ -8279,7 +8280,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); @@ -8319,12 +8320,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); @@ -8361,7 +8362,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); @@ -8422,12 +8423,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 { @@ -8473,7 +8474,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 { @@ -8588,12 +8589,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); @@ -8628,7 +8629,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); @@ -8668,12 +8669,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); @@ -8710,7 +8711,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); @@ -8771,12 +8772,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 { @@ -8822,7 +8823,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 { @@ -8937,12 +8938,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); @@ -8977,7 +8978,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); @@ -9017,12 +9018,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); @@ -9059,7 +9060,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); @@ -9120,12 +9121,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 { @@ -9171,7 +9172,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 { @@ -12402,8 +12403,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; @@ -12443,8 +12443,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; @@ -21881,7 +21880,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); @@ -22272,7 +22272,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); @@ -22621,7 +22622,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); @@ -23138,7 +23140,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); @@ -23599,7 +23602,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); @@ -23991,7 +23995,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); @@ -24305,7 +24310,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); @@ -24697,7 +24703,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); @@ -25132,7 +25139,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); @@ -25524,7 +25532,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); @@ -28891,8 +28900,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; @@ -28977,8 +28985,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; @@ -43280,7 +43287,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; @@ -43343,7 +43350,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; @@ -43409,6 +43416,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 028f455fa4c0c..1934162301bc1 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/Constructions/CasesOn.c b/stage0/stdlib/Lean/Meta/Constructions/CasesOn.c index 48d57035cf21d..ab387f47a7fec 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 f4a2c040ac25c..33d121655edf0 100644 --- a/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c +++ b/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c @@ -99,6 +99,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*); @@ -129,7 +130,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*, 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*); @@ -198,7 +198,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 0f34391a26644..a56d7a12e30ec 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 b0c1abf983afa..d3b32dce3df70 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 ed5a11710b363..94920cc3b483d 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; @@ -12953,7 +12951,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); @@ -13367,6 +13364,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 46c1312026bf1..35a102a5fd20d 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) { @@ -6937,7 +6939,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) { @@ -6971,7 +6972,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 ec3104e5e2134..5e08599c73372 100644 --- a/stage0/stdlib/Lean/Meta/ExprDefEq.c +++ b/stage0/stdlib/Lean/Meta/ExprDefEq.c @@ -243,7 +243,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*); @@ -3862,6 +3862,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; @@ -4188,6 +4189,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; @@ -4602,7 +4604,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; @@ -4697,7 +4700,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; @@ -4749,7 +4753,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; @@ -4829,7 +4834,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; @@ -4927,7 +4933,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; @@ -80383,7 +80390,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) { @@ -80394,7 +80400,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); @@ -80421,7 +80426,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); @@ -80433,7 +80437,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; } } @@ -80449,7 +80452,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) { @@ -80460,7 +80462,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); @@ -80489,7 +80490,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); @@ -80502,7 +80502,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; } } @@ -80531,6 +80530,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 @@ -80567,6 +80567,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 @@ -80625,6 +80626,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; } } @@ -83753,7 +83755,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; @@ -84226,7 +84229,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; @@ -84626,7 +84630,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 da8de97f424f1..cb60143b158bb 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 62ebe98f1005f..06e41b612ede2 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -23601,6 +23601,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; @@ -25135,7 +25136,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 c0c9bc45ae08c..edad423b7f064 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_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; static lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_InferType_0__Lean_Meta_checkInferTypeCache___spec__5___closed__1; @@ -4876,7 +4876,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; @@ -5879,7 +5880,8 @@ lean_dec(x_18); x_264 = lean_ctor_get(x_262, 0); lean_inc(x_264); lean_dec(x_262); -x_265 = lean_environment_find(x_264, x_16); +x_265 = l_Lean_Environment_find_x3f(x_264, x_16); +lean_dec(x_16); if (lean_obj_tag(x_265) == 0) { lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; diff --git a/stage0/stdlib/Lean/Meta/Injective.c b/stage0/stdlib/Lean/Meta/Injective.c index dfefe23799546..ee0cba68110bd 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 ffef57cbb0cb9..b4cff905e209e 100644 --- a/stage0/stdlib/Lean/Meta/Instances.c +++ b/stage0/stdlib/Lean/Meta/Instances.c @@ -101,7 +101,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_4656____ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Instances_erase___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_4656____lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___lambda__1___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3711____closed__6; @@ -15425,8 +15425,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; @@ -15475,8 +15474,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 4e760b232a322..52d7366cb6a64 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); @@ -16674,15 +16673,14 @@ return x_3; static lean_object* _init_l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__12() { _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_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__4; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__4; +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___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__13() { @@ -16779,7 +16777,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) { @@ -16833,7 +16831,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); @@ -16922,7 +16920,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); @@ -17844,7 +17842,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); @@ -18725,7 +18723,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); @@ -18748,7 +18746,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); @@ -19132,7 +19130,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); @@ -19589,7 +19587,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); @@ -19773,7 +19771,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); @@ -21171,7 +21169,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 caa9167901fae..c4c5f79fb684f 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 fa5cfb83e8ce9..5b1c786e2b483 100644 --- a/stage0/stdlib/Lean/Meta/Match/MatchEqs.c +++ b/stage0/stdlib/Lean/Meta/Match/MatchEqs.c @@ -189,7 +189,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; @@ -48744,6 +48746,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); @@ -48974,6 +48977,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); @@ -49226,6 +49230,7 @@ lean_inc(x_140); lean_dec(x_137); lean_inc(x_1); x_141 = l_Lean_mkPrivateName(x_140, x_1); +lean_dec(x_140); x_142 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5___closed__2; lean_inc(x_141); x_143 = l_Lean_Name_append(x_141, x_142); @@ -49517,6 +49522,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 bf2bfc0279321..453fdb59a99e9 100644 --- a/stage0/stdlib/Lean/Meta/SizeOf.c +++ b/stage0/stdlib/Lean/Meta/SizeOf.c @@ -111,7 +111,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_isInductiveHy static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__2___closed__2; 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; @@ -7574,7 +7574,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; @@ -7781,7 +7782,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; @@ -8325,7 +8327,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; @@ -8492,7 +8495,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; @@ -11286,7 +11290,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; @@ -11582,7 +11587,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; @@ -12058,8 +12064,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; @@ -12099,8 +12104,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; @@ -18353,7 +18357,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/Tactic/Apply.c b/stage0/stdlib/Lean/Meta/Tactic/Apply.c index 088ce11b55e7b..0b53bf61e8bdc 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*); @@ -6698,7 +6698,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 eef1a2d069de1..865cf560cb574 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 0a7ed8eccd853..4c10c13449534 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cases.c @@ -99,7 +99,7 @@ lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean 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*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10; 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*); 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*); @@ -627,7 +627,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 +704,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; @@ -2633,7 +2635,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_find(x_18, x_14); +x_19 = l_Lean_Environment_find_x3f(x_18, x_14); +lean_dec(x_14); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; @@ -3024,7 +3027,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 +3085,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 +3228,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 +3290,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; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Constructor.c b/stage0/stdlib/Lean/Meta/Tactic/Constructor.c index 521ffe434ecbf..4415f2f1b40b8 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 0ae3e87c199ff..2c2edae053a99 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c @@ -70,7 +70,7 @@ uint8_t lean_expr_has_loose_bvar(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Contradiction___hyg_4434____closed__9; 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_4434____closed__5; 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*); @@ -731,7 +731,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; @@ -803,7 +804,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; @@ -917,7 +919,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 5fe04abf72b2e..f988f2f89a02f 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*); 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; @@ -356,7 +357,8 @@ lean_dec(x_9); x_83 = lean_ctor_get(x_81, 0); lean_inc(x_83); lean_dec(x_81); -x_84 = lean_environment_find(x_83, x_7); +x_84 = l_Lean_Environment_find_x3f(x_83, x_7); +lean_dec(x_7); if (lean_obj_tag(x_84) == 0) { lean_object* x_85; lean_object* x_86; diff --git a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c index d9e1fa2973816..9c4ffbcd44fa3 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c +++ b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c @@ -240,7 +240,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*); @@ -2189,8 +2189,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; @@ -2234,8 +2233,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; @@ -28554,8 +28552,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; @@ -28603,8 +28600,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/RevertAll.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/RevertAll.c index 1d60b1dd1b88e..a09016473ca4e 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/RevertAll.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/RevertAll.c @@ -42,7 +42,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_MVarId_revertAll___lambda__2___closed__5; uint8_t l_Lean_LocalDecl_isAuxDecl(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_Array_filterMapM___at_Lean_MVarId_revertAll___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_revertAll___lambda__1___boxed(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*); @@ -394,7 +394,8 @@ if (x_86 == 0) lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; x_87 = lean_ctor_get(x_85, 0); x_88 = lean_ctor_get(x_85, 1); -x_89 = lean_environment_main_module(x_76); +x_89 = l_Lean_Environment_mainModule(x_76); +lean_dec(x_76); lean_ctor_set(x_85, 1, x_10); lean_ctor_set(x_85, 0, x_89); x_90 = lean_ctor_get(x_87, 1); @@ -795,7 +796,8 @@ x_206 = lean_ctor_get(x_85, 1); lean_inc(x_206); lean_inc(x_205); lean_dec(x_85); -x_207 = lean_environment_main_module(x_76); +x_207 = l_Lean_Environment_mainModule(x_76); +lean_dec(x_76); x_208 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_208, 0, x_207); lean_ctor_set(x_208, 1, x_10); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Revert.c b/stage0/stdlib/Lean/Meta/Tactic/Revert.c index 6253b86d74957..a82d74591fd4c 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 9c940870215e5..eb2099c27b8eb 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c @@ -201,7 +201,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*); @@ -1735,7 +1735,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; @@ -2570,7 +2571,8 @@ lean_dec(x_12); x_229 = lean_ctor_get(x_227, 0); lean_inc(x_229); lean_dec(x_227); -x_230 = lean_environment_find(x_229, x_11); +x_230 = l_Lean_Environment_find_x3f(x_229, x_11); +lean_dec(x_11); if (lean_obj_tag(x_230) == 0) { lean_object* x_231; lean_object* x_232; @@ -4092,7 +4094,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; @@ -4171,7 +4174,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 1c763642c466c..47da0f576592c 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 d57826b530216..b33d414d4f46e 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c @@ -50,7 +50,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpAll___hy 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*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpAll___hyg_2247____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_mkFreshId___at___private_Lean_Meta_Tactic_Simp_SimpAll_0__Lean_Meta_SimpAll_loop___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpAll___hyg_2247_(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -1372,8 +1372,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; @@ -1413,8 +1412,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 3618954b48c1f..6d82168eaa06b 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c @@ -126,7 +126,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*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreDSimprocs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -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 1ad549339ba1c..463f45224f3d7 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; @@ -180,7 +179,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*); @@ -365,12 +363,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*); @@ -8989,12 +8989,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); @@ -9029,7 +9029,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); @@ -9069,12 +9069,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); @@ -9111,7 +9111,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); @@ -9172,12 +9172,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 { @@ -9223,7 +9223,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 fde49928f595a..74bfabe002630 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; @@ -7542,7 +7542,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); @@ -7565,7 +7566,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); @@ -7587,7 +7589,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); @@ -7609,7 +7612,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; @@ -7804,7 +7808,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; @@ -8004,7 +8009,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); @@ -8032,7 +8038,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; @@ -8237,7 +8244,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); @@ -8265,7 +8273,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); @@ -8293,7 +8302,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; @@ -8503,7 +8513,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); @@ -8532,7 +8543,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); @@ -8560,7 +8572,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); @@ -8588,7 +8601,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 a560bc841f0b1..b46049f06f92d 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c +++ b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c @@ -140,14 +140,12 @@ 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_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_607____closed__6; 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; @@ -206,6 +204,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*); @@ -347,7 +346,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*); @@ -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*); @@ -1613,7 +1613,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) { @@ -1663,7 +1663,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); @@ -1693,7 +1693,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); @@ -5887,7 +5887,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) { @@ -5937,7 +5937,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); @@ -5967,7 +5967,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); @@ -7506,7 +7506,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); @@ -7565,7 +7566,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); @@ -7921,7 +7923,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); @@ -7989,7 +7992,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 141427f93420c..9e4a6ac3d6188 100644 --- a/stage0/stdlib/Lean/Meta/Transform.c +++ b/stage0/stdlib/Lean/Meta/Transform.c @@ -106,7 +106,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*); @@ -11304,7 +11304,8 @@ x_9 = l_Lean_Environment_contains(x_1, x_7); if (x_9 == 0) { lean_object* x_10; -x_10 = lean_environment_find(x_2, x_7); +x_10 = l_Lean_Environment_find_x3f(x_2, x_7); +lean_dec(x_7); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; diff --git a/stage0/stdlib/Lean/Meta/WHNF.c b/stage0/stdlib/Lean/Meta/WHNF.c index c19b817632802..05a603323ee5d 100644 --- a/stage0/stdlib/Lean/Meta/WHNF.c +++ b/stage0/stdlib/Lean/Meta/WHNF.c @@ -151,7 +151,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; @@ -1434,6 +1434,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; } } @@ -1460,6 +1461,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; @@ -3890,6 +3892,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; @@ -4087,6 +4090,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; @@ -4366,7 +4370,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) { @@ -4374,7 +4377,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; } @@ -4435,7 +4437,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); @@ -4466,7 +4467,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; @@ -4498,7 +4498,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) { @@ -4527,7 +4526,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) { @@ -4556,7 +4554,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) { @@ -4587,7 +4584,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) { @@ -4620,7 +4616,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) { @@ -4629,7 +4624,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); @@ -4697,7 +4691,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; @@ -4733,7 +4726,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); @@ -4764,7 +4756,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); @@ -4795,7 +4786,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)) { @@ -4824,7 +4814,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); @@ -4897,6 +4886,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; } } @@ -4974,6 +4964,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; } } @@ -5959,6 +5950,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; @@ -6239,6 +6231,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; @@ -13176,7 +13169,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) { @@ -18227,7 +18219,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; @@ -18366,7 +18359,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; @@ -23886,7 +23880,6 @@ lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -lean_inc(x_7); x_9 = l_Lean_Meta_getUnfoldableConstNoEx_x3f(x_7, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_9) == 0) { @@ -24309,7 +24302,8 @@ x_94 = l_Lean_ConstantInfo_name(x_70); x_95 = l_Lean_Meta_smartUnfoldingSuffix; lean_inc(x_94); x_96 = l_Lean_Name_str___override(x_94, x_95); -x_97 = lean_environment_find(x_93, x_96); +x_97 = l_Lean_Environment_find_x3f(x_93, x_96); +lean_dec(x_96); if (lean_obj_tag(x_97) == 0) { lean_object* x_98; lean_object* x_99; @@ -24713,7 +24707,8 @@ x_193 = l_Lean_ConstantInfo_name(x_167); x_194 = l_Lean_Meta_smartUnfoldingSuffix; lean_inc(x_193); x_195 = l_Lean_Name_str___override(x_193, x_194); -x_196 = lean_environment_find(x_192, x_195); +x_196 = l_Lean_Environment_find_x3f(x_192, x_195); +lean_dec(x_195); if (lean_obj_tag(x_196) == 0) { lean_object* x_197; lean_object* x_198; diff --git a/stage0/stdlib/Lean/Modifiers.c b/stage0/stdlib/Lean/Modifiers.c index 4b567111d6875..72ea42d3f0878 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 5352c0796e0c7..de040193bfa10 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 231af616156f5..e1aebfdb93414 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/Module.c b/stage0/stdlib/Lean/Parser/Module.c index a36e7ec4c971b..a5d02e1797c98 100644 --- a/stage0/stdlib/Lean/Parser/Module.c +++ b/stage0/stdlib/Lean/Parser/Module.c @@ -2434,15 +2434,14 @@ return x_5; static lean_object* _init_l_Lean_Parser_parseHeader___closed__6() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Parser_parseHeader___closed__5; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_parseHeader___closed__5; +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; } } LEAN_EXPORT lean_object* l_Lean_Parser_parseHeader(lean_object* x_1, lean_object* x_2) { diff --git a/stage0/stdlib/Lean/Parser/Tactic/Doc.c b/stage0/stdlib/Lean/Parser/Tactic/Doc.c index 0543533ef1c86..304341e735b09 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 384e5a9cd421d..eb92ec6dd7583 100644 --- a/stage0/stdlib/Lean/ParserCompiler.c +++ b/stage0/stdlib/Lean/ParserCompiler.c @@ -305,6 +305,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*); @@ -434,7 +435,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; @@ -22904,7 +22904,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -22993,7 +22993,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -23115,7 +23115,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -24167,7 +24167,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -24256,7 +24256,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -24378,7 +24378,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -25430,7 +25430,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -25519,7 +25519,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -25641,7 +25641,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -26693,7 +26693,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -26782,7 +26782,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -26904,7 +26904,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -27956,7 +27956,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -28045,7 +28045,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -28167,7 +28167,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -29274,7 +29274,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -29363,7 +29363,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -29485,7 +29485,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -30537,7 +30537,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -30626,7 +30626,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -30748,7 +30748,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -31800,7 +31800,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -31889,7 +31889,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -32011,7 +32011,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -33063,7 +33063,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -33152,7 +33152,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -33274,7 +33274,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); @@ -34326,7 +34326,7 @@ if (x_40 == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_39, 0); -x_42 = l_Lean_KernelException_toMessageData(x_41, x_26); +x_42 = l_Lean_Kernel_Exception_toMessageData(x_41, x_26); x_43 = lean_ctor_get(x_14, 5); lean_inc(x_43); x_44 = l_Lean_MessageData_toString(x_42, x_36); @@ -34415,7 +34415,7 @@ lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); lean_dec(x_39); -x_63 = l_Lean_KernelException_toMessageData(x_62, x_26); +x_63 = l_Lean_Kernel_Exception_toMessageData(x_62, x_26); x_64 = lean_ctor_get(x_14, 5); lean_inc(x_64); x_65 = l_Lean_MessageData_toString(x_63, x_36); @@ -34537,7 +34537,7 @@ if (lean_is_exclusive(x_88)) { lean_dec_ref(x_88); x_90 = lean_box(0); } -x_91 = l_Lean_KernelException_toMessageData(x_89, x_26); +x_91 = l_Lean_Kernel_Exception_toMessageData(x_89, x_26); x_92 = lean_ctor_get(x_14, 5); lean_inc(x_92); x_93 = l_Lean_MessageData_toString(x_91, x_85); diff --git a/stage0/stdlib/Lean/PrettyPrinter.c b/stage0/stdlib/Lean/PrettyPrinter.c index 75aa125d172b4..0930e120ca823 100644 --- a/stage0/stdlib/Lean/PrettyPrinter.c +++ b/stage0/stdlib/Lean/PrettyPrinter.c @@ -55,7 +55,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter__ static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1096____closed__1; 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_object* l_Lean_PrettyPrinter_format(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PPContext_runMetaM___rarg___closed__3; @@ -73,10 +73,8 @@ 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*); 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); @@ -103,6 +101,7 @@ static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_ static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1176____closed__12; static lean_object* l_Lean_PrettyPrinter_registerParserCompilers___closed__10; extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; +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_Lean_PrettyPrinter_ppTerm___closed__2; static lean_object* l_Lean_PrettyPrinter_ppExprLegacy___closed__3; @@ -228,6 +227,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext(lean_object*, lean_object*); 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; @@ -496,15 +496,14 @@ return x_2; static lean_object* _init_l_Lean_PPContext_runCoreM___rarg___closed__17() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_PPContext_runCoreM___rarg___closed__12; -x_3 = l_Lean_NameSet_empty; -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; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PPContext_runCoreM___rarg___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_PPContext_runCoreM___rarg___closed__18() { @@ -615,7 +614,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) { @@ -828,7 +827,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_52); +x_88 = l_Lean_Kernel_Environment_enableDiag(x_86, x_52); lean_ctor_set(x_83, 4, x_57); lean_ctor_set(x_83, 0, x_88); x_89 = lean_st_ref_set(x_74, x_83, x_84); @@ -942,7 +941,7 @@ lean_inc(x_113); lean_inc(x_112); lean_inc(x_111); lean_dec(x_83); -x_118 = l_Lean_Kernel_enableDiag(x_111, x_52); +x_118 = l_Lean_Kernel_Environment_enableDiag(x_111, x_52); x_119 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_119, 0, x_118); lean_ctor_set(x_119, 1, x_112); @@ -2144,8 +2143,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; @@ -2247,8 +2245,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; @@ -2440,7 +2437,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) { @@ -2494,7 +2491,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); @@ -2577,7 +2574,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); @@ -2773,7 +2770,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) { @@ -2848,7 +2845,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 { @@ -3154,7 +3151,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) { @@ -3372,7 +3369,7 @@ x_93 = lean_ctor_get(x_90, 0); x_94 = lean_ctor_get(x_90, 4); lean_dec(x_94); x_95 = l_Lean_PrettyPrinter_ppExprLegacy___closed__7; -x_96 = l_Lean_Kernel_enableDiag(x_93, x_95); +x_96 = l_Lean_Kernel_Environment_enableDiag(x_93, x_95); lean_ctor_set(x_90, 4, x_65); lean_ctor_set(x_90, 0, x_96); x_97 = lean_st_ref_set(x_81, x_90, x_91); @@ -3487,7 +3484,7 @@ lean_inc(x_120); lean_inc(x_119); lean_dec(x_90); x_126 = l_Lean_PrettyPrinter_ppExprLegacy___closed__7; -x_127 = l_Lean_Kernel_enableDiag(x_119, x_126); +x_127 = l_Lean_Kernel_Environment_enableDiag(x_119, x_126); x_128 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_128, 0, x_127); lean_ctor_set(x_128, 1, x_120); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index d4b246df7be54..cc96d123227e5 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -135,7 +135,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; @@ -144,7 +143,6 @@ extern lean_object* l_Lean_SubExpr_Pos_maxChildren; static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_4486____closed__13; 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*); @@ -205,6 +203,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; @@ -429,6 +428,7 @@ lean_object* l_Lean_KeyedDeclsAttribute_getValues___rarg(lean_object*, lean_obje 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_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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent___spec__1___boxed(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -11865,7 +11865,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) { @@ -11919,7 +11919,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); @@ -11949,7 +11949,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 354a23562b637..2a380f2f75d5e 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -347,7 +347,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*); @@ -499,7 +499,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*); @@ -515,7 +514,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*); @@ -726,6 +724,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*); @@ -1255,7 +1254,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; @@ -1599,6 +1598,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*); @@ -5855,6 +5855,7 @@ lean_inc(x_37); lean_dec(x_35); lean_inc(x_28); x_38 = l_Lean_mkPrivateName(x_37, x_28); +lean_dec(x_37); x_39 = lean_name_eq(x_11, x_38); lean_dec(x_38); if (x_39 == 0) @@ -23685,8 +23686,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; @@ -23726,8 +23726,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; @@ -47930,7 +47929,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) { @@ -47980,7 +47979,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); @@ -48010,7 +48009,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); @@ -48064,7 +48063,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); @@ -48083,7 +48083,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); @@ -53483,7 +53484,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) { @@ -53533,7 +53534,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); @@ -53563,7 +53564,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); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c index 196e9ebbe0295..a5f2740aee043 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; @@ -293,7 +293,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 b240407875304..880770f3c3f16 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 db5085170933e..c00c7a05415c1 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 a8332518c2b3d..6b63be702eb7a 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 1db627cdd7cf5..2301c19202242 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 3a74da7e77c8a..314d3c0a6dca9 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*); @@ -34,36 +28,25 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at_Lean_Envi LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_replay___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_isTodo___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3177_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedRecursors___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_getUsedConstantsAsSet(lean_object*); 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; 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*); @@ -75,10 +58,8 @@ 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*); @@ -87,58 +68,43 @@ lean_object* lean_st_ref_get(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*); +uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(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*); 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*); 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*); @@ -148,21 +114,18 @@ 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*); +uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Environment_Replay_replayConstant___closed__1; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_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*); -uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2780_(lean_object*, 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,610 +540,65 @@ lean_dec(x_2); return x_5; } } -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1() { +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_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: +lean_object* x_5; lean_object* x_6; lean_object* x_7; +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); +if (lean_obj_tag(x_7) == 0) { uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); +x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_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: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__12; -x_3 = l_Lean_NameSet_empty; -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__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_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; 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; uint8_t x_61; uint8_t x_62; -x_27 = lean_box(0); -x_28 = lean_box(0); -x_29 = lean_st_ref_get(x_3, x_4); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Environment_Replay_throwKernelException___closed__6; -x_34 = l_Lean_Environment_Replay_throwKernelException___closed__9; -x_35 = l_Lean_Environment_Replay_throwKernelException___closed__13; -x_36 = l_Lean_Environment_Replay_throwKernelException___closed__16; -x_37 = l_Lean_Environment_Replay_throwKernelException___closed__17; -x_38 = l_Lean_Environment_Replay_throwKernelException___closed__18; -x_39 = l_Lean_Environment_Replay_throwKernelException___closed__19; -x_40 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_40, 0, x_32); -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_io_get_num_heartbeats(x_31); -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_Environment_Replay_throwKernelException___closed__3; -x_45 = l_Lean_Environment_Replay_throwKernelException___closed__4; -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_unsigned_to_nat(1000u); -x_48 = lean_box(0); -x_49 = lean_box(0); -x_50 = l_Lean_Environment_Replay_throwKernelException___closed__5; -x_51 = l_Lean_firstFrontendMacroScope; -x_52 = 0; -x_53 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_53, 0, x_44); -lean_ctor_set(x_53, 1, x_45); -lean_ctor_set(x_53, 2, x_27); -lean_ctor_set(x_53, 3, x_46); -lean_ctor_set(x_53, 4, x_47); -lean_ctor_set(x_53, 5, x_48); -lean_ctor_set(x_53, 6, x_49); -lean_ctor_set(x_53, 7, x_27); -lean_ctor_set(x_53, 8, x_42); -lean_ctor_set(x_53, 9, x_50); -lean_ctor_set(x_53, 10, x_51); -lean_ctor_set(x_53, 11, x_28); -lean_ctor_set_uint8(x_53, sizeof(void*)*12, x_52); -lean_ctor_set_uint8(x_53, sizeof(void*)*12 + 1, x_52); -x_54 = lean_st_mk_ref(x_40, x_43); -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_st_ref_get(x_55, 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_58, 0); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Lean_Kernel_isDiagnosticsEnabled(x_60); -lean_dec(x_60); -if (x_61 == 0) -{ -uint8_t x_99; -x_99 = l_Lean_Environment_Replay_throwKernelException___closed__21; -if (x_99 == 0) -{ -uint8_t x_100; -x_100 = 1; -x_62 = x_100; -goto block_98; -} -else -{ -x_62 = x_52; -goto block_98; -} -} -else -{ -uint8_t x_101; -x_101 = l_Lean_Environment_Replay_throwKernelException___closed__21; -if (x_101 == 0) -{ -x_62 = x_52; -goto block_98; +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_102; -x_102 = 1; -x_62 = x_102; -goto block_98; -} +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; } -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_26: -{ -lean_object* x_13; lean_object* x_14; -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); -if (lean_obj_tag(x_14) == 0) -{ uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); +x_15 = !lean_is_exclusive(x_7); 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; -} -} -else -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_14); -if (x_22 == 0) -{ -x_5 = x_14; -goto block_10; -} -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); -x_5 = x_25; -goto block_10; -} -} -} -block_98: -{ -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_st_ref_take(x_55, x_59); -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_is_exclusive(x_64); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_67 = lean_ctor_get(x_64, 0); -x_68 = lean_ctor_get(x_64, 4); -lean_dec(x_68); -x_69 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_70 = l_Lean_Kernel_enableDiag(x_67, x_69); -lean_ctor_set(x_64, 4, x_36); -lean_ctor_set(x_64, 0, x_70); -x_71 = lean_st_ref_set(x_55, x_64, x_65); -x_72 = lean_ctor_get(x_71, 1); -lean_inc(x_72); -lean_dec(x_71); -x_73 = lean_box(0); -x_74 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_27, x_69, x_1, x_73, x_53, x_55, x_72); -lean_dec(x_55); -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_11 = x_75; -x_12 = x_76; -goto block_26; -} -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; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_77 = lean_ctor_get(x_64, 0); -x_78 = lean_ctor_get(x_64, 1); -x_79 = lean_ctor_get(x_64, 2); -x_80 = lean_ctor_get(x_64, 3); -x_81 = lean_ctor_get(x_64, 5); -x_82 = lean_ctor_get(x_64, 6); -x_83 = lean_ctor_get(x_64, 7); -lean_inc(x_83); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_64); -x_84 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_85 = l_Lean_Kernel_enableDiag(x_77, x_84); -x_86 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_78); -lean_ctor_set(x_86, 2, x_79); -lean_ctor_set(x_86, 3, x_80); -lean_ctor_set(x_86, 4, x_36); -lean_ctor_set(x_86, 5, x_81); -lean_ctor_set(x_86, 6, x_82); -lean_ctor_set(x_86, 7, x_83); -x_87 = lean_st_ref_set(x_55, x_86, x_65); -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = lean_box(0); -x_90 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_27, x_84, x_1, x_89, x_53, x_55, x_88); -lean_dec(x_55); -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_11 = x_91; -x_12 = x_92; -goto block_26; -} +return x_7; } else { -uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_93 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_94 = lean_box(0); -x_95 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_27, x_93, x_1, x_94, x_53, x_55, x_59); -lean_dec(x_55); -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_11 = x_96; -x_12 = x_97; -goto block_26; -} -} +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); +lean_dec(x_7); +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_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) { @@ -2180,7 +1598,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); @@ -2250,7 +1668,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) { @@ -3970,7 +3388,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; @@ -3993,7 +3411,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; } @@ -4043,7 +3461,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; } @@ -4077,6 +3495,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; } @@ -4129,7 +3555,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); @@ -4143,7 +3569,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); @@ -4182,7 +3608,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) { @@ -4195,7 +3621,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); @@ -4218,7 +3644,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_2780_(x_33, x_58); +x_59 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_33, x_58); lean_dec(x_58); lean_dec(x_33); if (x_59 == 0) @@ -4228,10 +3654,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); @@ -4258,7 +3684,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_2780_(x_33, x_69); +x_70 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_33, x_69); lean_dec(x_69); lean_dec(x_33); if (x_70 == 0) @@ -4268,10 +3694,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); @@ -4309,7 +3735,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); @@ -4327,7 +3753,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); @@ -4360,7 +3786,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) { @@ -4373,7 +3799,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); @@ -4399,7 +3825,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_2780_(x_99, x_124); +x_126 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_99, x_124); lean_dec(x_124); lean_dec(x_99); if (x_126 == 0) @@ -4409,10 +3835,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); @@ -4456,7 +3882,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); @@ -4488,7 +3914,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); @@ -4506,7 +3932,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); @@ -4528,7 +3954,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); @@ -4542,7 +3968,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); @@ -4586,7 +4012,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) { @@ -4599,7 +4025,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); @@ -4632,7 +4058,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_2780_(x_178, x_205); +x_207 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_178, x_205); lean_dec(x_205); lean_dec(x_178); if (x_207 == 0) @@ -4642,10 +4068,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); @@ -4689,7 +4115,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); @@ -4722,7 +4148,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); @@ -4762,7 +4188,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); @@ -4776,7 +4202,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); @@ -4824,7 +4250,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) { @@ -4837,7 +4263,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); @@ -4875,7 +4301,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_2780_(x_258, x_285); +x_287 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2906_(x_258, x_285); lean_dec(x_285); lean_dec(x_258); if (x_287 == 0) @@ -4885,10 +4311,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); @@ -4938,7 +4364,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); @@ -4976,7 +4402,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); @@ -5191,7 +4617,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); @@ -5205,7 +4631,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); @@ -5244,7 +4670,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) { @@ -5257,7 +4683,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); @@ -5280,7 +4706,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_3177_(x_33, x_58); +x_59 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_33, x_58); lean_dec(x_58); lean_dec(x_33); if (x_59 == 0) @@ -5293,7 +4719,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); @@ -5320,7 +4746,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_3177_(x_33, x_69); +x_70 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_33, x_69); lean_dec(x_69); lean_dec(x_33); if (x_70 == 0) @@ -5333,7 +4759,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); @@ -5371,7 +4797,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); @@ -5389,7 +4815,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); @@ -5422,7 +4848,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) { @@ -5435,7 +4861,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); @@ -5461,7 +4887,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_3177_(x_99, x_124); +x_126 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_99, x_124); lean_dec(x_124); lean_dec(x_99); if (x_126 == 0) @@ -5474,7 +4900,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); @@ -5518,7 +4944,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); @@ -5550,7 +4976,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); @@ -5568,7 +4994,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); @@ -5590,7 +5016,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); @@ -5604,7 +5030,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); @@ -5648,7 +5074,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) { @@ -5661,7 +5087,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); @@ -5694,7 +5120,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_3177_(x_178, x_205); +x_207 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_178, x_205); lean_dec(x_205); lean_dec(x_178); if (x_207 == 0) @@ -5707,7 +5133,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); @@ -5751,7 +5177,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); @@ -5784,7 +5210,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); @@ -5824,7 +5250,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); @@ -5838,7 +5264,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); @@ -5886,7 +5312,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) { @@ -5899,7 +5325,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); @@ -5937,7 +5363,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_3177_(x_258, x_285); +x_287 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3303_(x_258, x_285); lean_dec(x_285); lean_dec(x_258); if (x_287 == 0) @@ -5950,7 +5376,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); @@ -6000,7 +5426,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); @@ -6038,7 +5464,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); @@ -6522,49 +5948,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(); @@ -6595,6 +5978,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 6ce9eac03bbd2..35709df736b74 100644 --- a/stage0/stdlib/Lean/ResolveName.c +++ b/stage0/stdlib/Lean/ResolveName.c @@ -3222,7 +3222,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); @@ -3268,7 +3267,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); @@ -3342,7 +3340,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); @@ -3953,7 +3950,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/Completion/CompletionCollectors.c b/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c index ef96663dd8464..b98ed3f0a5a15 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); @@ -15524,7 +15523,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); @@ -15545,7 +15544,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); @@ -15579,7 +15578,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 b95076555333b..983636d5f2864 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; @@ -1811,7 +1811,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 334222cfdbe58..91bca48108eca 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/GoTo.c b/stage0/stdlib/Lean/Server/GoTo.c index 60fa1e80a30ce..131e07e2b8cfd 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 f9dc6204d4802..913bd515b7b21 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -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); @@ -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); diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index 4cc8b3337af02..5ca0c8166b09e 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 1e941b81077d7..9f7e8c1e6129b 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/Structure.c b/stage0/stdlib/Lean/Structure.c index 00debd671d389..20418a97a09dc 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*); @@ -332,6 +334,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*); @@ -3118,7 +3121,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; @@ -3161,7 +3164,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; @@ -3217,6 +3221,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: { @@ -4519,7 +4532,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; @@ -4608,6 +4621,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; } @@ -4647,7 +4661,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; @@ -4706,7 +4720,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; @@ -4800,12 +4815,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; @@ -4864,7 +4888,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; @@ -4929,6 +4954,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 cedbbc72b4d4a..4e80b96fc3af3 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 85ef253cd1c61..213e23c7a02ba 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; @@ -1096,7 +1096,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; @@ -1176,6 +1176,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/SafeExponentiation.c b/stage0/stdlib/Lean/Util/SafeExponentiation.c index 23d25ccc77ca8..f092f8ca6c2a2 100644 --- a/stage0/stdlib/Lean/Util/SafeExponentiation.c +++ b/stage0/stdlib/Lean/Util/SafeExponentiation.c @@ -14,7 +14,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_checkExponent___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_reportMessageKind(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_SafeExponentiation___hyg_5____closed__8; static lean_object* l_Lean_initFn____x40_Lean_Util_SafeExponentiation___hyg_5____closed__2; lean_object* l_Lean_log___at_Lean_Core_wrapAsyncAsSnapshot___spec__13(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -24,6 +23,7 @@ lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_checkExponent___closed__8; static lean_object* l_Lean_initFn____x40_Lean_Util_SafeExponentiation___hyg_5____closed__7; static lean_object* l_Lean_initFn____x40_Lean_Util_SafeExponentiation___hyg_5____closed__5; +lean_object* l_Lean_logMessageKind(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logWarning___at_Lean_checkExponent___spec__1___closed__1; 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*); static lean_object* l_Lean_checkExponent___closed__2; @@ -308,7 +308,7 @@ else { 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 = l_Lean_checkExponent___closed__3; -x_13 = l_Lean_reportMessageKind(x_12, x_2, x_3, x_4); +x_13 = l_Lean_logMessageKind(x_12, x_2, x_3, x_4); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index 001196586670e..625bbb8d74fd0 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -147,7 +147,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; @@ -513,7 +513,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*); @@ -14037,7 +14037,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); @@ -14263,7 +14264,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); @@ -14714,7 +14716,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; @@ -14763,7 +14766,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; @@ -15756,7 +15760,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); @@ -15938,7 +15943,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); @@ -18084,7 +18090,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; @@ -22115,9 +22122,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; @@ -22338,9 +22344,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;