diff --git a/stage0/src/kernel/instantiate.cpp b/stage0/src/kernel/instantiate.cpp index 48047433edc8..b845eda9129c 100644 --- a/stage0/src/kernel/instantiate.cpp +++ b/stage0/src/kernel/instantiate.cpp @@ -165,22 +165,38 @@ bool is_head_beta(expr const & t) { return is_app(t) && is_lambda(get_app_fn(t)); } -expr apply_beta(expr f, unsigned num_args, expr const * args) { - if (num_args == 0) { - return f; - } else if (!is_lambda(f)) { - return mk_rev_app(f, num_args, args); - } else { - unsigned m = 1; - while (is_lambda(binding_body(f)) && m < num_args) { - f = binding_body(f); - m++; +static expr apply_beta_rec(expr e, unsigned i, unsigned num_rev_args, expr const * rev_args, bool preserve_data, bool zeta) { + if (is_lambda(e)) { + if (i + 1 < num_rev_args) { + return apply_beta_rec(binding_body(e), i+1, num_rev_args, rev_args, preserve_data, zeta); + } else { + return instantiate(binding_body(e), num_rev_args, rev_args); + } + } else if (is_let(e)) { + if (zeta && i < num_rev_args) { + return apply_beta_rec(instantiate(let_body(e), let_value(e)), i, num_rev_args, rev_args, preserve_data, zeta); + } else { + unsigned n = num_rev_args - i; + return mk_rev_app(instantiate(e, i, rev_args + n), n, rev_args); + } + } else if (is_mdata(e)) { + if (preserve_data) { + unsigned n = num_rev_args - i; + return mk_rev_app(instantiate(e, i, rev_args + n), n, rev_args); + } else { + return apply_beta_rec(mdata_expr(e), i, num_rev_args, rev_args, preserve_data, zeta); } - lean_assert(m <= num_args); - return mk_rev_app(instantiate(binding_body(f), m, args + (num_args - m)), num_args - m, args); + } else { + unsigned n = num_rev_args - i; + return mk_rev_app(instantiate(e, i, rev_args + n), n, rev_args); } } +expr apply_beta(expr f, unsigned num_rev_args, expr const * rev_args, bool preserve_data, bool zeta) { + if (num_rev_args == 0) return f; + return apply_beta_rec(f, 0, num_rev_args, rev_args, preserve_data, zeta); +} + expr head_beta_reduce(expr const & t) { if (!is_head_beta(t)) { return t; diff --git a/stage0/src/kernel/instantiate.h b/stage0/src/kernel/instantiate.h index ac07bdde7e41..7ef2bd67ab2f 100644 --- a/stage0/src/kernel/instantiate.h +++ b/stage0/src/kernel/instantiate.h @@ -24,7 +24,7 @@ inline expr instantiate_rev(expr const & e, buffer const & s) { return instantiate_rev(e, s.size(), s.data()); } -expr apply_beta(expr f, unsigned num_rev_args, expr const * rev_args); +expr apply_beta(expr f, unsigned num_rev_args, expr const * rev_args, bool preserve_data = true, bool zeta = false); bool is_head_beta(expr const & t); expr head_beta_reduce(expr const & t); /* If `e` is of the form `(fun x, t) a` return `head_beta_const_fn(t)` if `t` does not depend on `x`, diff --git a/stage0/src/kernel/instantiate_mvars.cpp b/stage0/src/kernel/instantiate_mvars.cpp index 21ead4284057..b47f91da7db7 100644 --- a/stage0/src/kernel/instantiate_mvars.cpp +++ b/stage0/src/kernel/instantiate_mvars.cpp @@ -4,10 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura */ +#include #include +#include "util/name_set.h" #include "runtime/option_ref.h" +#include "runtime/array_ref.h" #include "kernel/instantiate.h" -#include "kernel/abstract.h" +#include "kernel/replace_fn.h" /* This module is not used by the kernel. It just provides an efficient implementation of @@ -15,7 +18,6 @@ This module is not used by the kernel. It just provides an efficient implementat */ namespace lean { - extern "C" object * lean_get_lmvar_assignment(obj_arg mctx, obj_arg mid); extern "C" object * lean_assign_lmvar(obj_arg mctx, obj_arg mid, obj_arg val); @@ -29,18 +31,19 @@ option_ref get_lmvar_assignment(metavar_ctx & mctx, name const & mid) { return option_ref(lean_get_lmvar_assignment(mctx.to_obj_arg(), mid.to_obj_arg())); } -class instantiate_lmvar_fn { +class instantiate_lmvars_fn { metavar_ctx & m_mctx; - std::unordered_map m_cache; + std::unordered_map m_cache; + std::vector m_saved; // Helper vector to prevent values from being garbagge collected - inline level cache(level const & l, level && r, bool shared) { + inline level cache(level const & l, level r, bool shared) { if (shared) { - m_cache.insert(mk_pair(l.raw(), r.raw())); + m_cache.insert(mk_pair(l.raw(), r)); } return r; } public: - instantiate_lmvar_fn(metavar_ctx & mctx):m_mctx(mctx) {} + instantiate_lmvars_fn(metavar_ctx & mctx):m_mctx(mctx) {} level visit(level const & l) { if (!has_mvar(l)) return l; @@ -48,7 +51,7 @@ class instantiate_lmvar_fn { if (is_shared(l)) { auto it = m_cache.find(l.raw()); if (it != m_cache.end()) { - return level(it->second, true); + return it->second; } shared = true; } @@ -70,6 +73,12 @@ class instantiate_lmvar_fn { } else { level a_new = visit(a); if (!is_eqp(a, a_new)) { + /* + We save `a` to ensure it will not be garbage collected + after we update `mctx`. This is necessary because `m_cache` + may contain references to its subterms. + */ + m_saved.push_back(a); assign_lmvar(m_mctx, mvar_id(l), a_new); } return a_new; @@ -82,14 +91,274 @@ class instantiate_lmvar_fn { extern "C" LEAN_EXPORT object * lean_instantiate_level_mvars(object * m, object * l) { metavar_ctx mctx(m); - level l_new = instantiate_lmvar_fn(mctx)(level(l)); + level l_new = instantiate_lmvars_fn(mctx)(level(l)); object * r = alloc_cnstr(0, 2, 0); cnstr_set(r, 0, mctx.steal()); cnstr_set(r, 1, l_new.steal()); return r; } -extern "C" LEAN_EXPORT object * lean_instantiate_expr_mvars(object *, object *) { - lean_internal_panic("not implemented yet"); +extern "C" object * lean_get_mvar_assignment(obj_arg mctx, obj_arg mid); +extern "C" object * lean_get_delayed_mvar_assignment(obj_arg mctx, obj_arg mid); +extern "C" object * lean_assign_mvar(obj_arg mctx, obj_arg mid, obj_arg val); +typedef object_ref delayed_assignment; + +void assign_mvar(metavar_ctx & mctx, name const & mid, expr const & e) { + object * r = lean_assign_mvar(mctx.steal(), mid.to_obj_arg(), e.to_obj_arg()); + mctx.set_box(r); +} + +option_ref get_mvar_assignment(metavar_ctx & mctx, name const & mid) { + return option_ref(lean_get_mvar_assignment(mctx.to_obj_arg(), mid.to_obj_arg())); +} + +option_ref get_delayed_mvar_assignment(metavar_ctx & mctx, name const & mid) { + return option_ref(lean_get_delayed_mvar_assignment(mctx.to_obj_arg(), mid.to_obj_arg())); +} + +expr replace_fvars(expr const & e, array_ref const & fvars, expr const * rev_args) { + size_t sz = fvars.size(); + if (sz == 0) + return e; + return replace(e, [=](expr const & m, unsigned offset) -> optional { + if (!has_fvar(m)) + return some_expr(m); // expression m does not contain free variables + if (is_fvar(m)) { + size_t i = sz; + name const & fid = fvar_name(m); + while (i > 0) { + --i; + if (fvar_name(fvars[i]) == fid) { + return some_expr(lift_loose_bvars(rev_args[sz - i - 1], offset)); + } + } + } + return none_expr(); + }); +} + +class instantiate_mvars_fn { + metavar_ctx & m_mctx; + instantiate_lmvars_fn m_level_fn; + name_set m_already_normalized; // Store metavariables whose assignment has already been normalized. + std::unordered_map m_cache; + std::vector m_saved; // Helper vector to prevent values from being garbagge collected + + level visit_level(level const & l) { + return m_level_fn(l); + } + + levels visit_levels(levels const & ls) { + buffer lsNew; + for (auto const & l : ls) + lsNew.push_back(visit_level(l)); + return levels(lsNew); + } + + inline expr cache(expr const & e, expr r, bool shared) { + if (shared) { + m_cache.insert(mk_pair(e.raw(), r)); + } + return r; + } + + optional get_assignment(name const & mid) { + option_ref r = get_mvar_assignment(m_mctx, mid); + if (!r) { + return optional(); + } else { + expr a(r.get_val()); + if (!has_mvar(a) || m_already_normalized.contains(mid)) { + return optional(a); + } else { + m_already_normalized.insert(mid); + expr a_new = visit(a); + if (!is_eqp(a, a_new)) { + /* + We save `a` to ensure it will not be garbage collected + after we update `mctx`. This is necessary because `m_cache` + may contain references to its subterms. + */ + m_saved.push_back(a); + assign_mvar(m_mctx, mid, a_new); + } + return optional(a_new); + } + } + } + + /* + Given `e` of the form `f a_1 ... a_n` where `f` is not a metavariable, + instantiate metavariables. + */ + expr visit_app_default(expr const & e) { + if (is_app(e)) { + return update_app(e, visit_app_default(app_fn(e)), visit(app_arg(e))); + } else { + lean_assert(!is_mvar(e)); + return visit(e); + } + } + + /* + Given `e` of the form `?m a_1 ... a_n`, return new application where + the metavariables in the arguments `a_i` have been instantiated. + */ + expr visit_mvar_app_args(expr const & e) { + if (is_app(e)) { + return update_app(e, visit_app_default(app_fn(e)), visit(app_arg(e))); + } else { + lean_assert(is_mvar(e)); + return e; + } + } + + /* + Given `e` of the form `f a_1 ... a_n`, return new application `f_new a_1' ... a_n'` + where `a_i'` is `visit(a_i)`. `args` is an accumulator for the new arguments. + */ + expr visit_args_and_beta(expr const & f_new, expr const & e, buffer & args) { + if (is_app(e)) { + args.push_back(visit(app_arg(e))); + return visit_args_and_beta(f_new, app_fn(e), args); + } else { + /* + Some of the arguments in `args` are irrelevant after we beta + reduce. Also, it may be a bug to not instantiate them, since they + may depend on free variables that are not in the context (see + issue #4375). So we pass `useZeta := true` to ensure that they are + instantiated. + */ + bool preserve_data = false; + bool zeta = true; + return apply_beta(f_new, args.size(), args.data(), preserve_data, zeta); + } + } + + /* + Helper function for delayed assignment case at `visit_app`. + `e` is a term of the form `?m t1 t2 t3` + Moreover, `?m` is delayed assigned + `?m #[x, y] := g x y` + where, `fvars := #[x, y]` and `val := g x y`. + `args` is an accumulator for `e`'s arguments. + + We want to return `g t1' t2' t3'` where + `ti'`s are `visit(ti)`. + */ + expr visit_delayed(array_ref const & fvars, expr const & val, expr const & e, buffer & args) { + if (is_app(e)) { + args.push_back(visit(app_arg(e))); + return visit_delayed(fvars, val, app_fn(e), args); + } else { + expr val_new = replace_fvars(val, fvars, args.data() + (args.size() - fvars.size())); + return mk_rev_app(val_new, args.size() - fvars.size(), args.data()); + } + } + + expr visit_app(expr const & e) { + expr const & f = get_app_fn(e); + if (!is_mvar(f)) { + return visit_app_default(e); + } else { + name const & mid = mvar_name(f); + option_ref d = get_delayed_mvar_assignment(m_mctx, mid); + if (!d) { + // mvar is not delayed assigned + expr f_new = visit(f); + if (is_eqp(f, f_new)) { + return visit_mvar_app_args(e); + } else { + buffer args; + return visit_args_and_beta(f_new, e, args); + } + } else { + /* + Apply "delayed substitution" (i.e., delayed assignment + application). + That is, `f` is some metavariable `?m`, that is delayed assigned to `val`. + If after instantiating `val`, we obtain `newVal`, and `newVal` does not contain + metavariables, we replace the free variables `fvars` in `newVal` with the first + `fvars.size` elements of `args`. + */ + array_ref fvars(cnstr_get(d.get_val().raw(), 0), true); + name mid_pending(cnstr_get(d.get_val().raw(), 1), true); + if (fvars.size() > get_app_num_args(e)) { + /* + We don't have sufficient arguments for instantiating the free variables `fvars`. + This can only happen if a tactic or elaboration function is not implemented correctly. + We decided to not use `panic!` here and report it as an error in the frontend + when we are checking for unassigned metavariables in an elaborated term. */ + return visit_mvar_app_args(e); + } + optional val = get_assignment(mid_pending); + if (!val) + // mid_pending has not been assigned yet. + return visit_mvar_app_args(e); + if (has_expr_mvar(*val)) + // mid_pending has been assigned, but assignment contains mvars. + return visit_mvar_app_args(e); + buffer args; + return visit_delayed(fvars, *val, e, args); + } + } + } + + expr visit_mvar(expr const & e) { + name const & mid = mvar_name(e); + if (auto r = get_assignment(mid)) { + return *r; + } else { + return e; + } + } + +public: + instantiate_mvars_fn(metavar_ctx & mctx):m_mctx(mctx), m_level_fn(mctx) {} + + expr visit(expr const & e) { + if (!has_mvar(e)) + return e; + bool shared = false; + if (is_shared(e)) { + auto it = m_cache.find(e.raw()); + if (it != m_cache.end()) { + return it->second; + } + shared = true; + } + + switch (e.kind()) { + case expr_kind::BVar: + case expr_kind::Lit: case expr_kind::FVar: + lean_unreachable(); + case expr_kind::Sort: + return cache(e, update_sort(e, visit_level(sort_level(e))), shared); + case expr_kind::Const: + return cache(e, update_const(e, visit_levels(const_levels(e))), shared); + case expr_kind::MVar: + return visit_mvar(e); + case expr_kind::MData: + return cache(e, update_mdata(e, visit(mdata_expr(e))), shared); + case expr_kind::Proj: + return cache(e, update_proj(e, visit(proj_expr(e))), shared); + case expr_kind::App: + return cache(e, visit_app(e), shared); + case expr_kind::Pi: case expr_kind::Lambda: + return cache(e, update_binding(e, visit(binding_domain(e)), visit(binding_body(e))), shared); + case expr_kind::Let: + return cache(e, update_let(e, visit(let_type(e)), visit(let_value(e)), visit(let_body(e))), shared); + } + } + + expr operator()(expr const & e) { return visit(e); } +}; + +extern "C" LEAN_EXPORT object * lean_instantiate_expr_mvars(object * m, object * e) { + metavar_ctx mctx(m); + expr e_new = instantiate_mvars_fn(mctx)(expr(e)); + object * r = alloc_cnstr(0, 2, 0); + cnstr_set(r, 0, mctx.steal()); + cnstr_set(r, 1, e_new.steal()); + return r; } } diff --git a/stage0/src/runtime/thread.cpp b/stage0/src/runtime/thread.cpp index 6aab3fbc6d17..cd53bae6adfd 100644 --- a/stage0/src/runtime/thread.cpp +++ b/stage0/src/runtime/thread.cpp @@ -9,6 +9,9 @@ Author: Leonardo de Moura #include #ifdef LEAN_WINDOWS #include +# ifdef LEAN_AUTO_THREAD_FINALIZATION +#include +# endif #else #include #endif diff --git a/stage0/src/util/shell.cpp b/stage0/src/util/shell.cpp index 779c682b8976..6310c9cc81c0 100644 --- a/stage0/src/util/shell.cpp +++ b/stage0/src/util/shell.cpp @@ -191,43 +191,44 @@ static void display_features(std::ostream & out) { static void display_help(std::ostream & out) { display_header(out); std::cout << "Miscellaneous:\n"; - std::cout << " --help -h display this message\n"; - std::cout << " --features -f display features compiler provides (eg. LLVM support)\n"; - std::cout << " --version -v display version number\n"; - std::cout << " --githash display the git commit hash number used to build this binary\n"; - std::cout << " --run call the 'main' definition in a file with the remaining arguments\n"; - std::cout << " --o=oname -o create olean file\n"; - std::cout << " --i=iname -i create ilean file\n"; - std::cout << " --c=fname -c name of the C output file\n"; - std::cout << " --bc=fname -b name of the LLVM bitcode file\n"; - std::cout << " --stdin take input from stdin\n"; - std::cout << " --root=dir set package root directory from which the module name of the input file is calculated\n" - << " (default: current working directory)\n"; - std::cout << " --trust=num -t trust level (default: max) 0 means do not trust any macro,\n" - << " and type check all imported modules\n"; - std::cout << " --quiet -q do not print verbose messages\n"; - std::cout << " --memory=num -M maximum amount of memory that should be used by Lean\n"; - std::cout << " (in megabytes)\n"; - std::cout << " --timeout=num -T maximum number of memory allocations per task\n"; - std::cout << " this is a deterministic way of interrupting long running tasks\n"; + std::cout << " -h, --help display this message\n"; + std::cout << " -f, --features display features compiler provides (eg. LLVM support)\n"; + std::cout << " -v, --version display version number\n"; + std::cout << " --githash display the git commit hash number used to build this binary\n"; + std::cout << " --run call the 'main' definition in a file with the remaining arguments\n"; + std::cout << " -o, --o=oname create olean file\n"; + std::cout << " -i, --i=iname create ilean file\n"; + std::cout << " -c, --c=fname name of the C output file\n"; + std::cout << " -b, --bc=fname name of the LLVM bitcode file\n"; + std::cout << " --stdin take input from stdin\n"; + std::cout << " --root=dir set package root directory from which the module name\n" + << " of the input file is calculated\n" + << " (default: current working directory)\n"; + std::cout << " -t, --trust=num trust level (default: max) 0 means do not trust any macro,\n" + << " and type check all imported modules\n"; + std::cout << " -q, --quiet do not print verbose messages\n"; + std::cout << " -M, --memory=num maximum amount of memory that should be used by Lean\n"; + std::cout << " (in megabytes)\n"; + std::cout << " -T, --timeout=num maximum number of memory allocations per task\n"; + std::cout << " this is a deterministic way of interrupting long running tasks\n"; #if defined(LEAN_MULTI_THREAD) - std::cout << " --threads=num -j number of threads used to process lean files\n"; - std::cout << " --tstack=num -s thread stack size in Kb\n"; - std::cout << " --server start lean in server mode\n"; - std::cout << " --worker start lean in server-worker mode\n"; + std::cout << " -j, --threads=num number of threads used to process lean files\n"; + std::cout << " -s, --tstack=num thread stack size in Kb\n"; + std::cout << " --server start lean in server mode\n"; + std::cout << " --worker start lean in server-worker mode\n"; #endif - std::cout << " --plugin=file load and initialize Lean shared library for registering linters etc.\n"; - std::cout << " --load-dynlib=file load shared library to make its symbols available to the interpreter\n"; - std::cout << " --json report Lean output (e.g., messages) as JSON (one per line)\n"; - std::cout << " --deps just print dependencies of a Lean input\n"; - std::cout << " --print-prefix print the installation prefix for Lean and exit\n"; - std::cout << " --print-libdir print the installation directory for Lean's built-in libraries and exit\n"; - std::cout << " --profile display elaboration/type checking time for each definition/theorem\n"; - std::cout << " --stats display environment statistics\n"; + std::cout << " --plugin=file load and initialize Lean shared library for registering linters etc.\n"; + std::cout << " --load-dynlib=file load shared library to make its symbols available to the interpreter\n"; + std::cout << " --json report Lean output (e.g., messages) as JSON (one per line)\n"; + std::cout << " --deps just print dependencies of a Lean input\n"; + std::cout << " --print-prefix print the installation prefix for Lean and exit\n"; + std::cout << " --print-libdir print the installation directory for Lean's built-in libraries and exit\n"; + std::cout << " --profile display elaboration/type checking time for each definition/theorem\n"; + std::cout << " --stats display environment statistics\n"; DEBUG_CODE( - std::cout << " --debug=tag enable assertions with the given tag\n"; + std::cout << " --debug=tag enable assertions with the given tag\n"; ) - std::cout << " -D name=value set a configuration option (see set_option command)\n"; + std::cout << " -D name=value set a configuration option (see set_option command)\n"; } static int print_prefix = 0; diff --git a/stage0/stdlib/Lake/Load/Lean/Elab.c b/stage0/stdlib/Lake/Load/Lean/Elab.c index 4c7ff7b965f9..661c36201d2e 100644 --- a/stage0/stdlib/Lake/Load/Lean/Elab.c +++ b/stage0/stdlib/Lake/Load/Lean/Elab.c @@ -2203,85 +2203,83 @@ return x_1; LEAN_EXPORT lean_object* l_Lake_elabConfigFile(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_79; lean_object* x_80; lean_object* x_318; -x_318 = l_IO_FS_readFile(x_4, x_6); -if (lean_obj_tag(x_318) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_73; lean_object* x_74; lean_object* x_280; +x_280 = l_IO_FS_readFile(x_4, x_6); +if (lean_obj_tag(x_280) == 0) { -uint8_t x_319; -x_319 = !lean_is_exclusive(x_318); -if (x_319 == 0) +uint8_t x_281; +x_281 = !lean_is_exclusive(x_280); +if (x_281 == 0) { -lean_object* x_320; -x_320 = lean_ctor_get(x_318, 1); -lean_ctor_set(x_318, 1, x_5); -x_79 = x_318; -x_80 = x_320; -goto block_317; +lean_object* x_282; +x_282 = lean_ctor_get(x_280, 1); +lean_ctor_set(x_280, 1, x_5); +x_73 = x_280; +x_74 = x_282; +goto block_279; } else { -lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_321 = lean_ctor_get(x_318, 0); -x_322 = lean_ctor_get(x_318, 1); -lean_inc(x_322); -lean_inc(x_321); -lean_dec(x_318); -x_323 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_323, 0, x_321); -lean_ctor_set(x_323, 1, x_5); -x_79 = x_323; -x_80 = x_322; -goto block_317; -} -} -else -{ -uint8_t x_324; -x_324 = !lean_is_exclusive(x_318); -if (x_324 == 0) -{ -lean_object* x_325; lean_object* x_326; lean_object* x_327; uint8_t x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; -x_325 = lean_ctor_get(x_318, 0); -x_326 = lean_ctor_get(x_318, 1); -x_327 = lean_io_error_to_string(x_325); -x_328 = 3; -x_329 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_329, 0, x_327); -lean_ctor_set_uint8(x_329, sizeof(void*)*1, x_328); -x_330 = lean_array_get_size(x_5); -x_331 = lean_array_push(x_5, x_329); -lean_ctor_set(x_318, 1, x_331); -lean_ctor_set(x_318, 0, x_330); -x_79 = x_318; -x_80 = x_326; -goto block_317; -} -else -{ -lean_object* x_332; lean_object* x_333; lean_object* x_334; uint8_t x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -x_332 = lean_ctor_get(x_318, 0); -x_333 = lean_ctor_get(x_318, 1); -lean_inc(x_333); -lean_inc(x_332); -lean_dec(x_318); -x_334 = lean_io_error_to_string(x_332); -x_335 = 3; -x_336 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_336, 0, x_334); -lean_ctor_set_uint8(x_336, sizeof(void*)*1, x_335); -x_337 = lean_array_get_size(x_5); -x_338 = lean_array_push(x_5, x_336); -x_339 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_339, 0, x_337); -lean_ctor_set(x_339, 1, x_338); -x_79 = x_339; -x_80 = x_333; -goto block_317; +lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_283 = lean_ctor_get(x_280, 0); +x_284 = lean_ctor_get(x_280, 1); +lean_inc(x_284); +lean_inc(x_283); +lean_dec(x_280); +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_283); +lean_ctor_set(x_285, 1, x_5); +x_73 = x_285; +x_74 = x_284; +goto block_279; } } -block_78: +else { -if (lean_obj_tag(x_7) == 0) +uint8_t x_286; +x_286 = !lean_is_exclusive(x_280); +if (x_286 == 0) +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_287 = lean_ctor_get(x_280, 0); +x_288 = lean_ctor_get(x_280, 1); +x_289 = lean_io_error_to_string(x_287); +x_290 = 3; +x_291 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set_uint8(x_291, sizeof(void*)*1, x_290); +x_292 = lean_array_get_size(x_5); +x_293 = lean_array_push(x_5, x_291); +lean_ctor_set(x_280, 1, x_293); +lean_ctor_set(x_280, 0, x_292); +x_73 = x_280; +x_74 = x_288; +goto block_279; +} +else +{ +lean_object* x_294; lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_294 = lean_ctor_get(x_280, 0); +x_295 = lean_ctor_get(x_280, 1); +lean_inc(x_295); +lean_inc(x_294); +lean_dec(x_280); +x_296 = lean_io_error_to_string(x_294); +x_297 = 3; +x_298 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_298, 0, x_296); +lean_ctor_set_uint8(x_298, sizeof(void*)*1, x_297); +x_299 = lean_array_get_size(x_5); +x_300 = lean_array_push(x_5, x_298); +x_301 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +x_73 = x_301; +x_74 = x_295; +goto block_279; +} +} +block_72: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_9 = lean_ctor_get(x_7, 0); @@ -2516,873 +2514,727 @@ return x_71; } } } -else +block_279: { -uint8_t x_72; -lean_dec(x_4); -x_72 = !lean_is_exclusive(x_7); -if (x_72 == 0) +if (lean_obj_tag(x_73) == 0) { -lean_object* x_73; -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_7); -lean_ctor_set(x_73, 1, x_8); -return x_73; -} -else +uint8_t x_75; +x_75 = !lean_is_exclusive(x_73); +if (x_75 == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_74 = lean_ctor_get(x_7, 0); -x_75 = lean_ctor_get(x_7, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_7); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_8); -return x_77; -} -} -} -block_317: -{ -if (lean_obj_tag(x_79) == 0) -{ -uint8_t x_81; -x_81 = !lean_is_exclusive(x_79); -if (x_81 == 0) -{ -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_223; -x_82 = lean_ctor_get(x_79, 0); -x_83 = lean_ctor_get(x_79, 1); -x_84 = 1; +lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_193; +x_76 = lean_ctor_get(x_73, 0); +x_77 = lean_ctor_get(x_73, 1); +x_78 = 1; lean_inc(x_4); -x_85 = l_Lean_Parser_mkInputContext(x_82, x_4, x_84); -lean_inc(x_85); -x_223 = l_Lean_Parser_parseHeader(x_85, x_80); -if (lean_obj_tag(x_223) == 0) -{ -lean_object* x_224; lean_object* x_225; -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); -lean_ctor_set(x_79, 0, x_224); -x_86 = x_79; -x_87 = x_225; -goto block_222; -} -else -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; uint8_t x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_226 = lean_ctor_get(x_223, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_223, 1); -lean_inc(x_227); -lean_dec(x_223); -x_228 = lean_io_error_to_string(x_226); -x_229 = 3; -x_230 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_230, 0, x_228); -lean_ctor_set_uint8(x_230, sizeof(void*)*1, x_229); -x_231 = lean_array_get_size(x_83); -x_232 = lean_array_push(x_83, x_230); -lean_ctor_set_tag(x_79, 1); -lean_ctor_set(x_79, 1, x_232); -lean_ctor_set(x_79, 0, x_231); -x_86 = x_79; -x_87 = x_227; -goto block_222; -} -block_222: -{ -if (lean_obj_tag(x_86) == 0) -{ -lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_88 = lean_ctor_get(x_86, 0); +x_79 = l_Lean_Parser_mkInputContext(x_76, x_4, x_78); +lean_inc(x_79); +x_193 = l_Lean_Parser_parseHeader(x_79, x_74); +if (lean_obj_tag(x_193) == 0) +{ +lean_object* x_194; lean_object* x_195; +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_193, 1); +lean_inc(x_195); +lean_dec(x_193); +lean_ctor_set(x_73, 0, x_194); +x_80 = x_73; +x_81 = x_195; +goto block_192; +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_196 = lean_ctor_get(x_193, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_193, 1); +lean_inc(x_197); +lean_dec(x_193); +x_198 = lean_io_error_to_string(x_196); +x_199 = 3; +x_200 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_200, 0, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*1, x_199); +x_201 = lean_array_get_size(x_77); +x_202 = lean_array_push(x_77, x_200); +lean_ctor_set_tag(x_73, 1); +lean_ctor_set(x_73, 1, x_202); +lean_ctor_set(x_73, 0, x_201); +x_80 = x_73; +x_81 = x_197; +goto block_192; +} +block_192: +{ +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_82 = lean_ctor_get(x_80, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +x_84 = !lean_is_exclusive(x_80); +if (x_84 == 0) +{ +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_133; +x_85 = lean_ctor_get(x_80, 1); +x_86 = lean_ctor_get(x_80, 0); +lean_dec(x_86); +x_87 = lean_ctor_get(x_82, 0); +lean_inc(x_87); +lean_dec(x_82); +x_88 = lean_ctor_get(x_83, 0); lean_inc(x_88); -x_89 = lean_ctor_get(x_88, 1); +x_89 = lean_ctor_get(x_83, 1); lean_inc(x_89); -x_90 = !lean_is_exclusive(x_86); -if (x_90 == 0) -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_155; -x_91 = lean_ctor_get(x_86, 1); -x_92 = lean_ctor_get(x_86, 0); -lean_dec(x_92); -x_93 = lean_ctor_get(x_88, 0); -lean_inc(x_93); -lean_dec(x_88); -x_94 = lean_ctor_get(x_89, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_89, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_96 = x_89; +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_90 = x_83; } else { - lean_dec_ref(x_89); - x_96 = lean_box(0); + lean_dec_ref(x_83); + x_90 = lean_box(0); } -lean_inc(x_85); +lean_inc(x_79); lean_inc(x_3); -x_155 = l_Lake_processHeader(x_93, x_3, x_85, x_95, x_87); -lean_dec(x_93); -if (lean_obj_tag(x_155) == 0) -{ -lean_object* x_156; lean_object* x_157; -x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_155, 1); -lean_inc(x_157); -lean_dec(x_155); -lean_ctor_set(x_86, 0, x_156); -x_97 = x_86; -x_98 = x_157; -goto block_154; -} -else -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_158 = lean_ctor_get(x_155, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_155, 1); -lean_inc(x_159); -lean_dec(x_155); -x_160 = lean_io_error_to_string(x_158); -x_161 = 3; -x_162 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set_uint8(x_162, sizeof(void*)*1, x_161); -x_163 = lean_array_get_size(x_91); -x_164 = lean_array_push(x_91, x_162); -lean_ctor_set_tag(x_86, 1); -lean_ctor_set(x_86, 1, x_164); -lean_ctor_set(x_86, 0, x_163); -x_97 = x_86; -x_98 = x_159; -goto block_154; -} -block_154: -{ -if (lean_obj_tag(x_97) == 0) +x_133 = l_Lake_processHeader(x_87, x_3, x_79, x_89, x_81); +lean_dec(x_87); +if (lean_obj_tag(x_133) == 0) { -uint8_t x_99; -lean_dec(x_96); -x_99 = !lean_is_exclusive(x_97); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_100 = lean_ctor_get(x_97, 0); -x_101 = lean_ctor_get(x_97, 1); -x_102 = lean_ctor_get(x_100, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_100, 1); -lean_inc(x_103); -lean_dec(x_100); -x_104 = l_Lake_configModuleName; -x_105 = lean_environment_set_main_module(x_102, x_104); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_1); -x_107 = l_Lake_elabConfigFile___closed__2; -x_108 = l_Lean_EnvExtension_setState___rarg(x_107, x_105, x_106); -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_2); -x_110 = l_Lake_elabConfigFile___closed__3; -x_111 = l_Lean_EnvExtension_setState___rarg(x_110, x_108, x_109); -x_112 = l_Lean_Elab_Command_mkState(x_111, x_103, x_3); -x_113 = l_Lean_Elab_IO_processCommands(x_85, x_94, x_112, x_98); -if (lean_obj_tag(x_113) == 0) -{ -lean_object* x_114; lean_object* x_115; -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -lean_ctor_set(x_97, 0, x_114); -x_7 = x_97; -x_8 = x_115; -goto block_78; -} -else -{ -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_116 = lean_ctor_get(x_113, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_113, 1); -lean_inc(x_117); -lean_dec(x_113); -x_118 = lean_io_error_to_string(x_116); -x_119 = 3; -x_120 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set_uint8(x_120, sizeof(void*)*1, x_119); -x_121 = lean_array_get_size(x_101); -x_122 = lean_array_push(x_101, x_120); -lean_ctor_set_tag(x_97, 1); -lean_ctor_set(x_97, 1, x_122); -lean_ctor_set(x_97, 0, x_121); -x_7 = x_97; -x_8 = x_117; -goto block_78; -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_123 = lean_ctor_get(x_97, 0); -x_124 = lean_ctor_get(x_97, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_97); -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -lean_dec(x_123); -x_127 = l_Lake_configModuleName; -x_128 = lean_environment_set_main_module(x_125, x_127); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_1); -x_130 = l_Lake_elabConfigFile___closed__2; -x_131 = l_Lean_EnvExtension_setState___rarg(x_130, x_128, x_129); -x_132 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_132, 0, x_2); -x_133 = l_Lake_elabConfigFile___closed__3; -x_134 = l_Lean_EnvExtension_setState___rarg(x_133, x_131, x_132); -x_135 = l_Lean_Elab_Command_mkState(x_134, x_126, x_3); -x_136 = l_Lean_Elab_IO_processCommands(x_85, x_94, x_135, x_98); -if (lean_obj_tag(x_136) == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_124); -x_7 = x_139; -x_8 = x_138; -goto block_78; +lean_object* x_134; lean_object* x_135; +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); +lean_inc(x_135); +lean_dec(x_133); +lean_ctor_set(x_80, 0, x_134); +x_91 = x_80; +x_92 = x_135; +goto block_132; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_140 = lean_ctor_get(x_136, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_136, 1); -lean_inc(x_141); -lean_dec(x_136); -x_142 = lean_io_error_to_string(x_140); -x_143 = 3; -x_144 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set_uint8(x_144, sizeof(void*)*1, x_143); -x_145 = lean_array_get_size(x_124); -x_146 = lean_array_push(x_124, x_144); -x_147 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -x_7 = x_147; -x_8 = x_141; -goto block_78; +lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_136 = lean_ctor_get(x_133, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_133, 1); +lean_inc(x_137); +lean_dec(x_133); +x_138 = lean_io_error_to_string(x_136); +x_139 = 3; +x_140 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set_uint8(x_140, sizeof(void*)*1, x_139); +x_141 = lean_array_get_size(x_85); +x_142 = lean_array_push(x_85, x_140); +lean_ctor_set_tag(x_80, 1); +lean_ctor_set(x_80, 1, x_142); +lean_ctor_set(x_80, 0, x_141); +x_91 = x_80; +x_92 = x_137; +goto block_132; } +block_132: +{ +if (lean_obj_tag(x_91) == 0) +{ +uint8_t x_93; +lean_dec(x_90); +x_93 = !lean_is_exclusive(x_91); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_94 = lean_ctor_get(x_91, 0); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +lean_dec(x_94); +x_97 = l_Lake_configModuleName; +x_98 = lean_environment_set_main_module(x_95, x_97); +x_99 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_99, 0, x_1); +x_100 = l_Lake_elabConfigFile___closed__2; +x_101 = l_Lean_EnvExtension_setState___rarg(x_100, x_98, x_99); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_2); +x_103 = l_Lake_elabConfigFile___closed__3; +x_104 = l_Lean_EnvExtension_setState___rarg(x_103, x_101, x_102); +x_105 = l_Lean_Elab_Command_mkState(x_104, x_96, x_3); +x_106 = l_Lean_Elab_IO_processCommands(x_79, x_88, x_105, x_92); +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); +lean_ctor_set(x_91, 0, x_107); +x_7 = x_91; +x_8 = x_108; +goto block_72; +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_109 = lean_ctor_get(x_91, 0); +x_110 = lean_ctor_get(x_91, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_91); +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_109, 1); +lean_inc(x_112); +lean_dec(x_109); +x_113 = l_Lake_configModuleName; +x_114 = lean_environment_set_main_module(x_111, x_113); +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_1); +x_116 = l_Lake_elabConfigFile___closed__2; +x_117 = l_Lean_EnvExtension_setState___rarg(x_116, x_114, x_115); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_2); +x_119 = l_Lake_elabConfigFile___closed__3; +x_120 = l_Lean_EnvExtension_setState___rarg(x_119, x_117, x_118); +x_121 = l_Lean_Elab_Command_mkState(x_120, x_112, x_3); +x_122 = l_Lean_Elab_IO_processCommands(x_79, x_88, x_121, x_92); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_110); +x_7 = x_125; +x_8 = x_124; +goto block_72; } } else { -uint8_t x_148; -lean_dec(x_94); -lean_dec(x_85); +uint8_t x_126; +lean_dec(x_88); +lean_dec(x_79); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_148 = !lean_is_exclusive(x_97); -if (x_148 == 0) +x_126 = !lean_is_exclusive(x_91); +if (x_126 == 0) { -lean_object* x_149; -if (lean_is_scalar(x_96)) { - x_149 = lean_alloc_ctor(0, 2, 0); +lean_object* x_127; +if (lean_is_scalar(x_90)) { + x_127 = lean_alloc_ctor(0, 2, 0); } else { - x_149 = x_96; + x_127 = x_90; } -lean_ctor_set(x_149, 0, x_97); -lean_ctor_set(x_149, 1, x_98); -return x_149; +lean_ctor_set(x_127, 0, x_91); +lean_ctor_set(x_127, 1, x_92); +return x_127; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_150 = lean_ctor_get(x_97, 0); -x_151 = lean_ctor_get(x_97, 1); -lean_inc(x_151); -lean_inc(x_150); -lean_dec(x_97); -x_152 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -if (lean_is_scalar(x_96)) { - x_153 = lean_alloc_ctor(0, 2, 0); +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_91, 0); +x_129 = lean_ctor_get(x_91, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_91); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +if (lean_is_scalar(x_90)) { + x_131 = lean_alloc_ctor(0, 2, 0); } else { - x_153 = x_96; + x_131 = x_90; } -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_98); -return x_153; +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_92); +return x_131; } } } } else { -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_204; -x_165 = lean_ctor_get(x_86, 1); -lean_inc(x_165); -lean_dec(x_86); -x_166 = lean_ctor_get(x_88, 0); -lean_inc(x_166); -lean_dec(x_88); -x_167 = lean_ctor_get(x_89, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_89, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_169 = x_89; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_174; +x_143 = lean_ctor_get(x_80, 1); +lean_inc(x_143); +lean_dec(x_80); +x_144 = lean_ctor_get(x_82, 0); +lean_inc(x_144); +lean_dec(x_82); +x_145 = lean_ctor_get(x_83, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_83, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_147 = x_83; } else { - lean_dec_ref(x_89); - x_169 = lean_box(0); + lean_dec_ref(x_83); + x_147 = lean_box(0); } -lean_inc(x_85); +lean_inc(x_79); lean_inc(x_3); -x_204 = l_Lake_processHeader(x_166, x_3, x_85, x_168, x_87); -lean_dec(x_166); -if (lean_obj_tag(x_204) == 0) -{ -lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_205 = lean_ctor_get(x_204, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_204, 1); -lean_inc(x_206); -lean_dec(x_204); -x_207 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_207, 0, x_205); -lean_ctor_set(x_207, 1, x_165); -x_170 = x_207; -x_171 = x_206; -goto block_203; -} -else +x_174 = l_Lake_processHeader(x_144, x_3, x_79, x_146, x_81); +lean_dec(x_144); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_208; lean_object* x_209; lean_object* x_210; uint8_t x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_208 = lean_ctor_get(x_204, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_204, 1); -lean_inc(x_209); -lean_dec(x_204); -x_210 = lean_io_error_to_string(x_208); -x_211 = 3; -x_212 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set_uint8(x_212, sizeof(void*)*1, x_211); -x_213 = lean_array_get_size(x_165); -x_214 = lean_array_push(x_165, x_212); -x_215 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_215, 0, x_213); -lean_ctor_set(x_215, 1, x_214); -x_170 = x_215; -x_171 = x_209; -goto block_203; -} -block_203: -{ -if (lean_obj_tag(x_170) == 0) -{ -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_169); -x_172 = lean_ctor_get(x_170, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_170, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_174 = x_170; -} else { - lean_dec_ref(x_170); - x_174 = lean_box(0); -} -x_175 = lean_ctor_get(x_172, 0); +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = lean_ctor_get(x_174, 0); lean_inc(x_175); -x_176 = lean_ctor_get(x_172, 1); +x_176 = lean_ctor_get(x_174, 1); lean_inc(x_176); -lean_dec(x_172); -x_177 = l_Lake_configModuleName; -x_178 = lean_environment_set_main_module(x_175, x_177); -x_179 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_179, 0, x_1); -x_180 = l_Lake_elabConfigFile___closed__2; -x_181 = l_Lean_EnvExtension_setState___rarg(x_180, x_178, x_179); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_2); -x_183 = l_Lake_elabConfigFile___closed__3; -x_184 = l_Lean_EnvExtension_setState___rarg(x_183, x_181, x_182); -x_185 = l_Lean_Elab_Command_mkState(x_184, x_176, x_3); -x_186 = l_Lean_Elab_IO_processCommands(x_85, x_167, x_185, x_171); -if (lean_obj_tag(x_186) == 0) -{ -lean_object* x_187; lean_object* x_188; lean_object* x_189; -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); -if (lean_is_scalar(x_174)) { - x_189 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_174); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set(x_177, 1, x_143); +x_148 = x_177; +x_149 = x_176; +goto block_173; +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_178 = lean_ctor_get(x_174, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_174, 1); +lean_inc(x_179); +lean_dec(x_174); +x_180 = lean_io_error_to_string(x_178); +x_181 = 3; +x_182 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set_uint8(x_182, sizeof(void*)*1, x_181); +x_183 = lean_array_get_size(x_143); +x_184 = lean_array_push(x_143, x_182); +x_185 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +x_148 = x_185; +x_149 = x_179; +goto block_173; +} +block_173: +{ +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_147); +x_150 = lean_ctor_get(x_148, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_148, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_152 = x_148; } else { - x_189 = x_174; -} -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_173); -x_7 = x_189; -x_8 = x_188; -goto block_78; + lean_dec_ref(x_148); + x_152 = lean_box(0); } -else -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_190 = lean_ctor_get(x_186, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_186, 1); -lean_inc(x_191); -lean_dec(x_186); -x_192 = lean_io_error_to_string(x_190); -x_193 = 3; -x_194 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set_uint8(x_194, sizeof(void*)*1, x_193); -x_195 = lean_array_get_size(x_173); -x_196 = lean_array_push(x_173, x_194); -if (lean_is_scalar(x_174)) { - x_197 = lean_alloc_ctor(1, 2, 0); +x_153 = lean_ctor_get(x_150, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_150, 1); +lean_inc(x_154); +lean_dec(x_150); +x_155 = l_Lake_configModuleName; +x_156 = lean_environment_set_main_module(x_153, x_155); +x_157 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_157, 0, x_1); +x_158 = l_Lake_elabConfigFile___closed__2; +x_159 = l_Lean_EnvExtension_setState___rarg(x_158, x_156, x_157); +x_160 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_160, 0, x_2); +x_161 = l_Lake_elabConfigFile___closed__3; +x_162 = l_Lean_EnvExtension_setState___rarg(x_161, x_159, x_160); +x_163 = l_Lean_Elab_Command_mkState(x_162, x_154, x_3); +x_164 = l_Lean_Elab_IO_processCommands(x_79, x_145, x_163, x_149); +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); +if (lean_is_scalar(x_152)) { + x_167 = lean_alloc_ctor(0, 2, 0); } else { - x_197 = x_174; - lean_ctor_set_tag(x_197, 1); -} -lean_ctor_set(x_197, 0, x_195); -lean_ctor_set(x_197, 1, x_196); -x_7 = x_197; -x_8 = x_191; -goto block_78; + x_167 = x_152; } +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_151); +x_7 = x_167; +x_8 = x_166; +goto block_72; } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_167); -lean_dec(x_85); +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_dec(x_145); +lean_dec(x_79); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_198 = lean_ctor_get(x_170, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_170, 1); -lean_inc(x_199); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_200 = x_170; +x_168 = lean_ctor_get(x_148, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_148, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_170 = x_148; } else { - lean_dec_ref(x_170); - x_200 = lean_box(0); + lean_dec_ref(x_148); + x_170 = lean_box(0); } -if (lean_is_scalar(x_200)) { - x_201 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_170)) { + x_171 = lean_alloc_ctor(1, 2, 0); } else { - x_201 = x_200; + x_171 = x_170; } -lean_ctor_set(x_201, 0, x_198); -lean_ctor_set(x_201, 1, x_199); -if (lean_is_scalar(x_169)) { - x_202 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_168); +lean_ctor_set(x_171, 1, x_169); +if (lean_is_scalar(x_147)) { + x_172 = lean_alloc_ctor(0, 2, 0); } else { - x_202 = x_169; + x_172 = x_147; } -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_171); -return x_202; +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_172, 1, x_149); +return x_172; } } } } else { -uint8_t x_216; -lean_dec(x_85); +uint8_t x_186; +lean_dec(x_79); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_216 = !lean_is_exclusive(x_86); -if (x_216 == 0) +x_186 = !lean_is_exclusive(x_80); +if (x_186 == 0) { -lean_object* x_217; -x_217 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_217, 0, x_86); -lean_ctor_set(x_217, 1, x_87); -return x_217; +lean_object* x_187; +x_187 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_187, 0, x_80); +lean_ctor_set(x_187, 1, x_81); +return x_187; } else { -lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_218 = lean_ctor_get(x_86, 0); -x_219 = lean_ctor_get(x_86, 1); -lean_inc(x_219); -lean_inc(x_218); -lean_dec(x_86); -x_220 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set(x_220, 1, x_219); -x_221 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_87); -return x_221; +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_188 = lean_ctor_get(x_80, 0); +x_189 = lean_ctor_get(x_80, 1); +lean_inc(x_189); +lean_inc(x_188); +lean_dec(x_80); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_81); +return x_191; } } } } else { -lean_object* x_233; lean_object* x_234; uint8_t x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_299; -x_233 = lean_ctor_get(x_79, 0); -x_234 = lean_ctor_get(x_79, 1); -lean_inc(x_234); -lean_inc(x_233); -lean_dec(x_79); -x_235 = 1; +lean_object* x_203; lean_object* x_204; uint8_t x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_261; +x_203 = lean_ctor_get(x_73, 0); +x_204 = lean_ctor_get(x_73, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_73); +x_205 = 1; lean_inc(x_4); -x_236 = l_Lean_Parser_mkInputContext(x_233, x_4, x_235); -lean_inc(x_236); -x_299 = l_Lean_Parser_parseHeader(x_236, x_80); -if (lean_obj_tag(x_299) == 0) -{ -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_299, 0); -lean_inc(x_300); -x_301 = lean_ctor_get(x_299, 1); -lean_inc(x_301); -lean_dec(x_299); -x_302 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_234); -x_237 = x_302; -x_238 = x_301; -goto block_298; -} -else -{ -lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -x_303 = lean_ctor_get(x_299, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_299, 1); -lean_inc(x_304); -lean_dec(x_299); -x_305 = lean_io_error_to_string(x_303); -x_306 = 3; -x_307 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_307, 0, x_305); -lean_ctor_set_uint8(x_307, sizeof(void*)*1, x_306); -x_308 = lean_array_get_size(x_234); -x_309 = lean_array_push(x_234, x_307); -x_310 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_310, 0, x_308); -lean_ctor_set(x_310, 1, x_309); -x_237 = x_310; -x_238 = x_304; -goto block_298; -} -block_298: -{ -if (lean_obj_tag(x_237) == 0) -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_281; -x_239 = lean_ctor_get(x_237, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_239, 1); -lean_inc(x_240); -x_241 = lean_ctor_get(x_237, 1); -lean_inc(x_241); -if (lean_is_exclusive(x_237)) { - lean_ctor_release(x_237, 0); - lean_ctor_release(x_237, 1); - x_242 = x_237; -} else { - lean_dec_ref(x_237); - x_242 = lean_box(0); -} -x_243 = lean_ctor_get(x_239, 0); -lean_inc(x_243); -lean_dec(x_239); -x_244 = lean_ctor_get(x_240, 0); -lean_inc(x_244); -x_245 = lean_ctor_get(x_240, 1); -lean_inc(x_245); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_246 = x_240; +x_206 = l_Lean_Parser_mkInputContext(x_203, x_4, x_205); +lean_inc(x_206); +x_261 = l_Lean_Parser_parseHeader(x_206, x_74); +if (lean_obj_tag(x_261) == 0) +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_262 = lean_ctor_get(x_261, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_261, 1); +lean_inc(x_263); +lean_dec(x_261); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_262); +lean_ctor_set(x_264, 1, x_204); +x_207 = x_264; +x_208 = x_263; +goto block_260; +} +else +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; uint8_t x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_265 = lean_ctor_get(x_261, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_261, 1); +lean_inc(x_266); +lean_dec(x_261); +x_267 = lean_io_error_to_string(x_265); +x_268 = 3; +x_269 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_269, 0, x_267); +lean_ctor_set_uint8(x_269, sizeof(void*)*1, x_268); +x_270 = lean_array_get_size(x_204); +x_271 = lean_array_push(x_204, x_269); +x_272 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_272, 0, x_270); +lean_ctor_set(x_272, 1, x_271); +x_207 = x_272; +x_208 = x_266; +goto block_260; +} +block_260: +{ +if (lean_obj_tag(x_207) == 0) +{ +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_243; +x_209 = lean_ctor_get(x_207, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_209, 1); +lean_inc(x_210); +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_212 = x_207; +} else { + lean_dec_ref(x_207); + x_212 = lean_box(0); +} +x_213 = lean_ctor_get(x_209, 0); +lean_inc(x_213); +lean_dec(x_209); +x_214 = lean_ctor_get(x_210, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_210, 1); +lean_inc(x_215); +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + x_216 = x_210; } else { - lean_dec_ref(x_240); - x_246 = lean_box(0); + lean_dec_ref(x_210); + x_216 = lean_box(0); } -lean_inc(x_236); +lean_inc(x_206); lean_inc(x_3); -x_281 = l_Lake_processHeader(x_243, x_3, x_236, x_245, x_238); -lean_dec(x_243); -if (lean_obj_tag(x_281) == 0) +x_243 = l_Lake_processHeader(x_213, x_3, x_206, x_215, x_208); +lean_dec(x_213); +if (lean_obj_tag(x_243) == 0) { -lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -lean_dec(x_281); -if (lean_is_scalar(x_242)) { - x_284 = lean_alloc_ctor(0, 2, 0); -} else { - x_284 = x_242; -} -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set(x_284, 1, x_241); -x_247 = x_284; -x_248 = x_283; -goto block_280; -} -else -{ -lean_object* x_285; lean_object* x_286; lean_object* x_287; uint8_t x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; -x_285 = lean_ctor_get(x_281, 0); -lean_inc(x_285); -x_286 = lean_ctor_get(x_281, 1); -lean_inc(x_286); -lean_dec(x_281); -x_287 = lean_io_error_to_string(x_285); -x_288 = 3; -x_289 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_289, 0, x_287); -lean_ctor_set_uint8(x_289, sizeof(void*)*1, x_288); -x_290 = lean_array_get_size(x_241); -x_291 = lean_array_push(x_241, x_289); -if (lean_is_scalar(x_242)) { - x_292 = lean_alloc_ctor(1, 2, 0); -} else { - x_292 = x_242; - lean_ctor_set_tag(x_292, 1); -} -lean_ctor_set(x_292, 0, x_290); -lean_ctor_set(x_292, 1, x_291); -x_247 = x_292; -x_248 = x_286; -goto block_280; -} -block_280: -{ -if (lean_obj_tag(x_247) == 0) -{ -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -lean_dec(x_246); -x_249 = lean_ctor_get(x_247, 0); -lean_inc(x_249); -x_250 = lean_ctor_get(x_247, 1); -lean_inc(x_250); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_251 = x_247; -} else { - lean_dec_ref(x_247); - x_251 = lean_box(0); -} -x_252 = lean_ctor_get(x_249, 0); -lean_inc(x_252); -x_253 = lean_ctor_get(x_249, 1); -lean_inc(x_253); -lean_dec(x_249); -x_254 = l_Lake_configModuleName; -x_255 = lean_environment_set_main_module(x_252, x_254); -x_256 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_256, 0, x_1); -x_257 = l_Lake_elabConfigFile___closed__2; -x_258 = l_Lean_EnvExtension_setState___rarg(x_257, x_255, x_256); -x_259 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_259, 0, x_2); -x_260 = l_Lake_elabConfigFile___closed__3; -x_261 = l_Lean_EnvExtension_setState___rarg(x_260, x_258, x_259); -x_262 = l_Lean_Elab_Command_mkState(x_261, x_253, x_3); -x_263 = l_Lean_Elab_IO_processCommands(x_236, x_244, x_262, x_248); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -lean_dec(x_263); -if (lean_is_scalar(x_251)) { - x_266 = lean_alloc_ctor(0, 2, 0); +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_243, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +lean_dec(x_243); +if (lean_is_scalar(x_212)) { + x_246 = lean_alloc_ctor(0, 2, 0); } else { - x_266 = x_251; + x_246 = x_212; } -lean_ctor_set(x_266, 0, x_264); -lean_ctor_set(x_266, 1, x_250); -x_7 = x_266; -x_8 = x_265; -goto block_78; +lean_ctor_set(x_246, 0, x_244); +lean_ctor_set(x_246, 1, x_211); +x_217 = x_246; +x_218 = x_245; +goto block_242; } else { -lean_object* x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_267 = lean_ctor_get(x_263, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_263, 1); -lean_inc(x_268); -lean_dec(x_263); -x_269 = lean_io_error_to_string(x_267); -x_270 = 3; -x_271 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_271, 0, x_269); -lean_ctor_set_uint8(x_271, sizeof(void*)*1, x_270); -x_272 = lean_array_get_size(x_250); -x_273 = lean_array_push(x_250, x_271); -if (lean_is_scalar(x_251)) { - x_274 = lean_alloc_ctor(1, 2, 0); +lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +x_247 = lean_ctor_get(x_243, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_243, 1); +lean_inc(x_248); +lean_dec(x_243); +x_249 = lean_io_error_to_string(x_247); +x_250 = 3; +x_251 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_251, 0, x_249); +lean_ctor_set_uint8(x_251, sizeof(void*)*1, x_250); +x_252 = lean_array_get_size(x_211); +x_253 = lean_array_push(x_211, x_251); +if (lean_is_scalar(x_212)) { + x_254 = lean_alloc_ctor(1, 2, 0); +} else { + x_254 = x_212; + lean_ctor_set_tag(x_254, 1); +} +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_217 = x_254; +x_218 = x_248; +goto block_242; +} +block_242: +{ +if (lean_obj_tag(x_217) == 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_dec(x_216); +x_219 = lean_ctor_get(x_217, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_217, 1); +lean_inc(x_220); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_221 = x_217; +} else { + lean_dec_ref(x_217); + x_221 = lean_box(0); +} +x_222 = lean_ctor_get(x_219, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_219, 1); +lean_inc(x_223); +lean_dec(x_219); +x_224 = l_Lake_configModuleName; +x_225 = lean_environment_set_main_module(x_222, x_224); +x_226 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_226, 0, x_1); +x_227 = l_Lake_elabConfigFile___closed__2; +x_228 = l_Lean_EnvExtension_setState___rarg(x_227, x_225, x_226); +x_229 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_229, 0, x_2); +x_230 = l_Lake_elabConfigFile___closed__3; +x_231 = l_Lean_EnvExtension_setState___rarg(x_230, x_228, x_229); +x_232 = l_Lean_Elab_Command_mkState(x_231, x_223, x_3); +x_233 = l_Lean_Elab_IO_processCommands(x_206, x_214, x_232, x_218); +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_221)) { + x_236 = lean_alloc_ctor(0, 2, 0); } else { - x_274 = x_251; - lean_ctor_set_tag(x_274, 1); -} -lean_ctor_set(x_274, 0, x_272); -lean_ctor_set(x_274, 1, x_273); -x_7 = x_274; -x_8 = x_268; -goto block_78; + x_236 = x_221; } +lean_ctor_set(x_236, 0, x_234); +lean_ctor_set(x_236, 1, x_220); +x_7 = x_236; +x_8 = x_235; +goto block_72; } else { -lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; -lean_dec(x_244); -lean_dec(x_236); +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_214); +lean_dec(x_206); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_275 = lean_ctor_get(x_247, 0); -lean_inc(x_275); -x_276 = lean_ctor_get(x_247, 1); -lean_inc(x_276); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_277 = x_247; +x_237 = lean_ctor_get(x_217, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_217, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_239 = x_217; } else { - lean_dec_ref(x_247); - x_277 = lean_box(0); + lean_dec_ref(x_217); + x_239 = lean_box(0); } -if (lean_is_scalar(x_277)) { - x_278 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); } else { - x_278 = x_277; + x_240 = x_239; } -lean_ctor_set(x_278, 0, x_275); -lean_ctor_set(x_278, 1, x_276); -if (lean_is_scalar(x_246)) { - x_279 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +if (lean_is_scalar(x_216)) { + x_241 = lean_alloc_ctor(0, 2, 0); } else { - x_279 = x_246; + x_241 = x_216; } -lean_ctor_set(x_279, 0, x_278); -lean_ctor_set(x_279, 1, x_248); -return x_279; +lean_ctor_set(x_241, 0, x_240); +lean_ctor_set(x_241, 1, x_218); +return x_241; } } } else { -lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; -lean_dec(x_236); +lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +lean_dec(x_206); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_293 = lean_ctor_get(x_237, 0); -lean_inc(x_293); -x_294 = lean_ctor_get(x_237, 1); -lean_inc(x_294); -if (lean_is_exclusive(x_237)) { - lean_ctor_release(x_237, 0); - lean_ctor_release(x_237, 1); - x_295 = x_237; +x_255 = lean_ctor_get(x_207, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_207, 1); +lean_inc(x_256); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_257 = x_207; } else { - lean_dec_ref(x_237); - x_295 = lean_box(0); + lean_dec_ref(x_207); + x_257 = lean_box(0); } -if (lean_is_scalar(x_295)) { - x_296 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_257)) { + x_258 = lean_alloc_ctor(1, 2, 0); } else { - x_296 = x_295; + x_258 = x_257; } -lean_ctor_set(x_296, 0, x_293); -lean_ctor_set(x_296, 1, x_294); -x_297 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_297, 0, x_296); -lean_ctor_set(x_297, 1, x_238); -return x_297; +lean_ctor_set(x_258, 0, x_255); +lean_ctor_set(x_258, 1, x_256); +x_259 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_259, 0, x_258); +lean_ctor_set(x_259, 1, x_208); +return x_259; } } } } else { -uint8_t x_311; +uint8_t x_273; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_311 = !lean_is_exclusive(x_79); -if (x_311 == 0) +x_273 = !lean_is_exclusive(x_73); +if (x_273 == 0) { -lean_object* x_312; -x_312 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_312, 0, x_79); -lean_ctor_set(x_312, 1, x_80); -return x_312; +lean_object* x_274; +x_274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_274, 0, x_73); +lean_ctor_set(x_274, 1, x_74); +return x_274; } else { -lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -x_313 = lean_ctor_get(x_79, 0); -x_314 = lean_ctor_get(x_79, 1); -lean_inc(x_314); -lean_inc(x_313); -lean_dec(x_79); -x_315 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_315, 0, x_313); -lean_ctor_set(x_315, 1, x_314); -x_316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_316, 0, x_315); -lean_ctor_set(x_316, 1, x_80); -return x_316; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_275 = lean_ctor_get(x_73, 0); +x_276 = lean_ctor_get(x_73, 1); +lean_inc(x_276); +lean_inc(x_275); +lean_dec(x_73); +x_277 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_74); +return x_278; } } } diff --git a/stage0/stdlib/Lean/Elab/Frontend.c b/stage0/stdlib/Lean/Elab/Frontend.c index 2e1823b7b7dd..9f235b001e8d 100644 --- a/stage0/stdlib/Lean/Elab/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Frontend.c @@ -13,10 +13,10 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Lean_Language_Lean_process(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_profileit(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommands___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setMessages___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Firefox_Profile_export(lean_object*, double, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_compress(lean_object*); @@ -26,44 +26,35 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Elab_IO_processCommandsInc LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommandsIncrementally(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getCommandState___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2(lean_object*, lean_object*, lean_object*, double, double, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); double lean_float_div(double, double); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Elab_process___closed__1; -lean_object* l_String_toNat_x3f(lean_object*); lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3; -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__1; +lean_object* l_Lean_Language_SnapshotTree_forM___at_Lean_Language_SnapshotTree_runAndReport___spec__1(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__2; uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setCommandState___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2; +lean_object* l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommands(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___closed__1; -lean_object* l_Lean_Elab_processHeader(lean_object*, lean_object*, lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_runFrontend___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos___boxed(lean_object*); -extern lean_object* l_Lean_maxRecDepth; -lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -uint8_t lean_string_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Language_SnapshotTree_getAll(lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM(lean_object*); static double l_Lean_Elab_runFrontend___closed__1; -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__6; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getCommandState___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); @@ -74,104 +65,84 @@ lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_Lean_processCommands(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_runFrontend___closed__3; -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; -lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2(size_t, size_t, lean_object*); -lean_object* l_Lean_Language_reportMessages(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; static lean_object* l_Lean_Elab_process___closed__2; -lean_object* l_Lean_MessageData_ofFormat(lean_object*); -lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setParserState___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_mkInputContext(lean_object*, lean_object*, uint8_t); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__7; -lean_object* l_Lean_getOptionDecls(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5(lean_object*, size_t, size_t, lean_object*); -lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_State_commands___default___closed__1; lean_object* l_Array_toPArray_x27___rarg(lean_object*); -lean_object* l_Lean_Name_getRoot(lean_object*); lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_set_main_module(lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; +lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_name_eq(lean_object*, lean_object*); -lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2; -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_task_get_own(lean_object*); lean_object* l___private_Lean_Util_Profiler_0__Lean_Firefox_toJsonProfile____x40_Lean_Util_Profiler___hyg_3604_(lean_object*); +lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_Server_ModuleRefs_toLspModuleRefs(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__1; -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___closed__2; LEAN_EXPORT lean_object* lean_run_frontend(lean_object*, lean_object*, lean_object*, lean_object*, uint32_t, lean_object*, uint8_t, lean_object*); lean_object* l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setMessages(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_KVMap_erase(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_NameSet_empty; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getCommandState(lean_object*); static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getInputContext(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommands(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go(lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; +lean_object* l_Lean_Language_Lean_waitForFinalCmdState_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommandsIncrementally_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_runFrontend___lambda__3___closed__1; static lean_object* l_Lean_Elab_runFrontend___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_State_commands___default; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommand(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__4(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1; extern lean_object* l_Lean_Elab_Command_instInhabitedScope; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setCommandState(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_process(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos___rarg(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +lean_object* l_Lean_Language_SnapshotTask_map___rarg(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_MessageLog_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getInputContext___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); lean_object* l_Lean_Language_SnapshotTask_get___rarg(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___rarg___boxed(lean_object*, lean_object*); lean_object* l_IO_FS_writeFile(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__5; +extern lean_object* l_Lean_Language_Lean_internal_minimalSnapshots; lean_object* lean_string_append(lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler_output; lean_object* lean_array_get_size(lean_object*); -lean_object* l_Lean_Parser_parseHeader(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setParserState(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -lean_object* l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(lean_object*, lean_object*); +lean_object* lean_mk_empty_environment(uint32_t, lean_object*); static lean_object* l_Lean_Elab_Frontend_processCommand___closed__1; -lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___boxed(lean_object*); lean_object* l_Lean_Parser_parseCommand(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_runFrontend___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_reparseOptions(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_updateCmdPos(lean_object*); static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; static lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__2; @@ -2065,77 +2036,6 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommands(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; -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -x_6 = l_Lean_Elab_Frontend_State_commands___default___closed__1; -x_7 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_2); -lean_ctor_set(x_7, 2, x_5); -lean_ctor_set(x_7, 3, x_6); -x_8 = lean_st_mk_ref(x_7, x_4); -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_Elab_Frontend_processCommands(x_1, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_9, x_12); -lean_dec(x_9); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -uint8_t x_18; -lean_dec(x_9); -x_18 = !lean_is_exclusive(x_11); -if (x_18 == 0) -{ -return x_11; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { @@ -2288,7 +2188,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommandsIncrementally_go(lean_obj _start: { lean_object* x_6; uint8_t x_7; -x_6 = l_Lean_Language_SnapshotTask_get___rarg(x_3); +x_6 = lean_task_get_own(x_3); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { @@ -2651,227 +2551,233 @@ return x_6; } else { -lean_object* x_97; +lean_object* x_97; lean_object* x_98; lean_free_object(x_6); lean_dec(x_8); x_97 = lean_ctor_get(x_9, 0); lean_inc(x_97); lean_dec(x_9); -x_3 = x_97; +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_3 = x_98; x_4 = x_11; goto _start; } } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_99 = lean_ctor_get(x_6, 0); -x_100 = lean_ctor_get(x_6, 1); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_100 = lean_ctor_get(x_6, 0); +x_101 = lean_ctor_get(x_6, 1); +lean_inc(x_101); lean_inc(x_100); -lean_inc(x_99); lean_dec(x_6); -x_101 = lean_ctor_get(x_99, 1); -lean_inc(x_101); -x_102 = lean_array_push(x_4, x_101); -if (lean_obj_tag(x_100) == 0) +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +x_103 = lean_array_push(x_4, x_102); +if (lean_obj_tag(x_101) == 0) { -lean_object* x_103; lean_object* x_104; size_t x_105; size_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_object* x_104; lean_object* x_105; size_t x_106; size_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_inc(x_2); -x_103 = l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go(x_2); -x_104 = l_Lean_Language_SnapshotTree_getAll(x_103); -x_105 = lean_array_size(x_104); -x_106 = 0; -lean_inc(x_104); -x_107 = l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__1(x_105, x_106, x_104); -x_108 = lean_array_get_size(x_107); -x_109 = lean_unsigned_to_nat(0u); -x_110 = lean_nat_dec_lt(x_109, x_108); -x_111 = l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2(x_105, x_106, x_104); -x_112 = lean_array_get_size(x_111); -x_113 = l_Array_filterMapM___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__3(x_111, x_109, x_112); -lean_dec(x_112); -lean_dec(x_111); -x_114 = l_Array_toPArray_x27___rarg(x_113); +x_104 = l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go(x_2); +x_105 = l_Lean_Language_SnapshotTree_getAll(x_104); +x_106 = lean_array_size(x_105); +x_107 = 0; +lean_inc(x_105); +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__1(x_106, x_107, x_105); +x_109 = lean_array_get_size(x_108); +x_110 = lean_unsigned_to_nat(0u); +x_111 = lean_nat_dec_lt(x_110, x_109); +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__2(x_106, x_107, x_105); +x_113 = lean_array_get_size(x_112); +x_114 = l_Array_filterMapM___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__3(x_112, x_110, x_113); lean_dec(x_113); -x_115 = lean_ctor_get(x_99, 4); -lean_inc(x_115); -x_116 = l_Lean_Language_SnapshotTask_get___rarg(x_115); -x_117 = lean_ctor_get(x_116, 1); -lean_inc(x_117); -lean_dec(x_116); -x_118 = lean_ctor_get(x_117, 6); +lean_dec(x_112); +x_115 = l_Array_toPArray_x27___rarg(x_114); +lean_dec(x_114); +x_116 = lean_ctor_get(x_100, 4); +lean_inc(x_116); +x_117 = l_Lean_Language_SnapshotTask_get___rarg(x_116); +x_118 = lean_ctor_get(x_117, 1); lean_inc(x_118); -x_119 = lean_ctor_get(x_117, 0); +lean_dec(x_117); +x_119 = lean_ctor_get(x_118, 6); lean_inc(x_119); -x_120 = lean_ctor_get(x_117, 2); +x_120 = lean_ctor_get(x_118, 0); lean_inc(x_120); -x_121 = lean_ctor_get(x_117, 3); +x_121 = lean_ctor_get(x_118, 2); lean_inc(x_121); -x_122 = lean_ctor_get(x_117, 4); +x_122 = lean_ctor_get(x_118, 3); lean_inc(x_122); -x_123 = lean_ctor_get(x_117, 5); +x_123 = lean_ctor_get(x_118, 4); lean_inc(x_123); -x_124 = lean_ctor_get(x_117, 7); +x_124 = lean_ctor_get(x_118, 5); lean_inc(x_124); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - lean_ctor_release(x_117, 2); - lean_ctor_release(x_117, 3); - lean_ctor_release(x_117, 4); - lean_ctor_release(x_117, 5); - lean_ctor_release(x_117, 6); - lean_ctor_release(x_117, 7); - x_125 = x_117; -} else { - lean_dec_ref(x_117); - x_125 = lean_box(0); -} -x_126 = lean_ctor_get_uint8(x_118, sizeof(void*)*2); -x_127 = lean_ctor_get(x_118, 0); -lean_inc(x_127); +x_125 = lean_ctor_get(x_118, 7); +lean_inc(x_125); if (lean_is_exclusive(x_118)) { lean_ctor_release(x_118, 0); lean_ctor_release(x_118, 1); - x_128 = x_118; + lean_ctor_release(x_118, 2); + lean_ctor_release(x_118, 3); + lean_ctor_release(x_118, 4); + lean_ctor_release(x_118, 5); + lean_ctor_release(x_118, 6); + lean_ctor_release(x_118, 7); + x_126 = x_118; } else { lean_dec_ref(x_118); - x_128 = lean_box(0); + x_126 = lean_box(0); +} +x_127 = lean_ctor_get_uint8(x_119, sizeof(void*)*2); +x_128 = lean_ctor_get(x_119, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_129 = x_119; +} else { + lean_dec_ref(x_119); + x_129 = lean_box(0); } -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(0, 2, 1); +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(0, 2, 1); } else { - x_129 = x_128; + x_130 = x_129; } -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_114); -lean_ctor_set_uint8(x_129, sizeof(void*)*2, x_126); -if (x_110 == 0) +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_115); +lean_ctor_set_uint8(x_130, sizeof(void*)*2, x_127); +if (x_111 == 0) { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_109); lean_dec(x_108); -lean_dec(x_107); -x_130 = lean_ctor_get(x_99, 2); -lean_inc(x_130); -lean_dec(x_99); -x_131 = lean_ctor_get(x_130, 0); +x_131 = lean_ctor_get(x_100, 2); lean_inc(x_131); -x_132 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; -if (lean_is_scalar(x_125)) { - x_133 = lean_alloc_ctor(0, 8, 0); +lean_dec(x_100); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; +if (lean_is_scalar(x_126)) { + x_134 = lean_alloc_ctor(0, 8, 0); } else { - x_133 = x_125; -} -lean_ctor_set(x_133, 0, x_119); -lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_120); -lean_ctor_set(x_133, 3, x_121); -lean_ctor_set(x_133, 4, x_122); -lean_ctor_set(x_133, 5, x_123); -lean_ctor_set(x_133, 6, x_129); -lean_ctor_set(x_133, 7, x_124); -x_134 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_130); -lean_ctor_set(x_134, 2, x_131); -lean_ctor_set(x_134, 3, x_102); -x_135 = lean_alloc_ctor(0, 3, 0); + x_134 = x_126; +} +lean_ctor_set(x_134, 0, x_120); +lean_ctor_set(x_134, 1, x_133); +lean_ctor_set(x_134, 2, x_121); +lean_ctor_set(x_134, 3, x_122); +lean_ctor_set(x_134, 4, x_123); +lean_ctor_set(x_134, 5, x_124); +lean_ctor_set(x_134, 6, x_130); +lean_ctor_set(x_134, 7, x_125); +x_135 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_1); -lean_ctor_set(x_135, 2, x_2); -x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 1, x_131); +lean_ctor_set(x_135, 2, x_132); +lean_ctor_set(x_135, 3, x_103); +x_136 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_5); -return x_136; +lean_ctor_set(x_136, 1, x_1); +lean_ctor_set(x_136, 2, x_2); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_5); +return x_137; } else { -lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_137 = lean_ctor_get(x_99, 2); -lean_inc(x_137); -lean_dec(x_99); -x_138 = lean_ctor_get(x_137, 0); +lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_138 = lean_ctor_get(x_100, 2); lean_inc(x_138); -x_139 = lean_nat_dec_le(x_108, x_108); -if (x_139 == 0) +lean_dec(x_100); +x_139 = lean_ctor_get(x_138, 0); +lean_inc(x_139); +x_140 = lean_nat_dec_le(x_109, x_109); +if (x_140 == 0) { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_109); lean_dec(x_108); -lean_dec(x_107); -x_140 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; -if (lean_is_scalar(x_125)) { - x_141 = lean_alloc_ctor(0, 8, 0); +x_141 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; +if (lean_is_scalar(x_126)) { + x_142 = lean_alloc_ctor(0, 8, 0); } else { - x_141 = x_125; -} -lean_ctor_set(x_141, 0, x_119); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_120); -lean_ctor_set(x_141, 3, x_121); -lean_ctor_set(x_141, 4, x_122); -lean_ctor_set(x_141, 5, x_123); -lean_ctor_set(x_141, 6, x_129); -lean_ctor_set(x_141, 7, x_124); -x_142 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_137); -lean_ctor_set(x_142, 2, x_138); -lean_ctor_set(x_142, 3, x_102); -x_143 = lean_alloc_ctor(0, 3, 0); + x_142 = x_126; +} +lean_ctor_set(x_142, 0, x_120); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_121); +lean_ctor_set(x_142, 3, x_122); +lean_ctor_set(x_142, 4, x_123); +lean_ctor_set(x_142, 5, x_124); +lean_ctor_set(x_142, 6, x_130); +lean_ctor_set(x_142, 7, x_125); +x_143 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_1); -lean_ctor_set(x_143, 2, x_2); -x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 1, x_138); +lean_ctor_set(x_143, 2, x_139); +lean_ctor_set(x_143, 3, x_103); +x_144 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_5); -return x_144; +lean_ctor_set(x_144, 1, x_1); +lean_ctor_set(x_144, 2, x_2); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_5); +return x_145; } else { -size_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_145 = lean_usize_of_nat(x_108); +size_t x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_146 = lean_usize_of_nat(x_109); +lean_dec(x_109); +x_147 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; +x_148 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5(x_108, x_107, x_146, x_147); lean_dec(x_108); -x_146 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; -x_147 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_IO_processCommandsIncrementally_go___spec__5(x_107, x_106, x_145, x_146); -lean_dec(x_107); -if (lean_is_scalar(x_125)) { - x_148 = lean_alloc_ctor(0, 8, 0); +if (lean_is_scalar(x_126)) { + x_149 = lean_alloc_ctor(0, 8, 0); } else { - x_148 = x_125; -} -lean_ctor_set(x_148, 0, x_119); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_120); -lean_ctor_set(x_148, 3, x_121); -lean_ctor_set(x_148, 4, x_122); -lean_ctor_set(x_148, 5, x_123); -lean_ctor_set(x_148, 6, x_129); -lean_ctor_set(x_148, 7, x_124); -x_149 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_137); -lean_ctor_set(x_149, 2, x_138); -lean_ctor_set(x_149, 3, x_102); -x_150 = lean_alloc_ctor(0, 3, 0); + x_149 = x_126; +} +lean_ctor_set(x_149, 0, x_120); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_121); +lean_ctor_set(x_149, 3, x_122); +lean_ctor_set(x_149, 4, x_123); +lean_ctor_set(x_149, 5, x_124); +lean_ctor_set(x_149, 6, x_130); +lean_ctor_set(x_149, 7, x_125); +x_150 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_1); -lean_ctor_set(x_150, 2, x_2); -x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_150, 1, x_138); +lean_ctor_set(x_150, 2, x_139); +lean_ctor_set(x_150, 3, x_103); +x_151 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_5); -return x_151; +lean_ctor_set(x_151, 1, x_1); +lean_ctor_set(x_151, 2, x_2); +x_152 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_5); +return x_152; } } } else { -lean_object* x_152; -lean_dec(x_99); -x_152 = lean_ctor_get(x_100, 0); -lean_inc(x_152); +lean_object* x_153; lean_object* x_154; lean_dec(x_100); -x_3 = x_152; -x_4 = x_102; +x_153 = lean_ctor_get(x_101, 0); +lean_inc(x_153); +lean_dec(x_101); +x_154 = lean_ctor_get(x_153, 1); +lean_inc(x_154); +lean_dec(x_153); +x_3 = x_154; +x_4 = x_103; goto _start; } } @@ -2956,7 +2862,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); lean_inc(x_8); -x_10 = l_Lean_Language_SnapshotTask_get___rarg(x_8); +x_10 = lean_task_get_own(x_8); x_11 = l_Lean_Elab_Frontend_State_commands___default___closed__1; x_12 = l_Lean_Elab_IO_processCommandsIncrementally_go(x_1, x_10, x_8, x_11, x_9); return x_12; @@ -3013,7 +2919,7 @@ x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); lean_inc(x_23); -x_25 = l_Lean_Language_SnapshotTask_get___rarg(x_23); +x_25 = lean_task_get_own(x_23); x_26 = l_Lean_Elab_Frontend_State_commands___default___closed__1; x_27 = l_Lean_Elab_IO_processCommandsIncrementally_go(x_1, x_25, x_23, x_26, x_24); return x_27; @@ -3069,7 +2975,7 @@ x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); lean_inc(x_38); -x_40 = l_Lean_Language_SnapshotTask_get___rarg(x_38); +x_40 = lean_task_get_own(x_38); x_41 = l_Lean_Elab_Frontend_State_commands___default___closed__1; x_42 = l_Lean_Elab_IO_processCommandsIncrementally_go(x_1, x_40, x_38, x_41, x_39); return x_42; @@ -3103,6 +3009,41 @@ return x_46; } } } +LEAN_EXPORT lean_object* l_Lean_Elab_IO_processCommands(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_box(0); +x_6 = l_Lean_Elab_IO_processCommandsIncrementally(x_1, x_2, x_3, x_5, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec(x_8); +lean_ctor_set(x_6, 0, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +} static lean_object* _init_l_Lean_Elab_process___closed__1() { _start: { @@ -3131,15 +3072,12 @@ x_6 = l_Lean_Elab_Frontend_elabCommandAtFrontend___closed__4; x_7 = l_Lean_Elab_Command_mkState(x_2, x_6, x_3); if (lean_obj_tag(x_4) == 0) { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; x_8 = l_Lean_Elab_process___closed__2; x_9 = 1; x_10 = l_Lean_Parser_mkInputContext(x_1, x_8, x_9); x_11 = l_Lean_Elab_process___closed__1; x_12 = l_Lean_Elab_IO_processCommands(x_10, x_11, x_7, x_5); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_13, 0); @@ -3184,1654 +3122,1352 @@ return x_24; } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_12); -if (x_25 == 0) -{ -return x_12; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_12, 0); -x_27 = lean_ctor_get(x_12, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_12); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; -} -} -} -else -{ -lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_29 = lean_ctor_get(x_4, 0); -lean_inc(x_29); +lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_25 = lean_ctor_get(x_4, 0); +lean_inc(x_25); lean_dec(x_4); -x_30 = 1; -x_31 = l_Lean_Parser_mkInputContext(x_1, x_29, x_30); -x_32 = l_Lean_Elab_process___closed__1; -x_33 = l_Lean_Elab_IO_processCommands(x_31, x_32, x_7, x_5); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = lean_ctor_get(x_33, 0); +x_26 = 1; +x_27 = l_Lean_Parser_mkInputContext(x_1, x_25, x_26); +x_28 = l_Lean_Elab_process___closed__1; +x_29 = l_Lean_Elab_IO_processCommands(x_27, x_28, x_7, x_5); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +lean_dec(x_30); +x_32 = !lean_is_exclusive(x_29); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_29, 0); +lean_dec(x_33); +x_34 = lean_ctor_get(x_31, 0); lean_inc(x_34); -x_35 = lean_ctor_get(x_34, 0); +x_35 = lean_ctor_get(x_31, 1); lean_inc(x_35); -lean_dec(x_34); -x_36 = !lean_is_exclusive(x_33); -if (x_36 == 0) +lean_dec(x_31); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_29, 0, x_36); +return x_29; +} +else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_33, 0); -lean_dec(x_37); -x_38 = lean_ctor_get(x_35, 0); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_29, 1); +lean_inc(x_37); +lean_dec(x_29); +x_38 = lean_ctor_get(x_31, 0); lean_inc(x_38); -x_39 = lean_ctor_get(x_35, 1); +x_39 = lean_ctor_get(x_31, 1); lean_inc(x_39); -lean_dec(x_35); +lean_dec(x_31); x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_33, 0, x_40); -return x_33; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_37); +return x_41; } -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_33, 1); -lean_inc(x_41); -lean_dec(x_33); -x_42 = lean_ctor_get(x_35, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_dec(x_35); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_41); -return x_45; } } -else +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1(lean_object* x_1, size_t x_2, size_t x_3) { +_start: { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_33); -if (x_46 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_array_uget(x_1, x_2); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_MessageLog_hasErrors(x_7); +lean_dec(x_7); +if (x_8 == 0) { -return x_33; +size_t x_9; size_t x_10; +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_2 = x_10; +goto _start; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_33, 0); -x_48 = lean_ctor_get(x_33, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_33); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +uint8_t x_12; +x_12 = 1; +return x_12; } } +else +{ +uint8_t x_13; +x_13 = 0; +return x_13; } } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1() { +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid -D parameter, unknown configuration option '", 52, 52); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("'\n \nIf the option is defined in this library, use '-D", 62, 62); -return x_1; +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; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get(x_6, 2); +lean_inc(x_7); +lean_dec(x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_11 = l_Array_append___rarg(x_4, x_10); +x_2 = x_9; +x_4 = x_11; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_7, 0); +lean_inc(x_13); +lean_dec(x_7); +x_14 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1; +x_15 = lean_array_push(x_14, x_13); +x_16 = l_Array_append___rarg(x_4, x_15); +lean_dec(x_15); +x_2 = x_9; +x_4 = x_16; +goto _start; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' to set it conditionally", 25, 25); -return x_1; +return x_4; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4() { +} +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object* x_1, lean_object* x_2, uint32_t 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("true", 4, 4); -return x_1; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_alloc_ctor(0, 2, 4); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_2); +lean_ctor_set_uint32(x_7, sizeof(void*)*2, x_3); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("false", 5, 5); -return x_1; +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = l_Lean_Language_SnapshotTree_getAll(x_1); +x_6 = lean_array_get_size(x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_lt(x_7, x_6); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_6); +lean_dec(x_5); +x_9 = lean_ctor_get(x_2, 0); +x_10 = 1; +x_11 = lean_box(x_10); +lean_inc(x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_4); +return x_13; } +else +{ +lean_object* x_14; size_t x_15; size_t x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_2, 0); +x_15 = 0; +x_16 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_17 = l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1(x_5, x_15, x_16); +lean_dec(x_5); +if (x_17 == 0) +{ +uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = 1; +x_19 = lean_box(x_18); +lean_inc(x_14); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_14); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_4); +return x_21; } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid -D parameter, invalid configuration option '", 52, 52); -return x_1; +uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = 0; +x_23 = lean_box(x_22); +lean_inc(x_14); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_14); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_4); +return x_25; +} } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7() { +} +static lean_object* _init_l_Lean_Elab_runFrontend___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("' value, it must be true/false", 30, 30); +x_1 = l_Lean_trace_profiler_output; return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8() { +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5, lean_object* x_6, lean_object* x_7) { _start: { -uint8_t x_1; lean_object* x_2; -x_1 = 0; -x_2 = lean_alloc_ctor(1, 0, 1); -lean_ctor_set_uint8(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9() { -_start: -{ -uint8_t x_1; lean_object* x_2; -x_1 = 1; -x_2 = lean_alloc_ctor(1, 0, 1); -lean_ctor_set_uint8(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid -D parameter, configuration option '", 44, 44); -return x_1; -} -} -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' cannot be set in the command line, use set_option command", 59, 59); -return x_1; -} -} -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' value, it must be a natural number", 36, 36); -return x_1; -} -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_8; +x_8 = l_Lean_Language_Lean_waitForFinalCmdState_x3f(x_1); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_box(0); -x_10 = l_Lean_Name_replacePrefix(x_1, x_2, x_9); -x_11 = l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(x_3, x_10); -if (lean_obj_tag(x_11) == 0) +uint32_t x_9; lean_object* x_10; +lean_dec(x_4); +lean_dec(x_2); +x_9 = 0; +x_10 = lean_mk_empty_environment(x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_dec(x_5); -if (x_4 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (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; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_6); -x_12 = 1; -lean_inc(x_10); -x_13 = l_Lean_Name_toString(x_10, x_12); -x_14 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1; -x_15 = lean_string_append(x_14, x_13); -lean_dec(x_13); -x_16 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2; -x_17 = lean_string_append(x_15, x_16); -x_18 = l_Lean_Name_append(x_2, x_10); -x_19 = l_Lean_Name_toString(x_18, x_12); -x_20 = lean_string_append(x_17, x_19); -lean_dec(x_19); -x_21 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3; -x_22 = lean_string_append(x_20, x_21); -x_23 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_8); -return x_24; +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_10, 0); +x_13 = 0; +x_14 = lean_box(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_12); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_10, 0, x_15); +return x_10; } else { -lean_object* x_25; lean_object* x_26; +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); +lean_inc(x_16); lean_dec(x_10); -lean_dec(x_2); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_6); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_8); -return x_26; +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_16); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_17); +return x_21; } } else { -uint8_t x_27; -lean_dec(x_2); -x_27 = !lean_is_exclusive(x_11); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_11, 0); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); -switch (lean_obj_tag(x_29)) { -case 0: -{ -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_29, 0); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_5); -x_32 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); -lean_ctor_set(x_11, 0, x_32); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_11); -lean_ctor_set(x_33, 1, x_8); -return x_33; +return x_10; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_29); -x_34 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_34, 0, x_5); -x_35 = l_Lean_KVMap_insertCore(x_6, x_10, x_34); -lean_ctor_set(x_11, 0, x_35); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_11); -lean_ctor_set(x_36, 1, x_8); -return x_36; -} -} -case 1: -{ -lean_object* x_37; uint8_t x_38; -lean_dec(x_29); -x_37 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; -x_38 = lean_string_dec_eq(x_5, x_37); -if (x_38 == 0) -{ -lean_object* x_39; uint8_t x_40; -x_39 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; -x_40 = lean_string_dec_eq(x_5, x_39); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); +lean_inc(x_23); lean_dec(x_10); -lean_dec(x_6); -x_41 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; -x_42 = lean_string_append(x_41, x_5); -lean_dec(x_5); -x_43 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; -x_44 = lean_string_append(x_42, x_43); -lean_ctor_set_tag(x_11, 18); -lean_ctor_set(x_11, 0, x_44); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_11); -lean_ctor_set(x_45, 1, x_8); -return x_45; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_5); -x_46 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; -x_47 = l_Lean_KVMap_insertCore(x_6, x_10, x_46); -lean_ctor_set(x_11, 0, x_47); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_11); -lean_ctor_set(x_48, 1, x_8); -return x_48; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_5); -x_49 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; -x_50 = l_Lean_KVMap_insertCore(x_6, x_10, x_49); -lean_ctor_set(x_11, 0, x_50); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_11); -lean_ctor_set(x_51, 1, x_8); -return x_51; -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_8, 0); +lean_inc(x_26); +lean_dec(x_8); +x_27 = l_Lean_Elab_runFrontend___lambda__3___closed__1; +x_28 = l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(x_3, x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_4); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_runFrontend___lambda__2(x_2, x_26, x_29, x_7); +lean_dec(x_26); +return x_30; } -case 3: +else { -uint8_t x_52; -x_52 = !lean_is_exclusive(x_29); -if (x_52 == 0) +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_28, 0); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_ctor_get(x_26, 7); +lean_inc(x_32); +x_33 = 1; +x_34 = l_Lean_Name_toString(x_4, x_33); +x_35 = l_Lean_Firefox_Profile_export(x_34, x_5, x_32, x_3, x_7); +lean_dec(x_32); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_29, 0); -lean_dec(x_53); -x_54 = l_String_toNat_x3f(x_5); -if (lean_obj_tag(x_54) == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l___private_Lean_Util_Profiler_0__Lean_Firefox_toJsonProfile____x40_Lean_Util_Profiler___hyg_3604_(x_36); +x_39 = l_Lean_Json_compress(x_38); +x_40 = l_IO_FS_writeFile(x_31, x_39, x_37); +lean_dec(x_39); +lean_dec(x_31); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_free_object(x_29); -lean_dec(x_10); -lean_dec(x_6); -x_55 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; -x_56 = lean_string_append(x_55, x_5); -lean_dec(x_5); -x_57 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; -x_58 = lean_string_append(x_56, x_57); -lean_ctor_set_tag(x_11, 18); -lean_ctor_set(x_11, 0, x_58); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_11); -lean_ctor_set(x_59, 1, x_8); -return x_59; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = l_Lean_Elab_runFrontend___lambda__2(x_2, x_26, x_41, x_42); +lean_dec(x_41); +lean_dec(x_26); +return x_43; } else { -uint8_t x_60; -lean_free_object(x_11); -lean_dec(x_5); -x_60 = !lean_is_exclusive(x_54); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_54, 0); -lean_ctor_set(x_29, 0, x_61); -x_62 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); -lean_ctor_set(x_54, 0, x_62); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_54); -lean_ctor_set(x_63, 1, x_8); -return x_63; +uint8_t x_44; +lean_dec(x_26); +lean_dec(x_2); +x_44 = !lean_is_exclusive(x_40); +if (x_44 == 0) +{ +return x_40; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = lean_ctor_get(x_54, 0); -lean_inc(x_64); -lean_dec(x_54); -lean_ctor_set(x_29, 0, x_64); -x_65 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_8); -return x_67; +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 0); +x_46 = lean_ctor_get(x_40, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_40); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -lean_object* x_68; -lean_dec(x_29); -x_68 = l_String_toNat_x3f(x_5); -if (lean_obj_tag(x_68) == 0) +uint8_t x_48; +lean_dec(x_31); +lean_dec(x_26); +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_35); +if (x_48 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_10); -lean_dec(x_6); -x_69 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; -x_70 = lean_string_append(x_69, x_5); -lean_dec(x_5); -x_71 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; -x_72 = lean_string_append(x_70, x_71); -lean_ctor_set_tag(x_11, 18); -lean_ctor_set(x_11, 0, x_72); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_11); -lean_ctor_set(x_73, 1, x_8); -return x_73; +return x_35; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_free_object(x_11); -lean_dec(x_5); -x_74 = lean_ctor_get(x_68, 0); -lean_inc(x_74); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - x_75 = x_68; -} else { - lean_dec_ref(x_68); - x_75 = lean_box(0); -} -x_76 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_76, 0, x_74); -x_77 = l_Lean_KVMap_insertCore(x_6, x_10, x_76); -if (lean_is_scalar(x_75)) { - x_78 = lean_alloc_ctor(1, 1, 0); -} else { - x_78 = x_75; -} -lean_ctor_set(x_78, 0, x_77); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_8); -return x_79; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } -default: -{ -uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_29); -lean_dec(x_6); -lean_dec(x_5); -x_80 = 1; -x_81 = l_Lean_Name_toString(x_10, x_80); -x_82 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; -x_83 = lean_string_append(x_82, x_81); -lean_dec(x_81); -x_84 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; -x_85 = lean_string_append(x_83, x_84); -lean_ctor_set_tag(x_11, 18); -lean_ctor_set(x_11, 0, x_85); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_11); -lean_ctor_set(x_86, 1, x_8); -return x_86; } } } -else -{ -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_11, 0); -lean_inc(x_87); -lean_dec(x_11); -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -switch (lean_obj_tag(x_88)) { -case 0: +static lean_object* _init_l_Lean_Elab_runFrontend___lambda__4___closed__1() { +_start: { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - x_89 = x_88; -} else { - lean_dec_ref(x_88); - x_89 = lean_box(0); -} -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(0, 1, 0); -} else { - x_90 = x_89; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go), 1, 0); +return x_1; } -lean_ctor_set(x_90, 0, x_5); -x_91 = l_Lean_KVMap_insertCore(x_6, x_10, x_90); -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_91); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_8); -return x_93; } -case 1: +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_94; uint8_t x_95; -lean_dec(x_88); -x_94 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4; -x_95 = lean_string_dec_eq(x_5, x_94); -if (x_95 == 0) +lean_object* x_4; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) { -lean_object* x_96; uint8_t x_97; -x_96 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5; -x_97 = lean_string_dec_eq(x_5, x_96); -if (x_97 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -lean_dec(x_10); +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 1); lean_dec(x_6); -x_98 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; -x_99 = lean_string_append(x_98, x_5); -lean_dec(x_5); -x_100 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7; -x_101 = lean_string_append(x_99, x_100); -x_102 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_8); -return x_103; +x_7 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_1, x_2); +lean_ctor_set(x_3, 1, x_7); +return x_3; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_5); -x_104 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8; -x_105 = l_Lean_KVMap_insertCore(x_6, x_10, x_104); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_105); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_8); -return x_107; +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_1, x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; } } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_5); -x_108 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9; -x_109 = l_Lean_KVMap_insertCore(x_6, x_10, x_108); -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_109); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_8); -return x_111; -} -} -case 3: -{ -lean_object* x_112; lean_object* x_113; -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - x_112 = x_88; -} else { - lean_dec_ref(x_88); - x_112 = lean_box(0); -} -x_113 = l_String_toNat_x3f(x_5); -if (lean_obj_tag(x_113) == 0) +uint8_t x_11; +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_3); +if (x_11 == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_112); -lean_dec(x_10); -lean_dec(x_6); -x_114 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6; -x_115 = lean_string_append(x_114, x_5); -lean_dec(x_5); -x_116 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12; -x_117 = lean_string_append(x_115, x_116); -x_118 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_118, 0, x_117); -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_8); -return x_119; +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_3, 1); +lean_dec(x_12); +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_ctor_get(x_4, 0); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = l_Lean_Elab_runFrontend___lambda__4___closed__1; +x_18 = 1; +x_19 = l_Lean_Language_SnapshotTask_map___rarg(x_15, x_17, x_16, x_18); +lean_ctor_set(x_4, 0, x_19); +x_20 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_4, x_2); +lean_ctor_set(x_3, 1, x_20); +return x_3; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -lean_dec(x_5); -x_120 = lean_ctor_get(x_113, 0); -lean_inc(x_120); -if (lean_is_exclusive(x_113)) { - lean_ctor_release(x_113, 0); - x_121 = x_113; -} else { - lean_dec_ref(x_113); - x_121 = lean_box(0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +lean_dec(x_4); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = l_Lean_Elab_runFrontend___lambda__4___closed__1; +x_25 = 1; +x_26 = l_Lean_Language_SnapshotTask_map___rarg(x_22, x_24, x_23, x_25); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_27, x_2); +lean_ctor_set(x_3, 1, x_28); +return x_3; } -if (lean_is_scalar(x_112)) { - x_122 = lean_alloc_ctor(3, 1, 0); -} else { - x_122 = x_112; } -lean_ctor_set(x_122, 0, x_120); -x_123 = l_Lean_KVMap_insertCore(x_6, x_10, x_122); -if (lean_is_scalar(x_121)) { - x_124 = lean_alloc_ctor(1, 1, 0); +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = lean_ctor_get(x_3, 0); +lean_inc(x_29); +lean_dec(x_3); +x_30 = lean_ctor_get(x_4, 0); +lean_inc(x_30); +if (lean_is_exclusive(x_4)) { + lean_ctor_release(x_4, 0); + x_31 = x_4; } else { - x_124 = x_121; -} -lean_ctor_set(x_124, 0, x_123); -x_125 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_8); -return x_125; + lean_dec_ref(x_4); + x_31 = lean_box(0); } +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = l_Lean_Elab_runFrontend___lambda__4___closed__1; +x_35 = 1; +x_36 = l_Lean_Language_SnapshotTask_map___rarg(x_32, x_34, x_33, x_35); +if (lean_is_scalar(x_31)) { + x_37 = lean_alloc_ctor(1, 1, 0); +} else { + x_37 = x_31; } -default: -{ -uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_88); -lean_dec(x_6); -lean_dec(x_5); -x_126 = 1; -x_127 = l_Lean_Name_toString(x_10, x_126); -x_128 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10; -x_129 = lean_string_append(x_128, x_127); -lean_dec(x_127); -x_130 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11; -x_131 = lean_string_append(x_129, x_130); -x_132 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_132, 0, x_131); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_8); -return x_133; +lean_ctor_set(x_37, 0, x_36); +x_38 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_37, x_2); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_29); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } } +static double _init_l_Lean_Elab_runFrontend___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1() { +static lean_object* _init_l_Lean_Elab_runFrontend___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("weak", 4, 4); +x_1 = l_Lean_Language_Lean_internal_minimalSnapshots; return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2() { +static lean_object* _init_l_Lean_Elab_runFrontend___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); +x_2 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_3 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* lean_run_frontend(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint32_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_3); -lean_ctor_set(x_5, 1, x_4); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_ctor_get(x_6, 0); -lean_inc(x_9); -lean_dec(x_6); -x_10 = lean_ctor_get(x_7, 0); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; double x_14; double x_15; double 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; uint8_t x_25; +x_9 = lean_io_mono_nanos_now(x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_7); -x_11 = l_Lean_Name_getRoot(x_9); -x_12 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2; -x_13 = lean_name_eq(x_11, x_12); -lean_dec(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -x_15 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_3, x_14, x_4); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 0); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = 0; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Float_ofScientific(x_10, x_12, x_13); +lean_dec(x_10); +x_15 = l_Lean_Elab_runFrontend___closed__1; +x_16 = lean_float_div(x_14, x_15); +x_17 = 1; +x_18 = l_Lean_Parser_mkInputContext(x_1, x_3, x_17); +x_19 = l_Lean_Elab_runFrontend___closed__2; +x_20 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_2, x_19, x_17); +x_21 = lean_box_uint32(x_5); +lean_inc(x_20); +lean_inc(x_4); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_runFrontend___lambda__1___boxed), 6, 3); +lean_closure_set(x_22, 0, x_4); +lean_closure_set(x_22, 1, x_20); +lean_closure_set(x_22, 2, x_21); +x_23 = lean_box(0); lean_inc(x_18); -lean_dec(x_16); -x_2 = x_8; -x_3 = x_18; -x_4 = x_17; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_8); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) -{ -return x_15; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -else +x_24 = l_Lean_Language_Lean_process(x_22, x_23, x_18, x_11); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_inc(x_9); -x_24 = l_Lean_KVMap_erase(x_3, x_9); -x_25 = lean_box(0); -x_26 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_24, x_25, x_4); -if (lean_obj_tag(x_26) == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_100; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +x_100 = lean_ctor_get(x_26, 3); +lean_inc(x_100); +if (lean_obj_tag(x_100) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_2 = x_8; -x_3 = x_29; -x_4 = x_28; -goto _start; +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_26, 0); +lean_inc(x_101); +x_102 = l_Lean_Elab_runFrontend___closed__3; +lean_ctor_set(x_24, 1, x_102); +lean_ctor_set(x_24, 0, x_101); +x_28 = x_24; +goto block_99; } else { -uint8_t x_31; -lean_dec(x_8); -x_31 = !lean_is_exclusive(x_26); -if (x_31 == 0) +lean_object* x_103; uint8_t x_104; +x_103 = lean_ctor_get(x_26, 0); +lean_inc(x_103); +x_104 = !lean_is_exclusive(x_100); +if (x_104 == 0) { -return x_26; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_105 = lean_ctor_get(x_100, 0); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_108 = lean_alloc_closure((void*)(l_Lean_Elab_runFrontend___lambda__4), 3, 2); +lean_closure_set(x_108, 0, x_23); +lean_closure_set(x_108, 1, x_107); +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +x_110 = l_Lean_Language_SnapshotTask_map___rarg(x_106, x_108, x_109, x_17); +lean_ctor_set(x_100, 0, x_110); +x_111 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_100, x_107); +lean_ctor_set(x_24, 1, x_111); +lean_ctor_set(x_24, 0, x_103); +x_28 = x_24; +goto block_99; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_26, 0); -x_33 = lean_ctor_get(x_26, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_26); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_112 = lean_ctor_get(x_100, 0); +lean_inc(x_112); +lean_dec(x_100); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +lean_dec(x_112); +x_114 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_115 = lean_alloc_closure((void*)(l_Lean_Elab_runFrontend___lambda__4), 3, 2); +lean_closure_set(x_115, 0, x_23); +lean_closure_set(x_115, 1, x_114); +x_116 = lean_ctor_get(x_113, 0); +lean_inc(x_116); +x_117 = l_Lean_Language_SnapshotTask_map___rarg(x_113, x_115, x_116, x_17); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_117); +x_119 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_118, x_114); +lean_ctor_set(x_24, 1, x_119); +lean_ctor_set(x_24, 0, x_103); +x_28 = x_24; +goto block_99; } } -else +block_99: { -lean_object* x_35; -lean_dec(x_7); -lean_dec(x_6); -x_35 = lean_ctor_get(x_2, 1); -lean_inc(x_35); -lean_dec(x_2); -x_2 = x_35; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_reparseOptions(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_29; +lean_inc(x_20); +x_29 = l_Lean_Language_SnapshotTree_forM___at_Lean_Language_SnapshotTree_runAndReport___spec__1(x_20, x_7, x_28, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = l_Lean_getOptionDecls(x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec(x_3); -lean_inc(x_1); -x_6 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(x_4, x_1, x_1, x_5); -lean_dec(x_4); if (lean_obj_tag(x_6) == 0) { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -return x_6; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_18); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_box(0); +x_32 = l_Lean_Elab_runFrontend___lambda__3(x_26, x_28, x_20, x_4, x_16, x_31, x_30); +lean_dec(x_20); +return x_32; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_6, 1); -lean_inc(x_9); -lean_inc(x_8); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_6, 0); +lean_inc(x_34); lean_dec(x_6); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -return x_10; -} -} -else +lean_inc(x_28); +x_35 = l_Lean_Language_SnapshotTree_getAll(x_28); +x_36 = lean_array_get_size(x_35); +x_37 = lean_nat_dec_lt(x_13, x_36); +if (x_37 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_6); -if (x_11 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_36); +lean_dec(x_35); +x_38 = lean_ctor_get(x_18, 2); +lean_inc(x_38); +lean_dec(x_18); +x_39 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_40 = l_Lean_Server_findModuleRefs(x_38, x_39, x_12, x_12); +x_41 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_40, x_33); +lean_dec(x_40); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_45 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_4); +lean_ctor_set(x_45, 2, x_42); +x_46 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_45); +x_47 = l_Lean_Json_compress(x_46); +x_48 = l_IO_FS_writeFile(x_34, x_47, x_43); +lean_dec(x_47); +lean_dec(x_34); +if (lean_obj_tag(x_48) == 0) { -return x_6; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Lean_Elab_runFrontend___lambda__3(x_26, x_28, x_20, x_4, x_16, x_49, x_50); +lean_dec(x_49); +lean_dec(x_20); +return x_51; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_6, 0); -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_6); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_4); +uint8_t x_52; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_4); -x_10 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = l_Lean_MessageLog_hasErrors(x_1); -if (x_5 == 0) +x_52 = !lean_is_exclusive(x_48); +if (x_52 == 0) { -uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = 1; -x_7 = lean_box(x_6); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_2); -lean_ctor_set(x_8, 1, x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_4); -return x_9; +return x_48; } else { -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = 0; -x_11 = lean_box(x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_2); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_4); -return x_13; -} -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_trace_profiler_output; -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Import", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_runFrontend___lambda__2___closed__2; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("importing", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_runFrontend___lambda__2___closed__4; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_runFrontend___lambda__2___closed__5; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_runFrontend___lambda__2___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_48, 0); +x_54 = lean_ctor_get(x_48, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_48); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, double x_4, double x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_runFrontend___lambda__2___closed__1; -x_11 = l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(x_3, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_7); -x_12 = lean_box(0); -x_13 = l_Lean_Elab_runFrontend___lambda__1(x_1, x_2, x_12, x_9); -return x_13; } else { -lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_14 = lean_ctor_get(x_11, 0); -lean_inc(x_14); -lean_dec(x_11); -x_15 = l_Lean_Elab_runFrontend___lambda__2___closed__3; -x_16 = 1; -x_17 = l_Lean_Elab_Frontend_runCommandElabM___rarg___closed__2; -x_18 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_18, 0, x_15); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set_float(x_18, sizeof(void*)*2, x_4); -lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_5); -lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_16); -x_19 = l_Lean_Elab_runFrontend___lambda__2___closed__6; -x_20 = l_Lean_Elab_Frontend_State_commands___default___closed__1; -x_21 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_19); -lean_ctor_set(x_21, 2, x_20); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_Elab_runFrontend___lambda__2___closed__7; -x_25 = lean_array_push(x_24, x_23); -x_26 = l_Array_toPArray_x27___rarg(x_25); -lean_dec(x_25); -x_27 = l_Lean_PersistentArray_append___rarg(x_26, x_6); -x_28 = l_Lean_Name_toString(x_7, x_16); -x_29 = l_Lean_Firefox_Profile_export(x_28, x_4, x_27, x_3, x_9); -lean_dec(x_27); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = l___private_Lean_Util_Profiler_0__Lean_Firefox_toJsonProfile____x40_Lean_Util_Profiler___hyg_3604_(x_30); -x_33 = l_Lean_Json_compress(x_32); -x_34 = l_IO_FS_writeFile(x_14, x_33, x_31); -lean_dec(x_33); -lean_dec(x_14); -if (lean_obj_tag(x_34) == 0) +lean_object* x_56; uint8_t x_57; +x_56 = lean_ctor_get(x_18, 2); +lean_inc(x_56); +lean_dec(x_18); +x_57 = lean_nat_dec_le(x_36, x_36); +if (x_57 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_Elab_runFrontend___lambda__1(x_1, x_2, x_35, x_36); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_36); lean_dec(x_35); -return x_37; -} -else -{ -uint8_t x_38; -lean_dec(x_2); -x_38 = !lean_is_exclusive(x_34); -if (x_38 == 0) -{ -return x_34; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_34, 0); -x_40 = lean_ctor_get(x_34, 1); -lean_inc(x_40); -lean_inc(x_39); +x_58 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_59 = l_Lean_Server_findModuleRefs(x_56, x_58, x_12, x_12); +x_60 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_59, x_33); +lean_dec(x_59); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_64 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_4); +lean_ctor_set(x_64, 2, x_61); +x_65 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_64); +x_66 = l_Lean_Json_compress(x_65); +x_67 = l_IO_FS_writeFile(x_34, x_66, x_62); +lean_dec(x_66); lean_dec(x_34); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} -} -} -else -{ -uint8_t x_42; -lean_dec(x_14); -lean_dec(x_2); -x_42 = !lean_is_exclusive(x_29); -if (x_42 == 0) -{ -return x_29; -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_29, 0); -x_44 = lean_ctor_get(x_29, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_29); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, double x_5, double x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -lean_inc(x_1); -x_12 = l_Lean_Elab_IO_processCommands(x_1, x_2, x_9, x_11); -if (lean_obj_tag(x_12) == 0) +if (lean_obj_tag(x_67) == 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; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec(x_13); -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); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -x_18 = lean_ctor_get(x_14, 6); -lean_inc(x_18); -x_19 = lean_ctor_get(x_14, 7); -lean_inc(x_19); -lean_dec(x_14); -lean_inc(x_3); -x_20 = l_Lean_Language_reportMessages(x_17, x_3, x_4, x_15); -if (lean_obj_tag(x_20) == 0) -{ -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_18); -lean_dec(x_1); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_Elab_runFrontend___lambda__3(x_26, x_28, x_20, x_4, x_16, x_68, x_69); +lean_dec(x_68); lean_dec(x_20); -x_22 = lean_box(0); -x_23 = l_Lean_Elab_runFrontend___lambda__2(x_17, x_16, x_3, x_5, x_6, x_19, x_7, x_22, x_21); -lean_dec(x_19); -lean_dec(x_3); -lean_dec(x_17); -return x_23; +return x_70; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_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_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_8, 0); -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); -lean_dec(x_18); -x_27 = l_Lean_PersistentArray_toArray___rarg(x_26); +uint8_t x_71; +lean_dec(x_28); lean_dec(x_26); -x_28 = lean_ctor_get(x_1, 2); -lean_inc(x_28); -lean_dec(x_1); -x_29 = 0; -x_30 = l_Lean_Server_findModuleRefs(x_28, x_27, x_29, x_29); -lean_dec(x_27); -x_31 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_30, x_24); -lean_dec(x_30); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_unsigned_to_nat(3u); -lean_inc(x_7); -x_35 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_7); -lean_ctor_set(x_35, 2, x_32); -x_36 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_35); -x_37 = l_Lean_Json_compress(x_36); -x_38 = l_IO_FS_writeFile(x_25, x_37, x_33); -lean_dec(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Elab_runFrontend___lambda__2(x_17, x_16, x_3, x_5, x_6, x_19, x_7, x_39, x_40); -lean_dec(x_39); -lean_dec(x_19); -lean_dec(x_3); -lean_dec(x_17); -return x_41; -} -else -{ -uint8_t x_42; -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_3); -x_42 = !lean_is_exclusive(x_38); -if (x_42 == 0) +lean_dec(x_20); +lean_dec(x_4); +x_71 = !lean_is_exclusive(x_67); +if (x_71 == 0) { -return x_38; +return x_67; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_38, 0); -x_44 = lean_ctor_get(x_38, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_38); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; -} +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_67, 0); +x_73 = lean_ctor_get(x_67, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_67); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } } } else { -uint8_t x_46; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_1); -x_46 = !lean_is_exclusive(x_20); -if (x_46 == 0) +size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_75 = 0; +x_76 = lean_usize_of_nat(x_36); +lean_dec(x_36); +x_77 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2(x_35, x_75, x_76, x_77); +lean_dec(x_35); +x_79 = l_Lean_Server_findModuleRefs(x_56, x_78, x_12, x_12); +lean_dec(x_78); +x_80 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_79, x_33); +lean_dec(x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_84 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_4); +lean_ctor_set(x_84, 2, x_81); +x_85 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_84); +x_86 = l_Lean_Json_compress(x_85); +x_87 = l_IO_FS_writeFile(x_34, x_86, x_82); +lean_dec(x_86); +lean_dec(x_34); +if (lean_obj_tag(x_87) == 0) { -return x_20; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Elab_runFrontend___lambda__3(x_26, x_28, x_20, x_4, x_16, x_88, x_89); +lean_dec(x_88); +lean_dec(x_20); +return x_90; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_20, 0); -x_48 = lean_ctor_get(x_20, 1); -lean_inc(x_48); -lean_inc(x_47); +uint8_t x_91; +lean_dec(x_28); +lean_dec(x_26); lean_dec(x_20); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_4); +x_91 = !lean_is_exclusive(x_87); +if (x_91 == 0) +{ +return x_87; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_87, 0); +x_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_87); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} } } } else { -uint8_t x_50; -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_1); -x_50 = !lean_is_exclusive(x_12); -if (x_50 == 0) +uint8_t x_95; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_18); +lean_dec(x_6); +lean_dec(x_4); +x_95 = !lean_is_exclusive(x_29); +if (x_95 == 0) { -return x_12; +return x_29; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_12, 0); -x_52 = lean_ctor_get(x_12, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_12); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -return x_53; -} -} +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_29, 0); +x_97 = lean_ctor_get(x_29, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_29); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } -static double _init_l_Lean_Elab_runFrontend___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; } } -static lean_object* _init_l_Lean_Elab_runFrontend___closed__2() { -_start: +else { -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; -} +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_194; +x_120 = lean_ctor_get(x_24, 0); +x_121 = lean_ctor_get(x_24, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_24); +x_194 = lean_ctor_get(x_120, 3); +lean_inc(x_194); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_195 = lean_ctor_get(x_120, 0); +lean_inc(x_195); +x_196 = l_Lean_Elab_runFrontend___closed__3; +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +x_122 = x_197; +goto block_193; } -static lean_object* _init_l_Lean_Elab_runFrontend___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_maxRecDepth; -return x_1; +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_198 = lean_ctor_get(x_120, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_194, 0); +lean_inc(x_199); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + x_200 = x_194; +} else { + lean_dec_ref(x_194); + x_200 = lean_box(0); +} +x_201 = lean_ctor_get(x_199, 1); +lean_inc(x_201); +lean_dec(x_199); +x_202 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_203 = lean_alloc_closure((void*)(l_Lean_Elab_runFrontend___lambda__4), 3, 2); +lean_closure_set(x_203, 0, x_23); +lean_closure_set(x_203, 1, x_202); +x_204 = lean_ctor_get(x_201, 0); +lean_inc(x_204); +x_205 = l_Lean_Language_SnapshotTask_map___rarg(x_201, x_203, x_204, x_17); +if (lean_is_scalar(x_200)) { + x_206 = lean_alloc_ctor(1, 1, 0); +} else { + x_206 = x_200; } +lean_ctor_set(x_206, 0, x_205); +x_207 = l___private_Lean_Language_Lean_Types_0__Lean_Language_Lean_pushOpt___rarg(x_206, x_202); +x_208 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_208, 0, x_198); +lean_ctor_set(x_208, 1, x_207); +x_122 = x_208; +goto block_193; } -LEAN_EXPORT lean_object* lean_run_frontend(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint32_t x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; double x_14; double x_15; double x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_9 = lean_io_mono_nanos_now(x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = 0; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Float_ofScientific(x_10, x_12, x_13); -lean_dec(x_10); -x_15 = l_Lean_Elab_runFrontend___closed__1; -x_16 = lean_float_div(x_14, x_15); -x_17 = 1; -x_18 = l_Lean_Parser_mkInputContext(x_1, x_3, x_17); -lean_inc(x_18); -x_19 = l_Lean_Parser_parseHeader(x_18, x_11); -if (lean_obj_tag(x_19) == 0) +block_193: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_19, 0); +lean_object* x_123; lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -lean_inc(x_18); -lean_inc(x_2); -x_26 = l_Lean_Elab_processHeader(x_23, x_2, x_25, x_18, x_5, x_17, x_22); -lean_dec(x_23); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = l_Lean_Elab_reparseOptions(x_2, x_28); -if (lean_obj_tag(x_31) == 0) +x_123 = l_Lean_Language_SnapshotTree_forM___at_Lean_Language_SnapshotTree_runAndReport___spec__1(x_20, x_7, x_122, x_121); +if (lean_obj_tag(x_123) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; double x_39; double x_40; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -lean_inc(x_4); -x_34 = lean_environment_set_main_module(x_29, x_4); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_34); -x_35 = l_Lean_Elab_Command_mkState(x_34, x_30, x_32); -x_36 = lean_io_mono_nanos_now(x_33); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Float_ofScientific(x_37, x_12, x_13); -lean_dec(x_37); -x_40 = lean_float_div(x_39, x_15); if (lean_obj_tag(x_6) == 0) { -lean_object* x_41; lean_object* x_42; -lean_dec(x_34); -lean_dec(x_30); -x_41 = lean_box(0); -x_42 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_41, x_38); -return x_42; -} -else -{ -uint8_t x_43; -x_43 = !lean_is_exclusive(x_35); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_44 = lean_ctor_get(x_35, 6); -x_45 = lean_ctor_get(x_35, 4); -lean_dec(x_45); -x_46 = lean_ctor_get(x_35, 3); -lean_dec(x_46); -x_47 = lean_ctor_get(x_35, 1); -lean_dec(x_47); -x_48 = lean_ctor_get(x_35, 0); -lean_dec(x_48); -x_49 = l_Lean_Elab_runFrontend___closed__3; -x_50 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_32, x_49); -x_51 = !lean_is_exclusive(x_44); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_ctor_set_uint8(x_44, sizeof(void*)*2, x_17); -x_52 = l_Lean_Elab_runFrontend___closed__2; -lean_ctor_set(x_35, 4, x_50); -lean_ctor_set(x_35, 3, x_52); -lean_ctor_set(x_35, 1, x_30); -lean_ctor_set(x_35, 0, x_34); -x_53 = lean_box(0); -x_54 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_53, x_38); -lean_dec(x_6); -return x_54; +lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_18); +x_124 = lean_ctor_get(x_123, 1); +lean_inc(x_124); +lean_dec(x_123); +x_125 = lean_box(0); +x_126 = l_Lean_Elab_runFrontend___lambda__3(x_120, x_122, x_20, x_4, x_16, x_125, x_124); +lean_dec(x_20); +return x_126; } 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; -x_55 = lean_ctor_get(x_44, 0); -x_56 = lean_ctor_get(x_44, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_44); -x_57 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set_uint8(x_57, sizeof(void*)*2, x_17); -x_58 = l_Lean_Elab_runFrontend___closed__2; -lean_ctor_set(x_35, 6, x_57); -lean_ctor_set(x_35, 4, x_50); -lean_ctor_set(x_35, 3, x_58); -lean_ctor_set(x_35, 1, x_30); -lean_ctor_set(x_35, 0, x_34); -x_59 = lean_box(0); -x_60 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_35, x_59, x_38); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_127 = lean_ctor_get(x_123, 1); +lean_inc(x_127); +lean_dec(x_123); +x_128 = lean_ctor_get(x_6, 0); +lean_inc(x_128); lean_dec(x_6); -return x_60; -} +lean_inc(x_122); +x_129 = l_Lean_Language_SnapshotTree_getAll(x_122); +x_130 = lean_array_get_size(x_129); +x_131 = lean_nat_dec_lt(x_13, 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; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_130); +lean_dec(x_129); +x_132 = lean_ctor_get(x_18, 2); +lean_inc(x_132); +lean_dec(x_18); +x_133 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_134 = l_Lean_Server_findModuleRefs(x_132, x_133, x_12, x_12); +x_135 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_134, x_127); +lean_dec(x_134); +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_138 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_139 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_4); +lean_ctor_set(x_139, 2, x_136); +x_140 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_139); +x_141 = l_Lean_Json_compress(x_140); +x_142 = l_IO_FS_writeFile(x_128, x_141, x_137); +lean_dec(x_141); +lean_dec(x_128); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +lean_dec(x_142); +x_145 = l_Lean_Elab_runFrontend___lambda__3(x_120, x_122, x_20, x_4, x_16, x_143, x_144); +lean_dec(x_143); +lean_dec(x_20); +return x_145; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_61 = lean_ctor_get(x_35, 2); -x_62 = lean_ctor_get(x_35, 5); -x_63 = lean_ctor_get(x_35, 6); -x_64 = lean_ctor_get(x_35, 7); -lean_inc(x_64); -lean_inc(x_63); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_35); -x_65 = l_Lean_Elab_runFrontend___closed__3; -x_66 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_32, x_65); -x_67 = lean_ctor_get(x_63, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_63, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_69 = x_63; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_122); +lean_dec(x_120); +lean_dec(x_20); +lean_dec(x_4); +x_146 = lean_ctor_get(x_142, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_142, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_148 = x_142; } else { - lean_dec_ref(x_63); - x_69 = lean_box(0); + lean_dec_ref(x_142); + x_148 = lean_box(0); } -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 2, 1); +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(1, 2, 0); } else { - x_70 = x_69; -} -lean_ctor_set(x_70, 0, x_67); -lean_ctor_set(x_70, 1, x_68); -lean_ctor_set_uint8(x_70, sizeof(void*)*2, x_17); -x_71 = l_Lean_Elab_runFrontend___closed__2; -x_72 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_72, 0, x_34); -lean_ctor_set(x_72, 1, x_30); -lean_ctor_set(x_72, 2, x_61); -lean_ctor_set(x_72, 3, x_71); -lean_ctor_set(x_72, 4, x_66); -lean_ctor_set(x_72, 5, x_62); -lean_ctor_set(x_72, 6, x_70); -lean_ctor_set(x_72, 7, x_64); -x_73 = lean_box(0); -x_74 = l_Lean_Elab_runFrontend___lambda__3(x_18, x_24, x_32, x_7, x_16, x_40, x_4, x_6, x_72, x_73, x_38); -lean_dec(x_6); -return x_74; + x_149 = x_148; } +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_147); +return x_149; } } else { -uint8_t x_75; -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_24); +lean_object* x_150; uint8_t x_151; +x_150 = lean_ctor_get(x_18, 2); +lean_inc(x_150); lean_dec(x_18); -lean_dec(x_6); -lean_dec(x_4); -x_75 = !lean_is_exclusive(x_31); -if (x_75 == 0) -{ -return x_31; +x_151 = lean_nat_dec_le(x_130, x_130); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_130); +lean_dec(x_129); +x_152 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_153 = l_Lean_Server_findModuleRefs(x_150, x_152, x_12, x_12); +x_154 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_153, x_127); +lean_dec(x_153); +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +lean_dec(x_154); +x_157 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_158 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_4); +lean_ctor_set(x_158, 2, x_155); +x_159 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_158); +x_160 = l_Lean_Json_compress(x_159); +x_161 = l_IO_FS_writeFile(x_128, x_160, x_156); +lean_dec(x_160); +lean_dec(x_128); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +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 = l_Lean_Elab_runFrontend___lambda__3(x_120, x_122, x_20, x_4, x_16, x_162, x_163); +lean_dec(x_162); +lean_dec(x_20); +return x_164; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_31, 0); -x_77 = lean_ctor_get(x_31, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_31); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +lean_dec(x_122); +lean_dec(x_120); +lean_dec(x_20); +lean_dec(x_4); +x_165 = lean_ctor_get(x_161, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_161, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_167 = x_161; +} else { + lean_dec_ref(x_161); + x_167 = lean_box(0); +} +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); +} else { + x_168 = x_167; } +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; } } else { -uint8_t x_79; -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -x_79 = !lean_is_exclusive(x_26); -if (x_79 == 0) -{ -return x_26; +size_t x_169; size_t x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_169 = 0; +x_170 = lean_usize_of_nat(x_130); +lean_dec(x_130); +x_171 = l_Lean_Elab_Frontend_State_commands___default___closed__1; +x_172 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2(x_129, x_169, x_170, x_171); +lean_dec(x_129); +x_173 = l_Lean_Server_findModuleRefs(x_150, x_172, x_12, x_12); +lean_dec(x_172); +x_174 = l_Lean_Server_ModuleRefs_toLspModuleRefs(x_173, x_127); +lean_dec(x_173); +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +lean_dec(x_174); +x_177 = lean_unsigned_to_nat(3u); +lean_inc(x_4); +x_178 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_4); +lean_ctor_set(x_178, 2, x_175); +x_179 = l___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1472_(x_178); +x_180 = l_Lean_Json_compress(x_179); +x_181 = l_IO_FS_writeFile(x_128, x_180, x_176); +lean_dec(x_180); +lean_dec(x_128); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = l_Lean_Elab_runFrontend___lambda__3(x_120, x_122, x_20, x_4, x_16, x_182, x_183); +lean_dec(x_182); +lean_dec(x_20); +return x_184; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_26, 0); -x_81 = lean_ctor_get(x_26, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_26); -x_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_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +lean_dec(x_122); +lean_dec(x_120); +lean_dec(x_20); +lean_dec(x_4); +x_185 = lean_ctor_get(x_181, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_181, 1); +lean_inc(x_186); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_187 = x_181; +} else { + lean_dec_ref(x_181); + x_187 = lean_box(0); +} +if (lean_is_scalar(x_187)) { + x_188 = lean_alloc_ctor(1, 2, 0); +} else { + x_188 = x_187; +} +lean_ctor_set(x_188, 0, x_185); +lean_ctor_set(x_188, 1, x_186); +return x_188; +} +} } } } else { -uint8_t x_83; +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec(x_122); +lean_dec(x_120); +lean_dec(x_20); lean_dec(x_18); lean_dec(x_6); lean_dec(x_4); -lean_dec(x_2); -x_83 = !lean_is_exclusive(x_19); -if (x_83 == 0) -{ -return x_19; +x_189 = lean_ctor_get(x_123, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_123, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_191 = x_123; +} else { + lean_dec_ref(x_123); + x_191 = lean_box(0); } -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_19, 0); -x_85 = lean_ctor_get(x_19, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_19); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -return x_86; +if (lean_is_scalar(x_191)) { + x_192 = lean_alloc_ctor(1, 2, 0); +} else { + x_192 = x_191; } +lean_ctor_set(x_192, 0, x_189); +lean_ctor_set(x_192, 1, x_190); +return x_192; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; -x_5 = l_Lean_Elab_runFrontend___lambda__1(x_1, x_2, x_3, x_4); +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); lean_dec(x_3); +x_6 = l_Array_anyMUnsafe_any___at_Lean_Elab_runFrontend___spec__1(x_1, x_4, x_5); lean_dec(x_1); -return x_5; +x_7 = lean_box(x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -double x_10; double x_11; lean_object* x_12; -x_10 = lean_unbox_float(x_4); -lean_dec(x_4); -x_11 = lean_unbox_float(x_5); -lean_dec(x_5); -x_12 = l_Lean_Elab_runFrontend___lambda__2(x_1, x_2, x_3, x_10, x_11, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_6); +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_Elab_runFrontend___spec__2(x_1, x_5, x_6, x_4); lean_dec(x_1); -return x_12; +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___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: { -uint8_t x_12; double x_13; double x_14; lean_object* x_15; -x_12 = lean_unbox(x_4); +uint32_t x_7; lean_object* x_8; +x_7 = lean_unbox_uint32(x_3); +lean_dec(x_3); +x_8 = l_Lean_Elab_runFrontend___lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); lean_dec(x_4); -x_13 = lean_unbox_float(x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___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_Lean_Elab_runFrontend___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_Lean_Elab_runFrontend___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +double x_8; lean_object* x_9; +x_8 = lean_unbox_float(x_5); lean_dec(x_5); -x_14 = lean_unbox_float(x_6); +x_9 = l_Lean_Elab_runFrontend___lambda__3(x_1, x_2, x_3, x_4, x_8, x_6, x_7); lean_dec(x_6); -x_15 = l_Lean_Elab_runFrontend___lambda__3(x_1, x_2, x_3, x_12, x_13, x_14, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_8); -return x_15; +lean_dec(x_3); +return x_9; } } LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -4889,48 +4525,12 @@ l_Lean_Elab_process___closed__1 = _init_l_Lean_Elab_process___closed__1(); lean_mark_persistent(l_Lean_Elab_process___closed__1); l_Lean_Elab_process___closed__2 = _init_l_Lean_Elab_process___closed__2(); lean_mark_persistent(l_Lean_Elab_process___closed__2); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__1); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__2); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__3); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__4); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__5); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__6); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__7); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__8); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__9); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__10); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__11); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___lambda__1___closed__12); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__1); -l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_reparseOptions___spec__1___closed__2); -l_Lean_Elab_runFrontend___lambda__2___closed__1 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__1); -l_Lean_Elab_runFrontend___lambda__2___closed__2 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__2); -l_Lean_Elab_runFrontend___lambda__2___closed__3 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__3); -l_Lean_Elab_runFrontend___lambda__2___closed__4 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__4); -l_Lean_Elab_runFrontend___lambda__2___closed__5 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__5); -l_Lean_Elab_runFrontend___lambda__2___closed__6 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__6); -l_Lean_Elab_runFrontend___lambda__2___closed__7 = _init_l_Lean_Elab_runFrontend___lambda__2___closed__7(); -lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__2___closed__7); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_runFrontend___spec__2___closed__1); +l_Lean_Elab_runFrontend___lambda__3___closed__1 = _init_l_Lean_Elab_runFrontend___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__3___closed__1); +l_Lean_Elab_runFrontend___lambda__4___closed__1 = _init_l_Lean_Elab_runFrontend___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_runFrontend___lambda__4___closed__1); l_Lean_Elab_runFrontend___closed__1 = _init_l_Lean_Elab_runFrontend___closed__1(); l_Lean_Elab_runFrontend___closed__2 = _init_l_Lean_Elab_runFrontend___closed__2(); lean_mark_persistent(l_Lean_Elab_runFrontend___closed__2); diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index cb14832b71f6..74752f54023c 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -16,7 +16,6 @@ extern "C" { static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkNumParams___spec__1___closed__1; static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__3; LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_isDomainDefEq___boxed(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_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__6(uint8_t, lean_object*, size_t, size_t); @@ -53,7 +52,6 @@ lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModi static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_replaceIndFVarsWithConsts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_computeFixedIndexBitMask_go___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__11(lean_object*); static lean_object* l_Lean_Elab_Command_checkResultingUniverse___closed__2; @@ -117,7 +115,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_acc static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_6587____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__8(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_List_forM___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_collectUniverses_go___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* l_Lean_Level_succ___override(lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_localDeclDependsOnPred___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___closed__1; @@ -129,12 +126,10 @@ lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__17; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___closed__2; -lean_object* l_Lean_mkLevelMax_x27(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs___spec__4___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___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*); static lean_object* l_Lean_Elab_Command_instInhabitedCtorView___closed__3; static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__1___closed__1; @@ -156,7 +151,6 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkLevelNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_withCtorRef___spec__2___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___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl(lean_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_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -227,7 +221,6 @@ static lean_object* l_Lean_Elab_Command_accLevelAtCtor___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidInductiveModifier(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___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_throwError___at_Lean_Elab_Command_accLevelAtCtor___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_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__5; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_48____closed__4; LEAN_EXPORT uint8_t l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__1(lean_object*, lean_object*); @@ -253,7 +246,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Induc LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_withCtorRef___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -360,6 +352,7 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__ lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeCompatible___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkParamsAndResultType___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_collectUniverses_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_occurs(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedCtorView___closed__2; lean_object* l_Lean_mkCasesOn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -385,6 +378,7 @@ lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_collectUniverses_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withInductiveLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_6587____closed__6; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -394,6 +388,7 @@ lean_object* l_Lean_Meta_isTypeFormerType(lean_object*, lean_object*, lean_objec static lean_object* l_List_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__8___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__9___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__10___closed__6; lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_instantiate_level_mvars(lean_object*, lean_object*); lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedElabHeaderResult___closed__1; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_48____closed__6; @@ -457,12 +452,12 @@ extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__7(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInductiveViews___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkInjectiveTheorems(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_collectUniverses(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withExplicitToImplicit___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___lambda__1___closed__2; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__1___lambda__3(uint8_t, uint8_t, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -514,12 +509,10 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabInd static lean_object* l_Lean_Elab_Command_instInhabitedElabHeaderResult___closed__5; LEAN_EXPORT uint8_t l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_simpLevelIMax_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_getArity___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__7; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__16; -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withExplicitToImplicit___spec__2(lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__9; @@ -600,7 +593,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_co lean_object* l_Lean_Elab_Term_addTermInfo_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors_elabCtorType___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___spec__1(lean_object*, lean_object*, size_t, size_t); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectFVars_visit___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkNumParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_erase___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs___spec__3___boxed(lean_object*, lean_object*); @@ -657,6 +649,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___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_ofLevel(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__4___closed__2; @@ -752,18 +745,17 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_is LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeCompatible___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkParamsAndResultType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__2___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withInductiveLocalDecls_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5_(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors_checkParamOccs___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* l_Lean_mkLevelIMax_x27(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__18; static lean_object* l_Lean_Elab_Command_accLevelAtCtor___closed__5; lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); @@ -827,6 +819,7 @@ lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Command_withCtorRef___rarg___lambda__1(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_reorderCtorArgs___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -837,6 +830,7 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__10; lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withExplicitToImplicit(lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___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_Array_forInUnsafe_loop___at_Lean_Elab_Command_withCtorRef___spec__2___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_isNeverZero(lean_object*); @@ -870,7 +864,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevelAtCtor___lambda__1___boxed( static lean_object* l_Lean_Elab_Command_instInhabitedElabHeaderResult___closed__6; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_accLevelAtCtor___lambda__2___closed__1; -lean_object* l_Lean_simpLevelMax_x27(lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_level_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -891,6 +884,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Induc LEAN_EXPORT lean_object* l_Lean_Elab_Command_withCtorRef___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses_check___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_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_inductive_autoPromoteIndices; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___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_Inductive_0__Lean_Elab_Command_isDomainDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_computeFixedIndexBitMask_go___spec__9___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___closed__2; @@ -919,7 +913,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_el static lean_object* l_List_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__8___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__9___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__10___closed__2; static lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___rarg___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___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_ofName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors_checkParamOccs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Constructions_BRecOn_0__Lean_mkBelowOrIBelow(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -10481,1023 +10474,454 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_st_ref_take(x_6, x_9); -x_11 = lean_ctor_get(x_10, 0); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_9 = lean_st_ref_get(x_5, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 0); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); lean_dec(x_10); -x_14 = !lean_is_exclusive(x_11); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_11, 0); -lean_dec(x_15); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) +x_13 = lean_instantiate_level_mvars(x_12, x_1); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_take(x_5, x_11); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_12, 6); -x_18 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_17, x_1, x_2); -lean_ctor_set(x_12, 6, x_18); -x_19 = lean_st_ref_set(x_6, x_11, x_13); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_17, 0); +lean_dec(x_20); +lean_ctor_set(x_17, 0, x_14); +x_21 = lean_st_ref_set(x_5, x_17, x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_19, 0); -lean_dec(x_21); -x_22 = lean_box(0); -lean_ctor_set(x_19, 0, x_22); -return x_19; +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_15); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_19, 1); -lean_inc(x_23); -lean_dec(x_19); -x_24 = lean_box(0); +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_24); -lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 0, x_15); +lean_ctor_set(x_25, 1, x_24); return x_25; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_26 = lean_ctor_get(x_12, 0); -x_27 = lean_ctor_get(x_12, 1); -x_28 = lean_ctor_get(x_12, 2); -x_29 = lean_ctor_get(x_12, 3); -x_30 = lean_ctor_get(x_12, 4); -x_31 = lean_ctor_get(x_12, 5); -x_32 = lean_ctor_get(x_12, 6); -x_33 = lean_ctor_get(x_12, 7); -x_34 = lean_ctor_get(x_12, 8); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_17, 2); +x_28 = lean_ctor_get(x_17, 3); +x_29 = lean_ctor_get(x_17, 4); lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_inc(x_26); -lean_dec(x_12); -x_35 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_32, x_1, x_2); -x_36 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_36, 0, x_26); -lean_ctor_set(x_36, 1, x_27); -lean_ctor_set(x_36, 2, x_28); -lean_ctor_set(x_36, 3, x_29); -lean_ctor_set(x_36, 4, x_30); -lean_ctor_set(x_36, 5, x_31); -lean_ctor_set(x_36, 6, x_35); -lean_ctor_set(x_36, 7, x_33); -lean_ctor_set(x_36, 8, x_34); -lean_ctor_set(x_11, 0, x_36); -x_37 = lean_st_ref_set(x_6, x_11, x_13); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -if (lean_is_exclusive(x_37)) { - lean_ctor_release(x_37, 0); - lean_ctor_release(x_37, 1); - x_39 = x_37; -} else { - lean_dec_ref(x_37); - x_39 = lean_box(0); -} -x_40 = lean_box(0); -if (lean_is_scalar(x_39)) { - x_41 = lean_alloc_ctor(0, 2, 0); -} else { - x_41 = x_39; -} -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_38); -return x_41; -} -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_42 = lean_ctor_get(x_11, 1); -x_43 = lean_ctor_get(x_11, 2); -x_44 = lean_ctor_get(x_11, 3); -x_45 = lean_ctor_get(x_11, 4); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_11); -x_46 = lean_ctor_get(x_12, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_12, 1); -lean_inc(x_47); -x_48 = lean_ctor_get(x_12, 2); -lean_inc(x_48); -x_49 = lean_ctor_get(x_12, 3); -lean_inc(x_49); -x_50 = lean_ctor_get(x_12, 4); -lean_inc(x_50); -x_51 = lean_ctor_get(x_12, 5); -lean_inc(x_51); -x_52 = lean_ctor_get(x_12, 6); -lean_inc(x_52); -x_53 = lean_ctor_get(x_12, 7); -lean_inc(x_53); -x_54 = lean_ctor_get(x_12, 8); -lean_inc(x_54); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - lean_ctor_release(x_12, 1); - lean_ctor_release(x_12, 2); - lean_ctor_release(x_12, 3); - lean_ctor_release(x_12, 4); - lean_ctor_release(x_12, 5); - lean_ctor_release(x_12, 6); - lean_ctor_release(x_12, 7); - lean_ctor_release(x_12, 8); - x_55 = x_12; -} else { - lean_dec_ref(x_12); - x_55 = lean_box(0); -} -x_56 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_52, x_1, x_2); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 9, 0); -} else { - x_57 = x_55; -} -lean_ctor_set(x_57, 0, x_46); -lean_ctor_set(x_57, 1, x_47); -lean_ctor_set(x_57, 2, x_48); -lean_ctor_set(x_57, 3, x_49); -lean_ctor_set(x_57, 4, x_50); -lean_ctor_set(x_57, 5, x_51); -lean_ctor_set(x_57, 6, x_56); -lean_ctor_set(x_57, 7, x_53); -lean_ctor_set(x_57, 8, x_54); -x_58 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_42); -lean_ctor_set(x_58, 2, x_43); -lean_ctor_set(x_58, 3, x_44); -lean_ctor_set(x_58, 4, x_45); -x_59 = lean_st_ref_set(x_6, x_58, x_13); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_61 = x_59; +lean_dec(x_17); +x_30 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_30, 0, x_14); +lean_ctor_set(x_30, 1, x_26); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_28); +lean_ctor_set(x_30, 4, x_29); +x_31 = lean_st_ref_set(x_5, x_30, x_18); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_33 = x_31; } else { - lean_dec_ref(x_59); - x_61 = lean_box(0); + lean_dec_ref(x_31); + x_33 = lean_box(0); } -x_62 = lean_box(0); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_61; + x_34 = x_33; } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +lean_ctor_set(x_34, 0, x_15); +lean_ctor_set(x_34, 1, x_32); +return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_inc(x_9); -x_10 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; size_t x_13; size_t x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ptr_addr(x_9); -lean_dec(x_9); -x_14 = lean_ptr_addr(x_12); -x_15 = lean_usize_dec_eq(x_13, x_14); -if (x_15 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_16; -lean_dec(x_1); -x_16 = l_Lean_Level_succ___override(x_12); -lean_ctor_set(x_10, 0, x_16); -return x_10; -} -else +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_dec(x_12); -lean_ctor_set(x_10, 0, x_1); -return x_10; -} +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; } else { -lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_10, 0); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_10); -x_19 = lean_ptr_addr(x_9); -lean_dec(x_9); -x_20 = lean_ptr_addr(x_17); -x_21 = lean_usize_dec_eq(x_19, x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -lean_dec(x_1); -x_22 = l_Lean_Level_succ___override(x_17); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_18); -return x_23; -} -else -{ -lean_object* x_24; -lean_dec(x_17); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_1); -lean_ctor_set(x_24, 1, x_18); -return x_24; -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -case 2: +else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_25 = lean_ctor_get(x_1, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 1); -lean_inc(x_26); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -x_27 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_30 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_29); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_34 = lean_ptr_addr(x_28); -x_35 = lean_usize_dec_eq(x_33, x_34); -if (x_35 == 0) -{ -lean_object* x_36; -lean_dec(x_26); -lean_dec(x_1); -x_36 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_36); -return x_30; +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); } -else -{ -size_t x_37; size_t x_38; uint8_t x_39; -x_37 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_38 = lean_ptr_addr(x_32); -x_39 = lean_usize_dec_eq(x_37, x_38); -if (x_39 == 0) -{ -lean_object* x_40; -lean_dec(x_1); -x_40 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_40); -return x_30; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); } -else -{ -lean_object* x_41; -x_41 = l_Lean_simpLevelMax_x27(x_28, x_32, x_1); -lean_dec(x_1); -lean_dec(x_32); -lean_dec(x_28); -lean_ctor_set(x_30, 0, x_41); -return x_30; +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; } } } -else -{ -lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; uint8_t x_46; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -x_44 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_45 = lean_ptr_addr(x_28); -x_46 = lean_usize_dec_eq(x_44, x_45); -if (x_46 == 0) +static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__1() { +_start: { -lean_object* x_47; lean_object* x_48; -lean_dec(x_26); -lean_dec(x_1); -x_47 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_43); -return x_48; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot infer resulting universe level of inductive datatype, given level contains metavariables ", 96, 96); +return x_1; } -else -{ -size_t x_49; size_t x_50; uint8_t x_51; -x_49 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_50 = lean_ptr_addr(x_42); -x_51 = lean_usize_dec_eq(x_49, x_50); -if (x_51 == 0) +} +static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__2() { +_start: { -lean_object* x_52; lean_object* x_53; -lean_dec(x_1); -x_52 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_43); -return x_53; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__3() { +_start: { -lean_object* x_54; lean_object* x_55; -x_54 = l_Lean_simpLevelMax_x27(x_28, x_42, x_1); -lean_dec(x_1); -lean_dec(x_42); -lean_dec(x_28); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_43); -return x_55; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", provide universe explicitly", 29, 29); +return x_1; } } +static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -case 3: +LEAN_EXPORT lean_object* l_Lean_Elab_Command_shouldInferResultUniverse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_56 = lean_ctor_get(x_1, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -lean_inc(x_56); -x_58 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -lean_inc(x_57); -x_61 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_60); -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_63; size_t x_64; size_t x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_61, 0); -x_64 = lean_ptr_addr(x_56); -lean_dec(x_56); -x_65 = lean_ptr_addr(x_59); -x_66 = lean_usize_dec_eq(x_64, x_65); -if (x_66 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = l_Lean_Level_hasMVar(x_11); +if (x_13 == 0) { -lean_object* x_67; -lean_dec(x_57); -lean_dec(x_1); -x_67 = l_Lean_mkLevelIMax_x27(x_59, x_63); -lean_ctor_set(x_61, 0, x_67); -return x_61; +lean_object* x_14; +lean_dec(x_11); +lean_dec(x_2); +x_14 = lean_box(0); +lean_ctor_set(x_9, 0, x_14); +return x_9; } else { -size_t x_68; size_t x_69; uint8_t x_70; -x_68 = lean_ptr_addr(x_57); -lean_dec(x_57); -x_69 = lean_ptr_addr(x_63); -x_70 = lean_usize_dec_eq(x_68, x_69); -if (x_70 == 0) +lean_object* x_15; +x_15 = l_Lean_Level_getLevelOffset(x_11); +if (lean_obj_tag(x_15) == 5) { -lean_object* x_71; -lean_dec(x_1); -x_71 = l_Lean_mkLevelIMax_x27(x_59, x_63); -lean_ctor_set(x_61, 0, x_71); -return x_61; +lean_object* x_16; lean_object* x_17; +lean_dec(x_11); +lean_dec(x_2); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_9, 0, x_17); +return x_9; } else { -lean_object* x_72; -x_72 = l_Lean_simpLevelIMax_x27(x_59, x_63, x_1); -lean_dec(x_1); -lean_ctor_set(x_61, 0, x_72); -return x_61; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_15); +lean_free_object(x_9); +x_18 = l_Lean_Expr_sort___override(x_11); +x_19 = l_Lean_MessageData_ofExpr(x_18); +x_20 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__2; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__4; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +return x_24; } } } else { -lean_object* x_73; lean_object* x_74; size_t x_75; size_t x_76; uint8_t x_77; -x_73 = lean_ctor_get(x_61, 0); -x_74 = lean_ctor_get(x_61, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_61); -x_75 = lean_ptr_addr(x_56); -lean_dec(x_56); -x_76 = lean_ptr_addr(x_59); -x_77 = lean_usize_dec_eq(x_75, x_76); -if (x_77 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_9, 0); +x_26 = lean_ctor_get(x_9, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_9); +x_27 = l_Lean_Level_hasMVar(x_25); +if (x_27 == 0) { -lean_object* x_78; lean_object* x_79; -lean_dec(x_57); -lean_dec(x_1); -x_78 = l_Lean_mkLevelIMax_x27(x_59, x_73); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_74); -return x_79; +lean_object* x_28; lean_object* x_29; +lean_dec(x_25); +lean_dec(x_2); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; } else { -size_t x_80; size_t x_81; uint8_t x_82; -x_80 = lean_ptr_addr(x_57); -lean_dec(x_57); -x_81 = lean_ptr_addr(x_73); -x_82 = lean_usize_dec_eq(x_80, x_81); -if (x_82 == 0) +lean_object* x_30; +x_30 = l_Lean_Level_getLevelOffset(x_25); +if (lean_obj_tag(x_30) == 5) { -lean_object* x_83; lean_object* x_84; -lean_dec(x_1); -x_83 = l_Lean_mkLevelIMax_x27(x_59, x_73); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_74); -return x_84; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_25); +lean_dec(x_2); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_26); +return x_33; } else { -lean_object* x_85; lean_object* x_86; -x_85 = l_Lean_simpLevelIMax_x27(x_59, x_73, x_1); -lean_dec(x_1); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_74); -return x_86; -} -} -} -} -case 5: -{ -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = lean_ctor_get(x_1, 0); -lean_inc(x_87); -x_88 = lean_st_ref_get(x_5, x_8); -x_89 = !lean_is_exclusive(x_88); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = lean_ctor_get(x_88, 0); -x_91 = lean_ctor_get(x_88, 1); -x_92 = lean_ctor_get(x_90, 0); -lean_inc(x_92); -lean_dec(x_90); -x_93 = lean_ctor_get(x_92, 6); -lean_inc(x_93); -lean_dec(x_92); -x_94 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_93, x_87); -if (lean_obj_tag(x_94) == 0) -{ -lean_dec(x_87); -lean_ctor_set(x_88, 0, x_1); -return x_88; -} -else -{ -lean_object* x_95; uint8_t x_96; -lean_dec(x_1); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -lean_dec(x_94); -x_96 = l_Lean_Level_hasMVar(x_95); -if (x_96 == 0) -{ -lean_dec(x_87); -lean_ctor_set(x_88, 0, x_95); -return x_88; -} -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -lean_free_object(x_88); -x_97 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_91); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -lean_inc(x_98); -x_100 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_87, x_98, x_2, x_3, x_4, x_5, x_6, x_7, x_99); -x_101 = !lean_is_exclusive(x_100); -if (x_101 == 0) -{ -lean_object* x_102; -x_102 = lean_ctor_get(x_100, 0); -lean_dec(x_102); -lean_ctor_set(x_100, 0, x_98); -return x_100; +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_dec(x_30); +x_34 = l_Lean_Expr_sort___override(x_25); +x_35 = l_Lean_MessageData_ofExpr(x_34); +x_36 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__2; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__4; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_26); +return x_40; } -else -{ -lean_object* x_103; lean_object* x_104; -x_103 = lean_ctor_get(x_100, 1); -lean_inc(x_103); -lean_dec(x_100); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_98); -lean_ctor_set(x_104, 1, x_103); -return x_104; } } } } -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_105 = lean_ctor_get(x_88, 0); -x_106 = lean_ctor_get(x_88, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_88); -x_107 = lean_ctor_get(x_105, 0); -lean_inc(x_107); -lean_dec(x_105); -x_108 = lean_ctor_get(x_107, 6); -lean_inc(x_108); -lean_dec(x_107); -x_109 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_108, x_87); -if (lean_obj_tag(x_109) == 0) +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_110; -lean_dec(x_87); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_1); -lean_ctor_set(x_110, 1, x_106); -return x_110; +lean_object* x_9; +x_9 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; } -else -{ -lean_object* x_111; uint8_t x_112; -lean_dec(x_1); -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -lean_dec(x_109); -x_112 = l_Lean_Level_hasMVar(x_111); -if (x_112 == 0) -{ -lean_object* x_113; -lean_dec(x_87); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_106); -return x_113; } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_114 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_111, x_2, x_3, x_4, x_5, x_6, x_7, x_106); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -lean_inc(x_115); -x_117 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_87, x_115, x_2, x_3, x_4, x_5, x_6, x_7, x_116); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_119 = x_117; -} else { - lean_dec_ref(x_117); - x_119 = lean_box(0); -} -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(0, 2, 0); -} else { - x_120 = x_119; -} -lean_ctor_set(x_120, 0, x_115); -lean_ctor_set(x_120, 1, x_118); -return x_120; -} -} +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; } } -default: +LEAN_EXPORT lean_object* l_Lean_Elab_Command_shouldInferResultUniverse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_121; -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_1); -lean_ctor_set(x_121, 1, x_8); -return x_121; -} -} +lean_object* x_9; +x_9 = l_Lean_Elab_Command_shouldInferResultUniverse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_6, 5); -x_10 = lean_ctor_get(x_2, 2); -lean_inc(x_10); -lean_inc(x_10); -x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -lean_dec(x_2); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 0); -lean_ctor_set(x_12, 1, x_18); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set_tag(x_16, 1); -lean_ctor_set(x_16, 0, x_12); -return x_16; +uint8_t x_3; +x_3 = 1; +return x_3; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_16); -lean_ctor_set(x_12, 1, x_19); -lean_ctor_set(x_12, 0, x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_20); -return x_21; +uint8_t x_4; +x_4 = 0; +return x_4; } } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_12, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_12); -x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); -lean_dec(x_2); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; -} else { - lean_dec_ref(x_24); - x_27 = lean_box(0); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; } -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_11); -lean_ctor_set(x_28, 1, x_25); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(1, 2, 0); -} else { - x_29 = x_27; - lean_ctor_set_tag(x_29, 1); +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_name_eq(x_6, x_7); +return x_8; } -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; } } } -static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__1() { +LEAN_EXPORT uint8_t l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("cannot infer resulting universe level of inductive datatype, given level contains metavariables ", 96, 96); -return x_1; +lean_object* x_3; uint8_t x_4; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +x_4 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); +lean_dec(x_3); +return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1___boxed), 2, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Lean_Elab_Term_levelMVarToParam(x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__3() { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(", provide universe explicitly", 29, 29); -return x_1; +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_shouldInferResultUniverse___closed__4() { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1(x_1, x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_shouldInferResultUniverse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = l_Lean_Level_hasMVar(x_11); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_11); -lean_dec(x_2); -x_14 = lean_box(0); -lean_ctor_set(x_9, 0, x_14); -return x_9; -} -else -{ -lean_object* x_15; -x_15 = l_Lean_Level_getLevelOffset(x_11); -if (lean_obj_tag(x_15) == 5) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_11); -lean_dec(x_2); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_9, 0, x_17); -return x_9; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_15); -lean_free_object(x_9); -x_18 = l_Lean_Expr_sort___override(x_11); -x_19 = l_Lean_MessageData_ofExpr(x_18); -x_20 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__2; -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__4; -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -return x_24; -} -} -} -else -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_ctor_get(x_9, 0); -x_26 = lean_ctor_get(x_9, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_9); -x_27 = l_Lean_Level_hasMVar(x_25); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -lean_dec(x_25); -lean_dec(x_2); -x_28 = lean_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; -} -else -{ -lean_object* x_30; -x_30 = l_Lean_Level_getLevelOffset(x_25); -if (lean_obj_tag(x_30) == 5) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_25); -lean_dec(x_2); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -lean_dec(x_30); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_26); -return x_33; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_30); -x_34 = l_Lean_Expr_sort___override(x_25); -x_35 = l_Lean_MessageData_ofExpr(x_34); -x_36 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__2; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Elab_Command_shouldInferResultUniverse___closed__4; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_26); -return x_40; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___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) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_shouldInferResultUniverse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Elab_Command_shouldInferResultUniverse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_name_eq(x_6, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT uint8_t l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; uint8_t x_4; -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -x_4 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1___boxed), 2, 1); -lean_closure_set(x_10, 0, x_1); -x_11 = l_Lean_Elab_Term_levelMVarToParam(x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___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_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1(x_1, x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; @@ -11767,1214 +11191,659 @@ x_49 = lean_box(0); lean_inc(x_1); x_50 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__1(x_1, x_44, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_48); 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); -if (lean_is_scalar(x_45)) { - x_53 = lean_alloc_ctor(0, 3, 0); -} else { - x_53 = x_45; -} -lean_ctor_set(x_53, 0, x_42); -lean_ctor_set(x_53, 1, x_47); -lean_ctor_set(x_53, 2, x_51); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_3); -x_2 = x_41; -x_3 = x_54; -x_10 = x_52; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_box(0); -x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2(x_2, x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkResultUniverse(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = l_Array_isEmpty___rarg(x_1); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_array_to_list(lean_box(0), x_1); -x_5 = l_Lean_Level_mkNaryMax(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_2, x_6); -if (x_7 == 0) -{ -lean_object* x_8; -x_8 = l_Lean_Level_normalize(x_5); -lean_dec(x_5); -return x_8; -} -else -{ -uint8_t x_9; -x_9 = l_Lean_Level_isZero(x_5); -if (x_9 == 0) -{ -uint8_t x_10; -x_10 = l_Lean_Level_isNeverZero(x_5); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_Lean_levelOne; -x_12 = l_Lean_Level_max___override(x_5, x_11); -x_13 = l_Lean_Level_normalize(x_12); -lean_dec(x_12); -return x_13; -} -else -{ -lean_object* x_14; -x_14 = l_Lean_Level_normalize(x_5); -lean_dec(x_5); -return x_14; -} -} -else -{ -lean_object* x_15; -x_15 = l_Lean_Level_normalize(x_5); -lean_dec(x_5); -return x_15; -} -} -} -else -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_2, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_to_list(lean_box(0), x_1); -x_19 = l_Lean_Level_mkNaryMax(x_18); -x_20 = l_Lean_Level_normalize(x_19); -lean_dec(x_19); -return x_20; -} -else -{ -lean_object* x_21; -lean_dec(x_1); -x_21 = l_Lean_levelOne; -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkResultUniverse___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Elab_Command_mkResultUniverse(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_level_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Elab_Command_accLevel_go___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_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_lt(x_4, x_3); -if (x_5 == 0) -{ -uint8_t x_6; -lean_dec(x_3); -x_6 = 0; -return x_6; -} -else -{ -size_t x_7; size_t x_8; uint8_t x_9; -x_7 = 0; -x_8 = lean_usize_of_nat(x_3); -lean_dec(x_3); -x_9 = l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2(x_2, x_1, x_7, x_8); -return x_9; -} -} -} -static lean_object* _init_l_Lean_Elab_Command_accLevel_go___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -switch (lean_obj_tag(x_2)) { -case 0: -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_3); -x_20 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_4); -return x_21; -} -case 1: -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_dec_eq(x_3, x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -lean_dec(x_2); -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_sub(x_3, x_25); -lean_dec(x_3); -x_2 = x_22; -x_3 = x_26; -goto _start; -} -else -{ -uint8_t x_28; -lean_dec(x_22); -lean_dec(x_3); -x_28 = lean_level_eq(x_2, x_1); -if (x_28 == 0) -{ -uint8_t x_29; -x_29 = l_Lean_Level_occurs(x_1, x_2); -if (x_29 == 0) -{ -uint8_t x_30; -x_30 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(x_4, x_2); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_array_push(x_4, x_2); -x_32 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; -} -else -{ -lean_object* x_34; lean_object* x_35; -lean_dec(x_2); -x_34 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_4); -return x_35; -} -} -else -{ -lean_object* x_36; lean_object* x_37; -lean_dec(x_2); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_4); -return x_37; -} -} -else -{ -lean_object* x_38; lean_object* x_39; -lean_dec(x_2); -x_38 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_4); -return x_39; -} -} -} -case 4: -{ -lean_object* x_40; uint8_t x_41; -x_40 = lean_unsigned_to_nat(0u); -x_41 = lean_nat_dec_eq(x_3, x_40); -if (x_41 == 0) -{ -lean_object* x_42; -x_42 = lean_box(0); -x_5 = x_42; -goto block_19; -} -else -{ -uint8_t x_43; -x_43 = lean_level_eq(x_2, x_1); -if (x_43 == 0) -{ -lean_object* x_44; -x_44 = lean_box(0); -x_5 = x_44; -goto block_19; -} -else -{ -lean_object* x_45; lean_object* x_46; -lean_dec(x_3); -lean_dec(x_2); -x_45 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_4); -return x_46; -} -} -} -case 5: -{ -lean_object* x_47; uint8_t x_48; -x_47 = lean_unsigned_to_nat(0u); -x_48 = lean_nat_dec_eq(x_3, x_47); -if (x_48 == 0) -{ -lean_object* x_49; -x_49 = lean_box(0); -x_5 = x_49; -goto block_19; -} -else -{ -uint8_t x_50; -x_50 = lean_level_eq(x_2, x_1); -if (x_50 == 0) -{ -lean_object* x_51; -x_51 = lean_box(0); -x_5 = x_51; -goto block_19; -} -else -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_3); -lean_dec(x_2); -x_52 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_4); -return x_53; -} -} -} -default: -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_2, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_2, 1); -lean_inc(x_55); -lean_dec(x_2); -lean_inc(x_3); -x_56 = l_Lean_Elab_Command_accLevel_go(x_1, x_54, x_3, x_4); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -if (lean_obj_tag(x_57) == 0) -{ -uint8_t x_58; -lean_dec(x_55); -lean_dec(x_3); -x_58 = !lean_is_exclusive(x_56); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_56, 0); -lean_dec(x_59); -x_60 = lean_box(0); -lean_ctor_set(x_56, 0, x_60); -return x_56; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_56, 1); -lean_inc(x_61); -lean_dec(x_56); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_61); -return x_63; -} -} -else -{ -lean_object* x_64; -lean_dec(x_57); -x_64 = lean_ctor_get(x_56, 1); -lean_inc(x_64); -lean_dec(x_56); -x_2 = x_55; -x_4 = x_64; -goto _start; -} -} -} -block_19: -{ -uint8_t x_6; -lean_dec(x_5); -x_6 = l_Lean_Level_occurs(x_1, x_2); -if (x_6 == 0) -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_lt(x_7, x_3); -lean_dec(x_3); -if (x_8 == 0) -{ -uint8_t x_9; -x_9 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(x_4, x_2); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_array_push(x_4, x_2); -x_11 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -x_13 = l_Lean_Elab_Command_accLevel_go___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_4); -return x_14; -} -} -else -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_15 = lean_box(0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_4); -return x_16; -} -} -else -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_3); -lean_dec(x_2); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_4); -return x_18; -} +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +if (lean_is_scalar(x_45)) { + x_53 = lean_alloc_ctor(0, 3, 0); +} else { + x_53 = x_45; } +lean_ctor_set(x_53, 0, x_42); +lean_ctor_set(x_53, 1, x_47); +lean_ctor_set(x_53, 2, x_51); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_3); +x_2 = x_41; +x_3 = x_54; +x_10 = x_52; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; } } -LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___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_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2(x_2, x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_5; -x_5 = l_Lean_Elab_Command_accLevel_go(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; +lean_object* x_11; +x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_5; -x_5 = l_Lean_Elab_Command_accLevel_go(x_2, x_1, x_3, x_4); -return x_5; +lean_object* x_11; +x_11 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_5; -x_5 = l_Lean_Elab_Command_accLevel(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; +lean_object* x_10; +x_10 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkResultUniverse(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_11 = lean_st_ref_take(x_7, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = !lean_is_exclusive(x_12); -if (x_15 == 0) +uint8_t x_3; +x_3 = l_Array_isEmpty___rarg(x_1); +if (x_3 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_12, 0); -lean_dec(x_16); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_array_to_list(lean_box(0), x_1); +x_5 = l_Lean_Level_mkNaryMax(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_2, x_6); +if (x_7 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_13, 6); -x_19 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_18, x_1, x_2); -lean_ctor_set(x_13, 6, x_19); -x_20 = lean_st_ref_set(x_7, x_12, x_14); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_8; +x_8 = l_Lean_Level_normalize(x_5); +lean_dec(x_5); +return x_8; +} +else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_20, 0, x_23); -return x_20; +uint8_t x_9; +x_9 = l_Lean_Level_isZero(x_5); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = l_Lean_Level_isNeverZero(x_5); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_Lean_levelOne; +x_12 = l_Lean_Level_max___override(x_5, x_11); +x_13 = l_Lean_Level_normalize(x_12); +lean_dec(x_12); +return x_13; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; +lean_object* x_14; +x_14 = l_Lean_Level_normalize(x_5); +lean_dec(x_5); +return x_14; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); -x_29 = lean_ctor_get(x_13, 2); -x_30 = lean_ctor_get(x_13, 3); -x_31 = lean_ctor_get(x_13, 4); -x_32 = lean_ctor_get(x_13, 5); -x_33 = lean_ctor_get(x_13, 6); -x_34 = lean_ctor_get(x_13, 7); -x_35 = lean_ctor_get(x_13, 8); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_13); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_33, x_1, x_2); -x_37 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_37, 0, x_27); -lean_ctor_set(x_37, 1, x_28); -lean_ctor_set(x_37, 2, x_29); -lean_ctor_set(x_37, 3, x_30); -lean_ctor_set(x_37, 4, x_31); -lean_ctor_set(x_37, 5, x_32); -lean_ctor_set(x_37, 6, x_36); -lean_ctor_set(x_37, 7, x_34); -lean_ctor_set(x_37, 8, x_35); -lean_ctor_set(x_12, 0, x_37); -x_38 = lean_st_ref_set(x_7, x_12, x_14); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; -} else { - lean_dec_ref(x_38); - x_40 = lean_box(0); -} -x_41 = lean_box(0); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_40; +lean_object* x_15; +x_15 = l_Lean_Level_normalize(x_5); +lean_dec(x_5); +return x_15; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_43 = lean_ctor_get(x_12, 1); -x_44 = lean_ctor_get(x_12, 2); -x_45 = lean_ctor_get(x_12, 3); -x_46 = lean_ctor_get(x_12, 4); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_12); -x_47 = lean_ctor_get(x_13, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_13, 1); -lean_inc(x_48); -x_49 = lean_ctor_get(x_13, 2); -lean_inc(x_49); -x_50 = lean_ctor_get(x_13, 3); -lean_inc(x_50); -x_51 = lean_ctor_get(x_13, 4); -lean_inc(x_51); -x_52 = lean_ctor_get(x_13, 5); -lean_inc(x_52); -x_53 = lean_ctor_get(x_13, 6); -lean_inc(x_53); -x_54 = lean_ctor_get(x_13, 7); -lean_inc(x_54); -x_55 = lean_ctor_get(x_13, 8); -lean_inc(x_55); -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); - lean_ctor_release(x_13, 5); - lean_ctor_release(x_13, 6); - lean_ctor_release(x_13, 7); - lean_ctor_release(x_13, 8); - x_56 = x_13; -} else { - lean_dec_ref(x_13); - x_56 = lean_box(0); -} -x_57 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_53, x_1, x_2); -if (lean_is_scalar(x_56)) { - x_58 = lean_alloc_ctor(0, 9, 0); -} else { - x_58 = x_56; -} -lean_ctor_set(x_58, 0, x_47); -lean_ctor_set(x_58, 1, x_48); -lean_ctor_set(x_58, 2, x_49); -lean_ctor_set(x_58, 3, x_50); -lean_ctor_set(x_58, 4, x_51); -lean_ctor_set(x_58, 5, x_52); -lean_ctor_set(x_58, 6, x_57); -lean_ctor_set(x_58, 7, x_54); -lean_ctor_set(x_58, 8, x_55); -x_59 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_43); -lean_ctor_set(x_59, 2, x_44); -lean_ctor_set(x_59, 3, x_45); -lean_ctor_set(x_59, 4, x_46); -x_60 = lean_st_ref_set(x_7, x_59, x_14); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_62 = x_60; -} else { - lean_dec_ref(x_60); - x_62 = lean_box(0); +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_2, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_array_to_list(lean_box(0), x_1); +x_19 = l_Lean_Level_mkNaryMax(x_18); +x_20 = l_Lean_Level_normalize(x_19); +lean_dec(x_19); +return x_20; } -x_63 = lean_box(0); -if (lean_is_scalar(x_62)) { - x_64 = lean_alloc_ctor(0, 2, 0); -} else { - x_64 = x_62; +else +{ +lean_object* x_21; +lean_dec(x_1); +x_21 = l_Lean_levelOne; +return x_21; } -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_61); -return x_64; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkResultUniverse___boxed(lean_object* x_1, lean_object* x_2) { _start: { -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_inc(x_10); -x_11 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_3; +x_3 = l_Lean_Elab_Command_mkResultUniverse(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: { -lean_object* x_13; size_t x_14; size_t x_15; uint8_t x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ptr_addr(x_10); -lean_dec(x_10); -x_15 = lean_ptr_addr(x_13); -x_16 = lean_usize_dec_eq(x_14, x_15); -if (x_16 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) { -lean_object* x_17; -lean_dec(x_1); -x_17 = l_Lean_Level_succ___override(x_13); -lean_ctor_set(x_11, 0, x_17); -return x_11; +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_level_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; } else { -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_1); +uint8_t x_11; +x_11 = 1; return x_11; } } else { -lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; uint8_t x_22; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_ptr_addr(x_10); -lean_dec(x_10); -x_21 = lean_ptr_addr(x_18); -x_22 = lean_usize_dec_eq(x_20, x_21); -if (x_22 == 0) +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_23; lean_object* x_24; -lean_dec(x_1); -x_23 = l_Lean_Level_succ___override(x_18); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_19); -return x_24; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_array_get_size(x_1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +uint8_t x_6; +lean_dec(x_3); +x_6 = 0; +return x_6; } else { -lean_object* x_25; -lean_dec(x_18); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_1); -lean_ctor_set(x_25, 1, x_19); -return x_25; +size_t x_7; size_t x_8; uint8_t x_9; +x_7 = 0; +x_8 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_9 = l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2(x_2, x_1, x_7, x_8); +return x_9; } } } -case 2: +static lean_object* _init_l_Lean_Elab_Command_accLevel_go___closed__1() { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_26 = lean_ctor_get(x_1, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_1, 1); -lean_inc(x_27); -lean_inc(x_26); -x_28 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_27); -x_31 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_33; size_t x_34; size_t x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_31, 0); -x_34 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_35 = lean_ptr_addr(x_29); -x_36 = lean_usize_dec_eq(x_34, x_35); -if (x_36 == 0) +lean_object* x_5; +switch (lean_obj_tag(x_2)) { +case 0: { -lean_object* x_37; -lean_dec(x_27); -lean_dec(x_1); -x_37 = l_Lean_mkLevelMax_x27(x_29, x_33); -lean_ctor_set(x_31, 0, x_37); -return x_31; +lean_object* x_20; lean_object* x_21; +lean_dec(x_3); +x_20 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_4); +return x_21; +} +case 1: +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_2, 0); +lean_inc(x_22); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_eq(x_3, x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_2); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_sub(x_3, x_25); +lean_dec(x_3); +x_2 = x_22; +x_3 = x_26; +goto _start; } else { -size_t x_38; size_t x_39; uint8_t x_40; -x_38 = lean_ptr_addr(x_27); -lean_dec(x_27); -x_39 = lean_ptr_addr(x_33); -x_40 = lean_usize_dec_eq(x_38, x_39); -if (x_40 == 0) +uint8_t x_28; +lean_dec(x_22); +lean_dec(x_3); +x_28 = lean_level_eq(x_2, x_1); +if (x_28 == 0) { -lean_object* x_41; -lean_dec(x_1); -x_41 = l_Lean_mkLevelMax_x27(x_29, x_33); -lean_ctor_set(x_31, 0, x_41); -return x_31; +uint8_t x_29; +x_29 = l_Lean_Level_occurs(x_1, x_2); +if (x_29 == 0) +{ +uint8_t x_30; +x_30 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(x_4, x_2); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_array_push(x_4, x_2); +x_32 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } else { -lean_object* x_42; -x_42 = l_Lean_simpLevelMax_x27(x_29, x_33, x_1); -lean_dec(x_1); -lean_dec(x_33); -lean_dec(x_29); -lean_ctor_set(x_31, 0, x_42); -return x_31; +lean_object* x_34; lean_object* x_35; +lean_dec(x_2); +x_34 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_4); +return x_35; +} } +else +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_2); +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_4); +return x_37; } } else { -lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; uint8_t x_47; -x_43 = lean_ctor_get(x_31, 0); -x_44 = lean_ctor_get(x_31, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_31); -x_45 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_46 = lean_ptr_addr(x_29); -x_47 = lean_usize_dec_eq(x_45, x_46); -if (x_47 == 0) +lean_object* x_38; lean_object* x_39; +lean_dec(x_2); +x_38 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_4); +return x_39; +} +} +} +case 4: { -lean_object* x_48; lean_object* x_49; -lean_dec(x_27); -lean_dec(x_1); -x_48 = l_Lean_mkLevelMax_x27(x_29, x_43); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_44); -return x_49; +lean_object* x_40; uint8_t x_41; +x_40 = lean_unsigned_to_nat(0u); +x_41 = lean_nat_dec_eq(x_3, x_40); +if (x_41 == 0) +{ +lean_object* x_42; +x_42 = lean_box(0); +x_5 = x_42; +goto block_19; } else { -size_t x_50; size_t x_51; uint8_t x_52; -x_50 = lean_ptr_addr(x_27); -lean_dec(x_27); -x_51 = lean_ptr_addr(x_43); -x_52 = lean_usize_dec_eq(x_50, x_51); -if (x_52 == 0) +uint8_t x_43; +x_43 = lean_level_eq(x_2, x_1); +if (x_43 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_1); -x_53 = l_Lean_mkLevelMax_x27(x_29, x_43); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_44); -return x_54; +lean_object* x_44; +x_44 = lean_box(0); +x_5 = x_44; +goto block_19; } else { -lean_object* x_55; lean_object* x_56; -x_55 = l_Lean_simpLevelMax_x27(x_29, x_43, x_1); -lean_dec(x_1); -lean_dec(x_43); -lean_dec(x_29); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_44); -return x_56; -} +lean_object* x_45; lean_object* x_46; +lean_dec(x_3); +lean_dec(x_2); +x_45 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_4); +return x_46; } } } -case 3: -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_57 = lean_ctor_get(x_1, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_1, 1); -lean_inc(x_58); -lean_inc(x_57); -x_59 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -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); -lean_inc(x_58); -x_62 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) +case 5: { -lean_object* x_64; size_t x_65; size_t x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_62, 0); -x_65 = lean_ptr_addr(x_57); -lean_dec(x_57); -x_66 = lean_ptr_addr(x_60); -x_67 = lean_usize_dec_eq(x_65, x_66); -if (x_67 == 0) +lean_object* x_47; uint8_t x_48; +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_nat_dec_eq(x_3, x_47); +if (x_48 == 0) { -lean_object* x_68; -lean_dec(x_58); -lean_dec(x_1); -x_68 = l_Lean_mkLevelIMax_x27(x_60, x_64); -lean_ctor_set(x_62, 0, x_68); -return x_62; +lean_object* x_49; +x_49 = lean_box(0); +x_5 = x_49; +goto block_19; } else { -size_t x_69; size_t x_70; uint8_t x_71; -x_69 = lean_ptr_addr(x_58); -lean_dec(x_58); -x_70 = lean_ptr_addr(x_64); -x_71 = lean_usize_dec_eq(x_69, x_70); -if (x_71 == 0) +uint8_t x_50; +x_50 = lean_level_eq(x_2, x_1); +if (x_50 == 0) { -lean_object* x_72; -lean_dec(x_1); -x_72 = l_Lean_mkLevelIMax_x27(x_60, x_64); -lean_ctor_set(x_62, 0, x_72); -return x_62; +lean_object* x_51; +x_51 = lean_box(0); +x_5 = x_51; +goto block_19; } else { -lean_object* x_73; -x_73 = l_Lean_simpLevelIMax_x27(x_60, x_64, x_1); -lean_dec(x_1); -lean_ctor_set(x_62, 0, x_73); -return x_62; +lean_object* x_52; lean_object* x_53; +lean_dec(x_3); +lean_dec(x_2); +x_52 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_4); +return x_53; } } } -else +default: { -lean_object* x_74; lean_object* x_75; size_t x_76; size_t x_77; uint8_t x_78; -x_74 = lean_ctor_get(x_62, 0); -x_75 = lean_ctor_get(x_62, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_62); -x_76 = lean_ptr_addr(x_57); -lean_dec(x_57); -x_77 = lean_ptr_addr(x_60); -x_78 = lean_usize_dec_eq(x_76, x_77); -if (x_78 == 0) +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_2, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_2, 1); +lean_inc(x_55); +lean_dec(x_2); +lean_inc(x_3); +x_56 = l_Lean_Elab_Command_accLevel_go(x_1, x_54, x_3, x_4); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_79; lean_object* x_80; -lean_dec(x_58); -lean_dec(x_1); -x_79 = l_Lean_mkLevelIMax_x27(x_60, x_74); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_75); -return x_80; +uint8_t x_58; +lean_dec(x_55); +lean_dec(x_3); +x_58 = !lean_is_exclusive(x_56); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 0); +lean_dec(x_59); +x_60 = lean_box(0); +lean_ctor_set(x_56, 0, x_60); +return x_56; } else { -size_t x_81; size_t x_82; uint8_t x_83; -x_81 = lean_ptr_addr(x_58); -lean_dec(x_58); -x_82 = lean_ptr_addr(x_74); -x_83 = lean_usize_dec_eq(x_81, x_82); -if (x_83 == 0) -{ -lean_object* x_84; lean_object* x_85; -lean_dec(x_1); -x_84 = l_Lean_mkLevelIMax_x27(x_60, x_74); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_75); -return x_85; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_dec(x_56); +x_62 = lean_box(0); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} } else { -lean_object* x_86; lean_object* x_87; -x_86 = l_Lean_simpLevelIMax_x27(x_60, x_74, x_1); -lean_dec(x_1); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_75); -return x_87; -} +lean_object* x_64; +lean_dec(x_57); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_2 = x_55; +x_4 = x_64; +goto _start; } } } -case 5: +block_19: { -lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_88 = lean_ctor_get(x_1, 0); -lean_inc(x_88); -x_89 = lean_st_ref_get(x_6, x_9); -x_90 = !lean_is_exclusive(x_89); -if (x_90 == 0) +uint8_t x_6; +lean_dec(x_5); +x_6 = l_Lean_Level_occurs(x_1, x_2); +if (x_6 == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_91 = lean_ctor_get(x_89, 0); -x_92 = lean_ctor_get(x_89, 1); -x_93 = lean_ctor_get(x_91, 0); -lean_inc(x_93); -lean_dec(x_91); -x_94 = lean_ctor_get(x_93, 6); -lean_inc(x_94); -lean_dec(x_93); -x_95 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_94, x_88); -if (lean_obj_tag(x_95) == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_lt(x_7, x_3); +lean_dec(x_3); +if (x_8 == 0) { -lean_dec(x_88); -lean_ctor_set(x_89, 0, x_1); -return x_89; +uint8_t x_9; +x_9 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(x_4, x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_push(x_4, x_2); +x_11 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_96; uint8_t x_97; -lean_dec(x_1); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -lean_dec(x_95); -x_97 = l_Lean_Level_hasMVar(x_96); -if (x_97 == 0) -{ -lean_dec(x_88); -lean_ctor_set(x_89, 0, x_96); -return x_89; +lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +x_13 = l_Lean_Elab_Command_accLevel_go___closed__1; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_4); +return x_14; +} } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; -lean_free_object(x_89); -x_98 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_92); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -lean_inc(x_99); -x_101 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(x_88, x_99, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_100); -x_102 = !lean_is_exclusive(x_101); -if (x_102 == 0) -{ -lean_object* x_103; -x_103 = lean_ctor_get(x_101, 0); -lean_dec(x_103); -lean_ctor_set(x_101, 0, x_99); -return x_101; +lean_object* x_15; lean_object* x_16; +lean_dec(x_2); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_4); +return x_16; +} } else { -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_101, 1); -lean_inc(x_104); -lean_dec(x_101); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_99); -lean_ctor_set(x_105, 1, x_104); -return x_105; +lean_object* x_17; lean_object* x_18; +lean_dec(x_3); +lean_dec(x_2); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_4); +return x_18; } } } } -else -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_106 = lean_ctor_get(x_89, 0); -x_107 = lean_ctor_get(x_89, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_89); -x_108 = lean_ctor_get(x_106, 0); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_ctor_get(x_108, 6); -lean_inc(x_109); -lean_dec(x_108); -x_110 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_109, x_88); -if (lean_obj_tag(x_110) == 0) +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Command_accLevel_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_111; -lean_dec(x_88); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_1); -lean_ctor_set(x_111, 1, x_107); -return x_111; +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_Elab_Command_accLevel_go___spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; } -else +} +LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_112; uint8_t x_113; +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_contains___at_Lean_Elab_Command_accLevel_go___spec__1(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); -x_112 = lean_ctor_get(x_110, 0); -lean_inc(x_112); -lean_dec(x_110); -x_113 = l_Lean_Level_hasMVar(x_112); -if (x_113 == 0) +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_114; -lean_dec(x_88); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_112); -lean_ctor_set(x_114, 1, x_107); -return x_114; +lean_object* x_5; +x_5 = l_Lean_Elab_Command_accLevel_go(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_115 = l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_107); -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); -lean_inc(x_116); -x_118 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(x_88, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_117); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_120 = x_118; -} else { - lean_dec_ref(x_118); - x_120 = lean_box(0); +lean_object* x_5; +x_5 = l_Lean_Elab_Command_accLevel_go(x_2, x_1, x_3, x_4); +return x_5; } -if (lean_is_scalar(x_120)) { - x_121 = lean_alloc_ctor(0, 2, 0); -} else { - x_121 = x_120; } -lean_ctor_set(x_121, 0, x_116); -lean_ctor_set(x_121, 1, x_119); -return x_121; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevel___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Elab_Command_accLevel(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} } +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_10 = lean_st_ref_get(x_6, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_instantiate_level_mvars(x_13, x_1); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_take(x_6, x_12); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +lean_ctor_set(x_18, 0, x_15); +x_22 = lean_st_ref_set(x_6, x_18, x_19); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_16); +return x_22; } +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_16); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } -default: +else { -lean_object* x_122; -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_1); -lean_ctor_set(x_122, 1, x_9); -return x_122; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_27 = lean_ctor_get(x_18, 1); +x_28 = lean_ctor_get(x_18, 2); +x_29 = lean_ctor_get(x_18, 3); +x_30 = lean_ctor_get(x_18, 4); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_31 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_31, 0, x_15); +lean_ctor_set(x_31, 1, x_27); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_31, 3, x_29); +lean_ctor_set(x_31, 4, x_30); +x_32 = lean_st_ref_set(x_6, x_31, x_19); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); +} +if (lean_is_scalar(x_34)) { + x_35 = lean_alloc_ctor(0, 2, 0); +} else { + x_35 = x_34; } +lean_ctor_set(x_35, 0, x_16); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -13016,7 +11885,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_accLevelAtCtor___lambda__1(lean_obj _start: { lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__3(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_11; } } @@ -14251,21 +13120,6 @@ return x_224; } } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; -} -} LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -14281,11 +13135,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_throwError___at_Lean_Elab_Command_accLevelAtCtor___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15994,7 +14848,201 @@ lean_dec(x_1); return x_13; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_take(x_6, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_12, 6); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_17, x_1, x_2); +lean_ctor_set(x_12, 6, x_18); +x_19 = lean_st_ref_set(x_6, x_11, x_13); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_dec(x_19); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_26 = lean_ctor_get(x_12, 0); +x_27 = lean_ctor_get(x_12, 1); +x_28 = lean_ctor_get(x_12, 2); +x_29 = lean_ctor_get(x_12, 3); +x_30 = lean_ctor_get(x_12, 4); +x_31 = lean_ctor_get(x_12, 5); +x_32 = lean_ctor_get(x_12, 6); +x_33 = lean_ctor_get(x_12, 7); +x_34 = lean_ctor_get(x_12, 8); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_12); +x_35 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_32, x_1, x_2); +x_36 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_36, 0, x_26); +lean_ctor_set(x_36, 1, x_27); +lean_ctor_set(x_36, 2, x_28); +lean_ctor_set(x_36, 3, x_29); +lean_ctor_set(x_36, 4, x_30); +lean_ctor_set(x_36, 5, x_31); +lean_ctor_set(x_36, 6, x_35); +lean_ctor_set(x_36, 7, x_33); +lean_ctor_set(x_36, 8, x_34); +lean_ctor_set(x_11, 0, x_36); +x_37 = lean_st_ref_set(x_6, x_11, x_13); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; +} else { + lean_dec_ref(x_37); + x_39 = lean_box(0); +} +x_40 = lean_box(0); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_39; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_42 = lean_ctor_get(x_11, 1); +x_43 = lean_ctor_get(x_11, 2); +x_44 = lean_ctor_get(x_11, 3); +x_45 = lean_ctor_get(x_11, 4); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_11); +x_46 = lean_ctor_get(x_12, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_12, 1); +lean_inc(x_47); +x_48 = lean_ctor_get(x_12, 2); +lean_inc(x_48); +x_49 = lean_ctor_get(x_12, 3); +lean_inc(x_49); +x_50 = lean_ctor_get(x_12, 4); +lean_inc(x_50); +x_51 = lean_ctor_get(x_12, 5); +lean_inc(x_51); +x_52 = lean_ctor_get(x_12, 6); +lean_inc(x_52); +x_53 = lean_ctor_get(x_12, 7); +lean_inc(x_53); +x_54 = lean_ctor_get(x_12, 8); +lean_inc(x_54); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + lean_ctor_release(x_12, 2); + lean_ctor_release(x_12, 3); + lean_ctor_release(x_12, 4); + lean_ctor_release(x_12, 5); + lean_ctor_release(x_12, 6); + lean_ctor_release(x_12, 7); + lean_ctor_release(x_12, 8); + x_55 = x_12; +} else { + lean_dec_ref(x_12); + x_55 = lean_box(0); +} +x_56 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_52, x_1, x_2); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 9, 0); +} else { + x_57 = x_55; +} +lean_ctor_set(x_57, 0, x_46); +lean_ctor_set(x_57, 1, x_47); +lean_ctor_set(x_57, 2, x_48); +lean_ctor_set(x_57, 3, x_49); +lean_ctor_set(x_57, 4, x_50); +lean_ctor_set(x_57, 5, x_51); +lean_ctor_set(x_57, 6, x_56); +lean_ctor_set(x_57, 7, x_53); +lean_ctor_set(x_57, 8, x_54); +x_58 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_42); +lean_ctor_set(x_58, 2, x_43); +lean_ctor_set(x_58, 3, x_44); +lean_ctor_set(x_58, 4, x_45); +x_59 = lean_st_ref_set(x_6, x_58, x_13); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_61 = x_59; +} else { + lean_dec_ref(x_59); + x_61 = lean_box(0); +} +x_62 = lean_box(0); +if (lean_is_scalar(x_61)) { + x_63 = lean_alloc_ctor(0, 2, 0); +} else { + x_63 = x_61; +} +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; +} +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (lean_obj_tag(x_1) == 0) @@ -16113,7 +15161,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (lean_obj_tag(x_1) == 0) @@ -16147,7 +15195,7 @@ x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); x_21 = lean_box(0); -x_22 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_17, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_22 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(x_17, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_20); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -16184,7 +15232,7 @@ x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); x_33 = lean_box(0); -x_34 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_29, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +x_34 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(x_29, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_32); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); x_36 = lean_ctor_get(x_34, 1); @@ -16237,7 +15285,7 @@ x_47 = lean_ctor_get(x_45, 1); lean_inc(x_47); lean_dec(x_45); x_48 = lean_box(0); -x_49 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_43, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +x_49 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(x_43, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_47); x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); @@ -16268,12 +15316,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_up lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_13 = l_Lean_Elab_Command_mkResultUniverse(x_1, x_2); x_14 = l_Lean_Level_mvarId_x21(x_3); -x_15 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_14, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_14, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); x_17 = lean_box(0); -x_18 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__2(x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +x_18 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__3(x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_16); return x_18; } } @@ -16649,11 +15697,11 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -16677,6 +15725,20 @@ lean_dec(x_3); return x_10; } } +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___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) { +_start: +{ +lean_object* x_10; +x_10 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -17736,7 +16798,7 @@ lean_dec(x_2); x_47 = lean_ctor_get(x_7, 0); lean_inc(x_47); lean_dec(x_7); -x_48 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_48 = l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -17756,7 +16818,7 @@ lean_dec(x_2); x_49 = lean_ctor_get(x_7, 0); lean_inc(x_49); lean_dec(x_7); -x_50 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_50 = l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index 225a0f2f3b24..6cddfb30dd3b 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -428,6 +428,7 @@ lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___spec__1 LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__1; +lean_object* l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_551____closed__24; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -847,7 +848,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion(lean_ob static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__4; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandOptDeclSig(lean_object*); lean_object* l_Lean_Name_getString_x21(lean_object*); @@ -22692,7 +22692,7 @@ lean_inc(x_19); lean_dec(x_17); x_20 = l_Lean_Elab_Command_mkResultUniverse(x_18, x_14); lean_dec(x_14); -x_21 = l_Lean_assignLevelMVar___at_Lean_Elab_Command_shouldInferResultUniverse___spec__2(x_16, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +x_21 = l_Lean_assignLevelMVar___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1(x_16, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_19); x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); lean_dec(x_21); diff --git a/stage0/stdlib/Lean/Language/Lean.c b/stage0/stdlib/Lean/Language/Lean.c index 72c31a349dc7..2d81a69b731f 100644 --- a/stage0/stdlib/Lean/Language/Lean.c +++ b/stage0/stdlib/Lean/Language/Lean.c @@ -13,194 +13,260 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__5; +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__8; +lean_object* lean_profileit(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1; LEAN_EXPORT uint32_t l_Lean_Language_Lean_SetupImportsResult_trustLevel___default; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7; +lean_object* l_EIO_toBaseIO___rarg(lean_object*, lean_object*); static lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, double, double, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__3; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; -lean_object* l_IO_mkRef___rarg(lean_object*, lean_object*); +static size_t l_Lean_Language_Lean_process_processHeader___lambda__5___closed__2; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10; static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8; lean_object* l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(lean_object*, lean_object*); +double lean_float_div(double, double); lean_object* lean_io_promise_new(lean_object*); +lean_object* l_String_toNat_x3f(lean_object*); lean_object* lean_get_set_stdout(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5; +lean_object* lean_io_as_task(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f(lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7; -static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2; lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_LeanProcessingM_run(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__1; lean_object* l_Lean_Language_SnapshotTask_ofIO___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__7; static lean_object* l_Lean_Language_Lean_process_parseCmd___closed__1; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10; static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__6; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_reparseOptions(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__1; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Language_SnapshotTask_bindIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11; lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Elab_processHeader(lean_object*, lean_object*, lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); lean_object* lean_io_promise_result(lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; extern lean_object* l_Lean_maxRecDepth; +lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8(lean_object*, lean_object*, lean_object*); +uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__2; lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_validate_utf8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader(lean_object*, lean_object*, lean_object*, 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_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; +lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_promise_resolve(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11; extern lean_object* l_Lean_Elab_instInhabitedInfoTree; uint8_t l_String_isEmpty(lean_object*); uint8_t l_Lean_Parser_isTerminalCommand(lean_object*); lean_object* l_IO_CancelToken_new(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f_goCmd(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__3; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_processCommands(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__5; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_ByteArray_empty; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toPArray_x27___rarg(lean_object*); lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1; +static double l_Lean_Language_Lean_process_processHeader___lambda__6___closed__1; +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2; LEAN_EXPORT lean_object* l_Lean_Language_Lean_instMonadLiftLeanProcessingMLeanProcessingTIO(lean_object*); +static lean_object* l_Lean_Language_Lean_process_doElab___closed__3; lean_object* l_Lean_MessageData_ofFormat(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*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__6; lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instMonadLiftProcessingTLeanProcessingT(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult(lean_object*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; lean_object* l_String_firstDiffPos(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9; LEAN_EXPORT lean_object* l_Lean_Language_Lean_LeanProcessingM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +lean_object* l_Lean_getOptionDecls(lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Name_append(lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__4; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_toPArray_x27___rarg(lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12; +lean_object* l_Lean_Name_getRoot(lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); lean_object* l_panic___at_String_fromUTF8_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__2; static lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1; lean_object* lean_environment_set_main_module(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__2; -static lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4; +static lean_object* l_Lean_Language_Lean_process_doElab___closed__4; lean_object* l_Lean_Language_diagnosticsOfHeaderError(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1; +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3; +extern lean_object* l_Task_Priority_default; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7; +uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_FS_Stream_ofBuffer(lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__12; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalCmdState_x3f_goCmd(lean_object*); extern lean_object* l_Lean_MessageLog_empty; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_set_stderr(lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1; static lean_object* l_Lean_Language_Lean_process_processHeader___closed__1; extern lean_object* l_Lean_Language_Snapshot_Diagnostics_empty; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__2(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Command_withLoggingExceptions(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10; lean_object* lean_string_from_utf8_unchecked(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__3; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__9; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___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*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__4; LEAN_EXPORT lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_KVMap_erase(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5; +static lean_object* l_Lean_Language_Lean_process_doElab___closed__1; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__10; +extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__10; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__7(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_Language_Lean_instMonadLiftLeanProcessingMLeanProcessingTIO___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__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_Language_Lean_isBeforeEditPos(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_stderrAsMessages; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5; +static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6; lean_object* l_Lean_Language_SnapshotTask_get_x3f___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalCmdState_x3f(lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9; lean_object* l_Lean_Language_SnapshotTask_bindIO___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__2; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1; static lean_object* l_Lean_Language_Lean_process_parseCmd___closed__3; +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_doElab___closed__2; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__11; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_set_stdin(lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__6; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); +lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Language_Lean_isBeforeEditPos___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_process_parseCmd___closed__2; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__3; +static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2; extern lean_object* l_Lean_Elab_Command_instInhabitedScope; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444_(lean_object*); lean_object* l_Lean_PersistentArray_get_x21___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__13; uint8_t l_Lean_Syntax_eqWithInfo(lean_object*, lean_object*); static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6; lean_object* l_Lean_Language_SnapshotTask_get___rarg(lean_object*); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -208,18 +274,26 @@ static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_internal_minimalSnapshots; +lean_object* lean_string_append(lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_output; lean_object* l_Lean_Parser_parseHeader(lean_object*, lean_object*); +lean_object* l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8; lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); +lean_object* lean_runtime_mark_persistent(lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7; +static lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9; lean_object* l_Lean_Parser_parseCommand(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions(lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__2; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__8; static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1; lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_SnapshotTask_pure___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Language_instTypeNameSnapshotLeaf; @@ -812,4305 +886,8151 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT lean_object* l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1() { _start: { -lean_object* x_4; -x_4 = lean_get_set_stdout(x_1, x_3); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -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_apply_1(x_2, x_6); -if (lean_obj_tag(x_7) == 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_dec(x_7); -x_10 = lean_get_set_stdout(x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -lean_dec(x_12); -lean_ctor_set(x_10, 0, x_8); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internal", 8, 8); +return x_1; } -else +} +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2() { +_start: { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_8); -lean_ctor_set(x_14, 1, x_13); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("minimalSnapshots", 16, 16); +return x_1; } } -else +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3() { +_start: { -uint8_t x_15; -lean_dec(x_8); -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4() { +_start: { -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } -else +} +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5() { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reduce information stored in snapshots to the minimum necessary for the cmdline driver: diagnostics per command and final full snapshot", 135, 135); +return x_1; } } +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 0; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_3 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} } -else +static lean_object* _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_7, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -lean_dec(x_7); -x_21 = lean_get_set_stdout(x_5, x_20); -if (lean_obj_tag(x_21) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; +x_3 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1; +x_4 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2; +x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444_(lean_object* x_1) { +_start: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3; +x_3 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6; +x_4 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7; +x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1() { +_start: { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 0, x_19); -return x_21; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, unknown configuration option '", 52, 52); +return x_1; } -else +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2() { +_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(1, 2, 0); -lean_ctor_set(x_25, 0, x_19); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'\n\nIf the option is defined in this library, use '-D", 52, 52); +return x_1; } } -else +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3() { +_start: { -uint8_t x_26; -lean_dec(x_19); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' to set it conditionally", 25, 25); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4() { +_start: { -return x_21; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; } -else +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5() { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; } } +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, invalid configuration option '", 52, 52); +return x_1; } } -else +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7() { +_start: { -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_4); -if (x_30 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' value, it must be true/false", 30, 30); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8() { +_start: { -return x_4; +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(1, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; } -else +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9() { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_4, 0); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_4); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(1, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; } } +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid -D parameter, configuration option '", 44, 44); +return x_1; } } -LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11() { _start: { -lean_object* x_4; -x_4 = lean_get_set_stdin(x_1, x_3); -if (lean_obj_tag(x_4) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' cannot be set in the command line, use set_option command", 59, 59); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12() { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -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_apply_1(x_2, x_6); -if (lean_obj_tag(x_7) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' value, it must be a natural number", 36, 36); +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_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_10 = lean_get_set_stdin(x_5, x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_box(0); +x_10 = l_Lean_Name_replacePrefix(x_1, x_2, x_9); +x_11 = l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(x_3, x_10); +if (lean_obj_tag(x_11) == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_dec(x_5); +if (x_4 == 0) { -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -lean_dec(x_12); -lean_ctor_set(x_10, 0, x_8); -return x_10; +uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_6); +x_12 = 1; +lean_inc(x_10); +x_13 = l_Lean_Name_toString(x_10, x_12); +x_14 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2; +x_17 = lean_string_append(x_15, x_16); +x_18 = l_Lean_Name_append(x_2, x_10); +x_19 = l_Lean_Name_toString(x_18, x_12); +x_20 = lean_string_append(x_17, x_19); +lean_dec(x_19); +x_21 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_8); +return x_24; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); +lean_object* x_25; lean_object* x_26; lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_8); -lean_ctor_set(x_14, 1, x_13); -return x_14; +lean_dec(x_2); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_6); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_8); +return x_26; } } else { -uint8_t x_15; -lean_dec(x_8); -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) +uint8_t x_27; +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_11); +if (x_27 == 0) { -return x_10; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_11, 0); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +switch (lean_obj_tag(x_29)) { +case 0: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +lean_ctor_set(x_29, 0, x_5); +x_32 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +lean_ctor_set(x_11, 0, x_32); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_11); +lean_ctor_set(x_33, 1, x_8); +return x_33; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_5); +x_35 = l_Lean_KVMap_insertCore(x_6, x_10, x_34); +lean_ctor_set(x_11, 0, x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_11); +lean_ctor_set(x_36, 1, x_8); +return x_36; } } -else +case 1: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_7, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -lean_dec(x_7); -x_21 = lean_get_set_stdin(x_5, x_20); -if (lean_obj_tag(x_21) == 0) +lean_object* x_37; uint8_t x_38; +lean_dec(x_29); +x_37 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4; +x_38 = lean_string_dec_eq(x_5, x_37); +if (x_38 == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_39; uint8_t x_40; +x_39 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5; +x_40 = lean_string_dec_eq(x_5, x_39); +if (x_40 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 0, x_19); -return x_21; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_10); +lean_dec(x_6); +x_41 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +x_42 = lean_string_append(x_41, x_5); +lean_dec(x_5); +x_43 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7; +x_44 = lean_string_append(x_42, x_43); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_44); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_11); +lean_ctor_set(x_45, 1, x_8); +return x_45; } else { -lean_object* x_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(1, 2, 0); -lean_ctor_set(x_25, 0, x_19); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_5); +x_46 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8; +x_47 = l_Lean_KVMap_insertCore(x_6, x_10, x_46); +lean_ctor_set(x_11, 0, x_47); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_11); +lean_ctor_set(x_48, 1, x_8); +return x_48; } } else { -uint8_t x_26; -lean_dec(x_19); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_5); +x_49 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9; +x_50 = l_Lean_KVMap_insertCore(x_6, x_10, x_49); +lean_ctor_set(x_11, 0, x_50); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_11); +lean_ctor_set(x_51, 1, x_8); +return x_51; +} +} +case 3: { -return x_21; +uint8_t x_52; +x_52 = !lean_is_exclusive(x_29); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_29, 0); +lean_dec(x_53); +x_54 = l_String_toNat_x3f(x_5); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_free_object(x_29); +lean_dec(x_10); +lean_dec(x_6); +x_55 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +x_56 = lean_string_append(x_55, x_5); +lean_dec(x_5); +x_57 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12; +x_58 = lean_string_append(x_56, x_57); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_58); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_11); +lean_ctor_set(x_59, 1, x_8); +return x_59; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +uint8_t x_60; +lean_free_object(x_11); +lean_dec(x_5); +x_60 = !lean_is_exclusive(x_54); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_54, 0); +lean_ctor_set(x_29, 0, x_61); +x_62 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +lean_ctor_set(x_54, 0, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_54); +lean_ctor_set(x_63, 1, x_8); +return x_63; } +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_54, 0); +lean_inc(x_64); +lean_dec(x_54); +lean_ctor_set(x_29, 0, x_64); +x_65 = l_Lean_KVMap_insertCore(x_6, x_10, x_29); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_8); +return x_67; } } } else { -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_4); -if (x_30 == 0) +lean_object* x_68; +lean_dec(x_29); +x_68 = l_String_toNat_x3f(x_5); +if (lean_obj_tag(x_68) == 0) { -return x_4; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_10); +lean_dec(x_6); +x_69 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +x_70 = lean_string_append(x_69, x_5); +lean_dec(x_5); +x_71 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12; +x_72 = lean_string_append(x_70, x_71); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_72); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_11); +lean_ctor_set(x_73, 1, x_8); +return x_73; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_4, 0); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_4); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_free_object(x_11); +lean_dec(x_5); +x_74 = lean_ctor_get(x_68, 0); +lean_inc(x_74); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + x_75 = x_68; +} else { + lean_dec_ref(x_68); + x_75 = lean_box(0); } +x_76 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_76, 0, x_74); +x_77 = l_Lean_KVMap_insertCore(x_6, x_10, x_76); +if (lean_is_scalar(x_75)) { + x_78 = lean_alloc_ctor(1, 1, 0); +} else { + x_78 = x_75; } +lean_ctor_set(x_78, 0, x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_8); +return x_79; } } -LEAN_EXPORT lean_object* l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_get_set_stderr(x_1, x_3); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -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_apply_1(x_2, x_6); -if (lean_obj_tag(x_7) == 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_dec(x_7); -x_10 = lean_get_set_stderr(x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -lean_dec(x_12); -lean_ctor_set(x_10, 0, x_8); -return x_10; } -else +default: { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_8); -lean_ctor_set(x_14, 1, x_13); -return x_14; +uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_29); +lean_dec(x_6); +lean_dec(x_5); +x_80 = 1; +x_81 = l_Lean_Name_toString(x_10, x_80); +x_82 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10; +x_83 = lean_string_append(x_82, x_81); +lean_dec(x_81); +x_84 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11; +x_85 = lean_string_append(x_83, x_84); +lean_ctor_set_tag(x_11, 18); +lean_ctor_set(x_11, 0, x_85); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_11); +lean_ctor_set(x_86, 1, x_8); +return x_86; } } -else -{ -uint8_t x_15; -lean_dec(x_8); -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) -{ -return x_10; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_11, 0); +lean_inc(x_87); +lean_dec(x_11); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +switch (lean_obj_tag(x_88)) { +case 0: +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + x_89 = x_88; +} else { + lean_dec_ref(x_88); + x_89 = lean_box(0); } +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(0, 1, 0); +} else { + x_90 = x_89; } +lean_ctor_set(x_90, 0, x_5); +x_91 = l_Lean_KVMap_insertCore(x_6, x_10, x_90); +x_92 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_92, 0, x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_8); +return x_93; } -else +case 1: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_7, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -lean_dec(x_7); -x_21 = lean_get_set_stderr(x_5, x_20); -if (lean_obj_tag(x_21) == 0) +lean_object* x_94; uint8_t x_95; +lean_dec(x_88); +x_94 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4; +x_95 = lean_string_dec_eq(x_5, x_94); +if (x_95 == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_96; uint8_t x_97; +x_96 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5; +x_97 = lean_string_dec_eq(x_5, x_96); +if (x_97 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 0, x_19); -return x_21; +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_10); +lean_dec(x_6); +x_98 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +x_99 = lean_string_append(x_98, x_5); +lean_dec(x_5); +x_100 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7; +x_101 = lean_string_append(x_99, x_100); +x_102 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_8); +return x_103; } else { -lean_object* x_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(1, 2, 0); -lean_ctor_set(x_25, 0, x_19); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_5); +x_104 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8; +x_105 = l_Lean_KVMap_insertCore(x_6, x_10, x_104); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_8); +return x_107; } -else -{ -uint8_t x_26; -lean_dec(x_19); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) -{ -return x_21; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_5); +x_108 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9; +x_109 = l_Lean_KVMap_insertCore(x_6, x_10, x_108); +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_8); +return x_111; } } -else +case 3: { -uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_4); -if (x_30 == 0) +lean_object* x_112; lean_object* x_113; +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + x_112 = x_88; +} else { + lean_dec_ref(x_88); + x_112 = lean_box(0); +} +x_113 = l_String_toNat_x3f(x_5); +if (lean_obj_tag(x_113) == 0) { -return x_4; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_112); +lean_dec(x_10); +lean_dec(x_6); +x_114 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6; +x_115 = lean_string_append(x_114, x_5); +lean_dec(x_5); +x_116 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12; +x_117 = lean_string_append(x_115, x_116); +x_118 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_118, 0, x_117); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_8); +return x_119; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_4, 0); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_4); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_5); +x_120 = lean_ctor_get(x_113, 0); +lean_inc(x_120); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + x_121 = x_113; +} else { + lean_dec_ref(x_113); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_112)) { + x_122 = lean_alloc_ctor(3, 1, 0); +} else { + x_122 = x_112; } +lean_ctor_set(x_122, 0, x_120); +x_123 = l_Lean_KVMap_insertCore(x_6, x_10, x_122); +if (lean_is_scalar(x_121)) { + x_124 = lean_alloc_ctor(1, 1, 0); +} else { + x_124 = x_121; } +lean_ctor_set(x_124, 0, x_123); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_8); +return x_125; } } -static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1() { -_start: +default: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_ByteArray_empty; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_dec(x_88); +lean_dec(x_6); +lean_dec(x_5); +x_126 = 1; +x_127 = l_Lean_Name_toString(x_10, x_126); +x_128 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10; +x_129 = lean_string_append(x_128, x_127); +lean_dec(x_127); +x_130 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11; +x_131 = lean_string_append(x_129, x_130); +x_132 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_132, 0, x_131); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_8); +return x_133; } } -static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Init.Data.String.Extra", 22, 22); -return x_1; } } -static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("String.fromUTF8!", 16, 16); -return x_1; } } -static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid UTF-8 string", 20, 20); +x_1 = lean_mk_string_unchecked("weak", 4, 4); return x_1; } } -static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2; -x_2 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__3; -x_3 = lean_unsigned_to_nat(92u); -x_4 = lean_unsigned_to_nat(47u); -x_5 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1(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; -x_4 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1; -x_5 = lean_st_mk_ref(x_4, x_3); -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_5, 0); +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_2, 0); lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); +x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_st_mk_ref(x_4, x_7); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_7) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_8, 0); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_ctor_get(x_6, 0); lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); +lean_dec(x_6); +x_10 = lean_ctor_get(x_7, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = l_IO_FS_Stream_ofBuffer(x_6); -lean_inc(x_9); -x_12 = l_IO_FS_Stream_ofBuffer(x_9); -if (x_2 == 0) +lean_dec(x_7); +x_11 = l_Lean_Name_getRoot(x_9); +x_12 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2; +x_13 = lean_name_eq(x_11, x_12); +lean_dec(x_11); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2), 3, 2); -lean_closure_set(x_13, 0, x_12); -lean_closure_set(x_13, 1, x_1); -x_14 = l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(x_11, x_13, x_10); -if (lean_obj_tag(x_14) == 0) +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_3, x_14, x_4); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_st_ref_get(x_9, x_16); -lean_dec(x_9); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_2 = x_8; +x_3 = x_18; +x_4 = x_17; +goto _start; +} +else { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -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_string_validate_utf8(x_20); -if (x_21 == 0) +uint8_t x_20; +lean_dec(x_8); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_20); -x_22 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; -x_23 = l_panic___at_String_fromUTF8_x21___spec__1(x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_15); -lean_ctor_set(x_17, 0, x_24); -return x_17; +return x_15; } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_string_from_utf8_unchecked(x_20); -lean_dec(x_20); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_15); -lean_ctor_set(x_17, 0, x_26); -return x_17; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_27 = lean_ctor_get(x_17, 0); -x_28 = lean_ctor_get(x_17, 1); -lean_inc(x_28); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_inc(x_9); +x_24 = l_Lean_KVMap_erase(x_3, x_9); +x_25 = lean_box(0); +x_26 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1(x_9, x_12, x_1, x_13, x_10, x_24, x_25, x_4); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_17); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_string_validate_utf8(x_29); -if (x_30 == 0) +x_2 = x_8; +x_3 = x_29; +x_4 = x_28; +goto _start; +} +else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_29); -x_31 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; -x_32 = l_panic___at_String_fromUTF8_x21___spec__1(x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_15); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_28); -return x_34; +uint8_t x_31; +lean_dec(x_8); +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) +{ +return x_26; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_string_from_utf8_unchecked(x_29); -lean_dec(x_29); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_15); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_28); -return x_37; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_26); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} } } } else { -uint8_t x_38; -lean_dec(x_15); -x_38 = !lean_is_exclusive(x_17); -if (x_38 == 0) +lean_object* x_35; +lean_dec(x_7); +lean_dec(x_6); +x_35 = lean_ctor_get(x_2, 1); +lean_inc(x_35); +lean_dec(x_2); +x_2 = x_35; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_reparseOptions(lean_object* x_1, lean_object* x_2) { +_start: { -return x_17; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Lean_getOptionDecls(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +lean_inc(x_1); +x_6 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1(x_4, x_1, x_1, x_5); +lean_dec(x_4); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +return x_6; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_17, 0); -x_40 = lean_ctor_get(x_17, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_17); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; } } else { -uint8_t x_42; -lean_dec(x_9); -x_42 = !lean_is_exclusive(x_14); -if (x_42 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) { -return x_14; +return x_6; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_14, 0); -x_44 = lean_ctor_get(x_14, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_14); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } } -else +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_inc(x_12); -x_46 = lean_alloc_closure((void*)(l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4), 3, 2); -lean_closure_set(x_46, 0, x_12); -lean_closure_set(x_46, 1, x_1); -x_47 = lean_alloc_closure((void*)(l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2), 3, 2); -lean_closure_set(x_47, 0, x_12); -lean_closure_set(x_47, 1, x_46); -x_48 = l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(x_11, x_47, x_10); -if (lean_obj_tag(x_48) == 0) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_4); +lean_dec(x_4); +x_10 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_st_ref_get(x_9, x_50); -lean_dec(x_9); -if (lean_obj_tag(x_51) == 0) +lean_object* x_5; +x_5 = l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_52; -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +lean_object* x_5; +x_5 = lean_get_set_stdout(x_1, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = lean_ctor_get(x_51, 0); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -lean_dec(x_53); -x_55 = lean_string_validate_utf8(x_54); -if (x_55 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; +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_apply_2(x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_54); -x_56 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; -x_57 = l_panic___at_String_fromUTF8_x21___spec__1(x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_49); -lean_ctor_set(x_51, 0, x_58); -return x_51; +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 = lean_get_set_stdout(x_6, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_9); +return x_11; } else { -lean_object* x_59; lean_object* x_60; -x_59 = lean_string_from_utf8_unchecked(x_54); -lean_dec(x_54); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_49); -lean_ctor_set(x_51, 0, x_60); -return x_51; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_51, 0); -x_62 = lean_ctor_get(x_51, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_51); -x_63 = lean_ctor_get(x_61, 0); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_string_validate_utf8(x_63); -if (x_64 == 0) +uint8_t x_16; +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -lean_dec(x_63); -x_65 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; -x_66 = l_panic___at_String_fromUTF8_x21___spec__1(x_65); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_49); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_62); -return x_68; +return x_11; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_string_from_utf8_unchecked(x_63); -lean_dec(x_63); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_49); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_62); -return x_71; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } } else { -uint8_t x_72; -lean_dec(x_49); -x_72 = !lean_is_exclusive(x_51); -if (x_72 == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_8, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); +lean_dec(x_8); +x_22 = lean_get_set_stdout(x_6, x_21); +if (lean_obj_tag(x_22) == 0) { -return x_51; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set_tag(x_22, 1); +lean_ctor_set(x_22, 0, x_20); +return x_22; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_51, 0); -x_74 = lean_ctor_get(x_51, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_51); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; -} +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_76; -lean_dec(x_9); -x_76 = !lean_is_exclusive(x_48); -if (x_76 == 0) -{ -return x_48; -} -else +uint8_t x_27; +lean_dec(x_20); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_48, 0); -x_78 = lean_ctor_get(x_48, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_48); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; -} -} -} +return x_22; } else { -uint8_t x_80; -lean_dec(x_6); -lean_dec(x_1); -x_80 = !lean_is_exclusive(x_8); -if (x_80 == 0) -{ -return x_8; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 0); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_22); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } -else -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_8, 0); -x_82 = lean_ctor_get(x_8, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_8); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -return x_83; } } } else { -uint8_t x_84; -lean_dec(x_1); -x_84 = !lean_is_exclusive(x_5); -if (x_84 == 0) +uint8_t x_31; +lean_dec(x_3); +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_5); +if (x_31 == 0) { return x_5; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_5, 0); -x_86 = lean_ctor_get(x_5, 1); -lean_inc(x_86); -lean_inc(x_85); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_5); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_3, x_4); -x_6 = lean_ctor_get(x_5, 1); +lean_object* x_5; +x_5 = lean_get_set_stdin(x_1, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +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_7 = l_Lean_Elab_Command_elabCommandTopLevel(x_1, x_2, x_3, x_6); -return x_7; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__1() { -_start: +x_8 = lean_apply_2(x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("process", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__2() { -_start: +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 = lean_get_set_stdin(x_6, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("doElab", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__3() { -_start: +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; -x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; -x_3 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__1; -x_4 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__2; -x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); -return x_5; -} +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_9); +return x_11; } -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__4() { -_start: +else { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__3; -x_2 = 1; -x_3 = l_Lean_Name_toString(x_1, x_2); -return x_3; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__5() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; -x_3 = l_Lean_Language_Snapshot_Diagnostics_empty; -x_4 = 0; -x_5 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_1); -lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__6() { -_start: +uint8_t x_16; +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Language_instTypeNameSnapshotLeaf; -x_2 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__5; -x_3 = l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(x_1, x_2); -return x_3; -} +return x_11; } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; 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_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 2); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 3); -lean_inc(x_8); -x_9 = lean_ctor_get(x_1, 4); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 5); -lean_inc(x_10); -x_11 = lean_ctor_get(x_1, 6); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 7); -lean_inc(x_12); -lean_dec(x_1); -lean_inc(x_11); -lean_inc(x_3); -x_13 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_3); -lean_ctor_set(x_13, 2, x_7); -lean_ctor_set(x_13, 3, x_8); -lean_ctor_set(x_13, 4, x_9); -lean_ctor_set(x_13, 5, x_10); -lean_ctor_set(x_13, 6, x_11); -lean_ctor_set(x_13, 7, x_12); -x_14 = lean_ctor_get(x_2, 1); -x_15 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__6; -x_16 = lean_io_promise_resolve(x_15, x_14, x_5); -if (lean_obj_tag(x_16) == 0) +else { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_3, x_17); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); lean_dec(x_11); -x_22 = lean_ctor_get(x_21, 2); -lean_inc(x_22); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_dec_lt(x_23, x_22); -lean_dec(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_21); -x_25 = l_Lean_Elab_instInhabitedInfoTree; -x_26 = l_outOfBounds___rarg(x_25); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; -x_29 = 0; -x_30 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_20); -lean_ctor_set(x_30, 2, x_27); -lean_ctor_set_uint8(x_30, sizeof(void*)*3, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_13); -lean_ctor_set(x_18, 0, x_31); -return x_18; +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; -x_32 = l_Lean_Elab_instInhabitedInfoTree; -x_33 = l_Lean_PersistentArray_get_x21___rarg(x_32, x_21, x_23); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -x_35 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; -x_36 = 0; -x_37 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_20); -lean_ctor_set(x_37, 2, x_34); -lean_ctor_set_uint8(x_37, sizeof(void*)*3, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_13); -lean_ctor_set(x_18, 0, x_38); -return x_18; } } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_39 = lean_ctor_get(x_18, 0); -x_40 = lean_ctor_get(x_18, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_18); -x_41 = lean_ctor_get(x_11, 1); -lean_inc(x_41); -lean_dec(x_11); -x_42 = lean_ctor_get(x_41, 2); -lean_inc(x_42); -x_43 = lean_unsigned_to_nat(0u); -x_44 = lean_nat_dec_lt(x_43, x_42); -lean_dec(x_42); -if (x_44 == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_8, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); +lean_dec(x_8); +x_22 = lean_get_set_stdin(x_6, x_21); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_41); -x_45 = l_Lean_Elab_instInhabitedInfoTree; -x_46 = l_outOfBounds___rarg(x_45); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; -x_49 = 0; -x_50 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_39); -lean_ctor_set(x_50, 2, x_47); -lean_ctor_set_uint8(x_50, sizeof(void*)*3, x_49); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_13); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_40); -return x_52; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set_tag(x_22, 1); +lean_ctor_set(x_22, 0, x_20); +return x_22; } else { -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; -x_53 = l_Lean_Elab_instInhabitedInfoTree; -x_54 = l_Lean_PersistentArray_get_x21___rarg(x_53, x_41, x_43); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_54); -x_56 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__4; -x_57 = 0; -x_58 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_39); -lean_ctor_set(x_58, 2, x_55); -lean_ctor_set_uint8(x_58, sizeof(void*)*3, x_57); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_13); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_40); -return x_60; -} +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_61; -lean_dec(x_13); -lean_dec(x_11); -x_61 = !lean_is_exclusive(x_18); -if (x_61 == 0) +uint8_t x_27; +lean_dec(x_20); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -return x_18; +return x_22; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_18, 0); -x_63 = lean_ctor_get(x_18, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_18); -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_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 0); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_22); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} } } } else { -uint8_t x_65; -lean_dec(x_13); -lean_dec(x_11); +uint8_t x_31; lean_dec(x_3); -x_65 = !lean_is_exclusive(x_16); -if (x_65 == 0) +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_5); +if (x_31 == 0) { -return x_16; +return x_5; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_16, 0); -x_67 = lean_ctor_get(x_16, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_16); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; -} -} -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Language_Lean_stderrAsMessages; -return x_1; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } -static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__2; -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_Language_Lean_process_doElab___lambda__3___closed__4() { +LEAN_EXPORT lean_object* l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_5; +x_5 = lean_get_set_stderr(x_1, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_9; -x_9 = lean_st_ref_get(x_1, x_8); -if (lean_obj_tag(x_9) == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; +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_apply_2(x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_9, 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); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_st_mk_ref(x_10, 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_2, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 2); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_box(0); -lean_inc(x_13); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_13); -lean_inc(x_3); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_3); -x_21 = lean_ctor_get(x_2, 3); -lean_inc(x_21); -lean_dec(x_2); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_firstFrontendMacroScope; -x_25 = lean_box(0); -x_26 = 0; -lean_inc(x_4); -lean_inc(x_17); -lean_inc(x_16); -x_27 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_27, 0, x_16); -lean_ctor_set(x_27, 1, x_17); -lean_ctor_set(x_27, 2, x_23); -lean_ctor_set(x_27, 3, x_4); -lean_ctor_set(x_27, 4, x_18); -lean_ctor_set(x_27, 5, x_24); -lean_ctor_set(x_27, 6, x_25); -lean_ctor_set(x_27, 7, x_19); -lean_ctor_set(x_27, 8, x_20); -lean_ctor_set(x_27, 9, x_22); -lean_ctor_set_uint8(x_27, sizeof(void*)*10, x_26); -x_28 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__1), 4, 1); -lean_closure_set(x_28, 0, x_5); -lean_inc(x_7); -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Command_withLoggingExceptions), 4, 3); -lean_closure_set(x_29, 0, x_28); -lean_closure_set(x_29, 1, x_27); -lean_closure_set(x_29, 2, x_7); -x_30 = lean_ctor_get(x_6, 1); -x_31 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; -x_32 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_30, x_31); -x_33 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_29, x_32, x_14); -if (lean_obj_tag(x_33) == 0) +lean_dec(x_8); +x_11 = lean_get_set_stderr(x_6, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = !lean_is_exclusive(x_34); -if (x_36 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_34, 0); -x_38 = lean_ctor_get(x_34, 1); -lean_dec(x_38); -x_39 = lean_st_ref_get(x_13, x_35); +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); lean_dec(x_13); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_st_ref_take(x_1, x_41); -if (lean_obj_tag(x_43) == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); -x_45 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; -lean_ctor_set(x_34, 1, x_45); -lean_ctor_set(x_34, 0, x_42); -x_46 = lean_st_ref_set(x_1, x_34, x_44); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = lean_st_ref_get(x_7, x_47); -lean_dec(x_7); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -x_52 = l_String_isEmpty(x_37); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_53 = l_Lean_FileMap_toPosition(x_17, x_4); -lean_dec(x_4); -x_54 = lean_box(0); -x_55 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_55, 0, x_37); -x_56 = l_Lean_MessageData_ofFormat(x_55); -x_57 = 0; -x_58 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; -x_59 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_59, 0, x_16); -lean_ctor_set(x_59, 1, x_53); -lean_ctor_set(x_59, 2, x_54); -lean_ctor_set(x_59, 3, x_58); -lean_ctor_set(x_59, 4, x_56); -lean_ctor_set_uint8(x_59, sizeof(void*)*5, x_26); -lean_ctor_set_uint8(x_59, sizeof(void*)*5 + 1, x_57); -x_60 = l_Lean_MessageLog_add(x_59, x_51); -x_61 = lean_box(0); -x_62 = l_Lean_Language_Lean_process_doElab___lambda__2(x_49, x_3, x_60, x_61, x_50); -lean_dec(x_3); -return x_62; +lean_ctor_set(x_11, 0, x_9); +return x_11; } else { -lean_object* x_63; lean_object* x_64; -lean_dec(x_37); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_4); -x_63 = lean_box(0); -x_64 = l_Lean_Language_Lean_process_doElab___lambda__2(x_49, x_3, x_51, x_63, x_50); -lean_dec(x_3); -return x_64; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } else { -uint8_t x_65; -lean_dec(x_37); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_4); -lean_dec(x_3); -x_65 = !lean_is_exclusive(x_48); -if (x_65 == 0) +uint8_t x_16; +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -return x_48; +return x_11; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_48, 0); -x_67 = lean_ctor_get(x_48, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_48); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } } else { -uint8_t x_69; -lean_dec(x_37); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -x_69 = !lean_is_exclusive(x_46); -if (x_69 == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_8, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); +lean_dec(x_8); +x_22 = lean_get_set_stderr(x_6, x_21); +if (lean_obj_tag(x_22) == 0) { -return x_46; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set_tag(x_22, 1); +lean_ctor_set(x_22, 0, x_20); +return x_22; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_46, 0); -x_71 = lean_ctor_get(x_46, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_46); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; -} +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_73; -lean_dec(x_42); -lean_free_object(x_34); -lean_dec(x_37); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -x_73 = !lean_is_exclusive(x_43); -if (x_73 == 0) +uint8_t x_27; +lean_dec(x_20); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -return x_43; +return x_22; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_43, 0); -x_75 = lean_ctor_get(x_43, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_43); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 0); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_22); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} } } } else { -uint8_t x_77; -lean_free_object(x_34); -lean_dec(x_37); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_4); +uint8_t x_31; lean_dec(x_3); -x_77 = !lean_is_exclusive(x_39); -if (x_77 == 0) +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_5); +if (x_31 == 0) { -return x_39; +return x_5; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_39, 0); -x_79 = lean_ctor_get(x_39, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_39); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } -else -{ -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_34, 0); -lean_inc(x_81); -lean_dec(x_34); -x_82 = lean_st_ref_get(x_13, x_35); -lean_dec(x_13); -if (lean_obj_tag(x_82) == 0) +} +static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1() { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_ctor_get(x_83, 1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_ByteArray_empty; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.String.Extra", 22, 22); +return x_1; +} +} +static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("String.fromUTF8!", 16, 16); +return x_1; +} +} +static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid UTF-8 string", 20, 20); +return x_1; +} +} +static lean_object* _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2; +x_2 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__3; +x_3 = lean_unsigned_to_nat(92u); +x_4 = lean_unsigned_to_nat(47u); +x_5 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4; +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_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1; +x_6 = lean_st_mk_ref(x_5, x_4); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_st_mk_ref(x_5, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_IO_FS_Stream_ofBuffer(x_7); +lean_inc(x_10); +x_13 = l_IO_FS_Stream_ofBuffer(x_10); +if (x_2 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_alloc_closure((void*)(l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2), 4, 2); +lean_closure_set(x_14, 0, x_13); +lean_closure_set(x_14, 1, x_1); +x_15 = l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(x_12, x_14, x_3, x_11); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_st_ref_get(x_10, x_17); +lean_dec(x_10); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_string_validate_utf8(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_21); +x_23 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; +x_24 = l_panic___at_String_fromUTF8_x21___spec__1(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_16); +lean_ctor_set(x_18, 0, x_25); +return x_18; +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_string_from_utf8_unchecked(x_21); +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_16); +lean_ctor_set(x_18, 0, x_27); +return x_18; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = lean_ctor_get(x_18, 0); +x_29 = lean_ctor_get(x_18, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_18); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_string_validate_utf8(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_30); +x_32 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; +x_33 = l_panic___at_String_fromUTF8_x21___spec__1(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_16); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_29); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_string_from_utf8_unchecked(x_30); +lean_dec(x_30); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_16); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_29); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_16); +x_39 = !lean_is_exclusive(x_18); +if (x_39 == 0) +{ +return x_18; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_18, 0); +x_41 = lean_ctor_get(x_18, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_18); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_10); +x_43 = !lean_is_exclusive(x_15); +if (x_43 == 0) +{ +return x_15; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_15, 0); +x_45 = lean_ctor_get(x_15, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_15); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_inc(x_13); +x_47 = lean_alloc_closure((void*)(l_IO_withStderr___at_Lean_Language_Lean_process_doElab___spec__4), 4, 2); +lean_closure_set(x_47, 0, x_13); +lean_closure_set(x_47, 1, x_1); +x_48 = lean_alloc_closure((void*)(l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2), 4, 2); +lean_closure_set(x_48, 0, x_13); +lean_closure_set(x_48, 1, x_47); +x_49 = l_IO_withStdin___at_Lean_Language_Lean_process_doElab___spec__3(x_12, x_48, x_3, x_11); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_st_ref_get(x_10, x_51); +lean_dec(x_10); +if (lean_obj_tag(x_52) == 0) +{ +uint8_t x_53; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_string_validate_utf8(x_55); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_55); +x_57 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; +x_58 = l_panic___at_String_fromUTF8_x21___spec__1(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_50); +lean_ctor_set(x_52, 0, x_59); +return x_52; +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_string_from_utf8_unchecked(x_55); +lean_dec(x_55); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_50); +lean_ctor_set(x_52, 0, x_61); +return x_52; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_62 = lean_ctor_get(x_52, 0); +x_63 = lean_ctor_get(x_52, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_52); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_string_validate_utf8(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_64); +x_66 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5; +x_67 = l_panic___at_String_fromUTF8_x21___spec__1(x_66); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_50); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_63); +return x_69; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_string_from_utf8_unchecked(x_64); +lean_dec(x_64); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_50); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_63); +return x_72; +} +} +} +else +{ +uint8_t x_73; +lean_dec(x_50); +x_73 = !lean_is_exclusive(x_52); +if (x_73 == 0) +{ +return x_52; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_52, 0); +x_75 = lean_ctor_get(x_52, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_52); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_10); +x_77 = !lean_is_exclusive(x_49); +if (x_77 == 0) +{ +return x_49; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_49, 0); +x_79 = lean_ctor_get(x_49, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_49); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +} +else +{ +uint8_t x_81; +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_1); +x_81 = !lean_is_exclusive(x_9); +if (x_81 == 0) +{ +return x_9; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_9, 0); +x_83 = lean_ctor_get(x_9, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_9); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; +} +} +} +else +{ +uint8_t x_85; +lean_dec(x_3); +lean_dec(x_1); +x_85 = !lean_is_exclusive(x_6); +if (x_85 == 0) +{ +return x_6; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_6, 0); +x_87 = lean_ctor_get(x_6, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_6); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_3, x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Elab_Command_elabCommandTopLevel(x_1, x_2, x_3, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_1, x_2, x_3, x_5); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_6, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set_tag(x_6, 0); +lean_ctor_set(x_6, 0, x_13); +return x_6; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_6, 1); +lean_inc(x_14); +lean_dec(x_6); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("process", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doElab", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; +x_3 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; +x_4 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__2; +x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; +x_2 = 1; +x_3 = l_Lean_Name_toString(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +x_3 = l_Lean_Language_Snapshot_Diagnostics_empty; +x_4 = 0; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_1); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_instTypeNameSnapshotLeaf; +x_2 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__5; +x_3 = l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 3); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 4); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 5); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 6); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 7); +lean_inc(x_13); +lean_dec(x_1); +lean_inc(x_12); +lean_inc(x_3); +x_14 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_3); +lean_ctor_set(x_14, 2, x_8); +lean_ctor_set(x_14, 3, x_9); +lean_ctor_set(x_14, 4, x_10); +lean_ctor_set(x_14, 5, x_11); +lean_ctor_set(x_14, 6, x_12); +lean_ctor_set(x_14, 7, x_13); +x_15 = lean_ctor_get(x_2, 1); +x_16 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__6; +x_17 = lean_io_promise_resolve(x_16, x_15, x_6); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_3, x_18); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_dec(x_12); +x_23 = lean_ctor_get(x_22, 2); +lean_inc(x_23); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_lt(x_24, x_23); +lean_dec(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_22); +x_26 = l_Lean_Elab_instInhabitedInfoTree; +x_27 = l_outOfBounds___rarg(x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +x_30 = 0; +x_31 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_21); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set_uint8(x_31, sizeof(void*)*3, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_14); +lean_ctor_set(x_19, 0, x_32); +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; +x_33 = l_Lean_Elab_instInhabitedInfoTree; +x_34 = l_Lean_PersistentArray_get_x21___rarg(x_33, x_22, x_24); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +x_37 = 0; +x_38 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_21); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_14); +lean_ctor_set(x_19, 0, x_39); +return x_19; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_40 = lean_ctor_get(x_19, 0); +x_41 = lean_ctor_get(x_19, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_19); +x_42 = lean_ctor_get(x_12, 1); +lean_inc(x_42); +lean_dec(x_12); +x_43 = lean_ctor_get(x_42, 2); +lean_inc(x_43); +x_44 = lean_unsigned_to_nat(0u); +x_45 = lean_nat_dec_lt(x_44, x_43); +lean_dec(x_43); +if (x_45 == 0) +{ +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_dec(x_42); +x_46 = l_Lean_Elab_instInhabitedInfoTree; +x_47 = l_outOfBounds___rarg(x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +x_50 = 0; +x_51 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_40); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*3, x_50); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_14); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_41); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_54 = l_Lean_Elab_instInhabitedInfoTree; +x_55 = l_Lean_PersistentArray_get_x21___rarg(x_54, x_42, x_44); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; +x_58 = 0; +x_59 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_40); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set_uint8(x_59, sizeof(void*)*3, x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_14); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_41); +return x_61; +} +} +} +else +{ +uint8_t x_62; +lean_dec(x_14); +lean_dec(x_12); +x_62 = !lean_is_exclusive(x_19); +if (x_62 == 0) +{ +return x_19; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_19, 0); +x_64 = lean_ctor_get(x_19, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_19); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +else +{ +uint8_t x_66; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_3); +x_66 = !lean_is_exclusive(x_17); +if (x_66 == 0) +{ +return x_17; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_17, 0); +x_68 = lean_ctor_get(x_17, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_17); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Language_Lean_internal_minimalSnapshots; +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Language_Lean_stderrAsMessages; +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_doElab___closed__3() { +_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_Language_Lean_process_doElab___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_doElab___closed__3; +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_Language_Lean_process_doElab(lean_object* x_1, 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_2); +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_2, 2); +x_10 = lean_ctor_get(x_2, 1); +lean_dec(x_10); +x_11 = l_Lean_Elab_Command_instInhabitedScope; +x_12 = l_List_head_x21___rarg(x_11, x_9); +x_13 = l_Lean_MessageLog_empty; +lean_ctor_set(x_2, 1, x_13); +x_14 = lean_st_mk_ref(x_2, x_7); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_5, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_st_mk_ref(x_18, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; 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; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_6, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 2); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +lean_inc(x_21); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_21); +x_28 = lean_ctor_get(x_12, 1); +lean_inc(x_28); +lean_dec(x_12); +x_29 = l_Lean_Language_Lean_process_doElab___closed__1; +x_30 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_29); +x_31 = lean_ctor_get(x_6, 3); +lean_inc(x_31); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__1), 4, 1); +lean_closure_set(x_33, 0, x_1); +x_34 = l_Lean_Language_Lean_process_doElab___closed__2; +x_35 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_34); +lean_dec(x_28); +if (x_30 == 0) +{ +lean_object* x_140; +lean_inc(x_4); +x_140 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_140, 0, x_4); +x_36 = x_140; +goto block_139; +} +else +{ +lean_object* x_141; +x_141 = lean_box(0); +x_36 = x_141; +goto block_139; +} +block_139: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_37 = lean_unsigned_to_nat(0u); +x_38 = l_Lean_firstFrontendMacroScope; +x_39 = lean_box(0); +x_40 = 0; +lean_inc(x_3); +lean_inc(x_25); +lean_inc(x_24); +x_41 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_41, 0, x_24); +lean_ctor_set(x_41, 1, x_25); +lean_ctor_set(x_41, 2, x_37); +lean_ctor_set(x_41, 3, x_3); +lean_ctor_set(x_41, 4, x_26); +lean_ctor_set(x_41, 5, x_38); +lean_ctor_set(x_41, 6, x_39); +lean_ctor_set(x_41, 7, x_27); +lean_ctor_set(x_41, 8, x_36); +lean_ctor_set(x_41, 9, x_32); +lean_ctor_set_uint8(x_41, sizeof(void*)*10, x_40); +lean_inc(x_15); +x_42 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__2___boxed), 5, 3); +lean_closure_set(x_42, 0, x_33); +lean_closure_set(x_42, 1, x_41); +lean_closure_set(x_42, 2, x_15); +lean_inc(x_6); +x_43 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_42, x_35, x_6, x_22); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; +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_is_exclusive(x_44); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_44, 0); +x_48 = lean_ctor_get(x_44, 1); +lean_dec(x_48); +x_49 = lean_st_ref_get(x_21, x_45); +lean_dec(x_21); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +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, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_st_ref_take(x_5, x_51); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = l_Lean_Language_Lean_process_doElab___closed__4; +lean_ctor_set(x_44, 1, x_55); +lean_ctor_set(x_44, 0, x_52); +x_56 = lean_st_ref_set(x_5, x_44, x_54); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_get(x_15, x_57); +lean_dec(x_15); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +x_62 = l_String_isEmpty(x_47); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_63 = l_Lean_FileMap_toPosition(x_25, x_3); +lean_dec(x_3); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_65, 0, x_47); +x_66 = l_Lean_MessageData_ofFormat(x_65); +x_67 = 0; +x_68 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_69 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_69, 0, x_24); +lean_ctor_set(x_69, 1, x_63); +lean_ctor_set(x_69, 2, x_64); +lean_ctor_set(x_69, 3, x_68); +lean_ctor_set(x_69, 4, x_66); +lean_ctor_set_uint8(x_69, sizeof(void*)*5, x_40); +lean_ctor_set_uint8(x_69, sizeof(void*)*5 + 1, x_67); +x_70 = l_Lean_MessageLog_add(x_69, x_61); +x_71 = lean_box(0); +x_72 = l_Lean_Language_Lean_process_doElab___lambda__3(x_59, x_4, x_70, x_71, x_6, x_60); +lean_dec(x_6); +lean_dec(x_4); +return x_72; +} +else +{ +lean_object* x_73; lean_object* x_74; +lean_dec(x_47); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_3); +x_73 = lean_box(0); +x_74 = l_Lean_Language_Lean_process_doElab___lambda__3(x_59, x_4, x_61, x_73, x_6, x_60); +lean_dec(x_6); +lean_dec(x_4); +return x_74; +} +} +else +{ +uint8_t x_75; +lean_dec(x_47); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_75 = !lean_is_exclusive(x_58); +if (x_75 == 0) +{ +return x_58; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_58, 0); +x_77 = lean_ctor_get(x_58, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_58); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +else +{ +uint8_t x_79; +lean_dec(x_47); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_79 = !lean_is_exclusive(x_56); +if (x_79 == 0) +{ +return x_56; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_56, 0); +x_81 = lean_ctor_get(x_56, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_56); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +} +} +else +{ +uint8_t x_83; +lean_dec(x_52); +lean_free_object(x_44); +lean_dec(x_47); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_83 = !lean_is_exclusive(x_53); +if (x_83 == 0) +{ +return x_53; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_53, 0); +x_85 = lean_ctor_get(x_53, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_53); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +} +else +{ +uint8_t x_87; +lean_free_object(x_44); +lean_dec(x_47); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_87 = !lean_is_exclusive(x_49); +if (x_87 == 0) +{ +return x_49; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_49, 0); +x_89 = lean_ctor_get(x_49, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_49); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_44, 0); +lean_inc(x_91); +lean_dec(x_44); +x_92 = lean_st_ref_get(x_21, x_45); +lean_dec(x_21); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_96 = lean_st_ref_take(x_5, x_94); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Language_Lean_process_doElab___closed__4; +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_95); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_st_ref_set(x_5, x_99, x_97); +if (lean_obj_tag(x_100) == 0) +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +lean_dec(x_100); +x_102 = lean_st_ref_get(x_15, x_101); +lean_dec(x_15); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +x_106 = l_String_isEmpty(x_91); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_107 = l_Lean_FileMap_toPosition(x_25, x_3); +lean_dec(x_3); +x_108 = lean_box(0); +x_109 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_109, 0, x_91); +x_110 = l_Lean_MessageData_ofFormat(x_109); +x_111 = 0; +x_112 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_113 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_113, 0, x_24); +lean_ctor_set(x_113, 1, x_107); +lean_ctor_set(x_113, 2, x_108); +lean_ctor_set(x_113, 3, x_112); +lean_ctor_set(x_113, 4, x_110); +lean_ctor_set_uint8(x_113, sizeof(void*)*5, x_40); +lean_ctor_set_uint8(x_113, sizeof(void*)*5 + 1, x_111); +x_114 = l_Lean_MessageLog_add(x_113, x_105); +x_115 = lean_box(0); +x_116 = l_Lean_Language_Lean_process_doElab___lambda__3(x_103, x_4, x_114, x_115, x_6, x_104); +lean_dec(x_6); +lean_dec(x_4); +return x_116; +} +else +{ +lean_object* x_117; lean_object* x_118; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_3); +x_117 = lean_box(0); +x_118 = l_Lean_Language_Lean_process_doElab___lambda__3(x_103, x_4, x_105, x_117, x_6, x_104); +lean_dec(x_6); +lean_dec(x_4); +return x_118; +} +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_119 = lean_ctor_get(x_102, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_102, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_121 = x_102; +} else { + lean_dec_ref(x_102); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_123 = lean_ctor_get(x_100, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_100, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_125 = x_100; +} else { + lean_dec_ref(x_100); + 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; +} +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +lean_dec(x_95); +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_127 = lean_ctor_get(x_96, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_96, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_129 = x_96; +} else { + lean_dec_ref(x_96); + x_129 = lean_box(0); +} +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(1, 2, 0); +} else { + x_130 = x_129; +} +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +return x_130; +} +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_91); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_131 = lean_ctor_get(x_92, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_92, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_133 = x_92; +} else { + lean_dec_ref(x_92); + x_133 = lean_box(0); +} +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(1, 2, 0); +} else { + x_134 = x_133; +} +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; +} +} +} +else +{ +uint8_t x_135; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_21); +lean_dec(x_15); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_135 = !lean_is_exclusive(x_43); +if (x_135 == 0) +{ +return x_43; +} +else +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_43, 0); +x_137 = lean_ctor_get(x_43, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_43); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); +return x_138; +} +} +} +} +else +{ +uint8_t x_142; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_142 = !lean_is_exclusive(x_20); +if (x_142 == 0) +{ +return x_20; +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_ctor_get(x_20, 0); +x_144 = lean_ctor_get(x_20, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_20); +x_145 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +return x_145; +} +} +} +else +{ +uint8_t x_146; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_146 = !lean_is_exclusive(x_17); +if (x_146 == 0) +{ +return x_17; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_17, 0); +x_148 = lean_ctor_get(x_17, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_17); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_150 = !lean_is_exclusive(x_14); +if (x_150 == 0) +{ +return x_14; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_14, 0); +x_152 = lean_ctor_get(x_14, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_14); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_154 = lean_ctor_get(x_2, 0); +x_155 = lean_ctor_get(x_2, 2); +x_156 = lean_ctor_get(x_2, 3); +x_157 = lean_ctor_get(x_2, 4); +x_158 = lean_ctor_get(x_2, 5); +x_159 = lean_ctor_get(x_2, 6); +x_160 = lean_ctor_get(x_2, 7); +lean_inc(x_160); +lean_inc(x_159); +lean_inc(x_158); +lean_inc(x_157); +lean_inc(x_156); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_2); +x_161 = l_Lean_Elab_Command_instInhabitedScope; +x_162 = l_List_head_x21___rarg(x_161, x_155); +x_163 = l_Lean_MessageLog_empty; +x_164 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_164, 0, x_154); +lean_ctor_set(x_164, 1, x_163); +lean_ctor_set(x_164, 2, x_155); +lean_ctor_set(x_164, 3, x_156); +lean_ctor_set(x_164, 4, x_157); +lean_ctor_set(x_164, 5, x_158); +lean_ctor_set(x_164, 6, x_159); +lean_ctor_set(x_164, 7, x_160); +x_165 = lean_st_mk_ref(x_164, x_7); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_165, 1); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_st_ref_get(x_5, x_167); +if (lean_obj_tag(x_168) == 0) +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_168, 1); +lean_inc(x_170); +lean_dec(x_168); +x_171 = lean_st_mk_ref(x_169, x_170); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; lean_object* x_187; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_6, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 2); +lean_inc(x_176); +lean_dec(x_174); +x_177 = lean_box(0); +lean_inc(x_172); +x_178 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_178, 0, x_172); +x_179 = lean_ctor_get(x_162, 1); +lean_inc(x_179); +lean_dec(x_162); +x_180 = l_Lean_Language_Lean_process_doElab___closed__1; +x_181 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_179, x_180); +x_182 = lean_ctor_get(x_6, 3); +lean_inc(x_182); +x_183 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_183, 0, x_182); +x_184 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__1), 4, 1); +lean_closure_set(x_184, 0, x_1); +x_185 = l_Lean_Language_Lean_process_doElab___closed__2; +x_186 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_179, x_185); +lean_dec(x_179); +if (x_181 == 0) +{ +lean_object* x_247; +lean_inc(x_4); +x_247 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_247, 0, x_4); +x_187 = x_247; +goto block_246; +} +else +{ +lean_object* x_248; +x_248 = lean_box(0); +x_187 = x_248; +goto block_246; +} +block_246: +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_188 = lean_unsigned_to_nat(0u); +x_189 = l_Lean_firstFrontendMacroScope; +x_190 = lean_box(0); +x_191 = 0; +lean_inc(x_3); +lean_inc(x_176); +lean_inc(x_175); +x_192 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_192, 0, x_175); +lean_ctor_set(x_192, 1, x_176); +lean_ctor_set(x_192, 2, x_188); +lean_ctor_set(x_192, 3, x_3); +lean_ctor_set(x_192, 4, x_177); +lean_ctor_set(x_192, 5, x_189); +lean_ctor_set(x_192, 6, x_190); +lean_ctor_set(x_192, 7, x_178); +lean_ctor_set(x_192, 8, x_187); +lean_ctor_set(x_192, 9, x_183); +lean_ctor_set_uint8(x_192, sizeof(void*)*10, x_191); +lean_inc(x_166); +x_193 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__2___boxed), 5, 3); +lean_closure_set(x_193, 0, x_184); +lean_closure_set(x_193, 1, x_192); +lean_closure_set(x_193, 2, x_166); +lean_inc(x_6); +x_194 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_193, x_186, x_6, x_173); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = lean_ctor_get(x_195, 0); +lean_inc(x_197); +if (lean_is_exclusive(x_195)) { + lean_ctor_release(x_195, 0); + lean_ctor_release(x_195, 1); + x_198 = x_195; +} else { + lean_dec_ref(x_195); + x_198 = lean_box(0); +} +x_199 = lean_st_ref_get(x_172, x_196); +lean_dec(x_172); +if (lean_obj_tag(x_199) == 0) +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_200 = lean_ctor_get(x_199, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_199, 1); +lean_inc(x_201); +lean_dec(x_199); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +lean_dec(x_200); +x_203 = lean_st_ref_take(x_5, x_201); +if (lean_obj_tag(x_203) == 0) +{ +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_204 = lean_ctor_get(x_203, 1); +lean_inc(x_204); +lean_dec(x_203); +x_205 = l_Lean_Language_Lean_process_doElab___closed__4; +if (lean_is_scalar(x_198)) { + x_206 = lean_alloc_ctor(0, 2, 0); +} else { + x_206 = x_198; +} +lean_ctor_set(x_206, 0, x_202); +lean_ctor_set(x_206, 1, x_205); +x_207 = lean_st_ref_set(x_5, x_206, x_204); +if (lean_obj_tag(x_207) == 0) +{ +lean_object* x_208; lean_object* x_209; +x_208 = lean_ctor_get(x_207, 1); +lean_inc(x_208); +lean_dec(x_207); +x_209 = lean_st_ref_get(x_166, x_208); +lean_dec(x_166); +if (lean_obj_tag(x_209) == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; +x_210 = lean_ctor_get(x_209, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_209, 1); +lean_inc(x_211); +lean_dec(x_209); +x_212 = lean_ctor_get(x_210, 1); +lean_inc(x_212); +x_213 = l_String_isEmpty(x_197); +if (x_213 == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_214 = l_Lean_FileMap_toPosition(x_176, x_3); +lean_dec(x_3); +x_215 = lean_box(0); +x_216 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_216, 0, x_197); +x_217 = l_Lean_MessageData_ofFormat(x_216); +x_218 = 0; +x_219 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_220 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_220, 0, x_175); +lean_ctor_set(x_220, 1, x_214); +lean_ctor_set(x_220, 2, x_215); +lean_ctor_set(x_220, 3, x_219); +lean_ctor_set(x_220, 4, x_217); +lean_ctor_set_uint8(x_220, sizeof(void*)*5, x_191); +lean_ctor_set_uint8(x_220, sizeof(void*)*5 + 1, x_218); +x_221 = l_Lean_MessageLog_add(x_220, x_212); +x_222 = lean_box(0); +x_223 = l_Lean_Language_Lean_process_doElab___lambda__3(x_210, x_4, x_221, x_222, x_6, x_211); +lean_dec(x_6); +lean_dec(x_4); +return x_223; +} +else +{ +lean_object* x_224; lean_object* x_225; +lean_dec(x_197); +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_3); +x_224 = lean_box(0); +x_225 = l_Lean_Language_Lean_process_doElab___lambda__3(x_210, x_4, x_212, x_224, x_6, x_211); +lean_dec(x_6); +lean_dec(x_4); +return x_225; +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_197); +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_226 = lean_ctor_get(x_209, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_209, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_209)) { + lean_ctor_release(x_209, 0); + lean_ctor_release(x_209, 1); + x_228 = x_209; +} else { + lean_dec_ref(x_209); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +return x_229; +} +} +else +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +lean_dec(x_197); +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_166); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_230 = lean_ctor_get(x_207, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_207, 1); +lean_inc(x_231); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_232 = x_207; +} else { + lean_dec_ref(x_207); + x_232 = lean_box(0); +} +if (lean_is_scalar(x_232)) { + x_233 = lean_alloc_ctor(1, 2, 0); +} else { + x_233 = x_232; +} +lean_ctor_set(x_233, 0, x_230); +lean_ctor_set(x_233, 1, x_231); +return x_233; +} +} +else +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +lean_dec(x_202); +lean_dec(x_198); +lean_dec(x_197); +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_166); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_234 = lean_ctor_get(x_203, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_203, 1); +lean_inc(x_235); +if (lean_is_exclusive(x_203)) { + lean_ctor_release(x_203, 0); + lean_ctor_release(x_203, 1); + x_236 = x_203; +} else { + lean_dec_ref(x_203); + x_236 = lean_box(0); +} +if (lean_is_scalar(x_236)) { + x_237 = lean_alloc_ctor(1, 2, 0); +} else { + x_237 = x_236; +} +lean_ctor_set(x_237, 0, x_234); +lean_ctor_set(x_237, 1, x_235); +return x_237; +} +} +else +{ +lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_198); +lean_dec(x_197); +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_166); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_238 = lean_ctor_get(x_199, 0); +lean_inc(x_238); +x_239 = lean_ctor_get(x_199, 1); +lean_inc(x_239); +if (lean_is_exclusive(x_199)) { + lean_ctor_release(x_199, 0); + lean_ctor_release(x_199, 1); + x_240 = x_199; +} else { + lean_dec_ref(x_199); + x_240 = lean_box(0); +} +if (lean_is_scalar(x_240)) { + x_241 = lean_alloc_ctor(1, 2, 0); +} else { + x_241 = x_240; +} +lean_ctor_set(x_241, 0, x_238); +lean_ctor_set(x_241, 1, x_239); +return x_241; +} +} +else +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_172); +lean_dec(x_166); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_242 = lean_ctor_get(x_194, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_194, 1); +lean_inc(x_243); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_244 = x_194; +} else { + lean_dec_ref(x_194); + x_244 = lean_box(0); +} +if (lean_is_scalar(x_244)) { + x_245 = lean_alloc_ctor(1, 2, 0); +} else { + x_245 = x_244; +} +lean_ctor_set(x_245, 0, x_242); +lean_ctor_set(x_245, 1, x_243); +return x_245; +} +} +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +lean_dec(x_166); +lean_dec(x_162); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_249 = lean_ctor_get(x_171, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_171, 1); +lean_inc(x_250); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_251 = x_171; +} else { + lean_dec_ref(x_171); + x_251 = lean_box(0); +} +if (lean_is_scalar(x_251)) { + x_252 = lean_alloc_ctor(1, 2, 0); +} else { + x_252 = x_251; +} +lean_ctor_set(x_252, 0, x_249); +lean_ctor_set(x_252, 1, x_250); +return x_252; +} +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_166); +lean_dec(x_162); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_253 = lean_ctor_get(x_168, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_168, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_168)) { + lean_ctor_release(x_168, 0); + lean_ctor_release(x_168, 1); + x_255 = x_168; +} else { + lean_dec_ref(x_168); + x_255 = lean_box(0); +} +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); +} else { + x_256 = x_255; +} +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; +} +} +else +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +lean_dec(x_162); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_257 = lean_ctor_get(x_165, 0); +lean_inc(x_257); +x_258 = lean_ctor_get(x_165, 1); +lean_inc(x_258); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_259 = x_165; +} else { + lean_dec_ref(x_165); + x_259 = lean_box(0); +} +if (lean_is_scalar(x_259)) { + x_260 = lean_alloc_ctor(1, 2, 0); +} else { + x_260 = x_259; +} +lean_ctor_set(x_260, 0, x_257); +lean_ctor_set(x_260, 1, x_258); +return x_260; +} +} +} +} +LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_1, x_5, x_3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Language_Lean_process_doElab___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Language_Lean_process_doElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___boxed(lean_object* x_1, 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_Language_Lean_process_doElab(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l_Lean_MessageLog_empty; +x_7 = l_Lean_Parser_parseCommand(x_5, x_2, x_3, x_6); +return x_7; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parseCmd", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; +x_3 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; +x_4 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1; +x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2; +x_2 = 1; +x_3 = l_Lean_Name_toString(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 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_Language_Lean_process_parseCmd___lambda__2___closed__5() { +_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_Language_Lean_process_parseCmd___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5; +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_Language_Lean_process_parseCmd___lambda__2___closed__7() { +_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_Language_Lean_process_parseCmd___lambda__2___closed__6; +x_3 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5; +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_Language_Lean_process_parseCmd___lambda__2___closed__8() { +_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__2___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; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9() { +_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_Language_Lean_process_parseCmd___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_3 = lean_box(0); +x_4 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_1); +lean_ctor_set(x_6, 4, x_1); +lean_ctor_set(x_6, 5, x_4); +lean_ctor_set(x_6, 6, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*7, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12() { +_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_Language_Lean_process_parseCmd___lambda__2___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_uniq", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14; +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_Language_Lean_process_parseCmd___lambda__2___closed__16() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_Language_Lean_process_doElab___closed__4; +x_3 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7; +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_Language_Lean_process_parseCmd___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_1, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +if (x_10 == 0) +{ +lean_object* x_78; +x_78 = lean_box(0); +x_17 = x_78; +goto block_77; +} +else +{ +uint8_t x_79; +lean_inc(x_2); +x_79 = l_Lean_Parser_isTerminalCommand(x_2); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_80 = lean_box(0); +x_81 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3; +x_82 = 0; +x_83 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_15); +lean_ctor_set(x_83, 2, x_80); +lean_ctor_set_uint8(x_83, sizeof(void*)*3, x_82); +x_84 = l_Lean_Syntax_getRange_x3f(x_2, x_82); +lean_dec(x_2); +x_85 = lean_io_promise_result(x_3); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = lean_ctor_get(x_4, 0); +lean_inc(x_87); +x_88 = !lean_is_exclusive(x_87); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_89 = lean_ctor_get(x_87, 2); +lean_dec(x_89); +x_90 = lean_ctor_get(x_87, 0); +lean_dec(x_90); +lean_ctor_set(x_87, 2, x_80); +lean_ctor_set(x_87, 0, x_81); +lean_ctor_set_uint8(x_87, sizeof(void*)*3, x_82); +x_91 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8; +x_92 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11; +x_93 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12; +x_94 = lean_unsigned_to_nat(0u); +x_95 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15; +x_96 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16; +x_97 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7; +lean_inc(x_9); +x_98 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_98, 0, x_9); +lean_ctor_set(x_98, 1, x_91); +lean_ctor_set(x_98, 2, x_92); +lean_ctor_set(x_98, 3, x_93); +lean_ctor_set(x_98, 4, x_94); +lean_ctor_set(x_98, 5, x_95); +lean_ctor_set(x_98, 6, x_96); +lean_ctor_set(x_98, 7, x_97); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_87); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_Language_SnapshotTask_pure___rarg(x_99); +x_101 = lean_box(0); +x_102 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4; +x_103 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_103, 0, x_83); +lean_ctor_set(x_103, 1, x_101); +lean_ctor_set(x_103, 2, x_102); +lean_ctor_set(x_103, 3, x_86); +lean_ctor_set(x_103, 4, x_100); +lean_ctor_set(x_103, 5, x_6); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_104; lean_object* x_105; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_80); +x_105 = lean_io_promise_resolve(x_104, x_7, x_16); +if (lean_obj_tag(x_105) == 0) +{ +uint8_t x_106; +x_106 = !lean_is_exclusive(x_105); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; +x_107 = lean_ctor_get(x_105, 0); +lean_dec(x_107); +x_108 = lean_box(0); +lean_ctor_set(x_105, 0, x_108); +return x_105; +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_105, 1); +lean_inc(x_109); +lean_dec(x_105); +x_110 = lean_box(0); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +else +{ +uint8_t x_112; +x_112 = !lean_is_exclusive(x_105); +if (x_112 == 0) +{ +return x_105; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_105, 0); +x_114 = lean_ctor_get(x_105, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_105); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +return x_115; +} +} +} +else +{ +uint8_t x_116; +x_116 = !lean_is_exclusive(x_11); +if (x_116 == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_117 = lean_ctor_get(x_11, 0); +x_118 = lean_ctor_get(x_5, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_8, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +lean_dec(x_119); +x_121 = lean_string_utf8_byte_size(x_120); +lean_dec(x_120); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_118); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_11, 0, x_122); +lean_inc(x_117); +x_123 = lean_io_promise_result(x_117); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_11); +lean_ctor_set(x_124, 1, x_123); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_124); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_103); +lean_ctor_set(x_126, 1, x_125); +x_127 = lean_io_promise_resolve(x_126, x_7, x_16); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +lean_dec(x_127); +x_129 = lean_ctor_get(x_4, 1); +lean_inc(x_129); +lean_dec(x_4); +x_130 = l_Lean_Language_Lean_process_parseCmd(x_80, x_5, x_129, x_9, x_117, x_8, x_128); +return x_130; +} +else +{ +uint8_t x_131; +lean_dec(x_117); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_131 = !lean_is_exclusive(x_127); +if (x_131 == 0) +{ +return x_127; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_127, 0); +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +lean_inc(x_132); +lean_dec(x_127); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +return x_134; +} +} +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_135 = lean_ctor_get(x_11, 0); +lean_inc(x_135); +lean_dec(x_11); +x_136 = lean_ctor_get(x_5, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_8, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +lean_dec(x_137); +x_139 = lean_string_utf8_byte_size(x_138); +lean_dec(x_138); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_136); +lean_ctor_set(x_140, 1, x_139); +x_141 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_141, 0, x_140); +lean_inc(x_135); +x_142 = lean_io_promise_result(x_135); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_143); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_103); +lean_ctor_set(x_145, 1, x_144); +x_146 = lean_io_promise_resolve(x_145, x_7, x_16); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +lean_dec(x_146); +x_148 = lean_ctor_get(x_4, 1); +lean_inc(x_148); +lean_dec(x_4); +x_149 = l_Lean_Language_Lean_process_parseCmd(x_80, x_5, x_148, x_9, x_135, x_8, x_147); +return x_149; +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_135); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_150 = lean_ctor_get(x_146, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_146, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_152 = x_146; +} else { + lean_dec_ref(x_146); + x_152 = lean_box(0); +} +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; +} +} +} +} +else +{ +lean_object* x_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; +x_154 = lean_ctor_get(x_87, 1); +lean_inc(x_154); +lean_dec(x_87); +x_155 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_155, 0, x_81); +lean_ctor_set(x_155, 1, x_154); +lean_ctor_set(x_155, 2, x_80); +lean_ctor_set_uint8(x_155, sizeof(void*)*3, x_82); +x_156 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8; +x_157 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11; +x_158 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12; +x_159 = lean_unsigned_to_nat(0u); +x_160 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15; +x_161 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16; +x_162 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7; +lean_inc(x_9); +x_163 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_163, 0, x_9); +lean_ctor_set(x_163, 1, x_156); +lean_ctor_set(x_163, 2, x_157); +lean_ctor_set(x_163, 3, x_158); +lean_ctor_set(x_163, 4, x_159); +lean_ctor_set(x_163, 5, x_160); +lean_ctor_set(x_163, 6, x_161); +lean_ctor_set(x_163, 7, x_162); +x_164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_164, 0, x_155); +lean_ctor_set(x_164, 1, x_163); +x_165 = l_Lean_Language_SnapshotTask_pure___rarg(x_164); +x_166 = lean_box(0); +x_167 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4; +x_168 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_168, 0, x_83); +lean_ctor_set(x_168, 1, x_166); +lean_ctor_set(x_168, 2, x_167); +lean_ctor_set(x_168, 3, x_86); +lean_ctor_set(x_168, 4, x_165); +lean_ctor_set(x_168, 5, x_6); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_169; lean_object* x_170; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_80); +x_170 = lean_io_promise_resolve(x_169, x_7, x_16); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_171 = lean_ctor_get(x_170, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_172 = x_170; +} else { + lean_dec_ref(x_170); + x_172 = lean_box(0); +} +x_173 = lean_box(0); +if (lean_is_scalar(x_172)) { + x_174 = lean_alloc_ctor(0, 2, 0); +} else { + x_174 = x_172; +} +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set(x_174, 1, x_171); +return x_174; +} +else +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_175 = lean_ctor_get(x_170, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_170, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_177 = x_170; +} else { + lean_dec_ref(x_170); + x_177 = lean_box(0); +} +if (lean_is_scalar(x_177)) { + x_178 = lean_alloc_ctor(1, 2, 0); +} else { + x_178 = x_177; +} +lean_ctor_set(x_178, 0, x_175); +lean_ctor_set(x_178, 1, x_176); +return x_178; +} +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_179 = lean_ctor_get(x_11, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + x_180 = x_11; +} else { + lean_dec_ref(x_11); + x_180 = lean_box(0); +} +x_181 = lean_ctor_get(x_5, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_8, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +lean_dec(x_182); +x_184 = lean_string_utf8_byte_size(x_183); +lean_dec(x_183); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_181); +lean_ctor_set(x_185, 1, x_184); +if (lean_is_scalar(x_180)) { + x_186 = lean_alloc_ctor(1, 1, 0); +} else { + x_186 = x_180; +} +lean_ctor_set(x_186, 0, x_185); +lean_inc(x_179); +x_187 = lean_io_promise_result(x_179); +x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +x_189 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_189, 0, x_188); +x_190 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_190, 0, x_168); +lean_ctor_set(x_190, 1, x_189); +x_191 = lean_io_promise_resolve(x_190, x_7, x_16); +if (lean_obj_tag(x_191) == 0) +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_191, 1); +lean_inc(x_192); +lean_dec(x_191); +x_193 = lean_ctor_get(x_4, 1); +lean_inc(x_193); +lean_dec(x_4); +x_194 = l_Lean_Language_Lean_process_parseCmd(x_80, x_5, x_193, x_9, x_179, x_8, x_192); +return x_194; +} +else +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_179); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_195 = lean_ctor_get(x_191, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_191, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_197 = x_191; +} else { + lean_dec_ref(x_191); + x_197 = lean_box(0); +} +if (lean_is_scalar(x_197)) { + x_198 = lean_alloc_ctor(1, 2, 0); +} else { + x_198 = x_197; +} +lean_ctor_set(x_198, 0, x_195); +lean_ctor_set(x_198, 1, x_196); +return x_198; +} +} +} +} +else +{ +lean_object* x_199; +x_199 = lean_box(0); +x_17 = x_199; +goto block_77; +} +} +block_77: +{ +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_dec(x_17); +x_18 = lean_box(0); +x_19 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3; +x_20 = 0; +x_21 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_15); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set_uint8(x_21, sizeof(void*)*3, x_20); +x_22 = l_Lean_Syntax_getRange_x3f(x_2, x_20); +x_23 = lean_io_promise_result(x_3); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +lean_inc(x_4); +x_25 = l_Lean_Language_SnapshotTask_pure___rarg(x_4); +lean_inc(x_5); +x_26 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_26, 1, x_2); +lean_ctor_set(x_26, 2, x_5); +lean_ctor_set(x_26, 3, x_24); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set(x_26, 5, x_6); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_18); +x_28 = lean_io_promise_resolve(x_27, x_7, x_16); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +x_31 = lean_box(0); +lean_ctor_set(x_28, 0, x_31); +return x_28; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_28, 1); +lean_inc(x_32); +lean_dec(x_28); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) +{ +return x_28; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_28, 0); +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_28); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +else +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_11); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_40 = lean_ctor_get(x_11, 0); +x_41 = lean_ctor_get(x_5, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_8, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_string_utf8_byte_size(x_43); +lean_dec(x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_11, 0, x_45); +lean_inc(x_40); +x_46 = lean_io_promise_result(x_40); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_11); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_26); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_io_promise_resolve(x_49, x_7, x_16); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = lean_ctor_get(x_4, 1); +lean_inc(x_52); +lean_dec(x_4); +x_53 = l_Lean_Language_Lean_process_parseCmd(x_18, x_5, x_52, x_9, x_40, x_8, x_51); +return x_53; +} +else +{ +uint8_t x_54; +lean_dec(x_40); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_54 = !lean_is_exclusive(x_50); +if (x_54 == 0) +{ +return x_50; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 0); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_50); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_58 = lean_ctor_get(x_11, 0); +lean_inc(x_58); +lean_dec(x_11); +x_59 = lean_ctor_get(x_5, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_8, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_string_utf8_byte_size(x_61); +lean_dec(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_59); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +lean_inc(x_58); +x_65 = lean_io_promise_result(x_58); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_66); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_26); +lean_ctor_set(x_68, 1, x_67); +x_69 = lean_io_promise_resolve(x_68, x_7, x_16); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = lean_ctor_get(x_4, 1); +lean_inc(x_71); +lean_dec(x_4); +x_72 = l_Lean_Language_Lean_process_parseCmd(x_18, x_5, x_71, x_9, x_58, x_8, x_70); +return x_72; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_58); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +x_73 = lean_ctor_get(x_69, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_69, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_75 = x_69; +} else { + lean_dec_ref(x_69); + x_75 = lean_box(0); +} +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(1, 2, 0); +} else { + x_76 = x_75; +} +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_74); +return x_76; +} +} +} +} +} +else +{ +uint8_t x_200; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_200 = !lean_is_exclusive(x_14); +if (x_200 == 0) +{ +return x_14; +} +else +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_201 = lean_ctor_get(x_14, 0); +x_202 = lean_ctor_get(x_14, 1); +lean_inc(x_202); +lean_inc(x_201); +lean_dec(x_14); +x_203 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_203, 0, x_201); +lean_ctor_set(x_203, 1, x_202); +return x_203; +} +} +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_doElab___closed__4; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = lean_io_promise_new(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; +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); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_63; lean_object* x_64; +x_63 = l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1; +x_64 = lean_st_mk_ref(x_63, x_16); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_17 = x_65; +x_18 = x_66; +goto block_62; +} +else +{ +uint8_t x_67; +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_64); +if (x_67 == 0) +{ +return x_64; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_64, 0); +x_69 = lean_ctor_get(x_64, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_64); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_10, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +lean_dec(x_71); +x_73 = lean_ctor_get(x_72, 5); +lean_inc(x_73); +lean_dec(x_72); +x_17 = x_73; +x_18 = x_16; +goto block_62; +} +block_62: +{ +lean_object* x_19; +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_43; +x_43 = lean_box(0); +x_19 = x_43; +goto block_42; +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_10); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; +x_45 = lean_ctor_get(x_10, 0); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_ctor_get(x_45, 1); +lean_dec(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_47, 3); +lean_inc(x_50); +lean_dec(x_47); +lean_ctor_set(x_45, 1, x_50); +lean_ctor_set(x_45, 0, x_49); +x_19 = x_10; +goto block_42; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_45, 0); +lean_inc(x_51); +lean_dec(x_45); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 3); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_10, 0, x_54); +x_19 = x_10; +goto block_42; +} +} +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; +x_55 = lean_ctor_get(x_10, 0); +lean_inc(x_55); +lean_dec(x_10); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_57 = x_55; +} else { + lean_dec_ref(x_55); + x_57 = lean_box(0); +} +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +x_59 = lean_ctor_get(x_56, 3); +lean_inc(x_59); +lean_dec(x_56); +if (lean_is_scalar(x_57)) { + x_60 = lean_alloc_ctor(0, 2, 0); +} else { + x_60 = x_57; +} +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_19 = x_61; +goto block_42; +} +} +block_42: +{ +lean_object* x_20; lean_object* x_21; +lean_inc(x_15); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +lean_inc(x_4); +lean_inc(x_1); +x_21 = l_Lean_Language_Lean_process_doElab(x_1, x_2, x_3, x_20, x_17, x_4, x_18); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Language_Lean_process_doElab___closed__1; +x_25 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_24); +lean_inc(x_1); +x_26 = l_Lean_Parser_isTerminalCommand(x_1); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_io_promise_new(x_23); +if (lean_obj_tag(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_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_28); +x_31 = l_Lean_Language_Lean_process_parseCmd___lambda__2(x_6, x_1, x_15, x_22, x_7, x_17, x_8, x_4, x_9, x_25, x_30, x_12, x_29); +return x_31; +} +else +{ +uint8_t x_32; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) +{ +return x_27; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_box(0); +x_37 = l_Lean_Language_Lean_process_parseCmd___lambda__2(x_6, x_1, x_15, x_22, x_7, x_17, x_8, x_4, x_9, x_25, x_36, x_12, x_23); +return x_37; +} +} +else +{ +uint8_t x_38; +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_21); +if (x_38 == 0) +{ +return x_21; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_21, 0); +x_40 = lean_ctor_get(x_21, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_21); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +} +} +else +{ +uint8_t x_74; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_74 = !lean_is_exclusive(x_14); +if (x_74 == 0) +{ +return x_14; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_14, 0); +x_76 = lean_ctor_get(x_14, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_14); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__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_6; +x_6 = lean_ctor_get(x_1, 2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_apply_3(x_2, x_7, x_4, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_6, 0); +x_10 = 1; +x_11 = lean_box(x_10); +x_12 = lean_st_ref_set(x_9, x_11, x_5); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_apply_3(x_2, x_13, x_4, x_14); +return x_15; +} +else +{ +uint8_t x_16; +lean_dec(x_4); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 0); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_12); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = l_Lean_Language_SnapshotTask_pure___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_6); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Lean_Language_Lean_process_parseCmd(x_8, x_2, x_9, x_3, x_4, x_5, x_7); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +x_13 = l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1; +lean_ctor_set(x_10, 0, x_13); +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) +{ +return x_10; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__5), 7, 5); +lean_closure_set(x_8, 0, x_6); +lean_closure_set(x_8, 1, x_1); +lean_closure_set(x_8, 2, x_2); +lean_closure_set(x_8, 3, x_3); +lean_closure_set(x_8, 4, x_4); +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +x_10 = 1; +x_11 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_5, x_8, x_9, x_10, x_7); +return x_11; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parsing", 7, 7); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); +x_13 = l_Lean_Elab_Command_instInhabitedScope; +x_14 = l_List_head_x21___rarg(x_13, x_12); +lean_dec(x_12); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 2); +lean_inc(x_16); +x_17 = lean_ctor_get(x_14, 3); +lean_inc(x_17); +lean_dec(x_14); +lean_inc(x_15); +x_18 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_15); +lean_ctor_set(x_18, 2, x_16); +lean_ctor_set(x_18, 3, x_17); +lean_inc(x_3); +x_19 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__1___boxed), 4, 3); +lean_closure_set(x_19, 0, x_3); +lean_closure_set(x_19, 1, x_18); +lean_closure_set(x_19, 2, x_1); +x_20 = l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1; +x_21 = lean_box(0); +x_22 = lean_profileit(x_20, x_15, x_19, x_21); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_box(0); +x_28 = l_Lean_Language_Lean_process_parseCmd___lambda__3(x_24, x_2, x_10, x_3, x_15, x_26, x_25, x_4, x_5, x_6, x_27, x_8, x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_15); +return x_28; +} +else +{ +lean_object* x_29; uint8_t x_30; +x_29 = lean_ctor_get(x_22, 0); +lean_inc(x_29); +lean_dec(x_22); +x_30 = !lean_is_exclusive(x_23); +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; uint8_t x_38; +x_31 = lean_ctor_get(x_23, 0); +x_32 = lean_ctor_get(x_23, 1); +x_33 = lean_ctor_get(x_6, 0); +lean_inc(x_33); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_31); +lean_inc(x_3); +lean_inc(x_29); +x_34 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed), 13, 10); +lean_closure_set(x_34, 0, x_29); +lean_closure_set(x_34, 1, x_2); +lean_closure_set(x_34, 2, x_10); +lean_closure_set(x_34, 3, x_3); +lean_closure_set(x_34, 4, x_15); +lean_closure_set(x_34, 5, x_32); +lean_closure_set(x_34, 6, x_31); +lean_closure_set(x_34, 7, x_4); +lean_closure_set(x_34, 8, x_5); +lean_closure_set(x_34, 9, x_6); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +x_38 = l_Lean_Syntax_eqWithInfo(x_29, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_33); +lean_free_object(x_23); +lean_dec(x_31); +lean_dec(x_5); +lean_dec(x_4); +x_39 = lean_box(0); +x_40 = l_Lean_Language_Lean_process_parseCmd___lambda__4(x_3, x_34, x_39, x_8, x_9); +lean_dec(x_3); +return x_40; +} +else +{ +lean_dec(x_34); +lean_dec(x_8); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_41; +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_31); +lean_dec(x_5); +lean_dec(x_3); +x_41 = lean_io_promise_resolve(x_33, x_4, x_9); +lean_dec(x_4); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +return x_41; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_41); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +else +{ +uint8_t x_46; +x_46 = !lean_is_exclusive(x_41); +if (x_46 == 0) +{ +return x_41; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_41, 0); +x_48 = lean_ctor_get(x_41, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_41); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_33); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = lean_ctor_get(x_33, 1); +lean_dec(x_51); +x_52 = lean_ctor_get(x_33, 0); +lean_dec(x_52); +x_53 = !lean_is_exclusive(x_36); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_36, 0); +x_55 = lean_io_promise_new(x_9); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +lean_inc(x_56); +x_58 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_58, 0, x_31); +lean_closure_set(x_58, 1, x_5); +lean_closure_set(x_58, 2, x_56); +lean_closure_set(x_58, 3, x_3); +lean_closure_set(x_58, 4, x_54); +x_59 = lean_ctor_get(x_35, 4); +lean_inc(x_59); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = 0; +x_62 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_59, x_58, x_60, x_61, x_57); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = lean_box(0); +x_65 = lean_io_promise_result(x_56); +lean_ctor_set(x_23, 1, x_65); +lean_ctor_set(x_23, 0, x_64); +lean_ctor_set(x_36, 0, x_23); +x_66 = lean_io_promise_resolve(x_33, x_4, x_63); +lean_dec(x_4); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) +{ +return x_66; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_66, 0); +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_66); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_66); +if (x_71 == 0) +{ +return x_66; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_66, 0); +x_73 = lean_ctor_get(x_66, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_66); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_56); +lean_free_object(x_36); +lean_free_object(x_33); +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_4); +x_75 = !lean_is_exclusive(x_62); +if (x_75 == 0) +{ +return x_62; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_62, 0); +x_77 = lean_ctor_get(x_62, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_62); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +else +{ +uint8_t x_79; +lean_free_object(x_36); +lean_dec(x_54); +lean_free_object(x_33); +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_31); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_79 = !lean_is_exclusive(x_55); +if (x_79 == 0) +{ +return x_55; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_55, 0); +x_81 = lean_ctor_get(x_55, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_55); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +} +} +else +{ +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_36, 0); +lean_inc(x_83); +lean_dec(x_36); +x_84 = lean_io_promise_new(x_9); +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; uint8_t x_90; lean_object* x_91; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); lean_inc(x_85); +x_87 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_87, 0, x_31); +lean_closure_set(x_87, 1, x_5); +lean_closure_set(x_87, 2, x_85); +lean_closure_set(x_87, 3, x_3); +lean_closure_set(x_87, 4, x_83); +x_88 = lean_ctor_get(x_35, 4); +lean_inc(x_88); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = 0; +x_91 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_88, x_87, x_89, x_90, x_86); +if (lean_obj_tag(x_91) == 0) +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +lean_dec(x_91); +x_93 = lean_box(0); +x_94 = lean_io_promise_result(x_85); +lean_ctor_set(x_23, 1, x_94); +lean_ctor_set(x_23, 0, x_93); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_23); +lean_ctor_set(x_33, 1, x_95); +x_96 = lean_io_promise_resolve(x_33, x_4, x_92); +lean_dec(x_4); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_99 = x_96; +} else { + lean_dec_ref(x_96); + x_99 = lean_box(0); +} +if (lean_is_scalar(x_99)) { + x_100 = lean_alloc_ctor(0, 2, 0); +} else { + x_100 = x_99; +} +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_98); +return x_100; +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_101 = lean_ctor_get(x_96, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_96, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_103 = x_96; +} else { + lean_dec_ref(x_96); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); +} else { + x_104 = x_103; +} +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_85); +lean_free_object(x_33); +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_4); +x_105 = lean_ctor_get(x_91, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_91, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_107 = x_91; +} else { + lean_dec_ref(x_91); + x_107 = lean_box(0); +} +if (lean_is_scalar(x_107)) { + x_108 = lean_alloc_ctor(1, 2, 0); +} else { + x_108 = x_107; +} +lean_ctor_set(x_108, 0, x_105); +lean_ctor_set(x_108, 1, x_106); +return x_108; +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_dec(x_83); -x_86 = lean_st_ref_take(x_1, x_84); -if (lean_obj_tag(x_86) == 0) +lean_free_object(x_33); +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_31); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_109 = lean_ctor_get(x_84, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_84, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_111 = x_84; +} else { + lean_dec_ref(x_84); + x_111 = lean_box(0); +} +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(1, 2, 0); +} else { + x_112 = x_111; +} +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_110); +return x_112; +} +} +} +else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_87 = lean_ctor_get(x_86, 1); -lean_inc(x_87); -lean_dec(x_86); -x_88 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_85); -lean_ctor_set(x_89, 1, x_88); -x_90 = lean_st_ref_set(x_1, x_89, x_87); -if (lean_obj_tag(x_90) == 0) +lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_33); +x_113 = lean_ctor_get(x_36, 0); +lean_inc(x_113); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + x_114 = x_36; +} else { + lean_dec_ref(x_36); + x_114 = lean_box(0); +} +x_115 = lean_io_promise_new(x_9); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_object* x_122; +x_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); +lean_inc(x_116); +x_118 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_118, 0, x_31); +lean_closure_set(x_118, 1, x_5); +lean_closure_set(x_118, 2, x_116); +lean_closure_set(x_118, 3, x_3); +lean_closure_set(x_118, 4, x_113); +x_119 = lean_ctor_get(x_35, 4); +lean_inc(x_119); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = 0; +x_122 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_119, x_118, x_120, x_121, x_117); +if (lean_obj_tag(x_122) == 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; +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = lean_box(0); +x_125 = lean_io_promise_result(x_116); +lean_ctor_set(x_23, 1, x_125); +lean_ctor_set(x_23, 0, x_124); +if (lean_is_scalar(x_114)) { + x_126 = lean_alloc_ctor(1, 1, 0); +} else { + x_126 = x_114; +} +lean_ctor_set(x_126, 0, x_23); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_35); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_io_promise_resolve(x_127, x_4, x_123); +lean_dec(x_4); +if (lean_obj_tag(x_128) == 0) { -lean_object* x_91; lean_object* x_92; -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_92 = lean_st_ref_get(x_7, x_91); -lean_dec(x_7); -if (lean_obj_tag(x_92) == 0) +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_128, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_131 = x_128; +} else { + lean_dec_ref(x_128); + x_131 = lean_box(0); +} +if (lean_is_scalar(x_131)) { + x_132 = lean_alloc_ctor(0, 2, 0); +} else { + x_132 = x_131; +} +lean_ctor_set(x_132, 0, x_129); +lean_ctor_set(x_132, 1, x_130); +return x_132; +} +else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -lean_dec(x_92); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -x_96 = l_String_isEmpty(x_81); -if (x_96 == 0) +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_133 = lean_ctor_get(x_128, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_128, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_135 = x_128; +} else { + lean_dec_ref(x_128); + x_135 = lean_box(0); +} +if (lean_is_scalar(x_135)) { + x_136 = lean_alloc_ctor(1, 2, 0); +} else { + x_136 = x_135; +} +lean_ctor_set(x_136, 0, x_133); +lean_ctor_set(x_136, 1, x_134); +return x_136; +} +} +else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_97 = l_Lean_FileMap_toPosition(x_17, x_4); +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_dec(x_116); +lean_dec(x_114); +lean_dec(x_35); +lean_free_object(x_23); lean_dec(x_4); -x_98 = lean_box(0); -x_99 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_99, 0, x_81); -x_100 = l_Lean_MessageData_ofFormat(x_99); -x_101 = 0; -x_102 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__4; -x_103 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_103, 0, x_16); -lean_ctor_set(x_103, 1, x_97); -lean_ctor_set(x_103, 2, x_98); -lean_ctor_set(x_103, 3, x_102); -lean_ctor_set(x_103, 4, x_100); -lean_ctor_set_uint8(x_103, sizeof(void*)*5, x_26); -lean_ctor_set_uint8(x_103, sizeof(void*)*5 + 1, x_101); -x_104 = l_Lean_MessageLog_add(x_103, x_95); -x_105 = lean_box(0); -x_106 = l_Lean_Language_Lean_process_doElab___lambda__2(x_93, x_3, x_104, x_105, x_94); -lean_dec(x_3); -return x_106; +x_137 = lean_ctor_get(x_122, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_122, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_139 = x_122; +} else { + lean_dec_ref(x_122); + x_139 = lean_box(0); +} +if (lean_is_scalar(x_139)) { + x_140 = lean_alloc_ctor(1, 2, 0); +} else { + x_140 = x_139; +} +lean_ctor_set(x_140, 0, x_137); +lean_ctor_set(x_140, 1, x_138); +return x_140; +} } else { -lean_object* x_107; lean_object* x_108; -lean_dec(x_81); -lean_dec(x_17); -lean_dec(x_16); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_114); +lean_dec(x_113); +lean_dec(x_35); +lean_free_object(x_23); +lean_dec(x_31); +lean_dec(x_5); lean_dec(x_4); -x_107 = lean_box(0); -x_108 = l_Lean_Language_Lean_process_doElab___lambda__2(x_93, x_3, x_95, x_107, x_94); lean_dec(x_3); -return x_108; +x_141 = lean_ctor_get(x_115, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_115, 1); +lean_inc(x_142); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_143 = x_115; +} else { + lean_dec_ref(x_115); + x_143 = lean_box(0); +} +if (lean_is_scalar(x_143)) { + x_144 = lean_alloc_ctor(1, 2, 0); +} else { + x_144 = x_143; +} +lean_ctor_set(x_144, 0, x_141); +lean_ctor_set(x_144, 1, x_142); +return x_144; +} +} +} } } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_81); -lean_dec(x_17); -lean_dec(x_16); +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; +x_145 = lean_ctor_get(x_23, 0); +x_146 = lean_ctor_get(x_23, 1); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_23); +x_147 = lean_ctor_get(x_6, 0); +lean_inc(x_147); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_145); +lean_inc(x_3); +lean_inc(x_29); +x_148 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed), 13, 10); +lean_closure_set(x_148, 0, x_29); +lean_closure_set(x_148, 1, x_2); +lean_closure_set(x_148, 2, x_10); +lean_closure_set(x_148, 3, x_3); +lean_closure_set(x_148, 4, x_15); +lean_closure_set(x_148, 5, x_146); +lean_closure_set(x_148, 6, x_145); +lean_closure_set(x_148, 7, x_4); +lean_closure_set(x_148, 8, x_5); +lean_closure_set(x_148, 9, x_6); +x_149 = lean_ctor_get(x_147, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_147, 1); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 1); +lean_inc(x_151); +x_152 = l_Lean_Syntax_eqWithInfo(x_29, x_151); +if (x_152 == 0) +{ +lean_object* x_153; lean_object* x_154; +lean_dec(x_150); +lean_dec(x_149); +lean_dec(x_147); +lean_dec(x_145); +lean_dec(x_5); lean_dec(x_4); +x_153 = lean_box(0); +x_154 = l_Lean_Language_Lean_process_parseCmd___lambda__4(x_3, x_148, x_153, x_8, x_9); lean_dec(x_3); -x_109 = lean_ctor_get(x_92, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_92, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_111 = x_92; +return x_154; +} +else +{ +lean_dec(x_148); +lean_dec(x_8); +if (lean_obj_tag(x_150) == 0) +{ +lean_object* x_155; +lean_dec(x_149); +lean_dec(x_145); +lean_dec(x_5); +lean_dec(x_3); +x_155 = lean_io_promise_resolve(x_147, x_4, x_9); +lean_dec(x_4); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_158 = x_155; } else { - lean_dec_ref(x_92); - x_111 = lean_box(0); + lean_dec_ref(x_155); + x_158 = lean_box(0); } -if (lean_is_scalar(x_111)) { - x_112 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(0, 2, 0); } else { - x_112 = x_111; + x_159 = x_158; } -lean_ctor_set(x_112, 0, x_109); -lean_ctor_set(x_112, 1, x_110); -return x_112; +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_157); +return x_159; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_160 = lean_ctor_get(x_155, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_155, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_162 = x_155; +} else { + lean_dec_ref(x_155); + x_162 = lean_box(0); +} +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); +} else { + x_163 = x_162; +} +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_81); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_164 = x_147; +} else { + lean_dec_ref(x_147); + x_164 = lean_box(0); +} +x_165 = lean_ctor_get(x_150, 0); +lean_inc(x_165); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + x_166 = x_150; +} else { + lean_dec_ref(x_150); + x_166 = lean_box(0); +} +x_167 = lean_io_promise_new(x_9); +if (lean_obj_tag(x_167) == 0) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; lean_object* x_174; +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); +lean_inc(x_169); +lean_dec(x_167); +lean_inc(x_168); +x_170 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_170, 0, x_145); +lean_closure_set(x_170, 1, x_5); +lean_closure_set(x_170, 2, x_168); +lean_closure_set(x_170, 3, x_3); +lean_closure_set(x_170, 4, x_165); +x_171 = lean_ctor_get(x_149, 4); +lean_inc(x_171); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = 0; +x_174 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_171, x_170, x_172, x_173, x_169); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); +x_176 = lean_box(0); +x_177 = lean_io_promise_result(x_168); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_176); +lean_ctor_set(x_178, 1, x_177); +if (lean_is_scalar(x_166)) { + x_179 = lean_alloc_ctor(1, 1, 0); +} else { + x_179 = x_166; +} +lean_ctor_set(x_179, 0, x_178); +if (lean_is_scalar(x_164)) { + x_180 = lean_alloc_ctor(0, 2, 0); +} else { + x_180 = x_164; +} +lean_ctor_set(x_180, 0, x_149); +lean_ctor_set(x_180, 1, x_179); +x_181 = lean_io_promise_resolve(x_180, x_4, x_175); lean_dec(x_4); -lean_dec(x_3); -x_113 = lean_ctor_get(x_90, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_90, 1); -lean_inc(x_114); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_115 = x_90; +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_184 = x_181; } else { - lean_dec_ref(x_90); - x_115 = lean_box(0); + lean_dec_ref(x_181); + x_184 = lean_box(0); } -if (lean_is_scalar(x_115)) { - x_116 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(0, 2, 0); } else { - x_116 = x_115; + x_185 = x_184; } -lean_ctor_set(x_116, 0, x_113); -lean_ctor_set(x_116, 1, x_114); -return x_116; +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +return x_185; +} +else +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_186 = lean_ctor_get(x_181, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_181, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_188 = x_181; +} else { + lean_dec_ref(x_181); + x_188 = lean_box(0); +} +if (lean_is_scalar(x_188)) { + x_189 = lean_alloc_ctor(1, 2, 0); +} else { + x_189 = x_188; +} +lean_ctor_set(x_189, 0, x_186); +lean_ctor_set(x_189, 1, x_187); +return x_189; } } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_85); -lean_dec(x_81); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_168); +lean_dec(x_166); +lean_dec(x_164); +lean_dec(x_149); lean_dec(x_4); -lean_dec(x_3); -x_117 = lean_ctor_get(x_86, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_86, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_119 = x_86; +x_190 = lean_ctor_get(x_174, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_174, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_192 = x_174; } else { - lean_dec_ref(x_86); - x_119 = lean_box(0); + lean_dec_ref(x_174); + x_192 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); } else { - x_120 = x_119; + x_193 = x_192; } -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_dec(x_81); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_7); +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_166); +lean_dec(x_165); +lean_dec(x_164); +lean_dec(x_149); +lean_dec(x_145); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_121 = lean_ctor_get(x_82, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_82, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_123 = x_82; +x_194 = lean_ctor_get(x_167, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_167, 1); +lean_inc(x_195); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_196 = x_167; } else { - lean_dec_ref(x_82); - x_123 = lean_box(0); + lean_dec_ref(x_167); + x_196 = lean_box(0); } -if (lean_is_scalar(x_123)) { - x_124 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_196)) { + x_197 = lean_alloc_ctor(1, 2, 0); } else { - x_124 = x_123; + x_197 = x_196; } -lean_ctor_set(x_124, 0, x_121); -lean_ctor_set(x_124, 1, x_122); -return x_124; +lean_ctor_set(x_197, 0, x_194); +lean_ctor_set(x_197, 1, x_195); +return x_197; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_8, x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_6, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_12); +x_14 = lean_box(0); +x_15 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9); +return x_15; } +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec(x_13); +lean_inc(x_17); +x_18 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_17, x_9); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_17); +lean_dec(x_16); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_21, x_8, x_20); +return x_22; } +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_19, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_dec(x_18); +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_27 = lean_ctor_get(x_24, 0); +x_28 = lean_ctor_get(x_24, 1); +lean_dec(x_28); +x_29 = lean_ctor_get(x_27, 2); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_Language_Lean_isBeforeEditPos(x_30, x_8, x_25); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_unbox(x_32); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_free_object(x_24); +lean_free_object(x_19); +lean_dec(x_17); +lean_dec(x_16); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_box(0); +x_36 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_35, x_8, x_34); +return x_36; } else { -uint8_t x_125; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_13); -lean_dec(x_7); +uint8_t x_37; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_31); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_31, 1); +x_39 = lean_ctor_get(x_31, 0); +lean_dec(x_39); +x_40 = lean_ctor_get(x_16, 2); +lean_inc(x_40); +x_41 = lean_io_promise_new(x_38); +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; uint8_t x_47; lean_object* x_48; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_42); +x_44 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_44, 0, x_40); +lean_closure_set(x_44, 1, x_5); +lean_closure_set(x_44, 2, x_42); +lean_closure_set(x_44, 3, x_3); +lean_closure_set(x_44, 4, x_17); +x_45 = lean_ctor_get(x_16, 4); +lean_inc(x_45); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = 0; +x_48 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_45, x_44, x_46, x_47, x_43); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = lean_box(0); +x_51 = lean_io_promise_result(x_42); +lean_ctor_set(x_31, 1, x_51); +lean_ctor_set(x_31, 0, x_50); +lean_ctor_set(x_19, 0, x_31); +lean_ctor_set(x_24, 1, x_19); +lean_ctor_set(x_24, 0, x_16); +x_52 = lean_io_promise_resolve(x_24, x_4, x_49); lean_dec(x_4); -lean_dec(x_3); -x_125 = !lean_is_exclusive(x_33); -if (x_125 == 0) +if (lean_obj_tag(x_52) == 0) { -return x_33; +uint8_t x_53; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +return x_52; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_33, 0); -x_127 = lean_ctor_get(x_33, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_33); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; -} +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_52); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } else { -uint8_t x_129; -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_129 = !lean_is_exclusive(x_12); -if (x_129 == 0) +uint8_t x_57; +x_57 = !lean_is_exclusive(x_52); +if (x_57 == 0) { -return x_12; +return x_52; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_12, 0); -x_131 = lean_ctor_get(x_12, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_12); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_52, 0); +x_59 = lean_ctor_get(x_52, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_52); +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_133; -lean_dec(x_7); -lean_dec(x_5); +uint8_t x_61; +lean_dec(x_42); +lean_free_object(x_31); +lean_free_object(x_24); +lean_free_object(x_19); +lean_dec(x_16); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_133 = !lean_is_exclusive(x_9); -if (x_133 == 0) +x_61 = !lean_is_exclusive(x_48); +if (x_61 == 0) { -return x_9; +return x_48; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_9, 0); -x_135 = lean_ctor_get(x_9, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_9); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -return x_136; -} -} +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_48, 0); +x_63 = lean_ctor_get(x_48, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_48); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab(lean_object* x_1, 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; uint8_t x_10; -x_8 = 0; -x_9 = l_Lean_Syntax_getTailPos_x3f(x_1, x_8); -x_10 = !lean_is_exclusive(x_2); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_11 = lean_ctor_get(x_2, 2); -x_12 = lean_ctor_get(x_2, 1); -lean_dec(x_12); -x_13 = l_Lean_Elab_Command_instInhabitedScope; -x_14 = l_List_head_x21___rarg(x_13, x_11); -x_15 = l_Lean_MessageLog_empty; -lean_ctor_set(x_2, 1, x_15); -x_16 = lean_alloc_closure((void*)(l_IO_mkRef___rarg), 2, 1); -lean_closure_set(x_16, 0, x_2); -lean_inc(x_3); -x_17 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__3___boxed), 8, 6); -lean_closure_set(x_17, 0, x_5); -lean_closure_set(x_17, 1, x_6); -lean_closure_set(x_17, 2, x_4); -lean_closure_set(x_17, 3, x_3); -lean_closure_set(x_17, 4, x_1); -lean_closure_set(x_17, 5, x_14); -x_18 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_18, 0, x_16); -lean_closure_set(x_18, 1, x_17); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_inc(x_3); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_19, 1, x_3); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_20, x_18, x_7); -return x_21; } else { -uint8_t x_22; +uint8_t x_65; +lean_dec(x_40); +lean_free_object(x_31); +lean_free_object(x_24); +lean_free_object(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_22 = !lean_is_exclusive(x_9); -if (x_22 == 0) +x_65 = !lean_is_exclusive(x_41); +if (x_65 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_9, 0); -lean_inc(x_23); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_9, 0, x_24); -x_25 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_9, x_18, x_7); -return x_25; +return x_41; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_9, 0); -lean_inc(x_26); -lean_dec(x_9); -lean_inc(x_26); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_28, x_18, x_7); -return x_29; -} +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_41, 0); +x_67 = lean_ctor_get(x_41, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_41); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; } } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_30 = lean_ctor_get(x_2, 0); -x_31 = lean_ctor_get(x_2, 2); -x_32 = lean_ctor_get(x_2, 3); -x_33 = lean_ctor_get(x_2, 4); -x_34 = lean_ctor_get(x_2, 5); -x_35 = lean_ctor_get(x_2, 6); -x_36 = lean_ctor_get(x_2, 7); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_2); -x_37 = l_Lean_Elab_Command_instInhabitedScope; -x_38 = l_List_head_x21___rarg(x_37, x_31); -x_39 = l_Lean_MessageLog_empty; -x_40 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_40, 0, x_30); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_31); -lean_ctor_set(x_40, 3, x_32); -lean_ctor_set(x_40, 4, x_33); -lean_ctor_set(x_40, 5, x_34); -lean_ctor_set(x_40, 6, x_35); -lean_ctor_set(x_40, 7, x_36); -x_41 = lean_alloc_closure((void*)(l_IO_mkRef___rarg), 2, 1); -lean_closure_set(x_41, 0, x_40); -lean_inc(x_3); -x_42 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__3___boxed), 8, 6); -lean_closure_set(x_42, 0, x_5); -lean_closure_set(x_42, 1, x_6); -lean_closure_set(x_42, 2, x_4); -lean_closure_set(x_42, 3, x_3); -lean_closure_set(x_42, 4, x_1); -lean_closure_set(x_42, 5, x_38); -x_43 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_43, 0, x_41); -lean_closure_set(x_43, 1, x_42); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_inc(x_3); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_3); -lean_ctor_set(x_44, 1, x_3); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_45, x_43, x_7); -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_dec(x_3); -x_47 = lean_ctor_get(x_9, 0); -lean_inc(x_47); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - x_48 = x_9; -} else { - lean_dec_ref(x_9); - x_48 = lean_box(0); -} -lean_inc(x_47); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_47); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(1, 1, 0); -} else { - x_50 = x_48; -} -lean_ctor_set(x_50, 0, x_49); -x_51 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_50, x_43, x_7); -return x_51; -} -} -} -} -LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_31, 1); +lean_inc(x_69); +lean_dec(x_31); +x_70 = lean_ctor_get(x_16, 2); +lean_inc(x_70); +x_71 = lean_io_promise_new(x_69); +if (lean_obj_tag(x_71) == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = lean_unbox(x_2); -lean_dec(x_2); -x_5 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_1, x_4, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +lean_inc(x_72); +x_74 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_74, 0, x_70); +lean_closure_set(x_74, 1, x_5); +lean_closure_set(x_74, 2, x_72); +lean_closure_set(x_74, 3, x_3); +lean_closure_set(x_74, 4, x_17); +x_75 = lean_ctor_get(x_16, 4); +lean_inc(x_75); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = 0; +x_78 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_75, x_74, x_76, x_77, x_73); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_6; -x_6 = l_Lean_Language_Lean_process_doElab___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_box(0); +x_81 = lean_io_promise_result(x_72); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_19, 0, x_82); +lean_ctor_set(x_24, 1, x_19); +lean_ctor_set(x_24, 0, x_16); +x_83 = lean_io_promise_resolve(x_24, x_4, x_79); lean_dec(x_4); -lean_dec(x_2); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Language_Lean_process_doElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_6); -lean_dec(x_1); -return x_9; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("parseCmd", 8, 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; -x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; -x_3 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__1; -x_4 = l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1; -x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3() { -_start: +if (lean_obj_tag(x_83) == 0) { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2; -x_2 = 1; -x_3 = l_Lean_Name_toString(x_1, x_2); -return x_3; +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_86 = x_83; +} else { + lean_dec_ref(x_83); + x_86 = lean_box(0); } +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_86; } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_1, x_8); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_box(0); -x_13 = l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3; -x_14 = 0; -x_15 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_11); -lean_ctor_set(x_15, 2, x_12); -lean_ctor_set_uint8(x_15, sizeof(void*)*3, x_14); -x_16 = l_Lean_Syntax_getRange_x3f(x_2, x_14); -x_17 = lean_io_promise_result(x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -x_19 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_4); -lean_ctor_set(x_19, 3, x_18); -lean_ctor_set(x_19, 4, x_5); -lean_ctor_set(x_19, 5, x_6); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_7); -lean_ctor_set(x_9, 0, x_20); -return x_9; +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_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_9, 0); -x_22 = lean_ctor_get(x_9, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_9); -x_23 = lean_box(0); -x_24 = l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3; -x_25 = 0; -x_26 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_21); -lean_ctor_set(x_26, 2, x_23); -lean_ctor_set_uint8(x_26, sizeof(void*)*3, x_25); -x_27 = l_Lean_Syntax_getRange_x3f(x_2, x_25); -x_28 = lean_io_promise_result(x_3); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_30, 0, x_26); -lean_ctor_set(x_30, 1, x_2); -lean_ctor_set(x_30, 2, x_4); -lean_ctor_set(x_30, 3, x_29); -lean_ctor_set(x_30, 4, x_5); -lean_ctor_set(x_30, 5, x_6); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_7); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_22); -return x_32; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_83, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_83, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_90 = x_83; +} else { + lean_dec_ref(x_83); + x_90 = lean_box(0); } +if (lean_is_scalar(x_90)) { + x_91 = lean_alloc_ctor(1, 2, 0); +} else { + x_91 = x_90; } -else -{ -uint8_t x_33; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_33 = !lean_is_exclusive(x_9); -if (x_33 == 0) -{ -return x_9; +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_89); +return x_91; } -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_9, 0); -x_35 = lean_ctor_get(x_9, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_9); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; } +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_72); +lean_free_object(x_24); +lean_free_object(x_19); +lean_dec(x_16); +lean_dec(x_4); +x_92 = lean_ctor_get(x_78, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_78, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_94 = x_78; +} else { + lean_dec_ref(x_78); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; } +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___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; -x_5 = lean_box(0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +lean_dec(x_70); +lean_free_object(x_24); +lean_free_object(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_7 = l_Lean_Language_Lean_process_parseCmd(x_5, x_1, x_6, x_2, x_4); -return x_7; +x_96 = lean_ctor_get(x_71, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_71, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_98 = x_71; +} else { + lean_dec_ref(x_71); + x_98 = lean_box(0); } +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; +lean_ctor_set(x_99, 0, x_96); +lean_ctor_set(x_99, 1, x_97); +return x_99; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = lean_io_promise_new(x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -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); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_60; lean_object* x_61; -x_60 = l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1; -x_61 = lean_st_mk_ref(x_60, x_12); -if (lean_obj_tag(x_61) == 0) +} +} +else { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_13 = x_62; -x_14 = x_63; -goto block_59; +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_100 = lean_ctor_get(x_24, 0); +lean_inc(x_100); +lean_dec(x_24); +x_101 = lean_ctor_get(x_100, 2); +lean_inc(x_101); +lean_dec(x_100); +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +lean_dec(x_101); +x_103 = l_Lean_Language_Lean_isBeforeEditPos(x_102, x_8, x_25); +lean_dec(x_102); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_unbox(x_104); +lean_dec(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_free_object(x_19); +lean_dec(x_17); +lean_dec(x_16); +x_106 = lean_ctor_get(x_103, 1); +lean_inc(x_106); +lean_dec(x_103); +x_107 = lean_box(0); +x_108 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_107, x_8, x_106); +return x_108; } else { -uint8_t x_64; -lean_dec(x_11); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_8); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_61); -if (x_64 == 0) -{ -return x_61; +x_109 = lean_ctor_get(x_103, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_110 = x_103; +} else { + lean_dec_ref(x_103); + x_110 = lean_box(0); } -else +x_111 = lean_ctor_get(x_16, 2); +lean_inc(x_111); +x_112 = lean_io_promise_new(x_109); +if (lean_obj_tag(x_112) == 0) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_61, 0); -x_66 = lean_ctor_get(x_61, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_61); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint8_t x_118; lean_object* x_119; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +lean_inc(x_113); +x_115 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_115, 0, x_111); +lean_closure_set(x_115, 1, x_5); +lean_closure_set(x_115, 2, x_113); +lean_closure_set(x_115, 3, x_3); +lean_closure_set(x_115, 4, x_17); +x_116 = lean_ctor_get(x_16, 4); +lean_inc(x_116); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = 0; +x_119 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_116, x_115, x_117, x_118, x_114); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = lean_box(0); +x_122 = lean_io_promise_result(x_113); +if (lean_is_scalar(x_110)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_110; +} +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_19, 0, x_123); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_16); +lean_ctor_set(x_124, 1, x_19); +x_125 = lean_io_promise_resolve(x_124, x_4, x_120); +lean_dec(x_4); +if (lean_obj_tag(x_125) == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_128 = x_125; +} else { + lean_dec_ref(x_125); + x_128 = lean_box(0); } +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_128; } +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_7, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -lean_dec(x_68); -x_70 = lean_ctor_get(x_69, 5); -lean_inc(x_70); -lean_dec(x_69); -x_13 = x_70; -x_14 = x_12; -goto block_59; +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_125, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_125, 1); +lean_inc(x_131); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_132 = x_125; +} else { + lean_dec_ref(x_125); + x_132 = lean_box(0); } -block_59: -{ -lean_object* x_15; -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_40; -x_40 = lean_box(0); -x_15 = x_40; -goto block_39; +if (lean_is_scalar(x_132)) { + x_133 = lean_alloc_ctor(1, 2, 0); +} else { + x_133 = x_132; +} +lean_ctor_set(x_133, 0, x_130); +lean_ctor_set(x_133, 1, x_131); +return x_133; } -else -{ -uint8_t x_41; -x_41 = !lean_is_exclusive(x_7); -if (x_41 == 0) -{ -lean_object* x_42; uint8_t x_43; -x_42 = lean_ctor_get(x_7, 0); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); -lean_dec(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -x_47 = lean_ctor_get(x_44, 3); -lean_inc(x_47); -lean_dec(x_44); -lean_ctor_set(x_42, 1, x_47); -lean_ctor_set(x_42, 0, x_46); -x_15 = x_7; -goto block_39; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_42, 0); -lean_inc(x_48); -lean_dec(x_42); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 3); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_7, 0, x_51); -x_15 = x_7; -goto block_39; +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_113); +lean_dec(x_110); +lean_free_object(x_19); +lean_dec(x_16); +lean_dec(x_4); +x_134 = lean_ctor_get(x_119, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_119, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_136 = x_119; +} else { + lean_dec_ref(x_119); + x_136 = lean_box(0); +} +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 2, 0); +} else { + x_137 = x_136; +} +lean_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_135); +return x_137; } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_52 = lean_ctor_get(x_7, 0); -lean_inc(x_52); -lean_dec(x_7); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_54 = x_52; +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_111); +lean_dec(x_110); +lean_free_object(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_138 = lean_ctor_get(x_112, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_112, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_140 = x_112; } else { - lean_dec_ref(x_52); - x_54 = lean_box(0); + lean_dec_ref(x_112); + x_140 = lean_box(0); } -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -x_56 = lean_ctor_get(x_53, 3); -lean_inc(x_56); -lean_dec(x_53); -if (lean_is_scalar(x_54)) { - x_57 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(1, 2, 0); } else { - x_57 = x_54; + x_141 = x_140; } -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_15 = x_58; -goto block_39; +lean_ctor_set(x_141, 0, x_138); +lean_ctor_set(x_141, 1, x_139); +return x_141; } } -block_39: +} +} +else { -lean_object* x_16; lean_object* x_17; -lean_inc(x_11); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -lean_inc(x_4); -lean_inc(x_13); -lean_inc(x_1); -x_17 = l_Lean_Language_Lean_process_doElab(x_1, x_2, x_3, x_16, x_13, x_4, x_14); -if (lean_obj_tag(x_17) == 0) +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; +x_142 = lean_ctor_get(x_19, 0); +lean_inc(x_142); +lean_dec(x_19); +x_143 = lean_ctor_get(x_18, 1); +lean_inc(x_143); +lean_dec(x_18); +x_144 = lean_ctor_get(x_142, 0); +lean_inc(x_144); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_145 = x_142; +} else { + lean_dec_ref(x_142); + x_145 = lean_box(0); +} +x_146 = lean_ctor_get(x_144, 2); +lean_inc(x_146); +lean_dec(x_144); +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +lean_dec(x_146); +x_148 = l_Lean_Language_Lean_isBeforeEditPos(x_147, x_8, x_143); +lean_dec(x_147); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_unbox(x_149); +lean_dec(x_149); +if (x_150 == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); +lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_145); lean_dec(x_17); -lean_inc(x_1); -x_20 = l_Lean_Parser_isTerminalCommand(x_1); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; -lean_inc(x_6); -x_21 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__2), 4, 2); -lean_closure_set(x_21, 0, x_6); -lean_closure_set(x_21, 1, x_4); -x_22 = lean_ctor_get(x_18, 0); -lean_inc(x_22); -x_23 = 0; -lean_inc(x_18); -x_24 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_18, x_21, x_22, x_23, x_19); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_25); -x_28 = l_Lean_Language_Lean_process_parseCmd___lambda__1(x_5, x_1, x_11, x_6, x_18, x_13, x_27, x_26); -return x_28; +lean_dec(x_16); +x_151 = lean_ctor_get(x_148, 1); +lean_inc(x_151); +lean_dec(x_148); +x_152 = lean_box(0); +x_153 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_152, x_8, x_151); +return x_153; } else { -uint8_t x_29; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_11); +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_8); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_2); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) +x_154 = lean_ctor_get(x_148, 1); +lean_inc(x_154); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_155 = x_148; +} else { + lean_dec_ref(x_148); + x_155 = lean_box(0); +} +x_156 = lean_ctor_get(x_16, 2); +lean_inc(x_156); +x_157 = lean_io_promise_new(x_154); +if (lean_obj_tag(x_157) == 0) { -return x_24; +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); +lean_inc(x_159); +lean_dec(x_157); +lean_inc(x_158); +x_160 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 7, 5); +lean_closure_set(x_160, 0, x_156); +lean_closure_set(x_160, 1, x_5); +lean_closure_set(x_160, 2, x_158); +lean_closure_set(x_160, 3, x_3); +lean_closure_set(x_160, 4, x_17); +x_161 = lean_ctor_get(x_16, 4); +lean_inc(x_161); +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = 0; +x_164 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_161, x_160, x_162, x_163, x_159); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_165 = lean_ctor_get(x_164, 1); +lean_inc(x_165); +lean_dec(x_164); +x_166 = lean_box(0); +x_167 = lean_io_promise_result(x_158); +if (lean_is_scalar(x_155)) { + x_168 = lean_alloc_ctor(0, 2, 0); +} else { + x_168 = x_155; +} +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +x_169 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_169, 0, x_168); +if (lean_is_scalar(x_145)) { + x_170 = lean_alloc_ctor(0, 2, 0); +} else { + x_170 = x_145; } -else +lean_ctor_set(x_170, 0, x_16); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_io_promise_resolve(x_170, x_4, x_165); +lean_dec(x_4); +if (lean_obj_tag(x_171) == 0) { -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_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +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_171); + x_174 = lean_box(0); } +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(0, 2, 0); +} else { + x_175 = x_174; } +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; } else { -lean_object* x_33; lean_object* x_34; -lean_dec(x_4); -x_33 = lean_box(0); -x_34 = l_Lean_Language_Lean_process_parseCmd___lambda__1(x_5, x_1, x_11, x_6, x_18, x_13, x_33, x_19); -return x_34; +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_176 = lean_ctor_get(x_171, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_171, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_178 = x_171; +} else { + lean_dec_ref(x_171); + x_178 = lean_box(0); } +if (lean_is_scalar(x_178)) { + x_179 = lean_alloc_ctor(1, 2, 0); +} else { + x_179 = x_178; +} +lean_ctor_set(x_179, 0, x_176); +lean_ctor_set(x_179, 1, x_177); +return x_179; } -else -{ -uint8_t x_35; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_35 = !lean_is_exclusive(x_17); -if (x_35 == 0) -{ -return x_17; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_17, 0); -x_37 = lean_ctor_get(x_17, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_17); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_145); +lean_dec(x_16); +lean_dec(x_4); +x_180 = lean_ctor_get(x_164, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_164, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_182 = x_164; +} else { + lean_dec_ref(x_164); + x_182 = lean_box(0); } +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 2, 0); +} else { + x_183 = x_182; } +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; } } else { -uint8_t x_71; -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_145); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_71 = !lean_is_exclusive(x_10); -if (x_71 == 0) -{ -return x_10; +x_184 = lean_ctor_get(x_157, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_157, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_186 = x_157; +} else { + lean_dec_ref(x_157); + x_186 = lean_box(0); } -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_10, 0); -x_73 = lean_ctor_get(x_10, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_10); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(1, 2, 0); +} else { + x_187 = x_186; } +lean_ctor_set(x_187, 0, x_184); +lean_ctor_set(x_187, 1, x_185); +return x_187; } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_ctor_get(x_1, 2); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_box(0); -x_7 = lean_apply_2(x_2, x_6, x_4); -return x_7; } -else -{ -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_5, 0); -x_9 = 1; -x_10 = lean_box(x_9); -x_11 = lean_st_ref_set(x_8, x_10, x_4); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_apply_2(x_2, x_12, x_13); -return x_14; } else { -uint8_t x_15; +uint8_t x_188; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +lean_dec(x_1); +x_188 = !lean_is_exclusive(x_18); +if (x_188 == 0) { -return x_11; +return x_18; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} -} +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_18, 0); +x_190 = lean_ctor_get(x_18, 1); +lean_inc(x_190); +lean_inc(x_189); +lean_dec(x_18); +x_191 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +return x_191; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_6, 0, x_4); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -lean_dec(x_1); -x_8 = l_Lean_Language_Lean_process_parseCmd(x_6, x_2, x_7, x_3, x_5); -return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7(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_Language_Lean_process_parseCmd___closed__1() { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; -x_6 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__6), 5, 3); -lean_closure_set(x_6, 0, x_4); -lean_closure_set(x_6, 1, x_1); -lean_closure_set(x_6, 2, x_2); -x_7 = lean_ctor_get(x_3, 0); -lean_inc(x_7); -x_8 = 1; -x_9 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_3, x_6, x_7, x_8, x_5); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3; +x_3 = l_Lean_Language_Snapshot_Diagnostics_empty; +x_4 = 0; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_1); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___closed__2() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_4, 0, x_2); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_4); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_3); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_instTypeNameSnapshotLeaf; +x_2 = l_Lean_Language_Lean_process_parseCmd___closed__1; +x_3 = l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1() { +static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___closed__3() { _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); +x_1 = l_Lean_Language_Lean_process_parseCmd___closed__2; +x_2 = l_Lean_Language_SnapshotTask_pure___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__5), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_8 = lean_ctor_get(x_1, 0); +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 3); lean_inc(x_8); -x_9 = lean_ctor_get(x_2, 0); -lean_inc(x_9); +x_9 = lean_st_ref_get(x_8, x_7); +lean_dec(x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; uint8_t x_11; x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -x_11 = lean_string_utf8_byte_size(x_10); +x_11 = lean_unbox(x_10); lean_dec(x_10); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_ctor_get(x_3, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -x_16 = l_Lean_Elab_Command_instInhabitedScope; -x_17 = l_List_head_x21___rarg(x_16, x_15); -lean_dec(x_15); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 2); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 3); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_21, 0, x_14); -lean_ctor_set(x_21, 1, x_18); -lean_ctor_set(x_21, 2, x_19); -lean_ctor_set(x_21, 3, x_20); -x_22 = l_Lean_MessageLog_empty; -x_23 = l_Lean_Parser_parseCommand(x_9, x_21, x_1, x_22); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 0); -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); -lean_inc(x_4); -lean_inc(x_26); -lean_inc(x_2); -lean_inc(x_25); -x_28 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed), 9, 7); -lean_closure_set(x_28, 0, x_25); -lean_closure_set(x_28, 1, x_3); -lean_closure_set(x_28, 2, x_8); -lean_closure_set(x_28, 3, x_2); -lean_closure_set(x_28, 4, x_27); -lean_closure_set(x_28, 5, x_26); -lean_closure_set(x_28, 6, x_4); -if (lean_obj_tag(x_4) == 0) +if (x_11 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_2); -x_29 = l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1; -x_30 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_30, 0, x_29); -lean_closure_set(x_30, 1, x_28); -x_31 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_30, x_7); -return x_31; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_box(0); +lean_inc(x_6); +x_14 = l_Lean_Language_Lean_process_parseCmd___lambda__8(x_2, x_3, x_6, x_5, x_4, x_1, x_13, x_6, x_12); +return x_14; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_32 = lean_ctor_get(x_4, 0); -lean_inc(x_32); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_6); lean_dec(x_4); -lean_inc(x_2); -x_33 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__4___boxed), 4, 2); -lean_closure_set(x_33, 0, x_2); -lean_closure_set(x_33, 1, x_28); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -x_37 = l_Lean_Syntax_eqWithInfo(x_25, x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_32); -lean_dec(x_26); -lean_dec(x_2); -x_38 = l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1; -x_39 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_39, 0, x_38); -lean_closure_set(x_39, 1, x_33); -x_40 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_39, x_7); -return x_40; -} -else -{ -lean_dec(x_33); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_34); -lean_dec(x_26); -lean_dec(x_2); -x_41 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_41, 0, x_32); -x_42 = l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2; -x_43 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_43, 0, x_41); -lean_closure_set(x_43, 1, x_42); -x_44 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_43, x_7); -return x_44; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_32); -x_45 = lean_ctor_get(x_35, 0); -lean_inc(x_45); -lean_dec(x_35); -x_46 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7), 5, 3); -lean_closure_set(x_46, 0, x_26); -lean_closure_set(x_46, 1, x_2); -lean_closure_set(x_46, 2, x_45); -x_47 = lean_ctor_get(x_34, 4); -lean_inc(x_47); -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -x_49 = 1; -x_50 = lean_box(x_49); -x_51 = lean_alloc_closure((void*)(l_Lean_Language_SnapshotTask_bindIO___rarg___boxed), 5, 4); -lean_closure_set(x_51, 0, x_47); -lean_closure_set(x_51, 1, x_46); -lean_closure_set(x_51, 2, x_48); -lean_closure_set(x_51, 3, x_50); -x_52 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__8), 3, 1); -lean_closure_set(x_52, 0, x_34); -x_53 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_53, 0, x_51); -lean_closure_set(x_53, 1, x_52); -x_54 = l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2; -x_55 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_55, 0, x_53); -lean_closure_set(x_55, 1, x_54); -x_56 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_55, x_7); -return x_56; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_box(0); -x_9 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_8, x_6, x_7); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_12, x_6, x_7); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_10, 0); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_ctor_get(x_11, 0); -lean_inc(x_15); -lean_dec(x_11); +lean_dec(x_1); +x_15 = lean_ctor_get(x_9, 1); lean_inc(x_15); -x_16 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_15, x_7); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); +lean_dec(x_9); +x_16 = l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1; +x_17 = lean_st_mk_ref(x_16, x_15); if (lean_obj_tag(x_17) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_15); -lean_dec(x_14); -x_18 = lean_ctor_get(x_16, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_box(0); -x_20 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_19, x_6, x_18); -return x_20; -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = l_Lean_Language_Lean_process_parseCmd___closed__1; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_3); +x_23 = l_Lean_Language_SnapshotTask_pure___rarg(x_22); +x_24 = lean_box(0); +x_25 = l_Lean_Language_Lean_process_parseCmd___closed__3; +x_26 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_26, 1, x_24); +lean_ctor_set(x_26, 2, x_2); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_26, 4, x_23); +lean_ctor_set(x_26, 5, x_18); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_20); +x_28 = lean_io_promise_resolve(x_27, x_5, x_19); +lean_dec(x_5); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_17, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_dec(x_16); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_25 = lean_ctor_get(x_22, 0); -x_26 = lean_ctor_get(x_22, 1); -lean_dec(x_26); -x_27 = lean_ctor_get(x_25, 2); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -lean_dec(x_27); -x_29 = l_Lean_Language_Lean_isBeforeEditPos(x_28, x_6, x_23); -lean_dec(x_28); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_unbox(x_30); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); lean_dec(x_30); -if (x_31 == 0) +x_31 = lean_box(0); +lean_ctor_set(x_28, 0, x_31); +return x_28; +} +else { lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_free_object(x_22); -lean_free_object(x_17); -lean_dec(x_15); -lean_dec(x_14); -x_32 = lean_ctor_get(x_29, 1); +x_32 = lean_ctor_get(x_28, 1); lean_inc(x_32); -lean_dec(x_29); +lean_dec(x_28); x_33 = lean_box(0); -x_34 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_33, x_6, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); return x_34; } +} else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_35 = lean_ctor_get(x_29, 1); -lean_inc(x_35); -lean_dec(x_29); -x_36 = lean_ctor_get(x_14, 2); -lean_inc(x_36); -x_37 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7), 5, 3); -lean_closure_set(x_37, 0, x_36); -lean_closure_set(x_37, 1, x_2); -lean_closure_set(x_37, 2, x_15); -x_38 = lean_ctor_get(x_14, 4); -lean_inc(x_38); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = 1; -x_41 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_38, x_37, x_39, x_40, x_35); -if (lean_obj_tag(x_41) == 0) -{ -uint8_t x_42; -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_41, 0); -lean_ctor_set(x_17, 0, x_43); -lean_ctor_set(x_22, 1, x_17); -lean_ctor_set(x_22, 0, x_14); -x_44 = l_Lean_Language_SnapshotTask_pure___rarg(x_22); -lean_ctor_set(x_41, 0, x_44); -return x_41; +return x_28; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_41, 0); -x_46 = lean_ctor_get(x_41, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_41); -lean_ctor_set(x_17, 0, x_45); -lean_ctor_set(x_22, 1, x_17); -lean_ctor_set(x_22, 0, x_14); -x_47 = l_Lean_Language_SnapshotTask_pure___rarg(x_22); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_28, 0); +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_28); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} } } else { -uint8_t x_49; -lean_free_object(x_22); -lean_free_object(x_17); -lean_dec(x_14); -x_49 = !lean_is_exclusive(x_41); -if (x_49 == 0) +uint8_t x_39; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_39 = !lean_is_exclusive(x_17); +if (x_39 == 0) { -return x_41; +return x_17; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_41, 0); -x_51 = lean_ctor_get(x_41, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_41); -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_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_17, 0); +x_41 = lean_ctor_get(x_17, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_17); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } -else -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_53 = lean_ctor_get(x_22, 0); -lean_inc(x_53); -lean_dec(x_22); -x_54 = lean_ctor_get(x_53, 2); -lean_inc(x_54); -lean_dec(x_53); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -lean_dec(x_54); -x_56 = l_Lean_Language_Lean_isBeforeEditPos(x_55, x_6, x_23); -lean_dec(x_55); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_unbox(x_57); -lean_dec(x_57); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_free_object(x_17); -lean_dec(x_15); -lean_dec(x_14); -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = lean_box(0); -x_61 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_60, x_6, x_59); -return x_61; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; +uint8_t x_43; +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_62 = lean_ctor_get(x_56, 1); -lean_inc(x_62); -lean_dec(x_56); -x_63 = lean_ctor_get(x_14, 2); -lean_inc(x_63); -x_64 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7), 5, 3); -lean_closure_set(x_64, 0, x_63); -lean_closure_set(x_64, 1, x_2); -lean_closure_set(x_64, 2, x_15); -x_65 = lean_ctor_get(x_14, 4); -lean_inc(x_65); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = 1; -x_68 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_65, x_64, x_66, x_67, x_62); -if (lean_obj_tag(x_68) == 0) +x_43 = !lean_is_exclusive(x_9); +if (x_43 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_71 = x_68; -} else { - lean_dec_ref(x_68); - x_71 = lean_box(0); -} -lean_ctor_set(x_17, 0, x_69); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_14); -lean_ctor_set(x_72, 1, x_17); -x_73 = l_Lean_Language_SnapshotTask_pure___rarg(x_72); -if (lean_is_scalar(x_71)) { - x_74 = lean_alloc_ctor(0, 2, 0); -} else { - x_74 = x_71; -} -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_70); -return x_74; +return x_9; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_free_object(x_17); -lean_dec(x_14); -x_75 = lean_ctor_get(x_68, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_68, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_77 = x_68; -} else { - lean_dec_ref(x_68); - x_77 = lean_box(0); -} -if (lean_is_scalar(x_77)) { - x_78 = lean_alloc_ctor(1, 2, 0); -} else { - x_78 = x_77; -} -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_76); -return x_78; -} -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_9, 0); +x_45 = lean_ctor_get(x_9, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_9); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_79 = lean_ctor_get(x_17, 0); -lean_inc(x_79); -lean_dec(x_17); -x_80 = lean_ctor_get(x_16, 1); -lean_inc(x_80); -lean_dec(x_16); -x_81 = lean_ctor_get(x_79, 0); -lean_inc(x_81); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_82 = x_79; -} else { - lean_dec_ref(x_79); - x_82 = lean_box(0); } -x_83 = lean_ctor_get(x_81, 2); -lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -lean_dec(x_83); -x_85 = l_Lean_Language_Lean_isBeforeEditPos(x_84, x_6, x_80); -lean_dec(x_84); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_unbox(x_86); -lean_dec(x_86); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_82); -lean_dec(x_15); -lean_dec(x_14); -x_88 = lean_ctor_get(x_85, 1); -lean_inc(x_88); -lean_dec(x_85); -x_89 = lean_box(0); -x_90 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_89, x_6, x_88); -return x_90; } -else +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; +lean_object* x_5; +x_5 = l_Lean_Language_Lean_process_parseCmd___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_91 = lean_ctor_get(x_85, 1); -lean_inc(x_91); -lean_dec(x_85); -x_92 = lean_ctor_get(x_14, 2); -lean_inc(x_92); -x_93 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7), 5, 3); -lean_closure_set(x_93, 0, x_92); -lean_closure_set(x_93, 1, x_2); -lean_closure_set(x_93, 2, x_15); -x_94 = lean_ctor_get(x_14, 4); -lean_inc(x_94); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = 1; -x_97 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_94, x_93, x_95, x_96, x_91); -if (lean_obj_tag(x_97) == 0) -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_100 = x_97; -} else { - lean_dec_ref(x_97); - x_100 = lean_box(0); -} -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_98); -if (lean_is_scalar(x_82)) { - x_102 = lean_alloc_ctor(0, 2, 0); -} else { - x_102 = x_82; -} -lean_ctor_set(x_102, 0, x_14); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Language_SnapshotTask_pure___rarg(x_102); -if (lean_is_scalar(x_100)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_100; +return x_5; } -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_99); -return x_104; } -else +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_82); -lean_dec(x_14); -x_105 = lean_ctor_get(x_97, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_97, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_107 = x_97; -} else { - lean_dec_ref(x_97); - x_107 = lean_box(0); -} -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); -} else { - x_108 = x_107; -} -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_106); -return x_108; -} +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_10); +lean_dec(x_10); +x_15 = l_Lean_Language_Lean_process_parseCmd___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_7); +return x_15; } } +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Language_Lean_process_parseCmd___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_5); +return x_14; } } -else +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4___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_109; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_4); +lean_object* x_6; +x_6 = l_Lean_Language_Lean_process_parseCmd___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_109 = !lean_is_exclusive(x_16); -if (x_109 == 0) -{ -return x_16; -} -else -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_16, 0); -x_111 = lean_ctor_get(x_16, 1); -lean_inc(x_111); -lean_inc(x_110); -lean_dec(x_16); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Language_Lean_process_parseCmd___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_7); +return x_10; } } +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Language_Lean_process_parseCmd___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_7); +return x_10; } } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___closed__1() { +static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3; -x_3 = l_Lean_Language_Snapshot_Diagnostics_empty; -x_4 = 0; -x_5 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_1); -lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("import", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___closed__2() { +static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Language_instTypeNameSnapshotLeaf; -x_2 = l_Lean_Language_Lean_process_parseCmd___closed__1; -x_3 = l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_instInhabitedDynamicSnapshot___spec__1(x_1, x_2); +x_1 = lean_box(0); +x_2 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_parseCmd___closed__3() { +static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Language_Lean_process_parseCmd___closed__2; -x_2 = l_Lean_Language_SnapshotTask_pure___rarg(x_1); +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +x_2 = l_Array_toPArray_x27___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(lean_object* x_1) { _start: { -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_4, 3); -lean_inc(x_6); -x_7 = lean_st_ref_get(x_6, x_5); -lean_dec(x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = lean_box(0); -lean_inc(x_4); -x_12 = l_Lean_Language_Lean_process_parseCmd___lambda__10(x_2, x_4, x_3, x_1, x_11, x_4, x_10); -lean_dec(x_4); -return x_12; +lean_object* x_2; +x_2 = lean_box(0); +return x_2; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_4); -lean_dec(x_1); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1; -x_15 = lean_st_mk_ref(x_14, x_13); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_box(0); -x_19 = l_Lean_Language_Lean_process_parseCmd___closed__1; -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -x_21 = l_Lean_Language_SnapshotTask_pure___rarg(x_20); -x_22 = lean_box(0); -x_23 = l_Lean_Language_Lean_process_parseCmd___closed__3; -x_24 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_24, 0, x_19); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_2); -lean_ctor_set(x_24, 3, x_23); -lean_ctor_set(x_24, 4, x_21); -lean_ctor_set(x_24, 5, x_17); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_18); -x_26 = l_Lean_Language_SnapshotTask_pure___rarg(x_25); -lean_ctor_set(x_15, 0, x_26); -return x_15; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +x_8 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3; +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +x_11 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_5); +lean_ctor_set(x_1, 1, x_11); +lean_ctor_set(x_1, 0, x_10); +return x_1; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_27 = lean_ctor_get(x_15, 0); -x_28 = lean_ctor_get(x_15, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_15); -x_29 = lean_box(0); -x_30 = l_Lean_Language_Lean_process_parseCmd___closed__1; -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_3); -x_32 = l_Lean_Language_SnapshotTask_pure___rarg(x_31); -x_33 = lean_box(0); -x_34 = l_Lean_Language_Lean_process_parseCmd___closed__3; -x_35 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_2); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_32); -lean_ctor_set(x_35, 5, x_27); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_29); -x_37 = l_Lean_Language_SnapshotTask_pure___rarg(x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_28); -return x_38; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_1); +x_14 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_12); +x_16 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_16, 0, x_15); +x_17 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3; +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_13); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} } } +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +lean_inc(x_3); +x_5 = lean_apply_2(x_1, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +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_apply_3(x_2, x_6, x_3, x_7); +return x_8; +} else { -uint8_t x_39; +uint8_t x_9; lean_dec(x_3); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_15); -if (x_39 == 0) +x_9 = !lean_is_exclusive(x_5); +if (x_9 == 0) { -return x_15; +return x_5; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_15, 0); -x_41 = lean_ctor_get(x_15, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_15); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_5, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_5); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; } } } } -else +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_43; -lean_dec(x_4); +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_43 = !lean_is_exclusive(x_7); -if (x_43 == 0) +x_6 = lean_apply_3(x_1, x_2, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -return x_7; +lean_object* x_8; uint8_t x_9; +x_8 = l_Lean_Language_Lean_process_parseCmd(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +return x_8; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_7, 0); -x_45 = lean_ctor_get(x_7, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_7); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1() { _start: { -lean_object* x_10; -x_10 = l_Lean_Language_Lean_process_parseCmd___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -return x_10; +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2() { _start: { -lean_object* x_5; -x_5 = l_Lean_Language_Lean_process_parseCmd___lambda__4(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Language_Lean_process_doElab___closed__4; +x_3 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set(x_3, 4, x_2); +lean_ctor_set(x_3, 5, x_2); +lean_ctor_set(x_3, 6, x_2); +lean_ctor_set(x_3, 7, x_2); +lean_ctor_set(x_3, 8, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_import", 7, 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4() { _start: { -lean_object* x_8; -x_8 = l_Lean_Language_Lean_process_parseCmd___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5() { _start: { -lean_object* x_8; -x_8 = l_Lean_Language_Lean_process_parseCmd___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("import", 6, 6); +x_1 = lean_mk_string_unchecked("header", 6, 6); return x_1; } } -static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1; +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); +x_1 = lean_unsigned_to_nat(1u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3; -x_2 = l_Array_toPArray_x27___rarg(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("processHeader", 13, 13); +return x_1; } } -LEAN_EXPORT lean_object* l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(lean_object* x_1) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_2; -x_2 = lean_box(0); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; +x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; +x_3 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; +x_4 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9; +x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); +return x_5; } -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2; -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_4); -x_8 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_8, 0, x_7); -x_9 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4; -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -x_11 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_5); -lean_ctor_set(x_1, 1, x_11); -lean_ctor_set(x_1, 0, x_10); -return x_1; } -else +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11() { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_1); -x_14 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2; -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_12); -x_16 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_16, 0, x_15); -x_17 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4; -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_13); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -} +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10; +x_2 = 1; +x_3 = l_Lean_Name_toString(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_5; +lean_object* x_13; +x_13 = l_Lean_Language_Lean_reparseOptions(x_1, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_14); lean_inc(x_3); -x_5 = lean_apply_2(x_1, x_3, x_4); -if (lean_obj_tag(x_5) == 0) +lean_inc(x_2); +x_16 = l_Lean_Elab_Command_mkState(x_2, x_3, x_14); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -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_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_18 = lean_ctor_get(x_16, 7); +lean_dec(x_18); +x_19 = lean_ctor_get(x_16, 6); +lean_dec(x_19); +x_20 = lean_ctor_get(x_16, 4); +lean_dec(x_20); +x_21 = lean_ctor_get(x_16, 3); +lean_dec(x_21); +x_22 = lean_ctor_get(x_16, 1); +lean_dec(x_22); +x_23 = lean_ctor_get(x_16, 0); +lean_dec(x_23); +x_24 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; +x_25 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_14, x_24); +lean_dec(x_14); +x_26 = lean_ctor_get(x_4, 2); +x_27 = lean_box(0); +x_28 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; +x_29 = lean_box(0); +x_30 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; +lean_inc(x_26); +lean_inc(x_2); +x_31 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_31, 0, x_2); +lean_ctor_set(x_31, 1, x_26); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_31, 3, x_27); +lean_ctor_set(x_31, 4, x_29); +lean_ctor_set(x_31, 5, x_27); +lean_ctor_set(x_31, 6, x_30); +x_32 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7; +lean_inc(x_5); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_5); +x_35 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = lean_unsigned_to_nat(1u); +x_37 = l_Lean_Syntax_getArg(x_5, x_36); lean_dec(x_5); -x_8 = lean_apply_3(x_2, x_6, x_3, x_7); -return x_8; +x_38 = l_Lean_Syntax_getArgs(x_37); +lean_dec(x_37); +x_39 = lean_array_to_list(lean_box(0), x_38); +x_40 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_39); +x_41 = l_List_toPArray_x27___rarg(x_40); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_35); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_32); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; +x_45 = lean_array_push(x_44, x_43); +x_46 = l_Array_toPArray_x27___rarg(x_45); +lean_dec(x_45); +x_47 = 1; +x_48 = l_Lean_Language_Lean_process_doElab___closed__4; +x_49 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +lean_ctor_set_uint8(x_49, sizeof(void*)*2, x_47); +x_50 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12; +lean_ctor_set(x_16, 7, x_9); +lean_ctor_set(x_16, 6, x_49); +lean_ctor_set(x_16, 4, x_25); +lean_ctor_set(x_16, 3, x_50); +lean_ctor_set(x_16, 1, x_3); +lean_ctor_set(x_16, 0, x_2); +x_51 = lean_io_promise_new(x_15); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_53 = lean_ctor_get(x_51, 0); +x_54 = lean_ctor_get(x_51, 1); +x_55 = lean_runtime_mark_persistent(x_6); +x_56 = lean_runtime_mark_persistent(x_16); +x_57 = lean_runtime_mark_persistent(x_7); +x_58 = lean_box(0); +x_59 = lean_ctor_get(x_56, 0); +lean_inc(x_59); +lean_inc(x_53); +lean_inc(x_56); +x_60 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__3), 7, 6); +lean_closure_set(x_60, 0, x_58); +lean_closure_set(x_60, 1, x_55); +lean_closure_set(x_60, 2, x_56); +lean_closure_set(x_60, 3, x_59); +lean_closure_set(x_60, 4, x_53); +lean_closure_set(x_60, 5, x_57); +x_61 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_61, 0, x_60); +x_62 = l_Task_Priority_default; +x_63 = lean_io_as_task(x_61, x_62, x_54); +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; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_65 = lean_ctor_get(x_63, 0); +lean_dec(x_65); +x_66 = lean_ctor_get(x_56, 6); +lean_inc(x_66); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = lean_ctor_get(x_67, 2); +lean_inc(x_68); +x_69 = lean_unsigned_to_nat(0u); +x_70 = lean_nat_dec_lt(x_69, x_68); +lean_dec(x_68); +x_71 = lean_io_promise_result(x_53); +lean_ctor_set(x_51, 1, x_71); +lean_ctor_set(x_51, 0, x_58); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_56); +lean_ctor_set(x_72, 1, x_51); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_72); +if (x_70 == 0) +{ +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_dec(x_67); +x_74 = l_Lean_Elab_instInhabitedInfoTree; +x_75 = l_outOfBounds___rarg(x_74); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_78 = 0; +x_79 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_8); +lean_ctor_set(x_79, 2, x_76); +lean_ctor_set_uint8(x_79, sizeof(void*)*3, x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_73); +lean_ctor_set(x_63, 0, x_80); +return x_63; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_81 = l_Lean_Elab_instInhabitedInfoTree; +x_82 = l_Lean_PersistentArray_get_x21___rarg(x_81, x_67, x_69); +x_83 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_83, 0, x_82); +x_84 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_85 = 0; +x_86 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_8); +lean_ctor_set(x_86, 2, x_83); +lean_ctor_set_uint8(x_86, sizeof(void*)*3, x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_73); +lean_ctor_set(x_63, 0, x_87); +return x_63; +} } else { -uint8_t x_9; -lean_dec(x_3); -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_5); -if (x_9 == 0) -{ -return x_5; +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; lean_object* x_95; lean_object* x_96; +x_88 = lean_ctor_get(x_63, 1); +lean_inc(x_88); +lean_dec(x_63); +x_89 = lean_ctor_get(x_56, 6); +lean_inc(x_89); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); +x_91 = lean_ctor_get(x_90, 2); +lean_inc(x_91); +x_92 = lean_unsigned_to_nat(0u); +x_93 = lean_nat_dec_lt(x_92, x_91); +lean_dec(x_91); +x_94 = lean_io_promise_result(x_53); +lean_ctor_set(x_51, 1, x_94); +lean_ctor_set(x_51, 0, x_58); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_56); +lean_ctor_set(x_95, 1, x_51); +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_95); +if (x_93 == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_90); +x_97 = l_Lean_Elab_instInhabitedInfoTree; +x_98 = l_outOfBounds___rarg(x_97); +x_99 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_99, 0, x_98); +x_100 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_101 = 0; +x_102 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_8); +lean_ctor_set(x_102, 2, x_99); +lean_ctor_set_uint8(x_102, sizeof(void*)*3, x_101); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_96); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_88); +return x_104; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_ctor_get(x_5, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_5); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = l_Lean_Elab_instInhabitedInfoTree; +x_106 = l_Lean_PersistentArray_get_x21___rarg(x_105, x_90, x_92); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_109 = 0; +x_110 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_8); +lean_ctor_set(x_110, 2, x_107); +lean_ctor_set_uint8(x_110, sizeof(void*)*3, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_96); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_88); +return x_112; } } } +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_113 = lean_ctor_get(x_51, 0); +x_114 = lean_ctor_get(x_51, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_51); +x_115 = lean_runtime_mark_persistent(x_6); +x_116 = lean_runtime_mark_persistent(x_16); +x_117 = lean_runtime_mark_persistent(x_7); +x_118 = lean_box(0); +x_119 = lean_ctor_get(x_116, 0); +lean_inc(x_119); +lean_inc(x_113); +lean_inc(x_116); +x_120 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__3), 7, 6); +lean_closure_set(x_120, 0, x_118); +lean_closure_set(x_120, 1, x_115); +lean_closure_set(x_120, 2, x_116); +lean_closure_set(x_120, 3, x_119); +lean_closure_set(x_120, 4, x_113); +lean_closure_set(x_120, 5, x_117); +x_121 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_121, 0, x_120); +x_122 = l_Task_Priority_default; +x_123 = lean_io_as_task(x_121, x_122, x_114); +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; +} else { + lean_dec_ref(x_123); + x_125 = lean_box(0); } -LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2(lean_object* x_1, lean_object* x_2) { -_start: +x_126 = lean_ctor_get(x_116, 6); +lean_inc(x_126); +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_ctor_get(x_127, 2); +lean_inc(x_128); +x_129 = lean_unsigned_to_nat(0u); +x_130 = lean_nat_dec_lt(x_129, x_128); +lean_dec(x_128); +x_131 = lean_io_promise_result(x_113); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_118); +lean_ctor_set(x_132, 1, x_131); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_116); +lean_ctor_set(x_133, 1, x_132); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_133); +if (x_130 == 0) { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg), 4, 0); -return x_3; +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_127); +x_135 = l_Lean_Elab_instInhabitedInfoTree; +x_136 = l_outOfBounds___rarg(x_135); +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_136); +x_138 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_139 = 0; +x_140 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_8); +lean_ctor_set(x_140, 2, x_137); +lean_ctor_set_uint8(x_140, sizeof(void*)*3, x_139); +x_141 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_134); +if (lean_is_scalar(x_125)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_125; } +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_124); +return x_142; } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_143 = l_Lean_Elab_instInhabitedInfoTree; +x_144 = l_Lean_PersistentArray_get_x21___rarg(x_143, x_127, x_129); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_144); +x_146 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_147 = 0; +x_148 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_8); +lean_ctor_set(x_148, 2, x_145); +lean_ctor_set_uint8(x_148, sizeof(void*)*3, x_147); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_134); +if (lean_is_scalar(x_125)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_125; } +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_124); +return x_150; } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -x_6 = lean_apply_3(x_1, x_2, x_5, x_4); -return x_6; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1() { -_start: +else { -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; -} +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_151 = lean_ctor_get(x_16, 2); +x_152 = lean_ctor_get(x_16, 5); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_16); +x_153 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; +x_154 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_14, x_153); +lean_dec(x_14); +x_155 = lean_ctor_get(x_4, 2); +x_156 = lean_box(0); +x_157 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; +x_158 = lean_box(0); +x_159 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; +lean_inc(x_155); +lean_inc(x_2); +x_160 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_160, 0, x_2); +lean_ctor_set(x_160, 1, x_155); +lean_ctor_set(x_160, 2, x_157); +lean_ctor_set(x_160, 3, x_156); +lean_ctor_set(x_160, 4, x_158); +lean_ctor_set(x_160, 5, x_156); +lean_ctor_set(x_160, 6, x_159); +x_161 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_161, 0, x_160); +x_162 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7; +lean_inc(x_5); +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_5); +x_164 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_164, 0, x_163); +x_165 = lean_unsigned_to_nat(1u); +x_166 = l_Lean_Syntax_getArg(x_5, x_165); +lean_dec(x_5); +x_167 = l_Lean_Syntax_getArgs(x_166); +lean_dec(x_166); +x_168 = lean_array_to_list(lean_box(0), x_167); +x_169 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_168); +x_170 = l_List_toPArray_x27___rarg(x_169); +x_171 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_171, 0, x_164); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_172, 0, x_161); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; +x_174 = lean_array_push(x_173, x_172); +x_175 = l_Array_toPArray_x27___rarg(x_174); +lean_dec(x_174); +x_176 = 1; +x_177 = l_Lean_Language_Lean_process_doElab___closed__4; +x_178 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_175); +lean_ctor_set_uint8(x_178, sizeof(void*)*2, x_176); +x_179 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12; +x_180 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_180, 0, x_2); +lean_ctor_set(x_180, 1, x_3); +lean_ctor_set(x_180, 2, x_151); +lean_ctor_set(x_180, 3, x_179); +lean_ctor_set(x_180, 4, x_154); +lean_ctor_set(x_180, 5, x_152); +lean_ctor_set(x_180, 6, x_178); +lean_ctor_set(x_180, 7, x_9); +x_181 = lean_io_promise_new(x_15); +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_184 = x_181; +} else { + lean_dec_ref(x_181); + x_184 = lean_box(0); +} +x_185 = lean_runtime_mark_persistent(x_6); +x_186 = lean_runtime_mark_persistent(x_180); +x_187 = lean_runtime_mark_persistent(x_7); +x_188 = lean_box(0); +x_189 = lean_ctor_get(x_186, 0); +lean_inc(x_189); +lean_inc(x_182); +lean_inc(x_186); +x_190 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__3), 7, 6); +lean_closure_set(x_190, 0, x_188); +lean_closure_set(x_190, 1, x_185); +lean_closure_set(x_190, 2, x_186); +lean_closure_set(x_190, 3, x_189); +lean_closure_set(x_190, 4, x_182); +lean_closure_set(x_190, 5, x_187); +x_191 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_191, 0, x_190); +x_192 = l_Task_Priority_default; +x_193 = lean_io_as_task(x_191, x_192, x_183); +x_194 = lean_ctor_get(x_193, 1); +lean_inc(x_194); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_195 = x_193; +} else { + lean_dec_ref(x_193); + x_195 = lean_box(0); } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_maxRecDepth; -return x_1; +x_196 = lean_ctor_get(x_186, 6); +lean_inc(x_196); +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +x_198 = lean_ctor_get(x_197, 2); +lean_inc(x_198); +x_199 = lean_unsigned_to_nat(0u); +x_200 = lean_nat_dec_lt(x_199, x_198); +lean_dec(x_198); +x_201 = lean_io_promise_result(x_182); +if (lean_is_scalar(x_184)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_184; +} +lean_ctor_set(x_202, 0, x_188); +lean_ctor_set(x_202, 1, x_201); +x_203 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_203, 0, x_186); +lean_ctor_set(x_203, 1, x_202); +x_204 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_204, 0, x_203); +if (x_200 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_197); +x_205 = l_Lean_Elab_instInhabitedInfoTree; +x_206 = l_outOfBounds___rarg(x_205); +x_207 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_207, 0, x_206); +x_208 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_209 = 0; +x_210 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_210, 0, x_208); +lean_ctor_set(x_210, 1, x_8); +lean_ctor_set(x_210, 2, x_207); +lean_ctor_set_uint8(x_210, sizeof(void*)*3, x_209); +x_211 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_211, 0, x_210); +lean_ctor_set(x_211, 1, x_204); +if (lean_is_scalar(x_195)) { + x_212 = lean_alloc_ctor(0, 2, 0); +} else { + x_212 = x_195; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_194); +return x_212; +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; uint8_t x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_213 = l_Lean_Elab_instInhabitedInfoTree; +x_214 = l_Lean_PersistentArray_get_x21___rarg(x_213, x_197, x_199); +x_215 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_215, 0, x_214); +x_216 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_217 = 0; +x_218 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_8); +lean_ctor_set(x_218, 2, x_215); +lean_ctor_set_uint8(x_218, sizeof(void*)*3, x_217); +x_219 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_219, 0, x_218); +lean_ctor_set(x_219, 1, x_204); +if (lean_is_scalar(x_195)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_195; } +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_194); +return x_220; } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___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_Language_Lean_process_doElab___lambda__3___closed__3; -x_3 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_1); -lean_ctor_set(x_3, 2, x_1); -lean_ctor_set(x_3, 3, x_2); -lean_ctor_set(x_3, 4, x_2); -lean_ctor_set(x_3, 5, x_2); -lean_ctor_set(x_3, 6, x_2); -lean_ctor_set(x_3, 7, x_2); -lean_ctor_set(x_3, 8, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_import", 7, 7); -return x_1; -} +uint8_t x_221; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_221 = !lean_is_exclusive(x_13); +if (x_221 == 0) +{ +return x_13; } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_13, 0); +x_223 = lean_ctor_get(x_13, 1); +lean_inc(x_223); +lean_inc(x_222); +lean_dec(x_13); +x_224 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_224, 0, x_222); +lean_ctor_set(x_224, 1, x_223); +return x_224; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("header", 6, 6); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8() { +static size_t _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; size_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_usize_of_nat(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; size_t x_4; lean_object* x_5; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1; +x_2 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__2; +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_3); +lean_ctor_set_usize(x_5, 4, x_4); +return x_5; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("processHeader", 13, 13); +x_1 = l_Lean_trace_profiler_output; return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; -x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; -x_3 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__1; -x_4 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10; -x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Import", 6, 6); +return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11; -x_2 = 1; -x_3 = l_Lean_Name_toString(x_1, x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7() { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_environment_set_main_module(x_2, x_12); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_13); -x_14 = l_Lean_Elab_Command_mkState(x_13, x_3, x_4); -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; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_16 = lean_ctor_get(x_14, 6); -lean_dec(x_16); -x_17 = lean_ctor_get(x_14, 4); -lean_dec(x_17); -x_18 = lean_ctor_get(x_14, 3); -lean_dec(x_18); -x_19 = lean_ctor_get(x_14, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_14, 0); -lean_dec(x_20); -x_21 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; -x_22 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_21); -lean_dec(x_4); -x_23 = lean_ctor_get(x_5, 2); -x_24 = lean_box(0); -x_25 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; -x_26 = lean_box(0); -x_27 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; -lean_inc(x_23); -lean_inc(x_13); -x_28 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_28, 0, x_13); -lean_ctor_set(x_28, 1, x_23); -lean_ctor_set(x_28, 2, x_25); -lean_ctor_set(x_28, 3, x_24); -lean_ctor_set(x_28, 4, x_26); -lean_ctor_set(x_28, 5, x_24); -lean_ctor_set(x_28, 6, x_27); -x_29 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8; -lean_inc(x_6); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -x_32 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_unsigned_to_nat(1u); -x_34 = l_Lean_Syntax_getArg(x_6, x_33); -lean_dec(x_6); -x_35 = l_Lean_Syntax_getArgs(x_34); -lean_dec(x_34); -x_36 = lean_array_to_list(lean_box(0), x_35); -x_37 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_36); -x_38 = l_List_toPArray_x27___rarg(x_37); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_32); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_29); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; -x_42 = lean_array_push(x_41, x_40); -x_43 = l_Array_toPArray_x27___rarg(x_42); -lean_dec(x_42); -x_44 = 1; -x_45 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; -lean_inc(x_43); -x_46 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_44); -x_47 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; -lean_ctor_set(x_14, 6, x_46); -lean_ctor_set(x_14, 4, x_22); -lean_ctor_set(x_14, 3, x_47); -lean_ctor_set(x_14, 1, x_3); -lean_ctor_set(x_14, 0, x_13); -x_48 = lean_box(0); -lean_inc(x_14); -x_49 = l_Lean_Language_Lean_process_parseCmd(x_48, x_7, x_14, x_10, x_11); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_43, 2); -lean_inc(x_52); -x_53 = lean_unsigned_to_nat(0u); -x_54 = lean_nat_dec_lt(x_53, x_52); -lean_dec(x_52); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_14); -lean_ctor_set(x_55, 1, x_51); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -if (x_54 == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_43); -x_57 = l_Lean_Elab_instInhabitedInfoTree; -x_58 = l_outOfBounds___rarg(x_57); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_60 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_61 = 0; -x_62 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_8); -lean_ctor_set(x_62, 2, x_59); -lean_ctor_set_uint8(x_62, sizeof(void*)*3, x_61); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_56); -lean_ctor_set(x_49, 0, x_63); -return x_49; -} -else -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; -x_64 = l_Lean_Elab_instInhabitedInfoTree; -x_65 = l_Lean_PersistentArray_get_x21___rarg(x_64, x_43, x_53); -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_68 = 0; -x_69 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_8); -lean_ctor_set(x_69, 2, x_66); -lean_ctor_set_uint8(x_69, sizeof(void*)*3, x_68); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_56); -lean_ctor_set(x_49, 0, x_70); -return x_49; -} -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; -x_71 = lean_ctor_get(x_49, 0); -x_72 = lean_ctor_get(x_49, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_49); -x_73 = lean_ctor_get(x_43, 2); -lean_inc(x_73); -x_74 = lean_unsigned_to_nat(0u); -x_75 = lean_nat_dec_lt(x_74, x_73); -lean_dec(x_73); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_14); -lean_ctor_set(x_76, 1, x_71); -x_77 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_77, 0, x_76); -if (x_75 == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_43); -x_78 = l_Lean_Elab_instInhabitedInfoTree; -x_79 = l_outOfBounds___rarg(x_78); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_79); -x_81 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_82 = 0; -x_83 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_8); -lean_ctor_set(x_83, 2, x_80); -lean_ctor_set_uint8(x_83, sizeof(void*)*3, x_82); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_77); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_72); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_86 = l_Lean_Elab_instInhabitedInfoTree; -x_87 = l_Lean_PersistentArray_get_x21___rarg(x_86, x_43, x_74); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_87); -x_89 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_90 = 0; -x_91 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_8); -lean_ctor_set(x_91, 2, x_88); -lean_ctor_set_uint8(x_91, sizeof(void*)*3, x_90); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_77); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_72); -return x_93; -} -} -} -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; -x_94 = lean_ctor_get(x_14, 2); -x_95 = lean_ctor_get(x_14, 5); -x_96 = lean_ctor_get(x_14, 7); -lean_inc(x_96); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_14); -x_97 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; -x_98 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_97); -lean_dec(x_4); -x_99 = lean_ctor_get(x_5, 2); -x_100 = lean_box(0); -x_101 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; -x_102 = lean_box(0); -x_103 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; -lean_inc(x_99); -lean_inc(x_13); -x_104 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_104, 0, x_13); -lean_ctor_set(x_104, 1, x_99); -lean_ctor_set(x_104, 2, x_101); -lean_ctor_set(x_104, 3, x_100); -lean_ctor_set(x_104, 4, x_102); -lean_ctor_set(x_104, 5, x_100); -lean_ctor_set(x_104, 6, x_103); -x_105 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8; -lean_inc(x_6); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_6); -x_108 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_108, 0, x_107); -x_109 = lean_unsigned_to_nat(1u); -x_110 = l_Lean_Syntax_getArg(x_6, x_109); -lean_dec(x_6); -x_111 = l_Lean_Syntax_getArgs(x_110); -lean_dec(x_110); -x_112 = lean_array_to_list(lean_box(0), x_111); -x_113 = l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1(x_112); -x_114 = l_List_toPArray_x27___rarg(x_113); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_108); -lean_ctor_set(x_115, 1, x_114); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_105); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; -x_118 = lean_array_push(x_117, x_116); -x_119 = l_Array_toPArray_x27___rarg(x_118); -lean_dec(x_118); -x_120 = 1; -x_121 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__3; -lean_inc(x_119); -x_122 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_119); -lean_ctor_set_uint8(x_122, sizeof(void*)*2, x_120); -x_123 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; -x_124 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_124, 0, x_13); -lean_ctor_set(x_124, 1, x_3); -lean_ctor_set(x_124, 2, x_94); -lean_ctor_set(x_124, 3, x_123); -lean_ctor_set(x_124, 4, x_98); -lean_ctor_set(x_124, 5, x_95); -lean_ctor_set(x_124, 6, x_122); -lean_ctor_set(x_124, 7, x_96); -x_125 = lean_box(0); -lean_inc(x_124); -x_126 = l_Lean_Language_Lean_process_parseCmd(x_125, x_7, x_124, x_10, x_11); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_126, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_126)) { - lean_ctor_release(x_126, 0); - lean_ctor_release(x_126, 1); - x_129 = x_126; -} else { - lean_dec_ref(x_126); - x_129 = lean_box(0); -} -x_130 = lean_ctor_get(x_119, 2); -lean_inc(x_130); -x_131 = lean_unsigned_to_nat(0u); -x_132 = lean_nat_dec_lt(x_131, x_130); -lean_dec(x_130); -x_133 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_133, 0, x_124); -lean_ctor_set(x_133, 1, x_127); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_133); -if (x_132 == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_119); -x_135 = l_Lean_Elab_instInhabitedInfoTree; -x_136 = l_outOfBounds___rarg(x_135); -x_137 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_137, 0, x_136); -x_138 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_139 = 0; -x_140 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_8); -lean_ctor_set(x_140, 2, x_137); -lean_ctor_set_uint8(x_140, sizeof(void*)*3, x_139); -x_141 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_134); -if (lean_is_scalar(x_129)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_129; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("importing", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; } -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_128); -return x_142; +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, double x_10, double x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +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_17 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4; +x_18 = l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1(x_3, x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3; +x_20 = lean_box(0); +x_21 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_3, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_19, x_20, x_13, x_14); +return x_21; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_143 = l_Lean_Elab_instInhabitedInfoTree; -x_144 = l_Lean_PersistentArray_get_x21___rarg(x_143, x_119, x_131); -x_145 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_145, 0, x_144); -x_146 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_147 = 0; -x_148 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_8); -lean_ctor_set(x_148, 2, x_145); -lean_ctor_set_uint8(x_148, sizeof(void*)*3, x_147); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_134); -if (lean_is_scalar(x_129)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_129; +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_18); +x_22 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6; +x_23 = 1; +x_24 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4; +x_25 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set_float(x_25, sizeof(void*)*2, x_10); +lean_ctor_set_float(x_25, sizeof(void*)*2 + 8, x_11); +lean_ctor_set_uint8(x_25, sizeof(void*)*2 + 16, x_23); +x_26 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9; +x_27 = l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9; +x_28 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_26); +lean_ctor_set(x_28, 2, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; +x_32 = lean_array_push(x_31, x_30); +x_33 = l_Array_toPArray_x27___rarg(x_32); +lean_dec(x_32); +x_34 = lean_box(0); +x_35 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_3, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_33, x_34, x_13, x_14); +return x_35; } -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_128); -return x_150; } } +static double _init_l_Lean_Language_Lean_process_processHeader___lambda__6___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; uint32_t x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -x_8 = lean_ctor_get_uint32(x_4, sizeof(void*)*2); -x_9 = l_Lean_MessageLog_empty; -x_10 = 1; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; double x_13; double x_14; double x_15; lean_object* x_16; uint32_t x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_8 = lean_io_mono_nanos_now(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = 0; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Float_ofScientific(x_9, x_11, x_12); +lean_dec(x_9); +x_14 = l_Lean_Language_Lean_process_processHeader___lambda__6___closed__1; +x_15 = lean_float_div(x_13, x_14); +x_16 = lean_ctor_get(x_5, 1); +lean_inc(x_16); +x_17 = lean_ctor_get_uint32(x_5, sizeof(void*)*2); +x_18 = l_Lean_MessageLog_empty; +x_19 = 1; lean_inc(x_2); -lean_inc(x_7); -x_11 = l_Lean_Elab_processHeader(x_1, x_7, x_9, x_2, x_8, x_10, x_6); -if (lean_obj_tag(x_11) == 0) +lean_inc(x_16); +x_20 = l_Lean_Elab_processHeader(x_1, x_16, x_18, x_2, x_17, x_19, x_10); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_io_mono_nanos_now(x_22); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_12, 0); -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -x_17 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_16, x_13); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_27; lean_object* x_28; double x_29; double x_30; lean_object* x_31; uint8_t x_32; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +x_29 = l_Float_ofScientific(x_27, x_11, x_12); +lean_dec(x_27); +x_30 = lean_float_div(x_29, x_14); +lean_inc(x_24); +x_31 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_24, x_28); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +x_35 = l_Lean_MessageLog_hasErrors(x_24); +if (x_35 == 0) { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -x_21 = l_Lean_MessageLog_hasErrors(x_16); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -lean_free_object(x_17); -lean_free_object(x_12); -x_22 = lean_box(0); -x_23 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_4, x_15, x_16, x_7, x_2, x_1, x_3, x_19, x_22, x_5, x_20); +lean_object* x_36; lean_object* x_37; +lean_free_object(x_31); +lean_free_object(x_25); +x_36 = lean_box(0); +x_37 = l_Lean_Language_Lean_process_processHeader___lambda__5(x_5, x_23, x_16, x_24, x_2, x_1, x_3, x_4, x_33, x_15, x_30, x_36, x_6, x_34); lean_dec(x_2); -return x_23; +return x_37; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(0); -x_25 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_26 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_19); -lean_ctor_set(x_26, 2, x_24); -lean_ctor_set_uint8(x_26, sizeof(void*)*3, x_10); -lean_ctor_set(x_12, 1, x_24); -lean_ctor_set(x_12, 0, x_26); -lean_ctor_set(x_17, 0, x_12); -return x_17; +x_38 = lean_box(0); +x_39 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_40 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_33); +lean_ctor_set(x_40, 2, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*3, x_19); +lean_ctor_set(x_25, 1, x_38); +lean_ctor_set(x_25, 0, x_40); +lean_ctor_set(x_31, 0, x_25); +return x_31; } } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_17, 0); -x_28 = lean_ctor_get(x_17, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_17); -x_29 = l_Lean_MessageLog_hasErrors(x_16); -if (x_29 == 0) +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_ctor_get(x_31, 0); +x_42 = lean_ctor_get(x_31, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_31); +x_43 = l_Lean_MessageLog_hasErrors(x_24); +if (x_43 == 0) { -lean_object* x_30; lean_object* x_31; -lean_free_object(x_12); -x_30 = lean_box(0); -x_31 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_4, x_15, x_16, x_7, x_2, x_1, x_3, x_27, x_30, x_5, x_28); +lean_object* x_44; lean_object* x_45; +lean_free_object(x_25); +x_44 = lean_box(0); +x_45 = l_Lean_Language_Lean_process_processHeader___lambda__5(x_5, x_23, x_16, x_24, x_2, x_1, x_3, x_4, x_41, x_15, x_30, x_44, x_6, x_42); lean_dec(x_2); -return x_31; +return x_45; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -x_33 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_34 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_27); -lean_ctor_set(x_34, 2, x_32); -lean_ctor_set_uint8(x_34, sizeof(void*)*3, x_10); -lean_ctor_set(x_12, 1, x_32); -lean_ctor_set(x_12, 0, x_34); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_12); -lean_ctor_set(x_35, 1, x_28); -return x_35; +x_46 = lean_box(0); +x_47 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_48 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_46); +lean_ctor_set_uint8(x_48, sizeof(void*)*3, x_19); +lean_ctor_set(x_25, 1, x_46); +lean_ctor_set(x_25, 0, x_48); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_25); +lean_ctor_set(x_49, 1, x_42); +return x_49; } } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_36 = lean_ctor_get(x_12, 0); -x_37 = lean_ctor_get(x_12, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_12); -lean_inc(x_37); -x_38 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_37, x_13); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_41 = x_38; +lean_object* x_50; lean_object* x_51; double x_52; double x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_50 = lean_ctor_get(x_25, 0); +x_51 = lean_ctor_get(x_25, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_25); +x_52 = l_Float_ofScientific(x_50, x_11, x_12); +lean_dec(x_50); +x_53 = lean_float_div(x_52, x_14); +lean_inc(x_24); +x_54 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_24, x_51); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_57 = x_54; } else { - lean_dec_ref(x_38); - x_41 = lean_box(0); + lean_dec_ref(x_54); + x_57 = lean_box(0); } -x_42 = l_Lean_MessageLog_hasErrors(x_37); -if (x_42 == 0) +x_58 = l_Lean_MessageLog_hasErrors(x_24); +if (x_58 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_41); -x_43 = lean_box(0); -x_44 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_4, x_36, x_37, x_7, x_2, x_1, x_3, x_39, x_43, x_5, x_40); +lean_object* x_59; lean_object* x_60; +lean_dec(x_57); +x_59 = lean_box(0); +x_60 = l_Lean_Language_Lean_process_processHeader___lambda__5(x_5, x_23, x_16, x_24, x_2, x_1, x_3, x_4, x_55, x_15, x_53, x_59, x_6, x_56); lean_dec(x_2); -return x_44; +return x_60; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -lean_dec(x_37); -lean_dec(x_36); -lean_dec(x_7); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_45 = lean_box(0); -x_46 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; -x_47 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_39); -lean_ctor_set(x_47, 2, x_45); -lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_10); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_45); -if (lean_is_scalar(x_41)) { - x_49 = lean_alloc_ctor(0, 2, 0); +x_61 = lean_box(0); +x_62 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_63 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_55); +lean_ctor_set(x_63, 2, x_61); +lean_ctor_set_uint8(x_63, sizeof(void*)*3, x_19); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +if (lean_is_scalar(x_57)) { + x_65 = lean_alloc_ctor(0, 2, 0); } else { - x_49 = x_41; + x_65 = x_57; } -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_40); -return x_49; +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_56); +return x_65; } } } else { -uint8_t x_50; -lean_dec(x_7); +uint8_t x_66; +lean_dec(x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_50 = !lean_is_exclusive(x_11); -if (x_50 == 0) +x_66 = !lean_is_exclusive(x_20); +if (x_66 == 0) { -return x_11; +return x_20; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_11, 0); -x_52 = lean_ctor_get(x_11, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_11); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -return x_53; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_20, 0); +x_68 = lean_ctor_get(x_20, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_20); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -if (lean_obj_tag(x_4) == 0) +if (lean_obj_tag(x_5) == 0) { -lean_object* x_7; lean_object* x_8; -lean_dec(x_5); +lean_object* x_8; lean_object* x_9; +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_7 = lean_ctor_get(x_4, 0); -lean_inc(x_7); -lean_dec(x_4); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_4, 0); -lean_inc(x_9); -lean_dec(x_4); -x_10 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_1, x_2, x_3, x_9, x_5, x_6); -return x_10; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +lean_dec(x_5); +x_11 = l_Lean_Language_Lean_process_processHeader___lambda__6(x_1, x_2, x_3, x_4, x_10, x_6, x_7); +return x_11; } } } @@ -5142,10 +9062,12 @@ lean_inc(x_2); x_12 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__2), 4, 2); lean_closure_set(x_12, 0, x_1); lean_closure_set(x_12, 1, x_2); -x_13 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__5), 6, 3); +lean_inc(x_4); +x_13 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__7___boxed), 7, 4); lean_closure_set(x_13, 0, x_2); lean_closure_set(x_13, 1, x_6); lean_closure_set(x_13, 2, x_3); +lean_closure_set(x_13, 3, x_4); x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg), 4, 2); lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_13); @@ -5158,14 +9080,48 @@ x_17 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_11, x_16, x_5); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_9); +lean_object* x_13; +x_13 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +double x_15; double x_16; lean_object* x_17; +x_15 = lean_unbox_float(x_10); +lean_dec(x_10); +x_16 = lean_unbox_float(x_11); +lean_dec(x_11); +x_17 = l_Lean_Language_Lean_process_processHeader___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15, x_16, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_5); -return x_12; +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Language_Lean_process_processHeader___lambda__6(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_Lean_Language_Lean_process_processHeader___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Language_Lean_process_processHeader___lambda__7(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_Lean_Language_Lean_process_parseHeader___lambda__1(lean_object* x_1, lean_object* x_2) { @@ -5205,7 +9161,7 @@ static lean_object* _init_l_Lean_Language_Lean_process_parseHeader___lambda__3__ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__6; x_2 = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_11____closed__7; -x_3 = l_Lean_Language_Lean_process_doElab___lambda__2___closed__1; +x_3 = l_Lean_Language_Lean_process_doElab___lambda__3___closed__1; x_4 = l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__1; x_5 = l_Lean_Name_mkStr5(x_1, x_2, x_1, x_3, x_4); return x_5; @@ -5385,82 +9341,142 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_5); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -lean_inc(x_8); -x_9 = l_Lean_Language_Lean_process_parseCmd(x_7, x_2, x_8, x_3, x_6); -if (lean_obj_tag(x_9) == 0) +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = lean_io_promise_new(x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_5); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_inc(x_8); +lean_inc(x_11); +x_13 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__3), 7, 6); +lean_closure_set(x_13, 0, x_10); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_11); +lean_closure_set(x_13, 3, x_12); +lean_closure_set(x_13, 4, x_8); +lean_closure_set(x_13, 5, x_3); +x_14 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = l_Task_Priority_default; +x_16 = lean_io_as_task(x_14, x_15, x_9); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +x_19 = lean_box(0); +x_20 = lean_io_promise_result(x_8); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_11); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Language_SnapshotTask_pure___rarg(x_24); +lean_ctor_set(x_16, 0, x_25); +return x_16; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_dec(x_16); +x_27 = lean_box(0); +x_28 = lean_io_promise_result(x_8); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_11); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Language_SnapshotTask_pure___rarg(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_26); +return x_34; +} +} +else { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +uint8_t x_35; +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_4); +x_35 = !lean_is_exclusive(x_16); +if (x_35 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_Language_SnapshotTask_pure___rarg(x_14); -lean_ctor_set(x_9, 0, x_15); -return x_9; +return x_16; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_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(0, 2, 0); -lean_ctor_set(x_18, 0, x_8); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_4); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Language_SnapshotTask_pure___rarg(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_17); -return x_22; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_16, 0); +x_37 = lean_ctor_get(x_16, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_16); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} } } else { -uint8_t x_23; -lean_dec(x_8); +uint8_t x_39; +lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_9); -if (x_23 == 0) +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_7); +if (x_39 == 0) { -return x_9; +return x_7; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_9, 0); -x_25 = lean_ctor_get(x_9, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_9); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_7, 0); +x_41 = lean_ctor_get(x_7, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_7); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } @@ -7545,66 +11561,256 @@ LEAN_EXPORT lean_object* l_Lean_Language_Lean_processCommands(lean_object* x_1, _start: { lean_object* x_6; -x_6 = lean_box(0); +x_6 = lean_io_promise_new(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; +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_3, 0); +lean_inc(x_9); +x_10 = lean_box(0); if (lean_obj_tag(x_4) == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 5, 3); -lean_closure_set(x_7, 0, x_6); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_3); -x_8 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_7, x_6, x_6, x_1, x_5); -return x_8; -} -else +lean_object* x_11; lean_object* x_12; +lean_inc(x_7); +x_11 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 7, 5); +lean_closure_set(x_11, 0, x_10); +lean_closure_set(x_11, 1, x_2); +lean_closure_set(x_11, 2, x_3); +lean_closure_set(x_11, 3, x_9); +lean_closure_set(x_11, 4, x_7); +x_12 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_11, x_10, x_10, x_1, x_8); +if (lean_obj_tag(x_12) == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_4); -if (x_9 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_10 = lean_ctor_get(x_4, 0); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_ctor_set(x_4, 0, x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 5, 3); -lean_closure_set(x_12, 0, x_4); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_3); -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -x_15 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_12, x_14, x_6, x_1, x_5); -return x_15; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +x_15 = lean_io_promise_result(x_7); +lean_ctor_set(x_12, 0, x_15); +return x_12; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_4, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 1); lean_inc(x_16); -lean_dec(x_4); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -x_18 = lean_alloc_ctor(1, 1, 0); +lean_dec(x_12); +x_17 = lean_io_promise_result(x_7); +x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); -x_19 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 5, 3); -lean_closure_set(x_19, 0, x_18); -lean_closure_set(x_19, 1, x_2); -lean_closure_set(x_19, 2, x_3); -x_20 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_7); +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) +{ +return x_12; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_19, x_21, x_6, x_1, x_5); +lean_dec(x_12); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); return x_22; } } } +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_4); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_24 = lean_ctor_get(x_4, 0); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_ctor_set(x_4, 0, x_25); +lean_inc(x_7); +x_26 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 7, 5); +lean_closure_set(x_26, 0, x_4); +lean_closure_set(x_26, 1, x_2); +lean_closure_set(x_26, 2, x_3); +lean_closure_set(x_26, 3, x_9); +lean_closure_set(x_26, 4, x_7); +x_27 = lean_ctor_get(x_24, 0); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_26, x_28, x_10, x_1, x_8); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +x_32 = lean_io_promise_result(x_7); +lean_ctor_set(x_29, 0, x_32); +return x_29; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_io_promise_result(x_7); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +else +{ +uint8_t x_36; +lean_dec(x_7); +x_36 = !lean_is_exclusive(x_29); +if (x_36 == 0) +{ +return x_29; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_29, 0); +x_38 = lean_ctor_get(x_29, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_29); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_40 = lean_ctor_get(x_4, 0); +lean_inc(x_40); +lean_dec(x_4); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +lean_inc(x_7); +x_43 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd), 7, 5); +lean_closure_set(x_43, 0, x_42); +lean_closure_set(x_43, 1, x_2); +lean_closure_set(x_43, 2, x_3); +lean_closure_set(x_43, 3, x_9); +lean_closure_set(x_43, 4, x_7); +x_44 = lean_ctor_get(x_40, 0); +lean_inc(x_44); +lean_dec(x_40); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_44); +x_46 = l_Lean_Language_Lean_LeanProcessingM_run___rarg(x_43, x_45, x_10, x_1, x_8); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_48 = x_46; +} else { + lean_dec_ref(x_46); + x_48 = lean_box(0); +} +x_49 = lean_io_promise_result(x_7); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_48; +} +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_47); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_7); +x_51 = lean_ctor_get(x_46, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_46, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_53 = x_46; +} else { + lean_dec_ref(x_46); + x_53 = lean_box(0); +} +if (lean_is_scalar(x_53)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_53; +} +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +return x_54; +} +} +} +} +else +{ +uint8_t x_55; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_6); +if (x_55 == 0) +{ +return x_6; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_6, 0); +x_57 = lean_ctor_get(x_6, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_6); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f_goCmd(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalCmdState_x3f_goCmd(lean_object* x_1) { _start: { lean_object* x_2; @@ -7612,7 +11818,7 @@ x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); if (lean_obj_tag(x_2) == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); @@ -7623,27 +11829,24 @@ x_5 = l_Lean_Language_SnapshotTask_get___rarg(x_4); x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } else { -lean_object* x_9; lean_object* x_10; +lean_object* x_8; lean_object* x_9; lean_dec(x_1); -x_9 = lean_ctor_get(x_2, 0); -lean_inc(x_9); +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); lean_dec(x_2); -x_10 = l_Lean_Language_SnapshotTask_get___rarg(x_9); -x_1 = x_10; +x_9 = l_Lean_Language_SnapshotTask_get___rarg(x_8); +x_1 = x_9; goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalCmdState_x3f(lean_object* x_1) { _start: { lean_object* x_2; @@ -7685,7 +11888,7 @@ x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); lean_dec(x_9); x_11 = l_Lean_Language_SnapshotTask_get___rarg(x_10); -x_12 = l_Lean_Language_Lean_waitForFinalEnv_x3f_goCmd(x_11); +x_12 = l_Lean_Language_Lean_waitForFinalCmdState_x3f_goCmd(x_11); return x_12; } } @@ -7758,6 +11961,53 @@ lean_mark_persistent(l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHe l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__12 = _init_l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__12(); lean_mark_persistent(l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__12); l_Lean_Language_Lean_SetupImportsResult_trustLevel___default = _init_l_Lean_Language_Lean_SetupImportsResult_trustLevel___default(); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__1); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__2); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__3); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__4); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__5); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__6); +l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7 = _init_l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7(); +lean_mark_persistent(l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444____closed__7); +if (builtin) {res = l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_444_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_Language_Lean_internal_minimalSnapshots = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_Language_Lean_internal_minimalSnapshots); +lean_dec_ref(res); +}l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__1); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__2); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__4); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__5); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__6); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__8); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__9); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__10); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__11); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__12); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__1); +l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Language_Lean_reparseOptions___spec__1___closed__2); l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1 = _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1(); lean_mark_persistent(l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1); l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2 = _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2(); @@ -7768,18 +12018,6 @@ l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___c lean_mark_persistent(l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4); l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5 = _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5(); lean_mark_persistent(l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__5); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__1 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__1); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__2 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__2); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__3 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__3); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__4 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__4); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__5 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__5); -l_Lean_Language_Lean_process_doElab___lambda__2___closed__6 = _init_l_Lean_Language_Lean_process_doElab___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__2___closed__6); l_Lean_Language_Lean_process_doElab___lambda__3___closed__1 = _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__3___closed__1); l_Lean_Language_Lean_process_doElab___lambda__3___closed__2 = _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__2(); @@ -7788,18 +12026,56 @@ l_Lean_Language_Lean_process_doElab___lambda__3___closed__3 = _init_l_Lean_Langu lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__3___closed__3); l_Lean_Language_Lean_process_doElab___lambda__3___closed__4 = _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__4(); lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__3___closed__4); -l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__1); -l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__2); -l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__1___closed__3); +l_Lean_Language_Lean_process_doElab___lambda__3___closed__5 = _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__3___closed__5); +l_Lean_Language_Lean_process_doElab___lambda__3___closed__6 = _init_l_Lean_Language_Lean_process_doElab___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___lambda__3___closed__6); +l_Lean_Language_Lean_process_doElab___closed__1 = _init_l_Lean_Language_Lean_process_doElab___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___closed__1); +l_Lean_Language_Lean_process_doElab___closed__2 = _init_l_Lean_Language_Lean_process_doElab___closed__2(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___closed__2); +l_Lean_Language_Lean_process_doElab___closed__3 = _init_l_Lean_Language_Lean_process_doElab___closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___closed__3); +l_Lean_Language_Lean_process_doElab___closed__4 = _init_l_Lean_Language_Lean_process_doElab___closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_process_doElab___closed__4); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__1); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__2); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__3); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__4); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__5); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__6 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__6); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__7); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__8); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__9); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__10 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__10); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__11); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__12); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__13 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__13(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__13); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__14); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__15); +l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__2___closed__16); l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__3___closed__1); -l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__1); -l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2(); -lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__9___closed__2); +l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__5___closed__1); +l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___lambda__7___closed__1); l_Lean_Language_Lean_process_parseCmd___closed__1 = _init_l_Lean_Language_Lean_process_parseCmd___closed__1(); lean_mark_persistent(l_Lean_Language_Lean_process_parseCmd___closed__1); l_Lean_Language_Lean_process_parseCmd___closed__2 = _init_l_Lean_Language_Lean_process_parseCmd___closed__2(); @@ -7812,32 +12088,46 @@ l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2 = lean_mark_persistent(l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2); l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3 = _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3(); lean_mark_persistent(l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3); -l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4 = _init_l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4(); -lean_mark_persistent(l_List_map___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11); -l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10); +l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__2 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__2(); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__3); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__4); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__5); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__6); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__7); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__8); +l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__9); +l_Lean_Language_Lean_process_processHeader___lambda__6___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__6___closed__1(); l_Lean_Language_Lean_process_processHeader___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___closed__1(); lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___closed__1); l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__1 = _init_l_Lean_Language_Lean_process_parseHeader___lambda__3___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index c58954d02470..bbcb9806b5d9 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -179,7 +179,6 @@ static lean_object* l_Lean_Meta_getLocalDeclFromUserName___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Diagnostics_unfoldCounter___default; -lean_object* l_Lean_Level_succ___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_FVarId_hasForwardDeps___spec__46___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isProp___default; @@ -193,7 +192,6 @@ lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLambdaFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -lean_object* l_Lean_mkLevelMax_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_FVarId_hasForwardDeps___spec__30___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_transparency___default; LEAN_EXPORT lean_object* l_Lean_Meta_withAssignableSyntheticOpaque___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -410,7 +408,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_withLocalDeclsD__ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_FVarId_hasForwardDeps___spec__36___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getParamNames___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Basic_0__Lean_Meta_withExistingLocalDeclsImp___spec__1(lean_object*, lean_object*); -size_t lean_ptr_addr(lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux(lean_object*); @@ -570,6 +567,7 @@ LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Basic_0_ LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_instantiate_level_mvars(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducibleAndInstances(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -652,11 +650,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_modifyDiag___boxed(lean_object*, lean_objec LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp_process(lean_object*); static lean_object* l_Lean_Meta_instAlternativeMetaM___closed__2; LEAN_EXPORT lean_object* l_Lean_MVarId_setType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Level_hasMVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_FVarId_hasForwardDeps___spec__12___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2351_(lean_object*); lean_object* l_Lean_LocalInstances_erase(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withExistingLocalDeclsImp(lean_object*); @@ -725,7 +721,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_State_diag___default; LEAN_EXPORT lean_object* l_Lean_Meta_getResetPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_simpLevelIMax_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_State_mctx___default; LEAN_EXPORT lean_object* l_Lean_Meta_resetZetaDeltaFVarIds___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16829____closed__7; @@ -851,7 +846,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_withLocalDecls LEAN_EXPORT lean_object* l_Lean_Meta_withReducibleAndInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedCache; -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_FVarId_hasForwardDeps___spec__20(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -989,7 +983,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_orelseMergeErr 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*); -lean_object* l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedDefEqCache___closed__1; lean_object* l_Lean_Core_withRestoreOrSaveFull___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; @@ -1016,7 +1009,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1386_(l LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_FVarId_hasForwardDeps___spec__20___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_FVarId_hasForwardDeps___spec__57___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_checkedAssign___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_processPostponed_loop___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1049,7 +1041,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_Meta_isLevelDefEq static lean_object* l_Lean_Meta_instOrElseMetaM___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_mkLevelIMax_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux___at_Lean_Meta_withNewLocalInstances___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_withErasedFVars___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1176,7 +1167,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getLocalDeclFromUserName___boxed(lean_objec LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_FVarId_hasForwardDeps___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyPostponed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_liftMkBindingM___rarg___closed__2; -lean_object* l_Lean_simpLevelMax_x27(lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_level_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducible___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -38044,652 +38034,97 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_fullApproxDefEq___rarg), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___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_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = lean_st_ref_take(x_4, x_7); -x_9 = lean_ctor_get(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; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_7 = lean_st_ref_get(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 0); +lean_dec(x_7); +x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); lean_dec(x_8); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_9, 0); -lean_dec(x_13); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +x_11 = lean_instantiate_level_mvars(x_10, x_1); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_take(x_3, x_9); +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_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_10, 6); -x_16 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_15, x_1, x_2); -lean_ctor_set(x_10, 6, x_16); -x_17 = lean_st_ref_set(x_4, x_9, x_11); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +lean_ctor_set(x_15, 0, x_12); +x_19 = lean_st_ref_set(x_3, x_15, x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_17, 0, x_20); -return x_17; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_13); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_17, 1); -lean_inc(x_21); -lean_dec(x_17); -x_22 = lean_box(0); +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 0, x_13); +lean_ctor_set(x_23, 1, x_22); return x_23; } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_24 = lean_ctor_get(x_10, 0); -x_25 = lean_ctor_get(x_10, 1); -x_26 = lean_ctor_get(x_10, 2); -x_27 = lean_ctor_get(x_10, 3); -x_28 = lean_ctor_get(x_10, 4); -x_29 = lean_ctor_get(x_10, 5); -x_30 = lean_ctor_get(x_10, 6); -x_31 = lean_ctor_get(x_10, 7); -x_32 = lean_ctor_get(x_10, 8); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_24 = lean_ctor_get(x_15, 1); +x_25 = lean_ctor_get(x_15, 2); +x_26 = lean_ctor_get(x_15, 3); +x_27 = lean_ctor_get(x_15, 4); lean_inc(x_27); lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); -lean_dec(x_10); -x_33 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_30, x_1, x_2); -x_34 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_25); -lean_ctor_set(x_34, 2, x_26); -lean_ctor_set(x_34, 3, x_27); -lean_ctor_set(x_34, 4, x_28); -lean_ctor_set(x_34, 5, x_29); -lean_ctor_set(x_34, 6, x_33); -lean_ctor_set(x_34, 7, x_31); -lean_ctor_set(x_34, 8, x_32); -lean_ctor_set(x_9, 0, x_34); -x_35 = lean_st_ref_set(x_4, x_9, x_11); -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -if (lean_is_exclusive(x_35)) { - lean_ctor_release(x_35, 0); - lean_ctor_release(x_35, 1); - x_37 = x_35; -} else { - lean_dec_ref(x_35); - x_37 = lean_box(0); -} -x_38 = lean_box(0); -if (lean_is_scalar(x_37)) { - x_39 = lean_alloc_ctor(0, 2, 0); -} else { - x_39 = x_37; -} -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -return x_39; -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_40 = lean_ctor_get(x_9, 1); -x_41 = lean_ctor_get(x_9, 2); -x_42 = lean_ctor_get(x_9, 3); -x_43 = lean_ctor_get(x_9, 4); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_9); -x_44 = lean_ctor_get(x_10, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_10, 1); -lean_inc(x_45); -x_46 = lean_ctor_get(x_10, 2); -lean_inc(x_46); -x_47 = lean_ctor_get(x_10, 3); -lean_inc(x_47); -x_48 = lean_ctor_get(x_10, 4); -lean_inc(x_48); -x_49 = lean_ctor_get(x_10, 5); -lean_inc(x_49); -x_50 = lean_ctor_get(x_10, 6); -lean_inc(x_50); -x_51 = lean_ctor_get(x_10, 7); -lean_inc(x_51); -x_52 = lean_ctor_get(x_10, 8); -lean_inc(x_52); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - lean_ctor_release(x_10, 2); - lean_ctor_release(x_10, 3); - lean_ctor_release(x_10, 4); - lean_ctor_release(x_10, 5); - lean_ctor_release(x_10, 6); - lean_ctor_release(x_10, 7); - lean_ctor_release(x_10, 8); - x_53 = x_10; -} else { - lean_dec_ref(x_10); - x_53 = lean_box(0); -} -x_54 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_50, x_1, x_2); -if (lean_is_scalar(x_53)) { - x_55 = lean_alloc_ctor(0, 9, 0); -} else { - x_55 = x_53; -} -lean_ctor_set(x_55, 0, x_44); -lean_ctor_set(x_55, 1, x_45); -lean_ctor_set(x_55, 2, x_46); -lean_ctor_set(x_55, 3, x_47); -lean_ctor_set(x_55, 4, x_48); -lean_ctor_set(x_55, 5, x_49); -lean_ctor_set(x_55, 6, x_54); -lean_ctor_set(x_55, 7, x_51); -lean_ctor_set(x_55, 8, x_52); -x_56 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_40); -lean_ctor_set(x_56, 2, x_41); -lean_ctor_set(x_56, 3, x_42); -lean_ctor_set(x_56, 4, x_43); -x_57 = lean_st_ref_set(x_4, x_56, x_11); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_59 = x_57; -} else { - lean_dec_ref(x_57); - x_59 = lean_box(0); -} -x_60 = lean_box(0); -if (lean_is_scalar(x_59)) { - x_61 = lean_alloc_ctor(0, 2, 0); -} else { - x_61 = x_59; -} -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_58); -return x_61; -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___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: -{ -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -lean_inc(x_7); -x_8 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; size_t x_11; size_t x_12; uint8_t x_13; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ptr_addr(x_7); -lean_dec(x_7); -x_12 = lean_ptr_addr(x_10); -x_13 = lean_usize_dec_eq(x_11, x_12); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_1); -x_14 = l_Lean_Level_succ___override(x_10); -lean_ctor_set(x_8, 0, x_14); -return x_8; -} -else -{ -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_1); -return x_8; -} -} -else -{ -lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; uint8_t x_19; -x_15 = lean_ctor_get(x_8, 0); -x_16 = lean_ctor_get(x_8, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_8); -x_17 = lean_ptr_addr(x_7); -lean_dec(x_7); -x_18 = lean_ptr_addr(x_15); -x_19 = lean_usize_dec_eq(x_17, x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_1); -x_20 = l_Lean_Level_succ___override(x_15); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_16); -return x_21; -} -else -{ -lean_object* x_22; lean_dec(x_15); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_1); -lean_ctor_set(x_22, 1, x_16); -return x_22; -} -} -} -case 2: -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_23 = lean_ctor_get(x_1, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -lean_inc(x_23); -x_25 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_23, x_2, x_3, x_4, x_5, x_6); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -lean_inc(x_24); -x_28 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_24, x_2, x_3, x_4, x_5, x_27); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; size_t x_31; size_t x_32; uint8_t x_33; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ptr_addr(x_23); -lean_dec(x_23); -x_32 = lean_ptr_addr(x_26); -x_33 = lean_usize_dec_eq(x_31, x_32); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_24); -lean_dec(x_1); -x_34 = l_Lean_mkLevelMax_x27(x_26, x_30); -lean_ctor_set(x_28, 0, x_34); -return x_28; -} -else -{ -size_t x_35; size_t x_36; uint8_t x_37; -x_35 = lean_ptr_addr(x_24); -lean_dec(x_24); -x_36 = lean_ptr_addr(x_30); -x_37 = lean_usize_dec_eq(x_35, x_36); -if (x_37 == 0) -{ -lean_object* x_38; -lean_dec(x_1); -x_38 = l_Lean_mkLevelMax_x27(x_26, x_30); -lean_ctor_set(x_28, 0, x_38); -return x_28; -} -else -{ -lean_object* x_39; -x_39 = l_Lean_simpLevelMax_x27(x_26, x_30, x_1); -lean_dec(x_1); -lean_dec(x_30); -lean_dec(x_26); -lean_ctor_set(x_28, 0, x_39); -return x_28; -} -} -} -else -{ -lean_object* x_40; lean_object* x_41; size_t x_42; size_t x_43; uint8_t x_44; -x_40 = lean_ctor_get(x_28, 0); -x_41 = lean_ctor_get(x_28, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_28); -x_42 = lean_ptr_addr(x_23); -lean_dec(x_23); -x_43 = lean_ptr_addr(x_26); -x_44 = lean_usize_dec_eq(x_42, x_43); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; -lean_dec(x_24); -lean_dec(x_1); -x_45 = l_Lean_mkLevelMax_x27(x_26, x_40); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_41); -return x_46; -} -else -{ -size_t x_47; size_t x_48; uint8_t x_49; -x_47 = lean_ptr_addr(x_24); -lean_dec(x_24); -x_48 = lean_ptr_addr(x_40); -x_49 = lean_usize_dec_eq(x_47, x_48); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_1); -x_50 = l_Lean_mkLevelMax_x27(x_26, x_40); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_41); -return x_51; -} -else -{ -lean_object* x_52; lean_object* x_53; -x_52 = l_Lean_simpLevelMax_x27(x_26, x_40, x_1); -lean_dec(x_1); -lean_dec(x_40); -lean_dec(x_26); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_41); -return x_53; -} -} -} -} -case 3: -{ -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; -x_54 = lean_ctor_get(x_1, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_1, 1); -lean_inc(x_55); -lean_inc(x_54); -x_56 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_54, x_2, x_3, x_4, x_5, x_6); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -lean_inc(x_55); -x_59 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_55, x_2, x_3, x_4, x_5, x_58); -x_60 = !lean_is_exclusive(x_59); -if (x_60 == 0) -{ -lean_object* x_61; size_t x_62; size_t x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_59, 0); -x_62 = lean_ptr_addr(x_54); -lean_dec(x_54); -x_63 = lean_ptr_addr(x_57); -x_64 = lean_usize_dec_eq(x_62, x_63); -if (x_64 == 0) -{ -lean_object* x_65; -lean_dec(x_55); -lean_dec(x_1); -x_65 = l_Lean_mkLevelIMax_x27(x_57, x_61); -lean_ctor_set(x_59, 0, x_65); -return x_59; -} -else -{ -size_t x_66; size_t x_67; uint8_t x_68; -x_66 = lean_ptr_addr(x_55); -lean_dec(x_55); -x_67 = lean_ptr_addr(x_61); -x_68 = lean_usize_dec_eq(x_66, x_67); -if (x_68 == 0) -{ -lean_object* x_69; -lean_dec(x_1); -x_69 = l_Lean_mkLevelIMax_x27(x_57, x_61); -lean_ctor_set(x_59, 0, x_69); -return x_59; -} -else -{ -lean_object* x_70; -x_70 = l_Lean_simpLevelIMax_x27(x_57, x_61, x_1); -lean_dec(x_1); -lean_ctor_set(x_59, 0, x_70); -return x_59; -} -} -} -else -{ -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; uint8_t x_75; -x_71 = lean_ctor_get(x_59, 0); -x_72 = lean_ctor_get(x_59, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_59); -x_73 = lean_ptr_addr(x_54); -lean_dec(x_54); -x_74 = lean_ptr_addr(x_57); -x_75 = lean_usize_dec_eq(x_73, x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; -lean_dec(x_55); -lean_dec(x_1); -x_76 = l_Lean_mkLevelIMax_x27(x_57, x_71); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_72); -return x_77; -} -else -{ -size_t x_78; size_t x_79; uint8_t x_80; -x_78 = lean_ptr_addr(x_55); -lean_dec(x_55); -x_79 = lean_ptr_addr(x_71); -x_80 = lean_usize_dec_eq(x_78, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_1); -x_81 = l_Lean_mkLevelIMax_x27(x_57, x_71); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_72); -return x_82; -} -else -{ -lean_object* x_83; lean_object* x_84; -x_83 = l_Lean_simpLevelIMax_x27(x_57, x_71, x_1); -lean_dec(x_1); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_72); -return x_84; -} -} -} -} -case 5: -{ -lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_85 = lean_ctor_get(x_1, 0); -lean_inc(x_85); -x_86 = lean_st_ref_get(x_3, x_6); -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_88 = lean_ctor_get(x_86, 0); -x_89 = lean_ctor_get(x_86, 1); -x_90 = lean_ctor_get(x_88, 0); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_ctor_get(x_90, 6); -lean_inc(x_91); -lean_dec(x_90); -x_92 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_91, x_85); -if (lean_obj_tag(x_92) == 0) -{ -lean_dec(x_85); -lean_ctor_set(x_86, 0, x_1); -return x_86; -} -else -{ -lean_object* x_93; uint8_t x_94; -lean_dec(x_1); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -lean_dec(x_92); -x_94 = l_Lean_Level_hasMVar(x_93); -if (x_94 == 0) -{ -lean_dec(x_85); -lean_ctor_set(x_86, 0, x_93); -return x_86; -} -else -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -lean_free_object(x_86); -x_95 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_93, x_2, x_3, x_4, x_5, x_89); -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); -lean_inc(x_96); -x_98 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_85, x_96, x_2, x_3, x_4, x_5, x_97); -x_99 = !lean_is_exclusive(x_98); -if (x_99 == 0) -{ -lean_object* x_100; -x_100 = lean_ctor_get(x_98, 0); -lean_dec(x_100); -lean_ctor_set(x_98, 0, x_96); -return x_98; -} -else -{ -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_98, 1); -lean_inc(x_101); -lean_dec(x_98); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_96); -lean_ctor_set(x_102, 1, x_101); -return x_102; -} -} -} -} -else -{ -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_103 = lean_ctor_get(x_86, 0); -x_104 = lean_ctor_get(x_86, 1); -lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_86); -x_105 = lean_ctor_get(x_103, 0); -lean_inc(x_105); -lean_dec(x_103); -x_106 = lean_ctor_get(x_105, 6); -lean_inc(x_106); -lean_dec(x_105); -x_107 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_106, x_85); -if (lean_obj_tag(x_107) == 0) -{ -lean_object* x_108; -lean_dec(x_85); -x_108 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_108, 0, x_1); -lean_ctor_set(x_108, 1, x_104); -return x_108; -} -else -{ -lean_object* x_109; uint8_t x_110; -lean_dec(x_1); -x_109 = lean_ctor_get(x_107, 0); -lean_inc(x_109); -lean_dec(x_107); -x_110 = l_Lean_Level_hasMVar(x_109); -if (x_110 == 0) -{ -lean_object* x_111; -lean_dec(x_85); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_104); -return x_111; -} -else -{ -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_112 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_109, x_2, x_3, x_4, x_5, x_104); -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -lean_dec(x_112); -lean_inc(x_113); -x_115 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_85, x_113, x_2, x_3, x_4, x_5, x_114); -x_116 = lean_ctor_get(x_115, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_117 = x_115; +x_28 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_28, 0, x_12); +lean_ctor_set(x_28, 1, x_24); +lean_ctor_set(x_28, 2, x_25); +lean_ctor_set(x_28, 3, x_26); +lean_ctor_set(x_28, 4, x_27); +x_29 = lean_st_ref_set(x_3, x_28, x_16); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_31 = x_29; } else { - lean_dec_ref(x_115); - x_117 = lean_box(0); + lean_dec_ref(x_29); + x_31 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_31)) { + x_32 = lean_alloc_ctor(0, 2, 0); } else { - x_118 = x_117; -} -lean_ctor_set(x_118, 0, x_113); -lean_ctor_set(x_118, 1, x_116); -return x_118; -} -} -} -} -default: -{ -lean_object* x_119; -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_1); -lean_ctor_set(x_119, 1, x_6); -return x_119; + x_32 = x_31; } +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_30); +return x_32; } } } @@ -38725,18 +38160,6 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_8; -} -} LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___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: { diff --git a/stage0/stdlib/Lean/Meta/LevelDefEq.c b/stage0/stdlib/Lean/Meta/LevelDefEq.c index a6eb81fda4b5..53b9a39916d6 100644 --- a/stage0/stdlib/Lean/Meta/LevelDefEq.c +++ b/stage0/stdlib/Lean/Meta/LevelDefEq.c @@ -42,6 +42,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_LevelDefEq___hyg_1929___ LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_isLevelDefEqAuxImpl___spec__4(lean_object*); lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg(lean_object*); static lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_postponeIsLevelDefEq___closed__7; +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_isLevelDefEqAuxImpl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxMaxMax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -110,12 +111,12 @@ uint8_t l_Lean_Level_isParam(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxSelfMax_solve___boxed(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_LevelDefEq___hyg_1929____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxSelfMax_solve(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_LevelDefEq___hyg_1929____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxSelfMax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_mkMaxArgsDiff___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxMaxMax_solve___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_isMax(lean_object*); static lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_postponeIsLevelDefEq___closed__2; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_isLevelDefEqAuxImpl___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -125,6 +126,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_mkMaxArgs LEAN_EXPORT lean_object* l_Lean_Meta_isLevelDefEqAuxImpl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_postponeIsLevelDefEq___closed__1; uint8_t l_Lean_Level_isMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_isMVarWithGreaterDepth___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -308,6 +310,200 @@ x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_st_ref_take(x_4, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_10, 6); +x_16 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_15, x_1, x_2); +lean_ctor_set(x_10, 6, x_16); +x_17 = lean_st_ref_set(x_4, x_9, x_11); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_24 = lean_ctor_get(x_10, 0); +x_25 = lean_ctor_get(x_10, 1); +x_26 = lean_ctor_get(x_10, 2); +x_27 = lean_ctor_get(x_10, 3); +x_28 = lean_ctor_get(x_10, 4); +x_29 = lean_ctor_get(x_10, 5); +x_30 = lean_ctor_get(x_10, 6); +x_31 = lean_ctor_get(x_10, 7); +x_32 = lean_ctor_get(x_10, 8); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_10); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_30, x_1, x_2); +x_34 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_34, 0, x_24); +lean_ctor_set(x_34, 1, x_25); +lean_ctor_set(x_34, 2, x_26); +lean_ctor_set(x_34, 3, x_27); +lean_ctor_set(x_34, 4, x_28); +lean_ctor_set(x_34, 5, x_29); +lean_ctor_set(x_34, 6, x_33); +lean_ctor_set(x_34, 7, x_31); +lean_ctor_set(x_34, 8, x_32); +lean_ctor_set(x_9, 0, x_34); +x_35 = lean_st_ref_set(x_4, x_9, x_11); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_37 = x_35; +} else { + lean_dec_ref(x_35); + x_37 = lean_box(0); +} +x_38 = lean_box(0); +if (lean_is_scalar(x_37)) { + x_39 = lean_alloc_ctor(0, 2, 0); +} else { + x_39 = x_37; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_40 = lean_ctor_get(x_9, 1); +x_41 = lean_ctor_get(x_9, 2); +x_42 = lean_ctor_get(x_9, 3); +x_43 = lean_ctor_get(x_9, 4); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_9); +x_44 = lean_ctor_get(x_10, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_10, 1); +lean_inc(x_45); +x_46 = lean_ctor_get(x_10, 2); +lean_inc(x_46); +x_47 = lean_ctor_get(x_10, 3); +lean_inc(x_47); +x_48 = lean_ctor_get(x_10, 4); +lean_inc(x_48); +x_49 = lean_ctor_get(x_10, 5); +lean_inc(x_49); +x_50 = lean_ctor_get(x_10, 6); +lean_inc(x_50); +x_51 = lean_ctor_get(x_10, 7); +lean_inc(x_51); +x_52 = lean_ctor_get(x_10, 8); +lean_inc(x_52); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + lean_ctor_release(x_10, 2); + lean_ctor_release(x_10, 3); + lean_ctor_release(x_10, 4); + lean_ctor_release(x_10, 5); + lean_ctor_release(x_10, 6); + lean_ctor_release(x_10, 7); + lean_ctor_release(x_10, 8); + x_53 = x_10; +} else { + lean_dec_ref(x_10); + x_53 = lean_box(0); +} +x_54 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_50, x_1, x_2); +if (lean_is_scalar(x_53)) { + x_55 = lean_alloc_ctor(0, 9, 0); +} else { + x_55 = x_53; +} +lean_ctor_set(x_55, 0, x_44); +lean_ctor_set(x_55, 1, x_45); +lean_ctor_set(x_55, 2, x_46); +lean_ctor_set(x_55, 3, x_47); +lean_ctor_set(x_55, 4, x_48); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_54); +lean_ctor_set(x_55, 7, x_51); +lean_ctor_set(x_55, 8, x_52); +x_56 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_40); +lean_ctor_set(x_56, 2, x_41); +lean_ctor_set(x_56, 3, x_42); +lean_ctor_set(x_56, 4, x_43); +x_57 = lean_st_ref_set(x_4, x_56, x_11); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_59 = x_57; +} else { + lean_dec_ref(x_57); + x_59 = lean_box(0); +} +x_60 = lean_box(0); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 2, 0); +} else { + x_61 = x_59; +} +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; +} +} +} static lean_object* _init_l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___closed__1() { _start: { @@ -387,7 +583,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_mkMaxArgsDiff(x_1, x_2, x_12); -x_15 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_1, x_14, x_3, x_4, x_5, x_6, x_13); +x_15 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_1, x_14, x_3, x_4, x_5, x_6, x_13); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -396,6 +592,18 @@ return x_15; } } } +LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_LevelDefEq_0__Lean_Meta_tryApproxSelfMax_solve(lean_object* x_1, lean_object* x_2, 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: { @@ -416,7 +624,7 @@ return x_12; else { lean_object* x_13; uint8_t x_14; -x_13 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_3, x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_3, x_1, x_4, x_5, x_6, x_7, x_8); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { @@ -603,7 +811,7 @@ return x_14; else { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_4, x_1, x_5, x_6, x_7, x_8, x_9); +x_15 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_4, x_1, x_5, x_6, x_7, x_8, x_9); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -634,7 +842,7 @@ else { lean_object* x_24; uint8_t x_25; lean_dec(x_1); -x_24 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_4, x_2, x_5, x_6, x_7, x_8, x_9); +x_24 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_4, x_2, x_5, x_6, x_7, x_8, x_9); x_25 = !lean_is_exclusive(x_24); if (x_25 == 0) { @@ -2020,7 +2228,7 @@ lean_object* x_176; lean_object* x_177; uint8_t x_178; lean_free_object(x_169); x_176 = l_Lean_Level_mvarId_x21(x_1); lean_dec(x_1); -x_177 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_176, x_2, x_3, x_4, x_5, x_6, x_173); +x_177 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_176, x_2, x_3, x_4, x_5, x_6, x_173); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -2159,7 +2367,7 @@ if (x_207 == 0) lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; uint8_t x_212; lean_object* x_213; lean_object* x_214; x_208 = l_Lean_Level_mvarId_x21(x_1); lean_dec(x_1); -x_209 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_208, x_2, x_3, x_4, x_5, x_6, x_206); +x_209 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_208, x_2, x_3, x_4, x_5, x_6, x_206); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -2292,7 +2500,7 @@ lean_inc(x_234); lean_dec(x_169); x_235 = l_Lean_Level_mvarId_x21(x_2); lean_dec(x_2); -x_236 = l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(x_235, x_1, x_3, x_4, x_5, x_6, x_234); +x_236 = l_Lean_assignLevelMVar___at___private_Lean_Meta_LevelDefEq_0__Lean_Meta_solveSelfMax___spec__2(x_235, x_1, x_3, x_4, x_5, x_6, x_234); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); diff --git a/stage0/stdlib/Lean/MetavarContext.c b/stage0/stdlib/Lean/MetavarContext.c index e50fcd63732a..748d249b526e 100644 --- a/stage0/stdlib/Lean/MetavarContext.c +++ b/stage0/stdlib/Lean/MetavarContext.c @@ -23,8 +23,9 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_abstractRange___boxed(lean_object LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_Context_mvarIdsToAbstract___default; LEAN_EXPORT lean_object* l_Lean_instantiateMVarDeclMVars___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___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_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__34(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_MetavarKind_isNatural(uint8_t); @@ -47,6 +48,7 @@ LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssignable___rarg___lambda__1(lean_ob LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_getDecl___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__26(lean_object*, lean_object*, size_t, size_t); @@ -54,8 +56,6 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_findDecl_x3f(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_MetavarContext_getDecl___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__33___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__23___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -68,11 +68,11 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_exprDependsOn___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__7(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__4(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__35___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_instantiateLCtxMVars___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__14(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); @@ -97,12 +97,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_Persisten LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__15(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__13___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__40(lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__34___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_DependsOn_instMonadMCtxM___closed__3; @@ -116,6 +114,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_MVarId_modifyLCtx___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__28___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateExprMVars___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); @@ -130,7 +129,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_instantiateMV static lean_object* l_Lean_MetavarContext_MkBinding_instToStringException___closed__3; LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLocalDeclMVars___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__13___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedLocalInstance___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__1(lean_object*, lean_object*, lean_object*); @@ -140,14 +138,13 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_getExprAssignmentDomain(lean_obje LEAN_EXPORT lean_object* l_Lean_getDelayedMVarRoot(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn_x27___spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MetavarContext_addLevelMVarDecl___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__36(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__10(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__21___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -157,11 +154,11 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBi LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__56___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAdapterExprStructEqExprM; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__59(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_getLevelMVarAssignment_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_MetavarContext_getExprAssignmentDomain___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -170,7 +167,6 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__29___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_isLevelMVarAssignable___spec__2(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAdapterExprStructEqExprM___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -180,17 +176,14 @@ LEAN_EXPORT lean_object* l_Lean_instMonadMCtxOfMonadLift(lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__101(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_MetavarContext_getExprAssignmentCore_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isDelayedAssigned___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__112(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_MetavarContext_setFVarKind___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignedMVar___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_elem___at___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_mkLambda(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__39(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__14(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_instantiateLCtxMVars___spec__3___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -226,12 +219,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore__ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_getLevelMVarAssignment_x3f___spec__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15___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_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__13___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(lean_object*, lean_object*); lean_object* l_Lean_Level_succ___override(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__92(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__29(lean_object*, lean_object*, lean_object*); @@ -243,9 +234,10 @@ LEAN_EXPORT lean_object* l_Lean_hasAssignedMVar___rarg___lambda__2(lean_object*, LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__13(lean_object*, lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14___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_MetavarContext_MkBinding_reduceLocalContext___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__6(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15___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_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__5___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__35(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__31___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -278,18 +270,19 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_co LEAN_EXPORT lean_object* l_Lean_MetavarContext_mkLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssignedOrDelayedAssigned___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_DependsOn_instMonadMCtxM___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__44___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_revert___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_exprDependsOn_x27___spec__15___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___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_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_toString___at_Lean_MetavarContext_MkBinding_instToStringException___spec__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__41___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__9___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__2(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__38___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__37___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_usize_dec_eq(size_t, size_t); @@ -297,22 +290,18 @@ LEAN_EXPORT uint8_t l_Lean_MetavarContext_MkBinding_Context_binderInfoForMVars__ extern uint8_t l_instInhabitedBool; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elim___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__32___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__20___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getInScope(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_revert___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__17___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_assignDelayedMVar___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__5___boxed(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_localDeclDependsOn_x27___spec__17(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__20(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_addLevelMVarDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__37___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__114___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg___boxed(lean_object*, lean_object*); @@ -339,18 +328,15 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_mkBinding___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__68(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findExprDependsOn(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_exprDependsOn___spec__13___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_exprDependsOn_x27___spec__5___boxed(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__61___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_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__27___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -373,12 +359,12 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_Mk lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__65___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__98___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_anyM___at_Lean_hasAssignedMVar___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_MetavarContext_findUserName_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_getExprAssignmentCore_x3f___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___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_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__30(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__33(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__94(lean_object*, lean_object*); @@ -387,7 +373,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_MVarId_isAssigned___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findExprDependsOn___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLocalDeclMVars___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__35___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVarsCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -399,44 +385,41 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___rarg(lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_MetavarDecl_numScopeArgs___default; LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___lambda__1(uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__13___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__97___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__7(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_LocalInstances_erase___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__40(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_instantiateExprMVars___spec__4___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_foldRevM_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_instMonadMCtxStateRefT_x27ST(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__17___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setMCtx___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__38(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__30___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_isSyntheticOpaque___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_instantiateLCtxMVars___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__27(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_hasAssignedLevelMVar___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__20___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__16(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__7(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instToStringException___closed__6; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__42(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_instantiateLCtxMVars___spec__11(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__9(lean_object*); @@ -447,12 +430,12 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_MVarId_isDelayedAssigned___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__38___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedLocalInstance; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__57___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__12___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_modify___at_Lean_instMonadMCtxStateMMetavarContext___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exprDependsOn_x27___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__16(lean_object*, lean_object*, size_t, size_t); @@ -462,7 +445,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_instantiateLCtx LEAN_EXPORT lean_object* l_Lean_exprDependsOn(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_elimMVarDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__16(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___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_Lean_instantiateLCtxMVars___spec__6___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__14(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_MVarId_isAssigned___spec__2(lean_object*, size_t, lean_object*); @@ -478,6 +460,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__53___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__39___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12(uint8_t, 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_instantiateLCtxMVars___spec__10___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_dAssignment___default; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__29___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -492,10 +475,10 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___sp static lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__37___boxed(lean_object*, lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__33___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__61(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_getLevelMVarAssignment_x3f___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_assignDelayedMVar___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__35___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); @@ -505,13 +488,11 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_modifyDecl___rarg(lean_object*, lean_obje static lean_object* l_Lean_isLevelMVarAssignable___rarg___lambda__1___closed__4; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__45(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_MkBinding_instToStringException___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___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_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_instantiateLCtxMVars___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_DependsOn_State_visited___default___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__41(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__31(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__20(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_instantiateMVarDeclMVars___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -520,7 +501,6 @@ static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__13___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__7(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_LocalDecl_index(lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__39(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__45(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_MetavarContext_LevelMVarToParam_main_visitApp___spec__2(lean_object*, lean_object*, lean_object*); @@ -538,8 +518,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_LevelMV static lean_object* l_Lean_MetavarContext_getDecl___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__16___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__11(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_MetavarContext_getExprAssignmentCore_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -559,7 +541,6 @@ size_t lean_ptr_addr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getInScope___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn_x27___spec__22(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_getFVar_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32___rarg___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_assignDelayedMVar(lean_object*); @@ -579,23 +560,28 @@ LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_instantiateLCtxMVa LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isDelayedAssigned___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__63(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_setFVarBinderInfo___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__30(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__77(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__75(lean_object*, lean_object*, lean_object*); lean_object* l_runST___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_elimMVarDeps(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__60___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_MetavarContext_collectForwardDeps(lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_LevelMVarToParam_main_visitApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_abstract(lean_object*, lean_object*); @@ -605,12 +591,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__49(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__10___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_exprDependsOn___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___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_Subarray_forInUnsafe_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getLocalDeclWithSmallestIdx___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_mkAppRangeAux(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_MetavarContext_getExprAssignmentDomain___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssigned(lean_object*); @@ -627,12 +611,14 @@ LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__15(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_instantiateLCtxMVars___spec__3___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__56(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_DependsOn_instMonadMCtxM___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_main(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__99(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__11___closed__2; @@ -652,9 +638,7 @@ uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__15(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_anyM___at_Lean_hasAssignedMVar___spec__1___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_instBEqLocalInstance(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_mvarCounter___default; LEAN_EXPORT lean_object* l_Lean_assignDelayedMVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -696,7 +680,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore__ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_instantiateMVarDeclMVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_MetavarContext_findUserName_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__21(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__20___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignableLevelMVar___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); @@ -711,13 +694,12 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_List_foldl___at_Lean_MetavarContext_MkBinding_instToStringException___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__26___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__102___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___rarg___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__105(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__47(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_LocalContext_setBinderInfo(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -728,13 +710,13 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_M LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__31(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___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_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__13(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__24(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__18___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn_x27___spec__4___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__25(lean_object*); @@ -742,10 +724,8 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___s LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__53(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__7(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__50(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__64___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__29___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateExprMVars___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -755,7 +735,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVar static lean_object* l_Lean_instMonadMCtxStateMMetavarContext___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__24(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__93___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_getDecl___closed__2; @@ -766,8 +745,8 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVar LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__35(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__4___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__34(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__20___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__19___closed__3; @@ -778,16 +757,18 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_instantiateLCtxMVars___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_modify___at_Lean_MetavarContext_instMonadMCtxStateRefT_x27ST___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__90(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Lean_MetavarKind_isSyntheticOpaque(uint8_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__51(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_panic___at_Lean_isLevelMVarAssignable___spec__4___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___boxed(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_instantiateExprMVars___spec__42___rarg___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_instantiate_level_mvars(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_getLevelMVarAssignment_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_MetavarContext_LevelMVarToParam_visitLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4(lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_State_cache___default___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -800,10 +781,14 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___sp LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__1; LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_MetavarContext_instMonadMCtxStateRefT_x27ST___spec__1___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_LocalInstances_erase___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__29___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__97(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType_abstractRangeAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_getDecl___closed__3; @@ -811,7 +796,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__18; lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__13___closed__1; static lean_object* l_Lean_MetavarContext_MkBinding_instToStringException___closed__2; @@ -821,13 +805,12 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDepen static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__4; LEAN_EXPORT lean_object* l_Lean_setMCtx___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__63___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__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_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__6(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__33(lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32(lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); @@ -835,17 +818,17 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__103(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit___spec__5(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__36___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__87(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVarDeclMVars___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__13(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_isLevelMVarAssigned___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_modifyExprMVarLCtx(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getLevelMVarAssignment_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_MVarId_isAssigned___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -861,22 +844,19 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn___s LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__10___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_instantiateLCtxMVars___spec__3(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn_x27___spec__38(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__32(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_addExprMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn_x27___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__31___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_instantiateLCtxMVars___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instMonadMCtxM___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__63___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_MetavarContext_MkBinding_instMonadMCtxM; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__18___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarDecl_userName___default; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___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_PersistentArray_foldlM___at_Lean_instantiateLCtxMVars___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__41___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -888,7 +868,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__1___boxed( lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType___spec__1___lambda__1___boxed(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_MetavarContext_MkBinding_collectForwardDeps___spec__100(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_hasAssignableLevelMVar___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -904,7 +883,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_isLevelMVarAssig LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__31___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__51___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_hasMVar(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__78___boxed(lean_object*, lean_object*); @@ -921,25 +900,27 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_setFVarBinderInfo___rarg(lean_object*, le LEAN_EXPORT lean_object* l_List_toString___at_Lean_MetavarContext_MkBinding_instToStringException___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_getExprAssignmentCore_x3f___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___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_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__19___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__37(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__12(lean_object*); uint64_t l_Lean_Expr_hash(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOnPred(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___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_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_isLevelMVarAssignable___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__36___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__58___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__5___boxed(lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_getExprAssignmentDomain___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_instantiateExprMVars___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__93(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_setFVarKind___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__81___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__7___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -948,8 +929,6 @@ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_levelMVarToParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadMCtxM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -960,12 +939,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_MetavarCon LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__35(lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__12___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_MetavarContext_getExprAssignmentCore_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__100___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__14___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -987,11 +964,11 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_mkParamName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__107(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__15(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_setFVarKind___rarg(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_MkBinding_instToStringException___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__76___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__26___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__10___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__43(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_exprDependsOn___spec__5(lean_object*, lean_object*); @@ -1001,9 +978,7 @@ uint8_t lean_expr_equal(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__109___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__25(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_collectForwardDeps(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLCtxMVars(lean_object*); @@ -1016,7 +991,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateLocalDeclMVars___rarg___lambda__3___b LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__21___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__32(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__12(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__106(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__18(lean_object*); @@ -1042,15 +1016,16 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_MkBi LEAN_EXPORT lean_object* l_Lean_DependsOn_instMonadMCtxM___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn_x27___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___boxed(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_instantiateExprMVars___spec__42___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_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__16(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__89(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_assign_mvar(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_lower_loose_bvars(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_MkBinding_instToStringException___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_MVarId_assign___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__92___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__64(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); @@ -1066,17 +1041,17 @@ LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_instantiateMVarsCore___spec_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__15(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__80___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_getExprAssignmentCore_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__110(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_instantiateExprMVars___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_abstractRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getLevelMVarAssignment_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignedLevelMVar___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___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_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__54(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__6(lean_object*, lean_object*, size_t, size_t); @@ -1103,16 +1078,13 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_co LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn___spec__10___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__11___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__15___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__49___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__8(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_MetavarContext_setMVarUserName___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_modifyExprMVarLCtx___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__37(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); @@ -1131,7 +1103,6 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_MetavarContext_LevelMV LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_State_cache___default; lean_object* l_Array_append___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_setFVarKind___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__38(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn_x27___spec__22___boxed(lean_object*, lean_object*); @@ -1147,7 +1118,6 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_isDelayedAssigned___rarg___lambda__1___bo LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__40___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elim___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13(lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__20; @@ -1160,6 +1130,7 @@ static size_t l_Lean_PersistentHashMap_findAux___at_Lean_getLevelMVarAssignment_ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__113___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_MetavarContext_instMonadMCtxStateRefT_x27ST___spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__31(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -1170,13 +1141,13 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM__ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__77___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssignable___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__18___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__24___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1192,10 +1163,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__23___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__116(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_findDecl_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123_(uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findExprDependsOn___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAdapterExprStructEqExprM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__103___boxed(lean_object*, lean_object*); @@ -1215,15 +1185,16 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarConte LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_preserveOrder___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__19(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_DependsOn_State_visited___default___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLocalDeclMVars___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_lAssignment___default; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_assignLevelMVar___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__70___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_setFVarBinderInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); uint8_t l_Lean_LocalContext_isSubPrefixOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__70(lean_object*, lean_object*); @@ -1245,23 +1216,22 @@ LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__14(lean_ob LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__72___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__40(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn_x27___spec__8(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__95(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findExprDependsOn___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__53___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn___spec__11___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__73___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_findUserName_x3f___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__9___closed__1; LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_setFVarBinderInfo(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_MetavarContext_setMVarUserName___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__16(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_setKind(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType_abstractRangeAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1275,7 +1245,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_co LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn_x27___spec__7___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_assignLevelMVar___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37(lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__91(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp___spec__1(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*); @@ -1286,7 +1255,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* lean_mk_metavar_ctx(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedMetavarContext___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_findLevelDepth_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_getExprAssignmentDomain___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_mkBinding(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*); @@ -1300,6 +1268,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarConte LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn_x27___spec__28(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_assignDelayedMVar___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__94___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedMetavarDecl___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__35___rarg___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_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__12___boxed(lean_object*, lean_object*, lean_object*); @@ -1316,7 +1285,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__3___boxed( LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateExprMVars___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__88___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_instantiateExprMVars___spec__4(lean_object*, lean_object*); static lean_object* l_Nat_foldRevM_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__10___boxed(lean_object*, lean_object*, lean_object*); @@ -1343,8 +1311,6 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn_x27 static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__9___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_mkBinding___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__7___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findExprDependsOn___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_MVarId_assign___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedMetavarContext; @@ -1377,11 +1343,9 @@ LEAN_EXPORT lean_object* l_Lean_instMonadMCtxOfMonadLift___rarg___lambda__1(lean LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__113(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__30___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3; static lean_object* l_Lean_instInhabitedMetavarDecl___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__39___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__47(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_instantiateMVarDeclMVars___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__37___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn___spec__37___boxed(lean_object*, lean_object*); @@ -1405,7 +1369,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT uint8_t l_Lean_instInhabitedMetavarKind; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__44___boxed(lean_object*, lean_object*); uint8_t l_ptrEqList___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn_x27___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__30___rarg___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_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn___spec__3(lean_object*, lean_object*, lean_object*); @@ -1423,11 +1386,8 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__23(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__15; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_dependsOnPred(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__115(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1442,7 +1402,8 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_main_visitApp(le lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_MetavarContext_addLevelMVarDecl___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___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_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_dependsOn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2(lean_object*, lean_object*, lean_object*); @@ -1452,10 +1413,8 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_MVarId_ass LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__33___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findLocalDeclDependsOn(lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instToStringException___closed__5; -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getLocalDeclWithSmallestIdx___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_setMVarType(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__2; @@ -1465,8 +1424,9 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAd static lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVarDeclMVars(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__6___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__42(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__40___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__23(lean_object*); @@ -1488,7 +1448,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_MetavarCon LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__24(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__23(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_isLevelMVarAssignable___rarg___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__33(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__48(lean_object*, lean_object*, size_t, size_t); @@ -1506,7 +1465,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___sp static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__16; static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__3; LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_MetavarContext_LevelMVarToParam_main_visitApp___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_MetavarContext_addExprMVarDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -1516,6 +1474,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBindi LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_exprDependsOn_x27___spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_isLevelMVarAssignable___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedMetavarDecl___closed__4; +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__10; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_MetavarContext_getDecl___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__18___boxed(lean_object*, lean_object*, lean_object*); @@ -1541,18 +1500,17 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__28___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getLocalDeclWithSmallestIdx___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isLevelMVarAssignable___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAdapterExprStructEqExprM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -lean_object* l_panic___at_Lean_Level_normalize___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__87___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__19___closed__2; @@ -1580,8 +1538,10 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_modifyLCtx___rarg___lambda__1(lean_object size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_instantiateLCtxMVars___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_MetavarContext_findLevelDepth_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27___spec__58(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLocalDeclMVars___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasAssignedLevelMVar___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); uint8_t l_Lean_Expr_hasFVar(lean_object*); @@ -1599,7 +1559,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_ LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_getLocalDeclWithSmallestIdx(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_isNatural___boxed(lean_object*); size_t lean_array_size(lean_object*); @@ -1631,6 +1590,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__71___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_MetavarContext_addExprMVarDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_MVarId_assign___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__67(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__99___boxed(lean_object*, lean_object*, lean_object*); @@ -1640,13 +1600,11 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_mkForall___boxed(lean_object*, le LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__34___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_MetavarContext_addLevelMVarDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__26(lean_object*); LEAN_EXPORT uint8_t l_Lean_HashSetImp_contains___at___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__22___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__19___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__33___boxed(lean_object*, lean_object*, lean_object*); @@ -1654,7 +1612,6 @@ static lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___closed__1; LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_setMVarKind(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn_x27___spec__62___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_exprDependsOn___spec__14(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1675,15 +1632,14 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn_x27 LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_abstractRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__51___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__52___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__48(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssigned___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssigned___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_x27___spec__55___boxed(lean_object*, lean_object*); @@ -1692,6 +1648,7 @@ lean_object* l_Lean_Expr_headBeta(lean_object*); LEAN_EXPORT lean_object* l_List_anyM___at_Lean_hasAssignableMVar___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__75___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); lean_object* lean_instantiate_expr_mvars(lean_object*, lean_object*); @@ -1699,7 +1656,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___spec__ static lean_object* l_Lean_MetavarContext_getExprAssignmentDomain___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_modifyDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_get_mvar_assignment(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_modifyLCtx(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1716,7 +1672,6 @@ lean_object* l_Lean_simpLevelMax_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__24(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Nat_foldRevM_loop___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkAuxMVarType___spec__1___closed__1; @@ -1724,6 +1679,7 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_eAssignment___default; LEAN_EXPORT lean_object* l_Lean_MetavarContext_revert(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_MetavarContext_LevelMVarToParam_instMonadMCtxM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_instantiateExprMVars___spec__7(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__15___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_localDeclDependsOn___spec__18(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1739,6 +1695,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MetavarContex LEAN_EXPORT lean_object* l_Lean_MetavarContext_lDepth___default; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36___rarg___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_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_get_delayed_mvar_assignment(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__106___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1763,9 +1720,8 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_localDeclDependsOn___s LEAN_EXPORT lean_object* l_Lean_MetavarContext_LevelMVarToParam_instMonadCacheExprStructEqExprM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__90___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_localDeclDependsOn___spec__44(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_MetavarContext_0__Lean_reprMetavarKind____x40_Lean_MetavarContext___hyg_123____closed__12; uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimMVar___spec__4___boxed(lean_object*, lean_object*, lean_object*); @@ -1773,6 +1729,7 @@ LEAN_EXPORT lean_object* l_Lean_MetavarContext_incDepth(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_toString___at_Lean_MetavarContext_MkBinding_instToStringException___spec__2(lean_object*); LEAN_EXPORT lean_object* l_modify___at_Lean_MetavarContext_instMonadMCtxStateRefT_x27ST___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssigned___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isDelayedAssigned___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_exprDependsOn_x27___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_instantiateLCtxMVars___spec__7___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1791,9 +1748,9 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_localDeclDependsOn_ LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn_x27___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_exprDependsOn___spec__12(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_localDeclDependsOn___spec__17(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__39___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_mkMVarApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadHashMapCacheAdapterExprStructEqExprM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -1808,12 +1765,12 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_isLevelMVarAssigned___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__74___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3(lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instMonadMCtxM___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarKind_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_exprDependsOn_x27___spec__21___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_instMonadMCtxM___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__38___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__78(lean_object*, lean_object*); @@ -1821,10 +1778,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars__ LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__9___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MetavarContext_MkBinding_instToStringException___closed__4; LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__25___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__19(lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__66___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___rarg(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_localDeclDependsOn_x27___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_abstract_range(lean_object*, lean_object*, lean_object*); @@ -1834,7 +1791,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__35___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_isWellFormed___rarg___lambda__4___boxed(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_instantiateExprMVars___spec__44___rarg___lambda__8(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__91___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1863,7 +1819,6 @@ lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_ob LEAN_EXPORT lean_object* l_Lean_MetavarContext_MkBinding_mkBinding___lambda__2(lean_object*, uint8_t, uint8_t, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasExprMVar(lean_object*); static lean_object* l_Lean_instantiateExprMVars___rarg___lambda__10___closed__1; -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_localDeclDependsOn___spec__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -8512,314 +8467,9 @@ x_3 = lean_instantiate_level_mvars(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Level", 10, 10); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Level.0.Lean.Level.updateSucc!Impl", 48, 48); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("succ level expected", 19, 19); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1; -x_2 = l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(541u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 1) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_ctor_get(x_4, 1); -lean_inc(x_5); -lean_dec(x_4); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_ptr_addr(x_6); -lean_dec(x_6); -x_8 = lean_ptr_addr(x_3); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_2); -x_10 = l_Lean_Level_succ___override(x_3); -x_11 = lean_apply_2(x_5, lean_box(0), x_10); -return x_11; -} -else -{ -lean_object* x_12; -lean_dec(x_3); -x_12 = lean_apply_2(x_5, lean_box(0), x_2); -return x_12; -} -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_3); -lean_dec(x_2); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4; -x_16 = l_panic___at_Lean_Level_normalize___spec__1(x_15); -x_17 = lean_apply_2(x_14, lean_box(0), x_16); -return x_17; -} -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Level.0.Lean.Level.updateMax!Impl", 47, 47); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("max level expected", 18, 18); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___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_instantiateLevelMVars___rarg___lambda__1___closed__1; -x_2 = l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1; -x_3 = lean_unsigned_to_nat(552u); -x_4 = lean_unsigned_to_nat(19u); -x_5 = l_Lean_instantiateLevelMVars___rarg___lambda__2___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_instantiateLevelMVars___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_2) == 2) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_ptr_addr(x_7); -x_10 = lean_ptr_addr(x_3); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_mkLevelMax_x27(x_3, x_4); -x_13 = lean_apply_2(x_6, lean_box(0), x_12); -return x_13; -} -else -{ -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = lean_ptr_addr(x_8); -x_15 = lean_ptr_addr(x_4); -x_16 = lean_usize_dec_eq(x_14, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = l_Lean_mkLevelMax_x27(x_3, x_4); -x_18 = lean_apply_2(x_6, lean_box(0), x_17); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = l_Lean_simpLevelMax_x27(x_3, x_4, x_2); -lean_dec(x_4); -lean_dec(x_3); -x_20 = lean_apply_2(x_6, lean_box(0), x_19); -return x_20; -} -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_4); -lean_dec(x_3); -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); -lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3; -x_24 = l_panic___at_Lean_Level_normalize___spec__1(x_23); -x_25 = lean_apply_2(x_22, lean_box(0), x_24); -return x_25; -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_inc(x_1); -x_7 = l_Lean_instantiateLevelMVars___rarg(x_1, x_2, x_3); -x_8 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_4); -lean_closure_set(x_8, 2, x_6); -x_9 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_7, x_8); -return x_9; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Level.0.Lean.Level.updateIMax!Impl", 48, 48); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("imax level expected", 19, 19); -return x_1; -} -} -static lean_object* _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___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_instantiateLevelMVars___rarg___lambda__1___closed__1; -x_2 = l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1; -x_3 = lean_unsigned_to_nat(563u); -x_4 = lean_unsigned_to_nat(20u); -x_5 = l_Lean_instantiateLevelMVars___rarg___lambda__4___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_instantiateLevelMVars___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_2) == 3) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_ptr_addr(x_7); -x_10 = lean_ptr_addr(x_3); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_mkLevelIMax_x27(x_3, x_4); -x_13 = lean_apply_2(x_6, lean_box(0), x_12); -return x_13; -} -else -{ -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = lean_ptr_addr(x_8); -x_15 = lean_ptr_addr(x_4); -x_16 = lean_usize_dec_eq(x_14, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = l_Lean_mkLevelIMax_x27(x_3, x_4); -x_18 = lean_apply_2(x_6, lean_box(0), x_17); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = l_Lean_simpLevelIMax_x27(x_3, x_4, x_2); -x_20 = lean_apply_2(x_6, lean_box(0), x_19); -return x_20; -} -} -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_4); -lean_dec(x_3); -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); -lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3; -x_24 = l_panic___at_Lean_Level_normalize___spec__1(x_23); -x_25 = lean_apply_2(x_22, lean_box(0), x_24); -return x_25; -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_inc(x_1); -x_7 = l_Lean_instantiateLevelMVars___rarg(x_1, x_2, x_3); -x_8 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__4___boxed), 4, 3); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_4); -lean_closure_set(x_8, 2, x_6); -x_9 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_7, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__6(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); @@ -8831,181 +8481,45 @@ x_6 = lean_apply_2(x_5, lean_box(0), x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_inc(x_5); -x_6 = l_Lean_assignLevelMVar___rarg(x_1, x_2, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); -lean_closure_set(x_7, 0, x_3); -lean_closure_set(x_7, 1, x_5); -x_8 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_6, x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__2(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_6) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_7 = lean_ctor_get(x_1, 0); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_instantiate_level_mvars(x_5, x_1); +x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_ctor_get(x_7, 1); +x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_apply_2(x_8, lean_box(0), x_2); -return x_9; -} -else -{ -lean_object* x_10; uint8_t x_11; -lean_dec(x_2); -x_10 = lean_ctor_get(x_6, 0); -lean_inc(x_10); lean_dec(x_6); -x_11 = l_Lean_Level_hasMVar(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_apply_2(x_13, lean_box(0), x_10); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_inc(x_3); -lean_inc(x_1); -x_15 = l_Lean_instantiateLevelMVars___rarg(x_1, x_3, x_10); -lean_inc(x_5); -x_16 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__7), 5, 4); -lean_closure_set(x_16, 0, x_3); -lean_closure_set(x_16, 1, x_4); -lean_closure_set(x_16, 2, x_1); -lean_closure_set(x_16, 3, x_5); -x_17 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_15, x_16); -return x_17; -} -} +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec(x_2); +x_10 = lean_alloc_closure((void*)(l_Lean_setMCtx___rarg___lambda__1___boxed), 2, 1); +lean_closure_set(x_10, 0, x_7); +x_11 = lean_apply_1(x_9, x_10); +x_12 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed), 3, 2); +lean_closure_set(x_12, 0, x_3); +lean_closure_set(x_12, 1, x_8); +x_13 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_11, x_12); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -switch (lean_obj_tag(x_3)) { -case 1: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_ctor_get(x_3, 0); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_2, 0); lean_inc(x_5); -lean_inc(x_1); -x_6 = l_Lean_instantiateLevelMVars___rarg(x_1, x_2, x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1), 3, 2); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_3); -x_8 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_6, x_7); -return x_8; -} -case 2: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = lean_ctor_get(x_3, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -lean_inc(x_2); -lean_inc(x_1); -x_12 = l_Lean_instantiateLevelMVars___rarg(x_1, x_2, x_9); -lean_inc(x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__3), 6, 5); -lean_closure_set(x_13, 0, x_1); -lean_closure_set(x_13, 1, x_2); -lean_closure_set(x_13, 2, x_10); -lean_closure_set(x_13, 3, x_3); -lean_closure_set(x_13, 4, x_11); -x_14 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_12, x_13); -return x_14; -} -case 3: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 1); -lean_inc(x_17); -lean_inc(x_2); -lean_inc(x_1); -x_18 = l_Lean_instantiateLevelMVars___rarg(x_1, x_2, x_15); -lean_inc(x_17); -x_19 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__5), 6, 5); -lean_closure_set(x_19, 0, x_1); -lean_closure_set(x_19, 1, x_2); -lean_closure_set(x_19, 2, x_16); -lean_closure_set(x_19, 3, x_3); -lean_closure_set(x_19, 4, x_17); -x_20 = lean_apply_4(x_17, lean_box(0), lean_box(0), x_18, x_19); -return x_20; -} -case 5: -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_21 = lean_ctor_get(x_3, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_1, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_2, 0); -lean_inc(x_23); -lean_inc(x_21); -lean_inc(x_1); -x_24 = lean_alloc_closure((void*)(l_Lean_getLevelMVarAssignment_x3f___rarg___lambda__1___boxed), 3, 2); -lean_closure_set(x_24, 0, x_1); -lean_closure_set(x_24, 1, x_21); -lean_inc(x_22); -x_25 = lean_apply_4(x_22, lean_box(0), lean_box(0), x_23, x_24); -lean_inc(x_22); -x_26 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__8), 6, 5); -lean_closure_set(x_26, 0, x_1); -lean_closure_set(x_26, 1, x_3); -lean_closure_set(x_26, 2, x_2); -lean_closure_set(x_26, 3, x_21); -lean_closure_set(x_26, 4, x_22); -x_27 = lean_apply_4(x_22, lean_box(0), lean_box(0), x_25, x_26); -return x_27; -} -default: -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_2); -x_28 = lean_ctor_get(x_1, 0); -lean_inc(x_28); -lean_dec(x_1); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); -x_30 = lean_apply_2(x_29, lean_box(0), x_3); -return x_30; -} -} +lean_inc(x_4); +x_6 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__2), 5, 4); +lean_closure_set(x_6, 0, x_3); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_1); +lean_closure_set(x_6, 3, x_4); +x_7 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_5, x_6); +return x_7; } } LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars(lean_object* x_1) { @@ -9016,29 +8530,11 @@ x_2 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___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_Lean_instantiateLevelMVars___rarg___lambda__2(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_instantiateLevelMVars___rarg___lambda__4(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___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_instantiateLevelMVars___rarg___lambda__6(x_1, x_2, x_3); +x_4 = l_Lean_instantiateLevelMVars___rarg___lambda__1(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -9478,823 +8974,545 @@ x_3 = lean_alloc_closure((void*)(l_Lean_MVarId_assign___at_Lean_instantiateExprM return x_3; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___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_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___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; lean_object* x_7; lean_object* x_8; -x_6 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___rarg___lambda__1), 3, 2); -lean_closure_set(x_6, 0, x_3); -lean_closure_set(x_6, 1, x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_instantiate_level_mvars(x_5, x_1); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_alloc_closure((void*)(l_Lean_setMCtx___rarg___lambda__1___boxed), 2, 1); +lean_closure_set(x_9, 0, x_7); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_dec(x_2); +x_11 = lean_apply_1(x_10, x_9); +x_12 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed), 3, 2); +lean_closure_set(x_12, 0, x_3); +lean_closure_set(x_12, 1, x_8); +x_13 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_11, x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___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; +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); x_7 = lean_ctor_get(x_1, 1); lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_apply_1(x_7, x_6); -return x_8; +lean_inc(x_7); +x_8 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1), 5, 4); +lean_closure_set(x_8, 0, x_4); +lean_closure_set(x_8, 1, x_2); +lean_closure_set(x_8, 2, x_1); +lean_closure_set(x_8, 3, x_7); +x_9 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_6, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___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; +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_inc(x_7); +x_8 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1), 5, 4); +lean_closure_set(x_8, 0, x_4); +lean_closure_set(x_8, 1, x_2); +lean_closure_set(x_8, 2, x_1); +lean_closure_set(x_8, 3, x_7); +x_9 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_6, x_8); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg___boxed), 5, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___rarg___boxed), 5, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -if (lean_obj_tag(x_1) == 1) +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_1); +x_9 = l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg(x_2, x_3, x_4, x_5, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_4; size_t x_5; size_t x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ptr_addr(x_4); -lean_dec(x_4); -x_6 = lean_ptr_addr(x_3); -x_7 = lean_usize_dec_eq(x_5, x_6); -if (x_7 == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +x_7 = l_List_reverse___rarg(x_5); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); lean_dec(x_1); -x_8 = l_Lean_Level_succ___override(x_3); -x_9 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); -lean_dec(x_2); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_8); -return x_11; +lean_dec(x_8); +x_10 = lean_apply_2(x_9, lean_box(0), x_7); +return x_10; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_3); -x_12 = lean_ctor_get(x_2, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_4, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 1); lean_inc(x_12); -lean_dec(x_2); -x_13 = lean_ctor_get(x_12, 1); +lean_dec(x_4); +x_13 = lean_ctor_get(x_1, 1); lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_apply_2(x_13, lean_box(0), x_1); -return x_14; +lean_inc(x_2); +lean_inc(x_1); +x_14 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___rarg(x_1, x_2, x_3, x_11, x_6); +x_15 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1), 7, 6); +lean_closure_set(x_15, 0, x_5); +lean_closure_set(x_15, 1, x_1); +lean_closure_set(x_15, 2, x_2); +lean_closure_set(x_15, 3, x_3); +lean_closure_set(x_15, 4, x_12); +lean_closure_set(x_15, 5, x_6); +x_16 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_14, x_15); +return x_16; } } -else +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_3); -lean_dec(x_1); -x_15 = l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4; -x_16 = l_panic___at_Lean_Level_normalize___spec__1(x_15); -x_17 = lean_ctor_get(x_2, 0); -lean_inc(x_17); -lean_dec(x_2); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_apply_2(x_18, lean_box(0), x_16); -return x_19; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg), 6, 0); +return x_3; +} } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = 1; +x_11 = lean_usize_add(x_1, x_10); +x_12 = lean_array_uset(x_2, x_1, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg(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: { -if (lean_obj_tag(x_1) == 2) -{ -lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ptr_addr(x_5); -x_8 = lean_ptr_addr(x_3); -x_9 = lean_usize_dec_eq(x_7, x_8); +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = l_Lean_mkLevelMax_x27(x_3, x_4); -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -lean_dec(x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_apply_2(x_12, lean_box(0), x_10); -return x_13; -} -else -{ -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = lean_ptr_addr(x_6); -x_15 = lean_ptr_addr(x_4); -x_16 = lean_usize_dec_eq(x_14, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = l_Lean_mkLevelMax_x27(x_3, x_4); -x_18 = lean_ctor_get(x_2, 0); -lean_inc(x_18); -lean_dec(x_2); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_apply_2(x_19, lean_box(0), x_17); -return x_20; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = l_Lean_simpLevelMax_x27(x_3, x_4, x_1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); lean_dec(x_2); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_apply_2(x_23, lean_box(0), x_21); -return x_24; -} -} +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_7); +return x_12; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_4); -lean_dec(x_3); -x_25 = l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3; -x_26 = l_panic___at_Lean_Level_normalize___spec__1(x_25); -x_27 = lean_ctor_get(x_2, 0); -lean_inc(x_27); -lean_dec(x_2); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_apply_2(x_28, lean_box(0), x_26); -return x_29; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_array_uget(x_7, x_6); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_7, x_6, x_14); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); +x_18 = lean_box_usize(x_6); +x_19 = lean_box_usize(x_5); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___lambda__1___boxed), 9, 8); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_15); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_3); +lean_closure_set(x_20, 5, x_4); +lean_closure_set(x_20, 6, x_19); +lean_closure_set(x_20, 7, x_8); +x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_inc(x_1); -x_9 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_4, x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_10, 0, x_6); -lean_closure_set(x_10, 1, x_1); -lean_closure_set(x_10, 2, x_8); -x_11 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_9, x_10); -return x_11; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___boxed), 8, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { -if (lean_obj_tag(x_1) == 3) -{ -lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ptr_addr(x_5); -x_8 = lean_ptr_addr(x_3); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = l_Lean_mkLevelIMax_x27(x_3, x_4); -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -lean_dec(x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_apply_2(x_12, lean_box(0), x_10); +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = 1; +x_11 = lean_usize_add(x_1, x_10); +x_12 = lean_array_uset(x_2, x_1, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } -else -{ -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = lean_ptr_addr(x_6); -x_15 = lean_ptr_addr(x_4); -x_16 = lean_usize_dec_eq(x_14, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = l_Lean_mkLevelIMax_x27(x_3, x_4); -x_18 = lean_ctor_get(x_2, 0); -lean_inc(x_18); -lean_dec(x_2); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_apply_2(x_19, lean_box(0), x_17); -return x_20; } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg(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: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = l_Lean_simpLevelIMax_x27(x_3, x_4, x_1); -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); -lean_dec(x_2); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_apply_2(x_23, lean_box(0), x_21); -return x_24; -} -} -} -else +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); +if (x_9 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_25 = l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3; -x_26 = l_panic___at_Lean_Level_normalize___spec__1(x_25); -x_27 = lean_ctor_get(x_2, 0); -lean_inc(x_27); lean_dec(x_2); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_apply_2(x_28, lean_box(0), x_26); -return x_29; -} -} +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_7); +return x_12; } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_array_uget(x_7, x_6); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_7, x_6, x_14); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_4, x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4___boxed), 4, 3); -lean_closure_set(x_10, 0, x_6); -lean_closure_set(x_10, 1, x_1); -lean_closure_set(x_10, 2, x_8); -x_11 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_9, x_10); -return x_11; +x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); +x_18 = lean_box_usize(x_6); +x_19 = lean_box_usize(x_5); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1___boxed), 9, 8); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_15); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_3); +lean_closure_set(x_20, 5, x_4); +lean_closure_set(x_20, 6, x_19); +lean_closure_set(x_20, 7, x_8); +x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_ctor_get(x_3, 6); -lean_inc(x_4); -lean_dec(x_3); -x_5 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_4, x_1); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_apply_2(x_7, lean_box(0), x_5); -return x_8; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___boxed), 8, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_inc(x_7); -x_8 = l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg(x_1, x_2, x_3, x_7, x_4); -x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); -lean_closure_set(x_9, 0, x_5); -lean_closure_set(x_9, 1, x_7); -x_10 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_8, x_9); -return x_10; +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = 1; +x_11 = lean_usize_add(x_1, x_10); +x_12 = lean_array_uset(x_2, x_1, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(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: { -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_2); -return x_11; -} -else +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); +if (x_9 == 0) { -lean_object* x_12; uint8_t x_13; -lean_dec(x_2); -x_12 = lean_ctor_get(x_8, 0); -lean_inc(x_12); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_dec(x_8); -x_13 = l_Lean_Level_hasMVar(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); +lean_dec(x_2); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_apply_2(x_15, lean_box(0), x_12); -return x_16; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_7); +return x_12; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_inc(x_5); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_array_uget(x_7, x_6); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_7, x_6, x_14); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_inc(x_8); lean_inc(x_4); lean_inc(x_3); +lean_inc(x_2); lean_inc(x_1); -x_17 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_3, x_4, x_12, x_5); -lean_inc(x_7); -x_18 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7___boxed), 7, 6); -lean_closure_set(x_18, 0, x_3); -lean_closure_set(x_18, 1, x_4); -lean_closure_set(x_18, 2, x_6); -lean_closure_set(x_18, 3, x_5); -lean_closure_set(x_18, 4, x_1); -lean_closure_set(x_18, 5, x_7); -x_19 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_17, x_18); -return x_19; -} +x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); +x_18 = lean_box_usize(x_6); +x_19 = lean_box_usize(x_5); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___lambda__1___boxed), 9, 8); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_15); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_3); +lean_closure_set(x_20, 5, x_4); +lean_closure_set(x_20, 6, x_19); +lean_closure_set(x_20, 7, x_8); +x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___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_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16(lean_object* x_1, lean_object* x_2) { _start: { -switch (lean_obj_tag(x_4)) { -case 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_4, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -lean_inc(x_1); -x_8 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_6, x_5); -x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1), 3, 2); -lean_closure_set(x_9, 0, x_4); -lean_closure_set(x_9, 1, x_1); -x_10 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_9); -return x_10; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___boxed), 8, 0); +return x_3; } -case 2: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_4, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_11, x_5); -lean_inc(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__3), 8, 7); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -lean_closure_set(x_15, 2, x_3); -lean_closure_set(x_15, 3, x_12); -lean_closure_set(x_15, 4, x_5); -lean_closure_set(x_15, 5, x_4); -lean_closure_set(x_15, 6, x_13); -x_16 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_14, x_15); -return x_16; } -case 3: +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_4, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_4, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 1); -lean_inc(x_19); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_20 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_17, x_5); -lean_inc(x_19); -x_21 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__5), 8, 7); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_2); -lean_closure_set(x_21, 2, x_3); -lean_closure_set(x_21, 3, x_18); -lean_closure_set(x_21, 4, x_5); -lean_closure_set(x_21, 5, x_4); -lean_closure_set(x_21, 6, x_19); -x_22 = lean_apply_4(x_19, lean_box(0), lean_box(0), x_20, x_21); -return x_22; +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = 1; +x_11 = lean_usize_add(x_1, x_10); +x_12 = lean_array_uset(x_2, x_1, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +return x_13; } -case 5: -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_4, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -lean_inc(x_1); -lean_inc(x_23); -x_26 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6___boxed), 3, 2); -lean_closure_set(x_26, 0, x_23); -lean_closure_set(x_26, 1, x_1); -lean_inc(x_25); -x_27 = lean_apply_4(x_25, lean_box(0), lean_box(0), x_24, x_26); -lean_inc(x_25); -x_28 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__8), 8, 7); -lean_closure_set(x_28, 0, x_1); -lean_closure_set(x_28, 1, x_4); -lean_closure_set(x_28, 2, x_2); -lean_closure_set(x_28, 3, x_3); -lean_closure_set(x_28, 4, x_5); -lean_closure_set(x_28, 5, x_23); -lean_closure_set(x_28, 6, x_25); -x_29 = lean_apply_4(x_25, lean_box(0), lean_box(0), x_27, x_28); -return x_29; } -default: +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(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: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_5); +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = lean_ctor_get(x_1, 0); -lean_inc(x_30); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); lean_dec(x_1); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = lean_apply_2(x_31, lean_box(0), x_4); -return x_32; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_7); +return x_12; } +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_array_uget(x_7, x_6); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_7, x_6, x_14); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); +x_18 = lean_box_usize(x_6); +x_19 = lean_box_usize(x_5); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___lambda__1___boxed), 9, 8); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_15); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_3); +lean_closure_set(x_20, 5, x_4); +lean_closure_set(x_20, 6, x_19); +lean_closure_set(x_20, 7, x_8); +x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg), 5, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___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_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(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; -x_6 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___rarg___lambda__1), 3, 2); -lean_closure_set(x_6, 0, x_3); -lean_closure_set(x_6, 1, x_4); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_3, x_1); +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_apply_2(x_6, lean_box(0), x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___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; +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +lean_dec(x_2); x_7 = lean_ctor_get(x_1, 1); lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_apply_1(x_7, x_6); -return x_8; +x_8 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1___boxed), 3, 2); +lean_closure_set(x_8, 0, x_4); +lean_closure_set(x_8, 1, x_1); +x_9 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_6, x_8); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___rarg___boxed), 5, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___boxed), 5, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_inc(x_1); -x_9 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_4, x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_10, 0, x_6); -lean_closure_set(x_10, 1, x_1); -lean_closure_set(x_10, 2, x_8); -x_11 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_9, x_10); -return x_11; +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = 1; +x_11 = lean_usize_add(x_1, x_10); +x_12 = lean_array_uset(x_2, x_1, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(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: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_7); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_array_uget(x_7, x_6); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_7, x_6, x_14); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_4, x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4___boxed), 4, 3); -lean_closure_set(x_10, 0, x_6); -lean_closure_set(x_10, 1, x_1); -lean_closure_set(x_10, 2, x_8); -x_11 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_9, x_10); -return x_11; +x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); +x_18 = lean_box_usize(x_6); +x_19 = lean_box_usize(x_5); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___lambda__1___boxed), 9, 8); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_15); +lean_closure_set(x_20, 2, x_1); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_3); +lean_closure_set(x_20, 5, x_4); +lean_closure_set(x_20, 6, x_19); +lean_closure_set(x_20, 7, x_8); +x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_inc(x_7); -x_8 = l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___rarg(x_1, x_2, x_3, x_7, x_4); -x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); -lean_closure_set(x_9, 0, x_5); -lean_closure_set(x_9, 1, x_7); -x_10 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_8, x_9); -return x_10; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___boxed), 8, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_2); -return x_11; -} -else -{ -lean_object* x_12; uint8_t x_13; -lean_dec(x_2); -x_12 = lean_ctor_get(x_8, 0); -lean_inc(x_12); -lean_dec(x_8); -x_13 = l_Lean_Level_hasMVar(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -lean_dec(x_1); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_apply_2(x_15, lean_box(0), x_12); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_17 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_3, x_4, x_12, x_5); -lean_inc(x_7); -x_18 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3___boxed), 7, 6); -lean_closure_set(x_18, 0, x_3); -lean_closure_set(x_18, 1, x_4); -lean_closure_set(x_18, 2, x_6); -lean_closure_set(x_18, 3, x_5); -lean_closure_set(x_18, 4, x_1); -lean_closure_set(x_18, 5, x_7); -x_19 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_17, x_18); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -switch (lean_obj_tag(x_4)) { -case 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_4, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -lean_inc(x_1); -x_8 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_6, x_5); -x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__1), 3, 2); -lean_closure_set(x_9, 0, x_4); -lean_closure_set(x_9, 1, x_1); -x_10 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_9); -return x_10; -} -case 2: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_4, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_11, x_5); -lean_inc(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__1), 8, 7); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -lean_closure_set(x_15, 2, x_3); -lean_closure_set(x_15, 3, x_12); -lean_closure_set(x_15, 4, x_5); -lean_closure_set(x_15, 5, x_4); -lean_closure_set(x_15, 6, x_13); -x_16 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_14, x_15); -return x_16; -} -case 3: -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_4, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_4, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 1); -lean_inc(x_19); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_20 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_17, x_5); -lean_inc(x_19); -x_21 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__2), 8, 7); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_2); -lean_closure_set(x_21, 2, x_3); -lean_closure_set(x_21, 3, x_18); -lean_closure_set(x_21, 4, x_5); -lean_closure_set(x_21, 5, x_4); -lean_closure_set(x_21, 6, x_19); -x_22 = lean_apply_4(x_19, lean_box(0), lean_box(0), x_20, x_21); -return x_22; -} -case 5: -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_4, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -lean_inc(x_1); -lean_inc(x_23); -x_26 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6___boxed), 3, 2); -lean_closure_set(x_26, 0, x_23); -lean_closure_set(x_26, 1, x_1); -lean_inc(x_25); -x_27 = lean_apply_4(x_25, lean_box(0), lean_box(0), x_24, x_26); -lean_inc(x_25); -x_28 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__4), 8, 7); -lean_closure_set(x_28, 0, x_1); -lean_closure_set(x_28, 1, x_4); -lean_closure_set(x_28, 2, x_2); -lean_closure_set(x_28, 3, x_3); -lean_closure_set(x_28, 4, x_5); -lean_closure_set(x_28, 5, x_23); -lean_closure_set(x_28, 6, x_25); -x_29 = lean_apply_4(x_25, lean_box(0), lean_box(0), x_27, x_28); -return x_29; -} -default: -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_30 = lean_ctor_get(x_1, 0); -lean_inc(x_30); -lean_dec(x_1); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = lean_apply_2(x_31, lean_box(0), x_4); -return x_32; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg), 5, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___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) { -_start: -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_1); -x_9 = l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg(x_2, x_3, x_4, x_5, x_8, x_6); -return x_9; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___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: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_2); -x_7 = l_List_reverse___rarg(x_5); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -lean_dec(x_8); -x_10 = lean_apply_2(x_9, lean_box(0), x_7); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_4, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_4, 1); -lean_inc(x_12); -lean_dec(x_4); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_inc(x_6); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg(x_1, x_2, x_3, x_11, x_6); -x_15 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1), 7, 6); -lean_closure_set(x_15, 0, x_5); -lean_closure_set(x_15, 1, x_1); -lean_closure_set(x_15, 2, x_2); -lean_closure_set(x_15, 3, x_3); -lean_closure_set(x_15, 4, x_12); -lean_closure_set(x_15, 5, x_6); -x_16 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_14, x_15); -return x_16; -} -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg), 6, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg(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; @@ -10331,7 +9549,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10345,26 +9563,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(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; @@ -10401,7 +9619,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10415,26 +9633,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(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; @@ -10471,7 +9689,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10485,26 +9703,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(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; @@ -10541,7 +9759,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10555,65 +9773,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_3, x_1); -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -lean_dec(x_2); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_apply_2(x_6, lean_box(0), x_4); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___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; -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1___boxed), 3, 2); -lean_closure_set(x_8, 0, x_4); -lean_closure_set(x_8, 1, x_1); -x_9 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_6, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___boxed), 5, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(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; @@ -10650,7 +9829,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10664,26 +9843,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(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; @@ -10720,7 +9899,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10734,26 +9913,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(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; @@ -10790,7 +9969,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10804,26 +9983,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(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; @@ -10860,7 +10039,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10874,26 +10053,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(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; @@ -10930,7 +10109,7 @@ lean_inc(x_1); x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); x_18 = lean_box_usize(x_6); x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___lambda__1___boxed), 9, 8); +x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_20, 0, x_18); lean_closure_set(x_20, 1, x_15); lean_closure_set(x_20, 2, x_1); @@ -10944,236 +10123,26 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg___boxed), 8, 0); +x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___boxed), 8, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; x_10 = 1; x_11 = lean_usize_add(x_1, x_10); x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); +x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(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_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_apply_2(x_11, lean_box(0), x_7); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_array_uget(x_7, x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_7, x_6, x_14); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); -x_18 = lean_box_usize(x_6); -x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___lambda__1___boxed), 9, 8); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_15); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_2); -lean_closure_set(x_20, 4, x_3); -lean_closure_set(x_20, 5, x_4); -lean_closure_set(x_20, 6, x_19); -lean_closure_set(x_20, 7, x_8); -x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 1; -x_11 = lean_usize_add(x_1, x_10); -x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(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_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_apply_2(x_11, lean_box(0), x_7); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_array_uget(x_7, x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_7, x_6, x_14); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); -x_18 = lean_box_usize(x_6); -x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___lambda__1___boxed), 9, 8); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_15); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_2); -lean_closure_set(x_20, 4, x_3); -lean_closure_set(x_20, 5, x_4); -lean_closure_set(x_20, 6, x_19); -lean_closure_set(x_20, 7, x_8); -x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 1; -x_11 = lean_usize_add(x_1, x_10); -x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(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_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_apply_2(x_11, lean_box(0), x_7); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_array_uget(x_7, x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_7, x_6, x_14); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); -x_18 = lean_box_usize(x_6); -x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___lambda__1___boxed), 9, 8); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_15); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_2); -lean_closure_set(x_20, 4, x_3); -lean_closure_set(x_20, 5, x_4); -lean_closure_set(x_20, 6, x_19); -lean_closure_set(x_20, 7, x_8); -x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 1; -x_11 = lean_usize_add(x_1, x_10); -x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(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; @@ -12072,147 +11041,7 @@ x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExp return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 1; -x_11 = lean_usize_add(x_1, x_10); -x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg(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_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_apply_2(x_11, lean_box(0), x_7); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_array_uget(x_7, x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_7, x_6, x_14); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); -x_18 = lean_box_usize(x_6); -x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 9, 8); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_15); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_2); -lean_closure_set(x_20, 4, x_3); -lean_closure_set(x_20, 5, x_4); -lean_closure_set(x_20, 6, x_19); -lean_closure_set(x_20, 7, x_8); -x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___lambda__1(size_t 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, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = 1; -x_11 = lean_usize_add(x_1, x_10); -x_12 = lean_array_uset(x_2, x_1, x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg(x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_8); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg(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_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_apply_2(x_11, lean_box(0), x_7); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_array_uget(x_7, x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_7, x_6, x_14); -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_17 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_13, x_8); -x_18 = lean_box_usize(x_6); -x_19 = lean_box_usize(x_5); -x_20 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___lambda__1___boxed), 9, 8); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_15); -lean_closure_set(x_20, 2, x_1); -lean_closure_set(x_20, 3, x_2); -lean_closure_set(x_20, 4, x_3); -lean_closure_set(x_20, 5, x_4); -lean_closure_set(x_20, 6, x_19); -lean_closure_set(x_20, 7, x_8); -x_21 = lean_apply_4(x_16, lean_box(0), lean_box(0), x_17, x_20); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___boxed), 8, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; @@ -12227,7 +11056,7 @@ x_7 = lean_apply_2(x_6, lean_box(0), x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12236,8 +11065,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12253,8 +11082,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12275,7 +11104,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__3(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__3(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12284,8 +11113,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12301,8 +11130,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12323,7 +11152,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__4(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__4(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12332,8 +11161,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12349,8 +11178,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12371,7 +11200,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__5(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_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; @@ -12392,7 +11221,7 @@ x_13 = lean_apply_2(x_12, lean_box(0), x_10); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; @@ -12404,8 +11233,8 @@ lean_dec(x_10); x_13 = lean_array_size(x_1); x_14 = 0; lean_inc(x_2); -x_15 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(x_2, x_3, x_4, x_5, x_13, x_14, x_1, x_6); -x_16 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__5___boxed), 5, 4); +x_15 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___rarg(x_2, x_3, x_4, x_5, x_13, x_14, x_1, x_6); +x_16 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__5___boxed), 5, 4); lean_closure_set(x_16, 0, x_11); lean_closure_set(x_16, 1, x_7); lean_closure_set(x_16, 2, x_8); @@ -12422,8 +11251,8 @@ lean_dec(x_7); x_18 = lean_array_size(x_1); x_19 = 0; lean_inc(x_2); -x_20 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(x_2, x_3, x_4, x_5, x_18, x_19, x_1, x_6); -x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_20 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__22___rarg(x_2, x_3, x_4, x_5, x_18, x_19, x_1, x_6); +x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_21, 0, x_10); lean_closure_set(x_21, 1, x_2); x_22 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_20, x_21); @@ -12431,7 +11260,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_10) == 0) @@ -12445,7 +11274,7 @@ lean_inc(x_1); x_11 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_5, x_6); x_12 = lean_box(x_7); lean_inc(x_9); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__4___boxed), 9, 8); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__4___boxed), 9, 8); lean_closure_set(x_13, 0, x_12); lean_closure_set(x_13, 1, x_8); lean_closure_set(x_13, 2, x_1); @@ -12483,7 +11312,7 @@ lean_inc(x_2); lean_inc(x_1); x_22 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_21, x_6); lean_inc(x_9); -x_23 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__6), 11, 10); +x_23 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__6), 11, 10); lean_closure_set(x_23, 0, x_8); lean_closure_set(x_23, 1, x_1); lean_closure_set(x_23, 2, x_2); @@ -12506,8 +11335,8 @@ lean_dec(x_16); x_25 = lean_array_size(x_8); x_26 = 0; lean_inc(x_1); -x_27 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(x_1, x_2, x_3, x_4, x_25, x_26, x_8, x_6); -x_28 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_27 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__23___rarg(x_1, x_2, x_3, x_4, x_25, x_26, x_8, x_6); +x_28 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_28, 0, x_5); lean_closure_set(x_28, 1, x_1); x_29 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_27, x_28); @@ -12516,7 +11345,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__8(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12525,8 +11354,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__24___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12542,8 +11371,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__25___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12564,7 +11393,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12573,8 +11402,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__28___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__26___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12590,8 +11419,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__29___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__27___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12612,7 +11441,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12621,8 +11450,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__30___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12638,8 +11467,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__33___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__31___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12660,7 +11489,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12669,8 +11498,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__34___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__32___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12686,8 +11515,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__35___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__33___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12708,7 +11537,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12717,8 +11546,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__34___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12734,8 +11563,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__37___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__35___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12756,7 +11585,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12765,8 +11594,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__38___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__36___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12782,8 +11611,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__39___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__37___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12804,7 +11633,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12813,8 +11642,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__40___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__38___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12830,8 +11659,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__41___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__39___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12852,7 +11681,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { if (x_1 == 0) @@ -12861,8 +11690,8 @@ size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_1 x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__40___rarg(x_3, x_4, x_5, x_6, x_10, x_11, x_2, x_7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_13, 0, x_9); lean_closure_set(x_13, 1, x_3); x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_12, x_13); @@ -12878,8 +11707,8 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_2 x_16 = lean_array_size(x_2); x_17 = 0; lean_inc(x_3); -x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); -x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed), 3, 2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__41___rarg(x_3, x_4, x_5, x_6, x_16, x_17, x_2, x_7); +x_19 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_19, 0, x_9); lean_closure_set(x_19, 1, x_3); x_20 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_18, x_19); @@ -12900,7 +11729,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { switch (lean_obj_tag(x_6)) { @@ -12917,7 +11746,7 @@ lean_inc(x_1); x_11 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_12 = lean_box(x_10); lean_inc(x_5); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__2___boxed), 9, 8); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__2___boxed), 9, 8); lean_closure_set(x_13, 0, x_12); lean_closure_set(x_13, 1, x_7); lean_closure_set(x_13, 2, x_1); @@ -12942,7 +11771,7 @@ lean_inc(x_1); x_16 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_17 = lean_box(x_15); lean_inc(x_5); -x_18 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__3___boxed), 9, 8); +x_18 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__3___boxed), 9, 8); lean_closure_set(x_18, 0, x_17); lean_closure_set(x_18, 1, x_7); lean_closure_set(x_18, 2, x_1); @@ -12963,10 +11792,10 @@ lean_inc(x_20); x_21 = l_Lean_Expr_isMVar(x_6); lean_inc(x_2); lean_inc(x_1); -x_22 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg(x_1, x_2, x_3, x_20, x_9); +x_22 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg(x_1, x_2, x_3, x_20, x_9); x_23 = lean_box(x_21); lean_inc(x_5); -x_24 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7___boxed), 10, 9); +x_24 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7___boxed), 10, 9); lean_closure_set(x_24, 0, x_1); lean_closure_set(x_24, 1, x_2); lean_closure_set(x_24, 2, x_3); @@ -12992,7 +11821,7 @@ lean_inc(x_1); x_27 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_28 = lean_box(x_26); lean_inc(x_5); -x_29 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__8___boxed), 9, 8); +x_29 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8___boxed), 9, 8); lean_closure_set(x_29, 0, x_28); lean_closure_set(x_29, 1, x_7); lean_closure_set(x_29, 2, x_1); @@ -13017,7 +11846,7 @@ lean_inc(x_1); x_32 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_33 = lean_box(x_31); lean_inc(x_5); -x_34 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9___boxed), 9, 8); +x_34 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9___boxed), 9, 8); lean_closure_set(x_34, 0, x_33); lean_closure_set(x_34, 1, x_7); lean_closure_set(x_34, 2, x_1); @@ -13059,7 +11888,7 @@ lean_inc(x_1); x_43 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_44 = lean_box(x_42); lean_inc(x_5); -x_45 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10___boxed), 9, 8); +x_45 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10___boxed), 9, 8); lean_closure_set(x_45, 0, x_44); lean_closure_set(x_45, 1, x_7); lean_closure_set(x_45, 2, x_1); @@ -13084,7 +11913,7 @@ lean_inc(x_1); x_48 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_49 = lean_box(x_47); lean_inc(x_5); -x_50 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11___boxed), 9, 8); +x_50 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11___boxed), 9, 8); lean_closure_set(x_50, 0, x_49); lean_closure_set(x_50, 1, x_7); lean_closure_set(x_50, 2, x_1); @@ -13109,7 +11938,7 @@ lean_inc(x_1); x_53 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_54 = lean_box(x_52); lean_inc(x_5); -x_55 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12___boxed), 9, 8); +x_55 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12___boxed), 9, 8); lean_closure_set(x_55, 0, x_54); lean_closure_set(x_55, 1, x_7); lean_closure_set(x_55, 2, x_1); @@ -13134,7 +11963,7 @@ lean_inc(x_1); x_58 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_59 = lean_box(x_57); lean_inc(x_5); -x_60 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13___boxed), 9, 8); +x_60 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13___boxed), 9, 8); lean_closure_set(x_60, 0, x_59); lean_closure_set(x_60, 1, x_7); lean_closure_set(x_60, 2, x_1); @@ -13159,7 +11988,7 @@ lean_inc(x_1); x_63 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_64 = lean_box(x_62); lean_inc(x_5); -x_65 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14___boxed), 9, 8); +x_65 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14___boxed), 9, 8); lean_closure_set(x_65, 0, x_64); lean_closure_set(x_65, 1, x_7); lean_closure_set(x_65, 2, x_1); @@ -13184,7 +12013,7 @@ lean_inc(x_1); x_68 = l_Lean_instantiateExprMVars___rarg(x_1, x_2, x_3, x_4, x_6, x_9); x_69 = lean_box(x_67); lean_inc(x_5); -x_70 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15___boxed), 9, 8); +x_70 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15___boxed), 9, 8); lean_closure_set(x_70, 0, x_69); lean_closure_set(x_70, 1, x_7); lean_closure_set(x_70, 2, x_1); @@ -13199,11 +12028,11 @@ return x_71; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg), 9, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg), 9, 0); return x_3; } } @@ -13267,7 +12096,7 @@ LEAN_EXPORT lean_object* l_Lean_instantiateExprMVars___rarg___lambda__5(lean_obj lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_inc(x_7); x_8 = l_Lean_MVarId_assign___at_Lean_instantiateExprMVars___spec__10___rarg(x_1, x_2, x_3, x_7, x_4); -x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); +x_9 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_9, 0, x_5); lean_closure_set(x_9, 1, x_7); x_10 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_8, x_9); @@ -13332,7 +12161,7 @@ x_8 = lean_alloc_closure((void*)(l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed), lean_closure_set(x_8, 0, x_2); lean_closure_set(x_8, 1, x_7); x_9 = lean_apply_2(x_3, lean_box(0), x_8); -x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); +x_10 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_10, 0, x_4); lean_closure_set(x_10, 1, x_6); x_11 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_9, x_10); @@ -14237,9 +13066,9 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean lean_dec(x_8); x_15 = lean_ctor_get(x_1, 0); lean_inc(x_15); -lean_inc(x_3); lean_inc(x_2); x_16 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_2, x_6, x_7, x_15, x_3); +lean_dec(x_7); lean_inc(x_1); lean_inc(x_2); x_17 = lean_alloc_closure((void*)(l_Lean_instantiateExprMVars___rarg___lambda__9), 3, 2); @@ -14266,7 +13095,7 @@ lean_inc(x_21); x_22 = lean_box(0); lean_inc(x_3); lean_inc(x_2); -x_23 = l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__15___rarg(x_2, x_6, x_7, x_21, x_22, x_3); +x_23 = l_List_mapM_loop___at_Lean_instantiateExprMVars___spec__13___rarg(x_2, x_6, x_7, x_21, x_22, x_3); lean_inc(x_1); lean_inc(x_2); x_24 = lean_alloc_closure((void*)(l_Lean_instantiateExprMVars___rarg___lambda__10), 3, 2); @@ -14301,7 +13130,7 @@ lean_inc(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_34 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg(x_2, x_6, x_7, x_4, x_5, x_1, x_31, x_33, x_3); +x_34 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg(x_2, x_6, x_7, x_4, x_5, x_1, x_31, x_33, x_3); lean_inc(x_5); x_35 = lean_alloc_closure((void*)(l_Lean_instantiateExprMVars___rarg___lambda__7), 6, 5); lean_closure_set(x_35, 0, x_1); @@ -14664,71 +13493,72 @@ lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___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_assignLevelMVar___at_Lean_instantiateExprMVars___spec__12___rarg(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___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; -x_5 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__2(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; +lean_object* x_6; +x_6 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__12___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___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_5; -x_5 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__4(x_1, x_2, x_3, x_4); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); lean_dec(x_1); -return x_5; +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_4; -x_4 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__6(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +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_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__14___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___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_8; -x_8 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__11___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); -lean_dec(x_2); -return x_8; +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___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_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_6; -x_6 = l_Lean_assignLevelMVar___at_Lean_instantiateExprMVars___spec__14___rarg(x_1, x_2, x_3, x_4, x_5); +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_2); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_instantiateLevelMVars___at_Lean_instantiateExprMVars___spec__13___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); -lean_dec(x_2); -return x_8; +x_10 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__15___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); +return x_11; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__16___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) { @@ -14779,28 +13609,23 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__17___rarg(x return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_1); +lean_object* x_4; +x_4 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(x_1, x_2, x_3); lean_dec(x_1); -x_11 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); -return x_12; +return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_5); +lean_object* x_6; +x_6 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__18___rarg(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); -x_10 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__18___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); -return x_11; +lean_dec(x_3); +return x_6; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___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) { @@ -14827,23 +13652,28 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__19___rarg(x return x_11; } } -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___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_4; -x_4 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(x_1, x_2, x_3); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); lean_dec(x_1); -return x_4; +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___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_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_6; -x_6 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateExprMVars___spec__20___rarg(x_1, x_2, x_3, x_4, x_5); +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_5); lean_dec(x_5); -lean_dec(x_3); -return x_6; +x_10 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__20___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); +return x_11; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__21___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) { @@ -15350,191 +14180,143 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__41___rarg(x return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___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: -{ -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_11 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -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_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__42___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___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: -{ -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_11 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_12 = l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___lambda__1(x_10, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -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_mapMUnsafe_map___at_Lean_instantiateExprMVars___spec__43___rarg(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___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_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__1(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___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: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__3(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__3(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___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, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__4(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__4(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__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_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__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_6; -x_6 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__5(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__5(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_7); lean_dec(x_7); -x_12 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9, x_10); +x_12 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_8, x_9, x_10); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__8(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__8(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__9(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__9(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__10(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__10(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__11(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__11(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__12(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__12(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__13(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__13(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__14(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__14(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__44___rarg___lambda__15(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Expr_withAppAux___at_Lean_instantiateExprMVars___spec__42___rarg___lambda__15(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -15604,1302 +14386,185 @@ uint8_t x_9; x_9 = !lean_is_exclusive(x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); -lean_dec(x_10); -x_11 = lean_box(0); -lean_ctor_set(x_8, 0, x_11); -return x_8; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_8, 1); -lean_inc(x_12); -lean_dec(x_8); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_8); -if (x_15 == 0) -{ -return x_8; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_8); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} -} -else -{ -uint8_t x_19; -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_4); -if (x_19 == 0) -{ -return x_4; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_4, 0); -x_21 = lean_ctor_get(x_4, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_4); -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_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_instMonadMCtxStateRefT_x27MetavarContextST(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_alloc_closure((void*)(l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg___boxed), 2, 0); -x_3 = lean_alloc_closure((void*)(l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed), 3, 0); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_instantiateMVarsCore___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_mkHashMapImp___rarg(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_st_ref_get(x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -x_8 = l_Lean_MetavarContext_getExprAssignmentCore_x3f(x_7, x_1); -lean_ctor_set(x_5, 0, x_8); -return x_5; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_5); -x_11 = l_Lean_MetavarContext_getExprAssignmentCore_x3f(x_9, x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) -{ -return x_5; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_ctor_get(x_5, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_5); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3___rarg___boxed), 4, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4___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; -x_6 = lean_st_ref_take(x_4, x_5); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = !lean_is_exclusive(x_7); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 7); -x_11 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_10, x_1, x_2); -lean_ctor_set(x_7, 7, x_11); -x_12 = lean_st_ref_set(x_4, x_7, x_8); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_12, 0, x_15); -return x_12; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -} -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) -{ -return x_12; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_23 = lean_ctor_get(x_7, 0); -x_24 = lean_ctor_get(x_7, 1); -x_25 = lean_ctor_get(x_7, 2); -x_26 = lean_ctor_get(x_7, 3); -x_27 = lean_ctor_get(x_7, 4); -x_28 = lean_ctor_get(x_7, 5); -x_29 = lean_ctor_get(x_7, 6); -x_30 = lean_ctor_get(x_7, 7); -x_31 = lean_ctor_get(x_7, 8); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_7); -x_32 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_30, x_1, x_2); -x_33 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_33, 0, x_23); -lean_ctor_set(x_33, 1, x_24); -lean_ctor_set(x_33, 2, x_25); -lean_ctor_set(x_33, 3, x_26); -lean_ctor_set(x_33, 4, x_27); -lean_ctor_set(x_33, 5, x_28); -lean_ctor_set(x_33, 6, x_29); -lean_ctor_set(x_33, 7, x_32); -lean_ctor_set(x_33, 8, x_31); -x_34 = lean_st_ref_set(x_4, x_33, x_8); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_36 = x_34; -} else { - lean_dec_ref(x_34); - x_36 = lean_box(0); -} -x_37 = lean_box(0); -if (lean_is_scalar(x_36)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_36; -} -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_35); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_34, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_34, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_41 = x_34; -} else { - lean_dec_ref(x_34); - x_41 = lean_box(0); -} -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(1, 2, 0); -} else { - x_42 = x_41; -} -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} -} -} -else -{ -uint8_t x_43; -lean_dec(x_2); -lean_dec(x_1); -x_43 = !lean_is_exclusive(x_6); -if (x_43 == 0) -{ -return x_6; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_6, 0); -x_45 = lean_ctor_get(x_6, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_6); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___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; -x_6 = lean_st_ref_take(x_4, x_5); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = !lean_is_exclusive(x_7); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 6); -x_11 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_10, x_1, x_2); -lean_ctor_set(x_7, 6, x_11); -x_12 = lean_st_ref_set(x_4, x_7, x_8); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_12, 0, x_15); -return x_12; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -} -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) -{ -return x_12; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_23 = lean_ctor_get(x_7, 0); -x_24 = lean_ctor_get(x_7, 1); -x_25 = lean_ctor_get(x_7, 2); -x_26 = lean_ctor_get(x_7, 3); -x_27 = lean_ctor_get(x_7, 4); -x_28 = lean_ctor_get(x_7, 5); -x_29 = lean_ctor_get(x_7, 6); -x_30 = lean_ctor_get(x_7, 7); -x_31 = lean_ctor_get(x_7, 8); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_7); -x_32 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_29, x_1, x_2); -x_33 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_33, 0, x_23); -lean_ctor_set(x_33, 1, x_24); -lean_ctor_set(x_33, 2, x_25); -lean_ctor_set(x_33, 3, x_26); -lean_ctor_set(x_33, 4, x_27); -lean_ctor_set(x_33, 5, x_28); -lean_ctor_set(x_33, 6, x_32); -lean_ctor_set(x_33, 7, x_30); -lean_ctor_set(x_33, 8, x_31); -x_34 = lean_st_ref_set(x_4, x_33, x_8); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_36 = x_34; -} else { - lean_dec_ref(x_34); - x_36 = lean_box(0); -} -x_37 = lean_box(0); -if (lean_is_scalar(x_36)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_36; -} -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_35); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_34, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_34, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_41 = x_34; -} else { - lean_dec_ref(x_34); - x_41 = lean_box(0); -} -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(1, 2, 0); -} else { - x_42 = x_41; -} -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} -} -} -else -{ -uint8_t x_43; -lean_dec(x_2); -lean_dec(x_1); -x_43 = !lean_is_exclusive(x_6); -if (x_43 == 0) -{ -return x_6; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_6, 0); -x_45 = lean_ctor_get(x_6, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_6); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_inc(x_5); -x_6 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_5, x_2, x_3, x_4); -if (lean_obj_tag(x_6) == 0) -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ptr_addr(x_5); -lean_dec(x_5); -x_10 = lean_ptr_addr(x_8); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_1); -x_12 = l_Lean_Level_succ___override(x_8); -lean_ctor_set(x_6, 0, x_12); -return x_6; -} -else -{ -lean_dec(x_8); -lean_ctor_set(x_6, 0, x_1); -return x_6; -} -} -else -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; uint8_t x_17; -x_13 = lean_ctor_get(x_6, 0); -x_14 = lean_ctor_get(x_6, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_6); -x_15 = lean_ptr_addr(x_5); -lean_dec(x_5); -x_16 = lean_ptr_addr(x_13); -x_17 = lean_usize_dec_eq(x_15, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_1); -x_18 = l_Lean_Level_succ___override(x_13); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_14); -return x_19; -} -else -{ -lean_object* x_20; -lean_dec(x_13); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_14); -return x_20; -} -} -} -else -{ -uint8_t x_21; -lean_dec(x_5); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) -{ -return x_6; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 0); -x_23 = lean_ctor_get(x_6, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_6); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -case 2: -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_1, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 1); -lean_inc(x_26); -lean_inc(x_25); -x_27 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_25, x_2, x_3, x_4); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_26); -x_30 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_26, x_2, x_3, x_29); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_34 = lean_ptr_addr(x_28); -x_35 = lean_usize_dec_eq(x_33, x_34); -if (x_35 == 0) -{ -lean_object* x_36; -lean_dec(x_26); -lean_dec(x_1); -x_36 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_36); -return x_30; -} -else -{ -size_t x_37; size_t x_38; uint8_t x_39; -x_37 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_38 = lean_ptr_addr(x_32); -x_39 = lean_usize_dec_eq(x_37, x_38); -if (x_39 == 0) -{ -lean_object* x_40; -lean_dec(x_1); -x_40 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_40); -return x_30; -} -else -{ -lean_object* x_41; -x_41 = l_Lean_simpLevelMax_x27(x_28, x_32, x_1); -lean_dec(x_1); -lean_dec(x_32); -lean_dec(x_28); -lean_ctor_set(x_30, 0, x_41); -return x_30; -} -} -} -else -{ -lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; uint8_t x_46; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -x_44 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_45 = lean_ptr_addr(x_28); -x_46 = lean_usize_dec_eq(x_44, x_45); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; -lean_dec(x_26); -lean_dec(x_1); -x_47 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_43); -return x_48; -} -else -{ -size_t x_49; size_t x_50; uint8_t x_51; -x_49 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_50 = lean_ptr_addr(x_42); -x_51 = lean_usize_dec_eq(x_49, x_50); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_1); -x_52 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_43); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; -x_54 = l_Lean_simpLevelMax_x27(x_28, x_42, x_1); -lean_dec(x_1); -lean_dec(x_42); -lean_dec(x_28); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_43); -return x_55; -} -} -} -} -else -{ -uint8_t x_56; -lean_dec(x_28); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_1); -x_56 = !lean_is_exclusive(x_30); -if (x_56 == 0) -{ -return x_30; -} -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_30, 0); -x_58 = lean_ctor_get(x_30, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_30); -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 -{ -uint8_t x_60; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_27); -if (x_60 == 0) -{ -return x_27; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_27, 0); -x_62 = lean_ctor_get(x_27, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_27); -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; -} -} -} -case 3: -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_1, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 1); -lean_inc(x_65); -lean_inc(x_64); -x_66 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_64, x_2, x_3, x_4); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -lean_inc(x_65); -x_69 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_65, x_2, x_3, x_68); -if (lean_obj_tag(x_69) == 0) -{ -uint8_t x_70; -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) -{ -lean_object* x_71; size_t x_72; size_t x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_69, 0); -x_72 = lean_ptr_addr(x_64); -lean_dec(x_64); -x_73 = lean_ptr_addr(x_67); -x_74 = lean_usize_dec_eq(x_72, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -lean_dec(x_65); -lean_dec(x_1); -x_75 = l_Lean_mkLevelIMax_x27(x_67, x_71); -lean_ctor_set(x_69, 0, x_75); -return x_69; -} -else -{ -size_t x_76; size_t x_77; uint8_t x_78; -x_76 = lean_ptr_addr(x_65); -lean_dec(x_65); -x_77 = lean_ptr_addr(x_71); -x_78 = lean_usize_dec_eq(x_76, x_77); -if (x_78 == 0) -{ -lean_object* x_79; -lean_dec(x_1); -x_79 = l_Lean_mkLevelIMax_x27(x_67, x_71); -lean_ctor_set(x_69, 0, x_79); -return x_69; -} -else -{ -lean_object* x_80; -x_80 = l_Lean_simpLevelIMax_x27(x_67, x_71, x_1); -lean_dec(x_1); -lean_ctor_set(x_69, 0, x_80); -return x_69; -} -} -} -else -{ -lean_object* x_81; lean_object* x_82; size_t x_83; size_t x_84; uint8_t x_85; -x_81 = lean_ctor_get(x_69, 0); -x_82 = lean_ctor_get(x_69, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_69); -x_83 = lean_ptr_addr(x_64); -lean_dec(x_64); -x_84 = lean_ptr_addr(x_67); -x_85 = lean_usize_dec_eq(x_83, x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; -lean_dec(x_65); -lean_dec(x_1); -x_86 = l_Lean_mkLevelIMax_x27(x_67, x_81); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_82); -return x_87; -} -else -{ -size_t x_88; size_t x_89; uint8_t x_90; -x_88 = lean_ptr_addr(x_65); -lean_dec(x_65); -x_89 = lean_ptr_addr(x_81); -x_90 = lean_usize_dec_eq(x_88, x_89); -if (x_90 == 0) -{ -lean_object* x_91; lean_object* x_92; -lean_dec(x_1); -x_91 = l_Lean_mkLevelIMax_x27(x_67, x_81); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_82); -return x_92; -} -else -{ -lean_object* x_93; lean_object* x_94; -x_93 = l_Lean_simpLevelIMax_x27(x_67, x_81, x_1); -lean_dec(x_1); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_82); -return x_94; -} -} -} -} -else -{ -uint8_t x_95; -lean_dec(x_67); -lean_dec(x_65); -lean_dec(x_64); -lean_dec(x_1); -x_95 = !lean_is_exclusive(x_69); -if (x_95 == 0) -{ -return x_69; -} -else -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_69, 0); -x_97 = lean_ctor_get(x_69, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_69); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; -} -} -} -else -{ -uint8_t x_99; -lean_dec(x_65); -lean_dec(x_64); -lean_dec(x_1); -x_99 = !lean_is_exclusive(x_66); -if (x_99 == 0) -{ -return x_66; -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_66, 0); -x_101 = lean_ctor_get(x_66, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_66); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; -} -} -} -case 5: -{ -lean_object* x_103; lean_object* x_104; -x_103 = lean_ctor_get(x_1, 0); -lean_inc(x_103); -x_104 = lean_st_ref_get(x_3, x_4); -if (lean_obj_tag(x_104) == 0) -{ -uint8_t x_105; -x_105 = !lean_is_exclusive(x_104); -if (x_105 == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_106 = lean_ctor_get(x_104, 0); -x_107 = lean_ctor_get(x_104, 1); -x_108 = lean_ctor_get(x_106, 6); -lean_inc(x_108); -lean_dec(x_106); -x_109 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_108, x_103); -if (lean_obj_tag(x_109) == 0) -{ -lean_dec(x_103); -lean_ctor_set(x_104, 0, x_1); -return x_104; -} -else -{ -lean_object* x_110; uint8_t x_111; -lean_dec(x_1); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -lean_dec(x_109); -x_111 = l_Lean_Level_hasMVar(x_110); -if (x_111 == 0) -{ -lean_dec(x_103); -lean_ctor_set(x_104, 0, x_110); -return x_104; -} -else -{ -lean_object* x_112; -lean_free_object(x_104); -x_112 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_110, x_2, x_3, x_107); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -lean_dec(x_112); -lean_inc(x_113); -x_115 = l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg(x_103, x_113, x_2, x_3, x_114); -if (lean_obj_tag(x_115) == 0) -{ -uint8_t x_116; -x_116 = !lean_is_exclusive(x_115); -if (x_116 == 0) -{ -lean_object* x_117; -x_117 = lean_ctor_get(x_115, 0); -lean_dec(x_117); -lean_ctor_set(x_115, 0, x_113); -return x_115; -} -else -{ -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_115, 1); -lean_inc(x_118); -lean_dec(x_115); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_113); -lean_ctor_set(x_119, 1, x_118); -return x_119; -} -} -else -{ -uint8_t x_120; -lean_dec(x_113); -x_120 = !lean_is_exclusive(x_115); -if (x_120 == 0) -{ -return x_115; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +x_11 = lean_box(0); +lean_ctor_set(x_8, 0, x_11); +return x_8; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_121 = lean_ctor_get(x_115, 0); -x_122 = lean_ctor_get(x_115, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_115); -x_123 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -return x_123; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_8, 1); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } } else { -uint8_t x_124; -lean_dec(x_103); -x_124 = !lean_is_exclusive(x_112); -if (x_124 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) { -return x_112; +return x_8; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_112, 0); -x_126 = lean_ctor_get(x_112, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_112); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; -} -} -} +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } -else -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_128 = lean_ctor_get(x_104, 0); -x_129 = lean_ctor_get(x_104, 1); -lean_inc(x_129); -lean_inc(x_128); -lean_dec(x_104); -x_130 = lean_ctor_get(x_128, 6); -lean_inc(x_130); -lean_dec(x_128); -x_131 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_130, x_103); -if (lean_obj_tag(x_131) == 0) -{ -lean_object* x_132; -lean_dec(x_103); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_1); -lean_ctor_set(x_132, 1, x_129); -return x_132; } else { -lean_object* x_133; uint8_t x_134; +uint8_t x_19; lean_dec(x_1); -x_133 = lean_ctor_get(x_131, 0); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Level_hasMVar(x_133); -if (x_134 == 0) +x_19 = !lean_is_exclusive(x_4); +if (x_19 == 0) { -lean_object* x_135; -lean_dec(x_103); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_129); -return x_135; +return x_4; } else { -lean_object* x_136; -x_136 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(x_133, x_2, x_3, x_129); -if (lean_obj_tag(x_136) == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -lean_inc(x_137); -x_139 = l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg(x_103, x_137, x_2, x_3, x_138); -if (lean_obj_tag(x_139) == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_139, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_141 = x_139; -} else { - lean_dec_ref(x_139); - x_141 = lean_box(0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_4, 0); +x_21 = lean_ctor_get(x_4, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_4); +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; } -if (lean_is_scalar(x_141)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_141; } -lean_ctor_set(x_142, 0, x_137); -lean_ctor_set(x_142, 1, x_140); -return x_142; } -else +} +LEAN_EXPORT lean_object* l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2(lean_object* x_1) { +_start: { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_137); -x_143 = lean_ctor_get(x_139, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_139, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_145 = x_139; -} else { - lean_dec_ref(x_139); - x_145 = lean_box(0); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed), 3, 0); +return x_2; } -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(1, 2, 0); -} else { - x_146 = x_145; } -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; +LEAN_EXPORT lean_object* l_Lean_instMonadMCtxStateRefT_x27MetavarContextST(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_alloc_closure((void*)(l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg___boxed), 2, 0); +x_3 = lean_alloc_closure((void*)(l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed), 3, 0); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; } } -else +LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_103); -x_147 = lean_ctor_get(x_136, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_136, 1); -lean_inc(x_148); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_149 = x_136; -} else { - lean_dec_ref(x_136); - x_149 = lean_box(0); -} -if (lean_is_scalar(x_149)) { - x_150 = lean_alloc_ctor(1, 2, 0); -} else { - x_150 = x_149; +lean_object* x_3; +x_3 = l_StateRefT_x27_get___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__1___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; } -lean_ctor_set(x_150, 0, x_147); -lean_ctor_set(x_150, 1, x_148); -return x_150; } +LEAN_EXPORT lean_object* l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_modify___at_Lean_instMonadMCtxStateRefT_x27MetavarContextST___spec__2___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; } } +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_instantiateMVarsCore___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_mkHashMapImp___rarg(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_151; -lean_dec(x_103); -lean_dec(x_1); -x_151 = !lean_is_exclusive(x_104); -if (x_151 == 0) +lean_object* x_5; +x_5 = lean_st_ref_get(x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -return x_104; +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +x_8 = l_Lean_MetavarContext_getExprAssignmentCore_x3f(x_7, x_1); +lean_ctor_set(x_5, 0, x_8); +return x_5; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_104, 0); -x_153 = lean_ctor_get(x_104, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_104); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = l_Lean_MetavarContext_getExprAssignmentCore_x3f(x_9, x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; } } +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +return x_5; } -default: +else { -lean_object* x_155; -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_1); -lean_ctor_set(x_155, 1, x_4); -return x_155; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_5); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } } } -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_getExprMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__3___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___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_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4___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; @@ -16916,9 +14581,9 @@ x_9 = !lean_is_exclusive(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 6); -x_11 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_10, x_1, x_2); -lean_ctor_set(x_7, 6, x_11); +x_10 = lean_ctor_get(x_7, 7); +x_11 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_10, x_1, x_2); +lean_ctor_set(x_7, 7, x_11); x_12 = lean_st_ref_set(x_4, x_7, x_8); if (lean_obj_tag(x_12) == 0) { @@ -16991,7 +14656,7 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_7); -x_32 = l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1(x_29, x_1, x_2); +x_32 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_30, x_1, x_2); x_33 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_33, 0, x_23); lean_ctor_set(x_33, 1, x_24); @@ -16999,8 +14664,8 @@ lean_ctor_set(x_33, 2, x_25); lean_ctor_set(x_33, 3, x_26); lean_ctor_set(x_33, 4, x_27); lean_ctor_set(x_33, 5, x_28); -lean_ctor_set(x_33, 6, x_32); -lean_ctor_set(x_33, 7, x_30); +lean_ctor_set(x_33, 6, x_29); +lean_ctor_set(x_33, 7, x_32); lean_ctor_set(x_33, 8, x_31); x_34 = lean_st_ref_set(x_4, x_33, x_8); if (lean_obj_tag(x_34) == 0) @@ -17046,790 +14711,313 @@ if (lean_is_scalar(x_41)) { } else { x_42 = x_41; } -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} -} -} -else -{ -uint8_t x_43; -lean_dec(x_2); -lean_dec(x_1); -x_43 = !lean_is_exclusive(x_6); -if (x_43 == 0) -{ -return x_6; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_6, 0); -x_45 = lean_ctor_get(x_6, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_6); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 1: -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_inc(x_5); -x_6 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_5, x_2, x_3, x_4); -if (lean_obj_tag(x_6) == 0) -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ptr_addr(x_5); -lean_dec(x_5); -x_10 = lean_ptr_addr(x_8); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_1); -x_12 = l_Lean_Level_succ___override(x_8); -lean_ctor_set(x_6, 0, x_12); -return x_6; -} -else -{ -lean_dec(x_8); -lean_ctor_set(x_6, 0, x_1); -return x_6; -} -} -else -{ -lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; uint8_t x_17; -x_13 = lean_ctor_get(x_6, 0); -x_14 = lean_ctor_get(x_6, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_6); -x_15 = lean_ptr_addr(x_5); -lean_dec(x_5); -x_16 = lean_ptr_addr(x_13); -x_17 = lean_usize_dec_eq(x_15, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_1); -x_18 = l_Lean_Level_succ___override(x_13); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_14); -return x_19; -} -else -{ -lean_object* x_20; -lean_dec(x_13); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_14); -return x_20; -} -} -} -else -{ -uint8_t x_21; -lean_dec(x_5); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_6); -if (x_21 == 0) -{ -return x_6; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_6, 0); -x_23 = lean_ctor_get(x_6, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_6); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -case 2: -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_1, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 1); -lean_inc(x_26); -lean_inc(x_25); -x_27 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_25, x_2, x_3, x_4); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_26); -x_30 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_26, x_2, x_3, x_29); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_34 = lean_ptr_addr(x_28); -x_35 = lean_usize_dec_eq(x_33, x_34); -if (x_35 == 0) -{ -lean_object* x_36; -lean_dec(x_26); -lean_dec(x_1); -x_36 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_36); -return x_30; -} -else -{ -size_t x_37; size_t x_38; uint8_t x_39; -x_37 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_38 = lean_ptr_addr(x_32); -x_39 = lean_usize_dec_eq(x_37, x_38); -if (x_39 == 0) -{ -lean_object* x_40; -lean_dec(x_1); -x_40 = l_Lean_mkLevelMax_x27(x_28, x_32); -lean_ctor_set(x_30, 0, x_40); -return x_30; -} -else -{ -lean_object* x_41; -x_41 = l_Lean_simpLevelMax_x27(x_28, x_32, x_1); -lean_dec(x_1); -lean_dec(x_32); -lean_dec(x_28); -lean_ctor_set(x_30, 0, x_41); -return x_30; -} -} -} -else -{ -lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; uint8_t x_46; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -x_44 = lean_ptr_addr(x_25); -lean_dec(x_25); -x_45 = lean_ptr_addr(x_28); -x_46 = lean_usize_dec_eq(x_44, x_45); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; -lean_dec(x_26); -lean_dec(x_1); -x_47 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_43); -return x_48; -} -else -{ -size_t x_49; size_t x_50; uint8_t x_51; -x_49 = lean_ptr_addr(x_26); -lean_dec(x_26); -x_50 = lean_ptr_addr(x_42); -x_51 = lean_usize_dec_eq(x_49, x_50); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_1); -x_52 = l_Lean_mkLevelMax_x27(x_28, x_42); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_43); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; -x_54 = l_Lean_simpLevelMax_x27(x_28, x_42, x_1); -lean_dec(x_1); -lean_dec(x_42); -lean_dec(x_28); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_43); -return x_55; -} -} -} -} -else -{ -uint8_t x_56; -lean_dec(x_28); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_1); -x_56 = !lean_is_exclusive(x_30); -if (x_56 == 0) -{ -return x_30; -} -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_30, 0); -x_58 = lean_ctor_get(x_30, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_30); -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 -{ -uint8_t x_60; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_27); -if (x_60 == 0) -{ -return x_27; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_27, 0); -x_62 = lean_ctor_get(x_27, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_27); -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; -} -} -} -case 3: -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_1, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 1); -lean_inc(x_65); -lean_inc(x_64); -x_66 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_64, x_2, x_3, x_4); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -lean_inc(x_65); -x_69 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_65, x_2, x_3, x_68); -if (lean_obj_tag(x_69) == 0) -{ -uint8_t x_70; -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) -{ -lean_object* x_71; size_t x_72; size_t x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_69, 0); -x_72 = lean_ptr_addr(x_64); -lean_dec(x_64); -x_73 = lean_ptr_addr(x_67); -x_74 = lean_usize_dec_eq(x_72, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -lean_dec(x_65); -lean_dec(x_1); -x_75 = l_Lean_mkLevelIMax_x27(x_67, x_71); -lean_ctor_set(x_69, 0, x_75); -return x_69; -} -else -{ -size_t x_76; size_t x_77; uint8_t x_78; -x_76 = lean_ptr_addr(x_65); -lean_dec(x_65); -x_77 = lean_ptr_addr(x_71); -x_78 = lean_usize_dec_eq(x_76, x_77); -if (x_78 == 0) -{ -lean_object* x_79; -lean_dec(x_1); -x_79 = l_Lean_mkLevelIMax_x27(x_67, x_71); -lean_ctor_set(x_69, 0, x_79); -return x_69; -} -else -{ -lean_object* x_80; -x_80 = l_Lean_simpLevelIMax_x27(x_67, x_71, x_1); -lean_dec(x_1); -lean_ctor_set(x_69, 0, x_80); -return x_69; -} -} -} -else -{ -lean_object* x_81; lean_object* x_82; size_t x_83; size_t x_84; uint8_t x_85; -x_81 = lean_ctor_get(x_69, 0); -x_82 = lean_ctor_get(x_69, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_69); -x_83 = lean_ptr_addr(x_64); -lean_dec(x_64); -x_84 = lean_ptr_addr(x_67); -x_85 = lean_usize_dec_eq(x_83, x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; -lean_dec(x_65); -lean_dec(x_1); -x_86 = l_Lean_mkLevelIMax_x27(x_67, x_81); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_82); -return x_87; -} -else -{ -size_t x_88; size_t x_89; uint8_t x_90; -x_88 = lean_ptr_addr(x_65); -lean_dec(x_65); -x_89 = lean_ptr_addr(x_81); -x_90 = lean_usize_dec_eq(x_88, x_89); -if (x_90 == 0) -{ -lean_object* x_91; lean_object* x_92; -lean_dec(x_1); -x_91 = l_Lean_mkLevelIMax_x27(x_67, x_81); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_82); -return x_92; -} -else -{ -lean_object* x_93; lean_object* x_94; -x_93 = l_Lean_simpLevelIMax_x27(x_67, x_81, x_1); -lean_dec(x_1); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_82); -return x_94; -} +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; } } } else { -uint8_t x_95; -lean_dec(x_67); -lean_dec(x_65); -lean_dec(x_64); +uint8_t x_43; +lean_dec(x_2); lean_dec(x_1); -x_95 = !lean_is_exclusive(x_69); -if (x_95 == 0) +x_43 = !lean_is_exclusive(x_6); +if (x_43 == 0) { -return x_69; +return x_6; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_69, 0); -x_97 = lean_ctor_get(x_69, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_69); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_6, 0); +x_45 = lean_ctor_get(x_6, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_6); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } -else -{ -uint8_t x_99; -lean_dec(x_65); -lean_dec(x_64); -lean_dec(x_1); -x_99 = !lean_is_exclusive(x_66); -if (x_99 == 0) -{ -return x_66; } -else +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4(lean_object* x_1) { +_start: { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_66, 0); -x_101 = lean_ctor_get(x_66, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_66); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; -} +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MVarId_assign___at_Lean_instantiateMVarsCore___spec__4___rarg___boxed), 5, 0); +return x_2; } } -case 5: +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_103; lean_object* x_104; -x_103 = lean_ctor_get(x_1, 0); -lean_inc(x_103); -x_104 = lean_st_ref_get(x_3, x_4); -if (lean_obj_tag(x_104) == 0) +lean_object* x_5; +x_5 = lean_st_ref_get(x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -uint8_t x_105; -x_105 = !lean_is_exclusive(x_104); -if (x_105 == 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_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_instantiate_level_mvars(x_6, x_1); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_st_ref_take(x_3, x_7); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_106 = lean_ctor_get(x_104, 0); -x_107 = lean_ctor_get(x_104, 1); -x_108 = lean_ctor_get(x_106, 6); -lean_inc(x_108); -lean_dec(x_106); -x_109 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_108, x_103); -if (lean_obj_tag(x_109) == 0) +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_set(x_3, x_9, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_dec(x_103); -lean_ctor_set(x_104, 0, x_1); -return x_104; +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_10); +return x_13; } else { -lean_object* x_110; uint8_t x_111; -lean_dec(x_1); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -lean_dec(x_109); -x_111 = l_Lean_Level_hasMVar(x_110); -if (x_111 == 0) -{ -lean_dec(x_103); -lean_ctor_set(x_104, 0, x_110); -return x_104; +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_10); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} } else { -lean_object* x_112; -lean_free_object(x_104); -x_112 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_110, x_2, x_3, x_107); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -lean_dec(x_112); -lean_inc(x_113); -x_115 = l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg(x_103, x_113, x_2, x_3, x_114); -if (lean_obj_tag(x_115) == 0) -{ -uint8_t x_116; -x_116 = !lean_is_exclusive(x_115); -if (x_116 == 0) +uint8_t x_18; +lean_dec(x_10); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -lean_object* x_117; -x_117 = lean_ctor_get(x_115, 0); -lean_dec(x_117); -lean_ctor_set(x_115, 0, x_113); -return x_115; +return x_13; } else { -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_115, 1); -lean_inc(x_118); -lean_dec(x_115); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_113); -lean_ctor_set(x_119, 1, x_118); -return x_119; +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; +} } } else { -uint8_t x_120; -lean_dec(x_113); -x_120 = !lean_is_exclusive(x_115); -if (x_120 == 0) +uint8_t x_22; +lean_dec(x_10); +lean_dec(x_9); +x_22 = !lean_is_exclusive(x_11); +if (x_22 == 0) { -return x_115; +return x_11; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_121 = lean_ctor_get(x_115, 0); -x_122 = lean_ctor_get(x_115, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_115); -x_123 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -return x_123; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_11, 0); +x_24 = lean_ctor_get(x_11, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_11); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -uint8_t x_124; -lean_dec(x_103); -x_124 = !lean_is_exclusive(x_112); -if (x_124 == 0) +uint8_t x_26; +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_5); +if (x_26 == 0) { -return x_112; +return x_5; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_112, 0); -x_126 = lean_ctor_get(x_112, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_112); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_5, 0); +x_28 = lean_ctor_get(x_5, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_5); +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; } } } } -else -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_128 = lean_ctor_get(x_104, 0); -x_129 = lean_ctor_get(x_104, 1); -lean_inc(x_129); -lean_inc(x_128); -lean_dec(x_104); -x_130 = lean_ctor_get(x_128, 6); -lean_inc(x_130); -lean_dec(x_128); -x_131 = l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(x_130, x_103); -if (lean_obj_tag(x_131) == 0) +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5(lean_object* x_1) { +_start: { -lean_object* x_132; -lean_dec(x_103); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_1); -lean_ctor_set(x_132, 1, x_129); -return x_132; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg___boxed), 4, 0); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_133; uint8_t x_134; -lean_dec(x_1); -x_133 = lean_ctor_get(x_131, 0); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Level_hasMVar(x_133); -if (x_134 == 0) +lean_object* x_5; +x_5 = lean_st_ref_get(x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_135; -lean_dec(x_103); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_129); -return x_135; -} -else +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_ctor_get(x_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_instantiate_level_mvars(x_6, x_1); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_st_ref_take(x_3, x_7); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_136; -x_136 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_133, x_2, x_3, x_129); -if (lean_obj_tag(x_136) == 0) +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_set(x_3, x_9, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -lean_inc(x_137); -x_139 = l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg(x_103, x_137, x_2, x_3, x_138); -if (lean_obj_tag(x_139) == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_139, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_141 = x_139; -} else { - lean_dec_ref(x_139); - x_141 = lean_box(0); -} -if (lean_is_scalar(x_141)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_141; -} -lean_ctor_set(x_142, 0, x_137); -lean_ctor_set(x_142, 1, x_140); -return x_142; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_10); +return x_13; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_137); -x_143 = lean_ctor_get(x_139, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_139, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_145 = x_139; -} else { - lean_dec_ref(x_139); - x_145 = lean_box(0); -} -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(1, 2, 0); -} else { - x_146 = x_145; +lean_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_10); +lean_ctor_set(x_17, 1, x_16); +return x_17; } -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; } +else +{ +uint8_t x_18; +lean_dec(x_10); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_13; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_103); -x_147 = lean_ctor_get(x_136, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_136, 1); -lean_inc(x_148); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_149 = x_136; -} else { - lean_dec_ref(x_136); - x_149 = lean_box(0); +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; } -if (lean_is_scalar(x_149)) { - x_150 = lean_alloc_ctor(1, 2, 0); -} else { - x_150 = x_149; } -lean_ctor_set(x_150, 0, x_147); -lean_ctor_set(x_150, 1, x_148); -return x_150; } +else +{ +uint8_t x_22; +lean_dec(x_10); +lean_dec(x_9); +x_22 = !lean_is_exclusive(x_11); +if (x_22 == 0) +{ +return x_11; } +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_11, 0); +x_24 = lean_ctor_get(x_11, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_11); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -uint8_t x_151; -lean_dec(x_103); +uint8_t x_26; lean_dec(x_1); -x_151 = !lean_is_exclusive(x_104); -if (x_151 == 0) +x_26 = !lean_is_exclusive(x_5); +if (x_26 == 0) { -return x_104; +return x_5; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_104, 0); -x_153 = lean_ctor_get(x_104, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_104); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; -} -} -} -default: -{ -lean_object* x_155; -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_1); -lean_ctor_set(x_155, 1, x_4); -return x_155; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_5, 0); +x_28 = lean_ctor_get(x_5, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_5); +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_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___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: { if (lean_obj_tag(x_2) == 0) @@ -17850,7 +15038,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_ctor_get(x_2, 0); x_11 = lean_ctor_get(x_2, 1); -x_12 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_10, x_4, x_5, x_6); +x_12 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg(x_10, x_4, x_5, x_6); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; @@ -17905,7 +15093,7 @@ x_21 = lean_ctor_get(x_2, 1); lean_inc(x_21); lean_inc(x_20); lean_dec(x_2); -x_22 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_20, x_4, x_5, x_6); +x_22 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg(x_20, x_4, x_5, x_6); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -17952,11 +15140,151 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_3, x_2); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_4); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_array_uget(x_4, x_3); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_4, x_3, x_11); +x_13 = l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2___rarg(x_10, x_5, x_6, x_7); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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 = 1; +x_17 = lean_usize_add(x_3, x_16); +x_18 = lean_array_uset(x_12, x_3, x_14); +x_3 = x_17; +x_4 = x_18; +x_7 = x_15; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_12); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_13); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg___boxed), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_3, x_2); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_4); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_array_uget(x_4, x_3); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_4, x_3, x_11); +x_13 = l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2___rarg(x_10, x_5, x_6, x_7); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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 = 1; +x_17 = lean_usize_add(x_3, x_16); +x_18 = lean_array_uset(x_12, x_3, x_14); +x_3 = x_17; +x_4 = x_18; +x_7 = x_15; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_12); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_13); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg___boxed), 6, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg___boxed), 7, 0); return x_2; } } @@ -18100,7 +15428,71 @@ x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVa return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_st_ref_get(x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +x_8 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_7, x_1); +lean_ctor_set(x_5, 0, x_8); +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_9, x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +return x_5; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_5); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -18162,15 +15554,15 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg___boxed), 7, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -18232,75 +15624,11 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg___boxed), 7, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_st_ref_get(x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -x_8 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_7, x_1); -lean_ctor_set(x_5, 0, x_8); -return x_5; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_5); -x_11 = l_Lean_MetavarContext_getDelayedMVarAssignmentCore_x3f(x_9, x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) -{ -return x_5; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_ctor_get(x_5, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_5); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed), 7, 0); return x_2; } } @@ -19774,147 +17102,7 @@ x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVa return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_3, x_2); -if (x_8 == 0) -{ -lean_object* x_9; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_4); -lean_ctor_set(x_9, 1, x_7); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_array_uget(x_4, x_3); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_4, x_3, x_11); -x_13 = l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2___rarg(x_10, x_5, x_6, x_7); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_18 = lean_array_uset(x_12, x_3, x_14); -x_3 = x_17; -x_4 = x_18; -x_7 = x_15; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_12); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) -{ -return x_13; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed), 7, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_3, x_2); -if (x_8 == 0) -{ -lean_object* x_9; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_4); -lean_ctor_set(x_9, 1, x_7); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_array_uget(x_4, x_3); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_4, x_3, x_11); -x_13 = l_Lean_instantiateExprMVars___at_Lean_instantiateMVarsCore___spec__2___rarg(x_10, x_5, x_6, x_7); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_18 = lean_array_uset(x_12, x_3, x_14); -x_3 = x_17; -x_4 = x_18; -x_7 = x_15; -goto _start; -} -else -{ -uint8_t x_20; -lean_dec(x_12); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) -{ -return x_13; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg___boxed), 7, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { switch (lean_obj_tag(x_2)) { @@ -19936,7 +17124,7 @@ lean_inc(x_11); lean_dec(x_9); x_12 = lean_array_size(x_3); x_13 = 0; -x_14 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg(x_1, x_12, x_13, x_3, x_5, x_6, x_11); +x_14 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg(x_1, x_12, x_13, x_3, x_5, x_6, x_11); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -20004,7 +17192,7 @@ if (x_28 == 0) size_t x_29; size_t x_30; lean_object* x_31; x_29 = lean_array_size(x_3); x_30 = 0; -x_31 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg(x_1, x_29, x_30, x_3, x_5, x_6, x_27); +x_31 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg(x_1, x_29, x_30, x_3, x_5, x_6, x_27); if (lean_obj_tag(x_31) == 0) { uint8_t x_32; @@ -20113,7 +17301,7 @@ lean_inc(x_55); lean_dec(x_53); x_56 = lean_array_size(x_3); x_57 = 0; -x_58 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg(x_1, x_56, x_57, x_3, x_5, x_6, x_55); +x_58 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg(x_1, x_56, x_57, x_3, x_5, x_6, x_55); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -20181,7 +17369,7 @@ if (x_72 == 0) size_t x_73; size_t x_74; lean_object* x_75; x_73 = lean_array_size(x_3); x_74 = 0; -x_75 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(x_1, x_73, x_74, x_3, x_5, x_6, x_71); +x_75 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg(x_1, x_73, x_74, x_3, x_5, x_6, x_71); if (lean_obj_tag(x_75) == 0) { uint8_t x_76; @@ -20279,7 +17467,7 @@ lean_dec(x_4); x_96 = lean_ctor_get(x_2, 0); lean_inc(x_96); x_97 = l_Lean_Expr_isMVar(x_2); -x_98 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg(x_96, x_5, x_6, x_7); +x_98 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg(x_96, x_5, x_6, x_7); lean_dec(x_96); if (lean_obj_tag(x_98) == 0) { @@ -20305,7 +17493,7 @@ lean_inc(x_103); lean_dec(x_101); x_104 = lean_array_size(x_3); x_105 = 0; -x_106 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__15___rarg(x_1, x_104, x_105, x_3, x_5, x_6, x_103); +x_106 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(x_1, x_104, x_105, x_3, x_5, x_6, x_103); if (lean_obj_tag(x_106) == 0) { uint8_t x_107; @@ -20373,7 +17561,7 @@ if (x_120 == 0) size_t x_121; size_t x_122; lean_object* x_123; x_121 = lean_array_size(x_3); x_122 = 0; -x_123 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__16___rarg(x_1, x_121, x_122, x_3, x_5, x_6, x_119); +x_123 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg(x_1, x_121, x_122, x_3, x_5, x_6, x_119); if (lean_obj_tag(x_123) == 0) { uint8_t x_124; @@ -20502,7 +17690,7 @@ size_t x_156; size_t x_157; lean_object* x_158; lean_dec(x_2); x_156 = lean_array_size(x_3); x_157 = 0; -x_158 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__17___rarg(x_1, x_156, x_157, x_3, x_5, x_6, x_154); +x_158 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__15___rarg(x_1, x_156, x_157, x_3, x_5, x_6, x_154); if (lean_obj_tag(x_158) == 0) { uint8_t x_159; @@ -20582,7 +17770,7 @@ lean_dec(x_149); lean_dec(x_146); x_178 = lean_array_size(x_3); x_179 = 0; -x_180 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__18___rarg(x_1, x_178, x_179, x_3, x_5, x_6, x_154); +x_180 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__16___rarg(x_1, x_178, x_179, x_3, x_5, x_6, x_154); if (lean_obj_tag(x_180) == 0) { uint8_t x_181; @@ -20672,7 +17860,7 @@ lean_dec(x_147); lean_dec(x_146); x_196 = lean_array_size(x_3); x_197 = 0; -x_198 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__19___rarg(x_1, x_196, x_197, x_3, x_5, x_6, x_145); +x_198 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__17___rarg(x_1, x_196, x_197, x_3, x_5, x_6, x_145); if (lean_obj_tag(x_198) == 0) { uint8_t x_199; @@ -20771,7 +17959,7 @@ lean_inc(x_217); lean_dec(x_215); x_218 = lean_array_size(x_3); x_219 = 0; -x_220 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__20___rarg(x_1, x_218, x_219, x_3, x_5, x_6, x_217); +x_220 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__18___rarg(x_1, x_218, x_219, x_3, x_5, x_6, x_217); if (lean_obj_tag(x_220) == 0) { uint8_t x_221; @@ -20839,7 +18027,7 @@ if (x_234 == 0) size_t x_235; size_t x_236; lean_object* x_237; x_235 = lean_array_size(x_3); x_236 = 0; -x_237 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__21___rarg(x_1, x_235, x_236, x_3, x_5, x_6, x_233); +x_237 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__19___rarg(x_1, x_235, x_236, x_3, x_5, x_6, x_233); if (lean_obj_tag(x_237) == 0) { uint8_t x_238; @@ -20948,7 +18136,7 @@ lean_inc(x_261); lean_dec(x_259); x_262 = lean_array_size(x_3); x_263 = 0; -x_264 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__22___rarg(x_1, x_262, x_263, x_3, x_5, x_6, x_261); +x_264 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__20___rarg(x_1, x_262, x_263, x_3, x_5, x_6, x_261); if (lean_obj_tag(x_264) == 0) { uint8_t x_265; @@ -21016,7 +18204,7 @@ if (x_278 == 0) size_t x_279; size_t x_280; lean_object* x_281; x_279 = lean_array_size(x_3); x_280 = 0; -x_281 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__23___rarg(x_1, x_279, x_280, x_3, x_5, x_6, x_277); +x_281 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__21___rarg(x_1, x_279, x_280, x_3, x_5, x_6, x_277); if (lean_obj_tag(x_281) == 0) { uint8_t x_282; @@ -21142,7 +18330,7 @@ lean_inc(x_311); lean_dec(x_309); x_312 = lean_array_size(x_3); x_313 = 0; -x_314 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__26___rarg(x_1, x_312, x_313, x_3, x_5, x_6, x_311); +x_314 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__24___rarg(x_1, x_312, x_313, x_3, x_5, x_6, x_311); if (lean_obj_tag(x_314) == 0) { uint8_t x_315; @@ -21210,7 +18398,7 @@ if (x_328 == 0) size_t x_329; size_t x_330; lean_object* x_331; x_329 = lean_array_size(x_3); x_330 = 0; -x_331 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__27___rarg(x_1, x_329, x_330, x_3, x_5, x_6, x_327); +x_331 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__25___rarg(x_1, x_329, x_330, x_3, x_5, x_6, x_327); if (lean_obj_tag(x_331) == 0) { uint8_t x_332; @@ -21319,7 +18507,7 @@ lean_inc(x_355); lean_dec(x_353); x_356 = lean_array_size(x_3); x_357 = 0; -x_358 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__28___rarg(x_1, x_356, x_357, x_3, x_5, x_6, x_355); +x_358 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__26___rarg(x_1, x_356, x_357, x_3, x_5, x_6, x_355); if (lean_obj_tag(x_358) == 0) { uint8_t x_359; @@ -21387,7 +18575,7 @@ if (x_372 == 0) size_t x_373; size_t x_374; lean_object* x_375; x_373 = lean_array_size(x_3); x_374 = 0; -x_375 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__29___rarg(x_1, x_373, x_374, x_3, x_5, x_6, x_371); +x_375 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__27___rarg(x_1, x_373, x_374, x_3, x_5, x_6, x_371); if (lean_obj_tag(x_375) == 0) { uint8_t x_376; @@ -21496,7 +18684,7 @@ lean_inc(x_399); lean_dec(x_397); x_400 = lean_array_size(x_3); x_401 = 0; -x_402 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__30___rarg(x_1, x_400, x_401, x_3, x_5, x_6, x_399); +x_402 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__28___rarg(x_1, x_400, x_401, x_3, x_5, x_6, x_399); if (lean_obj_tag(x_402) == 0) { uint8_t x_403; @@ -21564,7 +18752,7 @@ if (x_416 == 0) size_t x_417; size_t x_418; lean_object* x_419; x_417 = lean_array_size(x_3); x_418 = 0; -x_419 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__31___rarg(x_1, x_417, x_418, x_3, x_5, x_6, x_415); +x_419 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__29___rarg(x_1, x_417, x_418, x_3, x_5, x_6, x_415); if (lean_obj_tag(x_419) == 0) { uint8_t x_420; @@ -21673,7 +18861,7 @@ lean_inc(x_443); lean_dec(x_441); x_444 = lean_array_size(x_3); x_445 = 0; -x_446 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__32___rarg(x_1, x_444, x_445, x_3, x_5, x_6, x_443); +x_446 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__30___rarg(x_1, x_444, x_445, x_3, x_5, x_6, x_443); if (lean_obj_tag(x_446) == 0) { uint8_t x_447; @@ -21741,7 +18929,7 @@ if (x_460 == 0) size_t x_461; size_t x_462; lean_object* x_463; x_461 = lean_array_size(x_3); x_462 = 0; -x_463 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__33___rarg(x_1, x_461, x_462, x_3, x_5, x_6, x_459); +x_463 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__31___rarg(x_1, x_461, x_462, x_3, x_5, x_6, x_459); if (lean_obj_tag(x_463) == 0) { uint8_t x_464; @@ -21850,7 +19038,7 @@ lean_inc(x_487); lean_dec(x_485); x_488 = lean_array_size(x_3); x_489 = 0; -x_490 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__34___rarg(x_1, x_488, x_489, x_3, x_5, x_6, x_487); +x_490 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__32___rarg(x_1, x_488, x_489, x_3, x_5, x_6, x_487); if (lean_obj_tag(x_490) == 0) { uint8_t x_491; @@ -21918,7 +19106,7 @@ if (x_504 == 0) size_t x_505; size_t x_506; lean_object* x_507; x_505 = lean_array_size(x_3); x_506 = 0; -x_507 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__35___rarg(x_1, x_505, x_506, x_3, x_5, x_6, x_503); +x_507 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__33___rarg(x_1, x_505, x_506, x_3, x_5, x_6, x_503); if (lean_obj_tag(x_507) == 0) { uint8_t x_508; @@ -22027,7 +19215,7 @@ lean_inc(x_531); lean_dec(x_529); x_532 = lean_array_size(x_3); x_533 = 0; -x_534 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg(x_1, x_532, x_533, x_3, x_5, x_6, x_531); +x_534 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__34___rarg(x_1, x_532, x_533, x_3, x_5, x_6, x_531); if (lean_obj_tag(x_534) == 0) { uint8_t x_535; @@ -22095,7 +19283,7 @@ if (x_548 == 0) size_t x_549; size_t x_550; lean_object* x_551; x_549 = lean_array_size(x_3); x_550 = 0; -x_551 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg(x_1, x_549, x_550, x_3, x_5, x_6, x_547); +x_551 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__35___rarg(x_1, x_549, x_550, x_3, x_5, x_6, x_547); if (lean_obj_tag(x_551) == 0) { uint8_t x_552; @@ -22189,11 +19377,11 @@ return x_571; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed), 7, 0); return x_2; } } @@ -22480,7 +19668,7 @@ x_95 = lean_ctor_get(x_1, 1); lean_inc(x_95); x_96 = lean_box(0); lean_inc(x_95); -x_97 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg(x_5, x_95, x_96, x_2, x_3, x_48); +x_97 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg(x_5, x_95, x_96, x_2, x_3, x_48); if (lean_obj_tag(x_97) == 0) { lean_object* x_98; lean_object* x_99; uint8_t x_100; @@ -22547,7 +19735,7 @@ x_110 = lean_unsigned_to_nat(1u); x_111 = lean_nat_sub(x_107, x_110); lean_dec(x_107); lean_inc(x_1); -x_112 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg(x_5, x_1, x_109, x_111, x_2, x_3, x_48); +x_112 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg(x_5, x_1, x_109, x_111, x_2, x_3, x_48); if (lean_obj_tag(x_112) == 0) { lean_object* x_113; lean_object* x_114; lean_object* x_115; @@ -23648,7 +20836,7 @@ x_325 = lean_ctor_get(x_1, 1); lean_inc(x_325); x_326 = lean_box(0); lean_inc(x_325); -x_327 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg(x_5, x_325, x_326, x_2, x_3, x_278); +x_327 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg(x_5, x_325, x_326, x_2, x_3, x_278); if (lean_obj_tag(x_327) == 0) { lean_object* x_328; lean_object* x_329; uint8_t x_330; @@ -23717,7 +20905,7 @@ x_340 = lean_unsigned_to_nat(1u); x_341 = lean_nat_sub(x_337, x_340); lean_dec(x_337); lean_inc(x_1); -x_342 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg(x_5, x_1, x_339, x_341, x_2, x_3, x_278); +x_342 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg(x_5, x_1, x_339, x_341, x_2, x_3, x_278); if (lean_obj_tag(x_342) == 0) { lean_object* x_343; lean_object* x_344; lean_object* x_345; @@ -25010,16 +22198,6 @@ lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___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_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__6___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -25030,38 +22208,28 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___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_assignLevelMVar___at_Lean_instantiateMVarsCore___spec__8___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__7___rarg(x_1, x_2, x_3, x_4); +x_5 = l_Lean_instantiateLevelMVars___at_Lean_instantiateMVarsCore___spec__6___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_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___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_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___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_7; -x_7 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__9___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_List_mapM_loop___at_Lean_instantiateMVarsCore___spec__7___rarg(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25069,14 +22237,14 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__8___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25084,14 +22252,14 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__9___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25099,14 +22267,14 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__12___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__10___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25114,24 +22282,54 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__11___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__14___rarg(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getDelayedMVarAssignment_x3f___at_Lean_instantiateMVarsCore___spec__12___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__13___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__14___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_10; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__15___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { @@ -25447,41 +22645,11 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_9 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__36___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_9 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_instantiateMVarsCore___spec__37___rarg(x_1, x_8, x_9, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__38___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Expr_withAppAux___at_Lean_instantiateMVarsCore___spec__36___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); @@ -25514,7 +22682,7 @@ lean_dec(x_2); x_10 = lean_alloc_closure((void*)(l_Lean_setMCtx___rarg___lambda__1___boxed), 2, 1); lean_closure_set(x_10, 0, x_8); x_11 = lean_apply_1(x_9, x_10); -x_12 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__6___boxed), 3, 2); +x_12 = lean_alloc_closure((void*)(l_Lean_instantiateLevelMVars___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_12, 0, x_3); lean_closure_set(x_12, 1, x_7); x_13 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_11, x_12); @@ -51899,7 +49067,7 @@ static lean_object* _init_l_Lean_MetavarContext_getLevelDepth___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_isLevelMVarAssignable___rarg___lambda__1___closed__1; x_2 = l_Lean_MetavarContext_getLevelDepth___closed__1; -x_3 = lean_unsigned_to_nat(946u); +x_3 = lean_unsigned_to_nat(936u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_MetavarContext_getDecl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -78660,26 +75828,6 @@ l_Lean_MetavarContext_getDecl___closed__3 = _init_l_Lean_MetavarContext_getDecl_ lean_mark_persistent(l_Lean_MetavarContext_getDecl___closed__3); l_Lean_PersistentHashMap_insertAux___at_Lean_assignLevelMVar___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_assignLevelMVar___spec__2___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_assignLevelMVar___spec__2___closed__1); -l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__1); -l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__2); -l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__3); -l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__1___closed__4); -l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__1); -l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__2 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__2); -l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__2___closed__3); -l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__1); -l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__2 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__2); -l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3 = _init_l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_instantiateLevelMVars___rarg___lambda__4___closed__3); l_Lean_instantiateExprMVars___rarg___lambda__9___closed__1 = _init_l_Lean_instantiateExprMVars___rarg___lambda__9___closed__1(); lean_mark_persistent(l_Lean_instantiateExprMVars___rarg___lambda__9___closed__1); l_Lean_instantiateExprMVars___rarg___lambda__9___closed__2 = _init_l_Lean_instantiateExprMVars___rarg___lambda__9___closed__2();