From 452ee6f2d2b372519f2658c3e795084436c3a2ed Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 17 Jun 2021 17:25:17 -0400 Subject: [PATCH 1/6] AST tracking and export --- .gitignore | 1 + src/frontends/lean/brackets.cpp | 111 +++- src/frontends/lean/builtin_cmds.cpp | 177 ++++-- src/frontends/lean/builtin_exprs.cpp | 433 ++++++++++---- src/frontends/lean/calc.cpp | 32 +- src/frontends/lean/calc.h | 2 +- src/frontends/lean/cmd_table.h | 9 +- src/frontends/lean/decl_attributes.cpp | 36 +- src/frontends/lean/decl_attributes.h | 6 +- src/frontends/lean/decl_cmds.cpp | 226 ++++--- src/frontends/lean/decl_cmds.h | 4 +- src/frontends/lean/decl_util.cpp | 54 +- src/frontends/lean/decl_util.h | 8 +- src/frontends/lean/definition_cmds.cpp | 79 ++- src/frontends/lean/definition_cmds.h | 4 +- src/frontends/lean/elaborator.cpp | 2 +- src/frontends/lean/inductive_cmds.cpp | 59 +- src/frontends/lean/inductive_cmds.h | 4 +- src/frontends/lean/json.cpp | 2 + src/frontends/lean/json.h | 6 +- src/frontends/lean/match_expr.cpp | 22 +- src/frontends/lean/module_parser.cpp | 2 +- src/frontends/lean/module_parser.h | 1 + src/frontends/lean/notation_cmd.cpp | 267 ++++++--- src/frontends/lean/notation_cmd.h | 4 +- src/frontends/lean/parse_table.cpp | 26 +- src/frontends/lean/parse_table.h | 18 +- src/frontends/lean/parser.cpp | 561 +++++++++++++----- src/frontends/lean/parser.h | 138 +++-- src/frontends/lean/parser_config.cpp | 45 +- src/frontends/lean/parser_config.h | 8 +- src/frontends/lean/prenum.cpp | 6 + src/frontends/lean/print_cmd.cpp | 101 +++- src/frontends/lean/print_cmd.h | 2 +- src/frontends/lean/structure_cmd.cpp | 101 +++- src/frontends/lean/structure_cmd.h | 4 +- src/frontends/lean/structure_instance.cpp | 9 + src/frontends/lean/tactic_notation.cpp | 120 ++-- src/frontends/lean/type_util.cpp | 13 +- src/frontends/lean/type_util.h | 3 +- src/frontends/lean/user_command.cpp | 15 +- src/frontends/lean/user_notation.cpp | 13 +- src/frontends/lean/util.cpp | 31 +- src/frontends/lean/util.h | 7 +- src/kernel/expr.h | 5 + src/library/CMakeLists.txt | 2 +- src/library/abstract_parser.h | 28 +- src/library/annotation.cpp | 6 + src/library/ast_exporter.cpp | 188 ++++++ src/library/ast_exporter.h | 21 + src/library/attribute_manager.cpp | 32 +- src/library/attribute_manager.h | 20 +- src/library/choice.cpp | 3 + src/library/compiler/nat_value.cpp | 6 + src/library/compiler/rec_fn_macro.cpp | 7 + src/library/constructions/projection.cpp | 15 + src/library/delayed_abstraction.cpp | 7 + src/library/equations_compiler/equations.cpp | 35 +- src/library/module_mgr.cpp | 9 +- src/library/module_mgr.h | 3 + src/library/quote.cpp | 7 + src/library/scoped_ext.h | 1 - src/library/sorry.cpp | 6 + src/library/string.cpp | 6 + src/library/tactic/ac_tactics.cpp | 8 + .../tactic/backward/backward_lemmas.cpp | 10 +- src/library/tactic/smt/util.cpp | 4 + src/library/tactic/user_attribute.cpp | 10 +- src/library/typed_expr.cpp | 7 +- src/library/typed_expr.h | 2 +- src/library/vm/vm_override.cpp | 11 +- src/library/vm/vm_parser.cpp | 54 +- src/library/vm/vm_parser.h | 4 +- src/shell/lean.cpp | 7 + src/util/lean_json.h | 20 + 75 files changed, 2409 insertions(+), 907 deletions(-) create mode 100644 src/library/ast_exporter.cpp create mode 100644 src/library/ast_exporter.h create mode 100644 src/util/lean_json.h diff --git a/.gitignore b/.gitignore index 78e13f4ac4..721799fd95 100644 --- a/.gitignore +++ b/.gitignore @@ -3,6 +3,7 @@ .#* *.olean *.tlean +*.ast.json *.lock build GPATH diff --git a/src/frontends/lean/brackets.cpp b/src/frontends/lean/brackets.cpp index 62b138c274..0a6ab46670 100644 --- a/src/frontends/lean/brackets.cpp +++ b/src/frontends/lean/brackets.cpp @@ -15,11 +15,12 @@ Author: Leonardo de Moura namespace lean { /* Parse rest of the subtype expression prefix '{' id ':' expr '\\' ... */ -static expr parse_subtype(parser & p, pos_info const & pos, expr const & local) { +static expr parse_subtype(parser & p, pos_info const & pos, ast_data & parent, expr const & local) { parser::local_scope scope(p); p.add_local(local); expr pred = p.parse_expr(); p.check_token_next(get_rcurly_tk(), "invalid subtype, '}' expected"); + parent.push(p.get_id(pred)); bool use_cache = false; pred = p.save_pos(Fun(local, pred, use_cache), pos); expr subtype = p.save_pos(mk_constant(get_subtype_name()), pos); @@ -27,10 +28,11 @@ static expr parse_subtype(parser & p, pos_info const & pos, expr const & local) } /* Parse rest of the set_of expression prefix '{' id ':' expr '|' ... */ -static expr parse_set_of(parser & p, pos_info const & pos, expr const & local) { +static expr parse_set_of(parser & p, pos_info const & pos, ast_data & parent, expr const & local) { parser::local_scope scope(p); p.add_local(local); expr pred = p.parse_expr(); + parent.push(p.get_id(pred)); p.check_token_next(get_rcurly_tk(), "invalid set_of, '}' expected"); bool use_cache = false; pred = p.save_pos(Fun(local, pred, use_cache), pos); @@ -45,7 +47,8 @@ static expr parse_set_replacement(parser & p, pos_info const & pos, expr const & // So we parse the binders, assemble an expression for the characteristic predicate, // then use `patexpr_to_expr` to do the updating. buffer params; - auto env = p.parse_binders(params, false); + auto& bis = p.new_ast("binders", p.pos()); + auto env = p.parse_binders(&bis, params, false); p.check_token_next(get_rcurly_tk(), "invalid set_replacement, '}' expected"); // The predicate will look like `λ _x, ∃ p_1, ∃ p_2, ..., hole_expr[p_1, p_2, ...] = _x`. @@ -65,7 +68,9 @@ static expr parse_set_replacement(parser & p, pos_info const & pos, expr const & // Update identifiers so globals are actually globals. pred = p.patexpr_to_expr_core(pred); // `{_x | ∃ p_1, ∃ p_2, ..., hole_expr[p_1, p_2, ...] = _x}` - return p.mk_app(mk_constant(get_set_of_name()), pred, pos); + pred = p.mk_app(mk_constant(get_set_of_name()), pred, pos); + p.set_ast_pexpr(p.new_ast("set_replacement", pos).push(p.get_id(hole_expr)).push(bis.m_id).m_id, pred); + return pred; } /* Create singletoncollection for '{' expr '}' */ @@ -77,12 +82,14 @@ static expr mk_singleton(parser & p, pos_info const & pos, expr const & e) { static expr parse_fin_set(parser & p, pos_info const & pos, expr const & e) { lean_assert(p.curr_is_token(get_comma_tk()) || p.curr_is_token(get_rcurly_tk())); buffer> stack; + auto& data = p.new_ast("fin_set", pos).push(p.get_id(e)); stack.emplace_back(pos, e); while (p.curr_is_token(get_comma_tk())) { auto ins_pos = p.pos(); p.next(); if (p.curr_is_token(get_rcurly_tk())) break; expr e2 = p.parse_expr(); + data.push(p.get_id(e2)); stack.emplace_back(ins_pos, e2); } p.check_token_next(get_rcurly_tk(), "invalid explicit finite collection, '}' expected"); @@ -94,45 +101,61 @@ static expr parse_fin_set(parser & p, pos_info const & pos, expr const & e) { expr insert = p.save_pos(mk_constant(get_has_insert_insert_name()), e2.first); r = p.rec_save_pos(mk_app(insert, e2.second, r), e2.first); } + p.set_ast_pexpr(data.m_id, r); return r; } /* Parse rest of the sep expression '{' id '∈' ... */ -static expr parse_sep(parser & p, pos_info const & pos, name const & id) { +static expr parse_sep(parser & p, pos_info const & pos, name const & id, ast_data & id_ast) { expr s = p.parse_expr(); p.check_token_next(get_bar_tk(), "invalid sep expression, '|' expected"); parser::local_scope scope(p); expr local = p.save_pos(mk_local(id, p.save_pos(mk_expr_placeholder(), pos)), pos); p.add_local(local); expr pred = p.parse_expr(); + auto& data = p.new_ast("sep", pos).push(id_ast.m_id).push(p.get_id(s)).push(p.get_id(pred)); p.check_token_next(get_rcurly_tk(), "invalid sep expression, '}' expected"); bool use_cache = false; pred = p.rec_save_pos(Fun(local, pred, use_cache), pos); - return p.rec_save_pos(mk_app(mk_constant(get_has_sep_sep_name()), pred, s), pos); + expr r = p.rec_save_pos(mk_app(mk_constant(get_has_sep_sep_name()), pred, s), pos); + p.set_ast_pexpr(data.m_id, r); + return r; } -static expr parse_structure_instance_core(parser & p, optional const & src = {}, name const & S = {}, - buffer fns = {}, buffer fvs = {}) { +static expr parse_structure_instance_core(parser & p, pos_info const & pos, optional const & src = {}, + name const & S = {}, ast_id S_ast = 0, + buffer fns = {}, buffer fvs = {}, ast_data * _fields = nullptr) { buffer sources; - bool catchall = false; - if (src) + ast_id catchall = 0; + auto& fields = _fields ? *_fields : p.new_ast("fields", pos); + auto& sources_ast = p.new_ast("sources", pos); + auto& data = p.new_ast("structinst", pos).push(S_ast); + ast_id src_id = 0; + if (src) { sources.push_back(*src); + src_id = p.get_id(*src); + } + data.push(src_id).push(fields.m_id).push(sources_ast.m_id); bool read_source = false; while (!p.curr_is_token(get_rcurly_tk())) { if (p.curr_is_token(get_dotdot_tk())) { p.next(); if (p.curr_is_token(get_rcurly_tk())) { - // ", ...}" - catchall = true; + // ", ..}" + catchall = p.new_ast("..", p.pos()).m_id; break; } - // ", ...src" + // ", ..src" sources.push_back(p.parse_expr()); + sources_ast.push(p.get_id(sources.back())); read_source = true; } else if (!read_source) { - fns.push_back(p.check_id_next("invalid structure instance, identifier expected")); + ast_id id; name n; + std::tie(id, n) = p.check_id_next("invalid structure instance, identifier expected"); + fns.push_back(n); p.check_token_next(get_assign_tk(), "invalid structure instance, ':=' expected"); fvs.push_back(p.parse_expr()); + fields.push(p.new_ast(":=", p.pos()).push(id).push(p.get_id(fvs.back())).m_id); } else { break; } @@ -143,26 +166,31 @@ static expr parse_structure_instance_core(parser & p, optional const & src break; } } + data.push(catchall); p.check_token_next(get_rcurly_tk(), "invalid structure instance, '}' expected"); - return mk_structure_instance(S, fns, fvs, sources, catchall); + expr r = mk_structure_instance(S, fns, fvs, sources, catchall); + p.set_ast_pexpr(data.m_id, r); + return r; } /* Parse rest of the qualified structure instance prefix '{' S '.' ... */ -static expr parse_qualified_structure_instance(parser & p, name S, pos_info const & S_pos) { - S = p.to_constant(S, "invalid structure name", S_pos); - return parse_structure_instance_core(p, none_expr(), S); +static expr parse_qualified_structure_instance(parser & p, pos_info const & pos, name S, ast_data & S_ast) { + S = p.to_constant(S, "invalid structure name", S_ast.m_start); + return parse_structure_instance_core(p, pos, none_expr(), S, S_ast.m_id); } /* Parse rest of the structure instance prefix '{' fname ... */ -static expr parse_structure_instance(parser & p, name const & fname) { +static expr parse_structure_instance(parser & p, pos_info const & pos, name const & fname, ast_data & fn_ast) { buffer fns; buffer fvs; fns.push_back(fname); p.check_token_next(get_assign_tk(), "invalid structure instance, ':=' expected"); fvs.push_back(p.parse_expr()); + auto& fields = p.new_ast("fields", pos) + .push(p.new_ast(":=", fn_ast.m_start).push(fn_ast.m_id).push(p.get_id(fvs.back())).m_id); if (p.curr_is_token(get_comma_tk())) p.next(); - return parse_structure_instance_core(p, none_expr(), name(), fns, fvs); + return parse_structure_instance_core(p, pos, none_expr(), name(), 0, fns, fvs, &fields); } static name * g_emptyc_or_emptys = nullptr; @@ -175,42 +203,58 @@ expr parse_curly_bracket(parser & p, unsigned, expr const *, pos_info const & po expr e; if (p.curr_is_token(get_rcurly_tk())) { p.next(); - return p.save_pos(mk_constant(*g_emptyc_or_emptys), pos); + expr r = p.save_pos(mk_constant(*g_emptyc_or_emptys), pos); + p.set_ast_pexpr(p.new_ast("{}", pos).m_id, r); + return r; } else if (p.curr_is_identifier()) { auto id_pos = p.pos(); name id = p.get_name_val(); + auto& id_ast = p.new_ast("ident", id_pos, id); p.next(); if (p.curr_is_token(get_dslash_tk())) { expr type = p.save_pos(mk_expr_placeholder(), id_pos); expr local = p.save_pos(mk_local(id, type), id_pos); p.next(); - return parse_subtype(p, pos, local); + auto& data = p.new_ast("subtype", pos).push(id_ast.m_id).push(0); + expr r = parse_subtype(p, pos, data, local); + p.set_ast_pexpr(data.m_id, r); + return r; } else if (p.curr_is_token(get_bar_tk())) { expr type = p.save_pos(mk_expr_placeholder(), id_pos); expr local = p.save_pos(mk_local(id, type), id_pos); p.next(); - return parse_set_of(p, pos, local); + auto& data = p.new_ast("set_of", pos).push(id_ast.m_id).push(0); + expr r = parse_set_of(p, pos, data, local); + p.set_ast_pexpr(data.m_id, r); + return r; + } else if (p.curr_is_token(get_colon_tk())) { p.next(); expr type = p.parse_expr(); expr local = p.save_pos(mk_local(id, type), id_pos); if (p.curr_is_token(get_bar_tk())) { p.next(); - return parse_set_of(p, pos, local); + auto& data = p.new_ast("set_of", pos).push(id_ast.m_id).push(p.get_id(type)); + expr r = parse_set_of(p, pos, data, local); + p.set_ast_pexpr(data.m_id, r); + return r; } else { p.check_token_next(get_dslash_tk(), "invalid expression, '//' or '|' expected"); - return parse_subtype(p, pos, local); + auto& data = p.new_ast("subtype", pos).push(id_ast.m_id).push(p.get_id(type)); + expr r = parse_subtype(p, pos, data, local); + p.set_ast_pexpr(data.m_id, r); + return r; } } else if (p.curr_is_token(get_period_tk())) { p.next(); - return parse_qualified_structure_instance(p, id, id_pos); + return parse_qualified_structure_instance(p, pos, id, id_ast); } else if (p.curr_is_token(get_assign_tk()) || p.curr_is_token(get_fieldarrow_tk())) { - return parse_structure_instance(p, id); + return parse_structure_instance(p, pos, id, id_ast); } else if (p.curr_is_token(get_membership_tk()) || p.curr_is_token(get_in_tk())) { p.next(); - return parse_sep(p, pos, id); + return parse_sep(p, pos, id, id_ast); } else { - expr left = p.id_to_expr(id, id_pos); + expr left = p.id_to_expr(id, id_ast); unsigned rbp = 0; while (rbp < p.curr_lbp()) { left = p.parse_led(left); @@ -231,14 +275,17 @@ expr parse_curly_bracket(parser & p, unsigned, expr const *, pos_info const & po return parse_fin_set(p, pos, p.patexpr_to_expr_core(e)); } else { p.maybe_throw_error({"invalid set replacement notation, ',', '}', or `|` expected", p.pos()}); - return mk_singleton(p, pos, p.patexpr_to_expr_core(e)); + ast_id id = p.new_ast("fin_set", pos).push(p.get_id(e)).m_id; + expr r = mk_singleton(p, pos, p.patexpr_to_expr_core(e)); + p.set_ast_pexpr(id, r); + return r; } } else if (p.curr_is_token(get_period_tk())) { p.next(); p.check_token_next(get_rcurly_tk(), "invalid empty structure instance, '}' expected"); return p.save_pos(mk_structure_instance(), pos); } else if (p.curr_is_token(get_dotdot_tk())) { - return parse_structure_instance_core(p); + return parse_structure_instance_core(p, pos); } else { e = p.parse_expr(); } @@ -249,7 +296,7 @@ expr parse_curly_bracket(parser & p, unsigned, expr const *, pos_info const & po return parse_fin_set(p, pos, e); } else if (p.curr_is_token(get_with_tk())) { p.next(); - return parse_structure_instance_core(p, some_expr(e)); + return parse_structure_instance_core(p, pos, some_expr(e)); } else { return p.parser_error_or_expr({"invalid '{' expression, ',', '}', '..', `//` or `|` expected", p.pos()}); } diff --git a/src/frontends/lean/builtin_cmds.cpp b/src/frontends/lean/builtin_cmds.cpp index 89100e67ea..dde33ba5c3 100644 --- a/src/frontends/lean/builtin_cmds.cpp +++ b/src/frontends/lean/builtin_cmds.cpp @@ -55,10 +55,11 @@ Author: Leonardo de Moura #include "frontends/lean/decl_attributes.h" namespace lean { -environment section_cmd(parser & p) { - name n; +environment section_cmd(parser & p, ast_id & cmd_id) { + ast_id id = 0; name n; if (p.curr_is_identifier()) - n = p.check_atomic_id_next("invalid section, atomic identifier expected"); + std::tie(id, n) = p.check_atomic_id_next("invalid section, atomic identifier expected"); + p.get_ast(cmd_id).push(id); p.push_local_scope(); return push_scope(p.env(), p.ios(), scope_kind::Section, n); } @@ -106,8 +107,10 @@ environment execute_open(environment env, io_state const & ios, export_decl cons return replay_export_decls_core(env, ios, old_export_decls_sz); } -environment namespace_cmd(parser & p) { - name n = p.check_decl_id_next("invalid namespace declaration, identifier expected"); +environment namespace_cmd(parser & p, ast_id & cmd_id) { + ast_id id = 0; name n; + std::tie(id, n) = p.check_decl_id_next("invalid namespace declaration, identifier expected"); + p.get_ast(cmd_id).push(id); p.push_local_scope(); unsigned old_export_decls_sz = length(get_active_export_decls(p.env())); environment env = push_scope(p.env(), p.ios(), scope_kind::Namespace, n); @@ -150,7 +153,7 @@ static environment redeclare_aliases(environment env, parser & p, return env; } -environment end_scoped_cmd(parser & p) { +environment end_scoped_cmd(parser & p, ast_id & cmd_id) { local_level_decls level_decls = p.get_local_level_decls(); list> entries = p.get_local_entries(); if (!p.has_local_scopes()) { @@ -160,11 +163,13 @@ environment end_scoped_cmd(parser & p) { try { p.check_break_before(); if (p.curr_is_identifier()) { - name n; - n = p.check_id_next("invalid end of scope, identifier expected"); + ast_id id; name n; + std::tie(id, n) = p.check_id_next("invalid end of scope, identifier expected"); + p.get_ast(cmd_id).push(id); environment env = pop_scope(p.env(), p.ios(), n); return redeclare_aliases(env, p, level_decls, entries); } else { + p.get_ast(cmd_id).push(0); environment env = pop_scope(p.env(), p.ios()); return redeclare_aliases(env, p, level_decls, entries); } @@ -191,17 +196,18 @@ class transient_cmd_scope { } }; -environment check_cmd(parser & p) { +environment check_cmd(parser & p, ast_id & cmd_id) { expr e; level_param_names ls; transient_cmd_scope cmd_scope(p); - std::tie(e, ls) = parse_local_expr(p, "_check"); + auto& data = p.get_ast(cmd_id); + std::tie(e, ls) = parse_local_expr(p, data, "_check"); type_checker tc(p.env(), true, false); expr type = tc.check(e, ls); if (is_synthetic_sorry(e) && (is_synthetic_sorry(type) || is_metavar(type))) { // do not show useless type-checking results such as ?? : ?M_1 return p.env(); } - auto out = p.mk_message(p.cmd_pos(), p.pos(), INFORMATION); + auto out = p.mk_message(data.m_start, p.pos(), INFORMATION); formatter fmt = out.get_formatter(); unsigned indent = get_pp_indent(p.get_options()); format e_fmt = fmt(e); @@ -212,15 +218,18 @@ environment check_cmd(parser & p) { return p.env(); } -environment reduce_cmd(parser & p) { +environment reduce_cmd(parser & p, ast_id & cmd_id) { transient_cmd_scope cmd_scope(p); bool whnf = false; + ast_id id = 0; if (p.curr_is_token(get_whnf_tk())) { + id = p.new_ast("whnf", p.pos()).m_id; p.next(); whnf = true; } + auto& data = p.get_ast(cmd_id).push(id); expr e; level_param_names ls; - std::tie(e, ls) = parse_local_expr(p, "_reduce"); + std::tie(e, ls) = parse_local_expr(p, data, "_reduce"); expr r; type_context_old ctx(p.env(), p.get_options(), metavar_context(), local_context(), transparency_mode::All); if (whnf) { @@ -229,38 +238,49 @@ environment reduce_cmd(parser & p) { bool eta = false; r = normalize(ctx, e, eta); } - auto out = p.mk_message(p.cmd_pos(), p.pos(), INFORMATION); + auto out = p.mk_message(data.m_start, p.pos(), INFORMATION); out.set_caption("reduce result") << r; out.report(); return p.env(); } -environment exit_cmd(parser & p) { - (p.mk_message(p.cmd_pos(), WARNING) << "using 'exit' to interrupt Lean").report(); +environment exit_cmd(parser & p, ast_id & cmd_id) { + (p.mk_message(p.get_ast(cmd_id).m_start, WARNING) << "using 'exit' to interrupt Lean").report(); throw interrupt_parser(); } -environment set_option_cmd(parser & p) { - auto id_kind = parse_option_name(p, "invalid set option, identifier (i.e., option name) expected"); +environment set_option_cmd(parser & p, ast_id & cmd_id) { + auto& data = p.get_ast(cmd_id); + auto id_kind = parse_option_name(p, data, "invalid set option, identifier (i.e., option name) expected"); name id = id_kind.first; option_kind k = id_kind.second; if (k == BoolOption) { - if (p.curr_is_token_or_id(get_true_tk())) + if (p.curr_is_token_or_id(get_true_tk())) { + data.push(p.new_ast("bool", p.pos(), get_true_tk()).m_id); p.set_option(id, true); - else if (p.curr_is_token_or_id(get_false_tk())) + } else if (p.curr_is_token_or_id(get_false_tk())) { + data.push(p.new_ast("bool", p.pos(), get_false_tk()).m_id); p.set_option(id, false); - else + } else { throw parser_error("invalid Boolean option value, 'true' or 'false' expected", p.pos()); + } p.next(); } else if (k == StringOption) { if (!p.curr_is_string()) throw parser_error("invalid option value, given option is not a string", p.pos()); - p.set_option(id, p.get_str_val()); + auto s = p.get_str_val(); + data.push(p.new_ast("string", p.pos(), s).m_id); + p.set_option(id, s); p.next(); } else if (k == DoubleOption) { - p.set_option(id, p.parse_double()); + ast_id d_id; double d; + std::tie(d_id, d) = p.parse_double(); + data.push(d_id); + p.set_option(id, d); } else if (k == UnsignedOption || k == IntOption) { - p.set_option(id, p.parse_small_nat()); + auto n = p.parse_small_nat(); + data.push(n.first); + p.set_option(id, n.second); } else { throw parser_error("invalid option value, 'true', 'false', string, integer or decimal value expected", p.pos()); } @@ -275,51 +295,71 @@ static void check_identifier(parser & p, environment const & env, name const & n } // open/export id (as id)? (id ...) (renaming id->id id->id) (hiding id ... id) -environment open_export_cmd(parser & p, bool open) { +environment open_export_cmd(parser & p, ast_id cmd_id, bool open) { environment env = p.env(); + auto& data = p.get_ast(cmd_id); while (true) { + auto& group = p.new_ast("group", p.pos()); + data.push(group.m_id); auto pos = p.pos(); - name ns = p.check_id_next("invalid 'open/export' command, identifier expected", - break_at_pos_exception::token_context::namespc); + ast_id ns_id; name ns; + std::tie(ns_id, ns) = p.check_id_next("invalid 'open/export' command, identifier expected", + break_at_pos_exception::token_context::namespc); + group.push(ns_id); optional real_ns = to_valid_namespace_name(env, ns); if (!real_ns) throw parser_error(sstream() << "invalid namespace name '" << ns << "'", pos); ns = *real_ns; - name as; + ast_id as_id = 0; name as; if (p.curr_is_token_or_id(get_as_tk())) { p.next(); - as = p.check_id_next("invalid 'open/export' command, identifier expected"); + std::tie(as_id, as) = p.check_id_next("invalid 'open/export' command, identifier expected"); } + group.push(as_id); buffer exception_names; buffer> renames; bool found_explicit = false; // Remark: we currently to not allow renaming and hiding of universe levels env = mark_namespace_as_open(env, ns); while (p.curr_is_token(get_lparen_tk())) { + auto pos = p.pos(); p.next(); if (p.curr_is_token_or_id(get_renaming_tk())) { + auto& renames_ast = p.new_ast(get_renaming_tk(), pos); + group.push(renames_ast.m_id); p.next(); while (p.curr_is_identifier()) { name from_id = p.get_name_val(); + auto& rename = p.new_ast(get_arrow_tk(), p.pos()); + renames_ast.push(rename.m_id); + rename.push(p.new_ast("ident", p.pos(), from_id).m_id); p.next(); p.check_token_next(get_arrow_tk(), "invalid 'open/export' command renaming, '->' expected"); - name to_id = p.check_id_next("invalid 'open/export' command renaming, identifier expected"); + ast_id to_ast; name to_id; + std::tie(to_ast, to_id) = p.check_id_next("invalid 'open/export' command renaming, identifier expected"); + rename.push(to_ast); check_identifier(p, env, ns, from_id); exception_names.push_back(from_id); renames.emplace_back(as+to_id, ns+from_id); } } else if (p.curr_is_token_or_id(get_hiding_tk())) { + auto& hides = p.new_ast(get_hiding_tk(), pos); + group.push(hides.m_id); p.next(); while (p.curr_is_identifier()) { name id = p.get_name_val(); + hides.push(p.new_ast("ident", p.pos(), id).m_id); p.next(); check_identifier(p, env, ns, id); exception_names.push_back(id); } } else if (p.curr_is_identifier()) { + auto& explicit_ast = p.new_ast("explicit", pos); + group.push(explicit_ast.m_id); found_explicit = true; while (p.curr_is_identifier()) { name id = p.get_name_val(); + explicit_ast.push(p.new_ast("ident", p.pos(), id).m_id); p.next(); check_identifier(p, env, ns, id); renames.emplace_back(as+id, ns+id); @@ -343,21 +383,23 @@ environment open_export_cmd(parser & p, bool open) { } return env; } -environment open_cmd(parser & p) { return open_export_cmd(p, true); } -static environment export_cmd(parser & p) { return open_export_cmd(p, false); } +environment open_cmd(parser & p, ast_id & cmd_id) { return open_export_cmd(p, cmd_id, true); } +static environment export_cmd(parser & p, ast_id & cmd_id) { return open_export_cmd(p, cmd_id, false); } -static environment local_cmd(parser & p, cmd_meta const & meta) { +static environment local_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { if (p.curr_is_token_or_id(get_attribute_tk())) { p.next(); - return local_attribute_cmd(p, meta); + return local_attribute_cmd(p, cmd_id, meta); } else { - return local_notation_cmd(p); + return local_notation_cmd(p, cmd_id); } } -static environment help_cmd(parser & p) { - auto rep = p.mk_message(p.cmd_pos(), INFORMATION); +static environment help_cmd(parser & p, ast_id & cmd_id) { + auto& data = p.get_ast(cmd_id); + auto rep = p.mk_message(data.m_start, INFORMATION); if (p.curr_is_token_or_id(get_options_tk())) { + data.push(p.new_ast(get_options_tk(), p.pos()).m_id); p.next(); rep.set_end_pos(p.pos()); auto decls = get_option_declarations(); @@ -366,6 +408,7 @@ static environment help_cmd(parser & p) { << opt.get_description() << " (default: " << opt.get_default_value() << ")\n"; }); } else if (p.curr_is_token_or_id(get_commands_tk())) { + data.push(p.new_ast(get_commands_tk(), p.pos()).m_id); p.next(); buffer ns; cmd_table const & cmds = p.cmds(); @@ -385,7 +428,7 @@ static environment help_cmd(parser & p) { return p.env(); } -static environment init_quotient_cmd(parser & p) { +static environment init_quotient_cmd(parser & p, ast_id &) { return module::declare_quotient(p.env()); } @@ -413,19 +456,20 @@ static expr convert_metavars(metavar_context & mctx, expr const & e) { return convert(e); } -static environment unify_cmd(parser & p) { +static environment unify_cmd(parser & p, ast_id & cmd_id) { transient_cmd_scope cmd_scope(p); environment const & env = p.env(); + auto& data = p.get_ast(cmd_id); expr e1; level_param_names ls1; - std::tie(e1, ls1) = parse_local_expr(p, "_unify"); + std::tie(e1, ls1) = parse_local_expr(p, data, "_unify"); p.check_token_next(get_comma_tk(), "invalid #unify command, proper usage \"#unify e1, e2\""); expr e2; level_param_names ls2; - std::tie(e2, ls2) = parse_local_expr(p, "_unify"); + std::tie(e2, ls2) = parse_local_expr(p, data, "_unify"); metavar_context mctx; local_context lctx; e1 = convert_metavars(mctx, e1); e2 = convert_metavars(mctx, e2); - auto rep = p.mk_message(p.cmd_pos(), p.pos(), INFORMATION); + auto rep = p.mk_message(data.m_start, p.pos(), INFORMATION); rep << e1 << " =?= " << e2 << "\n"; type_context_old ctx(env, p.get_options(), mctx, lctx, transparency_mode::Semireducible); bool success = ctx.is_def_eq(e1, e2); @@ -436,20 +480,23 @@ static environment unify_cmd(parser & p) { return env; } -static environment compile_cmd(parser & p) { +static environment compile_cmd(parser & p, ast_id & cmd_id) { auto pos = p.pos(); - name n = p.check_constant_next("invalid #compile command, constant expected"); + ast_id id = 0; name n; + std::tie(id, n) = p.check_constant_next("invalid #compile command, constant expected"); + p.get_ast(cmd_id).push(id); declaration d = p.env().get(n); if (!d.is_definition()) throw parser_error("invalid #compile command, declaration is not a definition", pos); return vm_compile(p.env(), p.get_options(), d); } -static environment eval_cmd(parser & p) { +static environment eval_cmd(parser & p, ast_id & cmd_id) { transient_cmd_scope cmd_scope(p); auto pos = p.pos(); + auto& data = p.get_ast(cmd_id); expr e; level_param_names ls; - std::tie(e, ls) = parse_local_expr(p, "_eval", /* relaxed */ false); + std::tie(e, ls) = parse_local_expr(p, data, "_eval", /* relaxed */ false); if (has_synthetic_sorry(e)) return p.env(); @@ -480,9 +527,9 @@ static environment eval_cmd(parser & p) { name fn_name = "_main"; auto new_env = compile_expr(p.env(), p.get_options(), fn_name, ls, type, e, pos); - auto out = p.mk_message(p.cmd_pos(), p.pos(), INFORMATION); + auto out = p.mk_message(data.m_start, p.pos(), INFORMATION); out.set_caption("eval result"); - scope_traces_as_messages scope_traces(p.get_stream_name(), p.cmd_pos()); + scope_traces_as_messages scope_traces(p.get_stream_name(), data.m_start); bool should_report = false; auto run = [&] { @@ -492,7 +539,7 @@ static environment eval_cmd(parser & p) { auto r = fn.invoke_fn(); should_report = true; if (!has_repr_inst) { - (p.mk_message(p.cmd_pos(), WARNING) << "result type does not have an instance of type class 'has_repr', dumping internal representation").report(); + (p.mk_message(data.m_start, WARNING) << "result type does not have an instance of type class 'has_repr', dumping internal representation").report(); } if (is_constant(fn.get_type(), get_string_name())) { out << to_string(r); @@ -501,7 +548,7 @@ static environment eval_cmd(parser & p) { } } } catch (throwable & t) { - p.mk_message(p.cmd_pos(), ERROR).set_exception(t).report(); + p.mk_message(data.m_start, ERROR).set_exception(t).report(); } if (fn.get_profiler().enabled()) { if (fn.get_profiler().get_snapshots().display("#eval", p.get_options(), out.get_text_stream().get_stream())) @@ -544,24 +591,28 @@ struct declare_trace_modification : public modification { } }; -environment declare_trace_cmd(parser & p) { - name cls = p.check_id_next("invalid declare_trace command, identifier expected"); - return module::add_and_perform(p.env(), std::make_shared(cls)); +environment declare_trace_cmd(parser & p, ast_id & cmd_id) { + auto cls = p.check_id_next("invalid declare_trace command, identifier expected"); + p.get_ast(cmd_id).push(cls.first); + return module::add_and_perform(p.env(), std::make_shared(cls.second)); } -environment add_key_equivalence_cmd(parser & p) { - name h1 = p.check_constant_next("invalid add_key_equivalence command, constant expected"); - name h2 = p.check_constant_next("invalid add_key_equivalence command, constant expected"); +environment add_key_equivalence_cmd(parser & p, ast_id & cmd_id) { + ast_id id1, id2; name h1, h2; + std::tie(id1, h1) = p.check_constant_next("invalid add_key_equivalence command, constant expected"); + std::tie(id2, h2) = p.check_constant_next("invalid add_key_equivalence command, constant expected"); + p.get_ast(cmd_id).push(id1).push(id2); return add_key_equivalence(p.env(), h1, h2); } -static environment run_command_cmd(parser & p) { +static environment run_command_cmd(parser & p, ast_id & cmd_id) { transient_cmd_scope cmd_scope(p); module::scope_pos_info scope_pos(p.pos()); environment env = p.env(); options opts = p.get_options(); metavar_context mctx; expr tactic = p.parse_expr(); + p.get_ast(cmd_id).push(p.get_id(tactic)); expr try_triv = mk_app(mk_constant(get_tactic_try_name()), mk_constant(get_tactic_triv_name())); tactic = mk_app(mk_constant(get_has_bind_and_then_name()), tactic, try_triv); tactic = mk_typed_expr(mk_tactic_unit(), tactic); @@ -574,19 +625,21 @@ static environment run_command_cmd(parser & p) { return env; } -environment import_cmd(parser & p) { - throw parser_error("invalid 'import' command, it must be used in the beginning of the file", p.cmd_pos()); +environment import_cmd(parser & p, ast_id & cmd_id) { + throw parser_error("invalid 'import' command, it must be used in the beginning of the file", p.get_ast(cmd_id).m_start); } -environment hide_cmd(parser & p) { +environment hide_cmd(parser & p, ast_id & cmd_id) { buffer ids; + auto& data = p.get_ast(cmd_id); while (p.curr_is_identifier()) { name id = p.get_name_val(); + data.push(p.new_ast("ident", p.pos(), id).m_id); p.next(); ids.push_back(id); } if (ids.empty()) - throw parser_error("invalid 'hide' command, identifier expected", p.cmd_pos()); + throw parser_error("invalid 'hide' command, identifier expected", data.m_start); environment new_env = p.env(); for (name id : ids) { if (get_expr_aliases(new_env, id)) { @@ -594,7 +647,7 @@ environment hide_cmd(parser & p) { } else { /* TODO(Leo): check if `id` is a declaration and hide it too. */ throw parser_error(sstream() << "invalid 'hide' command, '" << id << "' is not an alias", - p.cmd_pos()); + data.m_start); } } return new_env; diff --git a/src/frontends/lean/builtin_exprs.cpp b/src/frontends/lean/builtin_exprs.cpp index 9610042d09..0d1624400e 100644 --- a/src/frontends/lean/builtin_exprs.cpp +++ b/src/frontends/lean/builtin_exprs.cpp @@ -50,76 +50,107 @@ bool get_parser_checkpoint_have(options const & opts) { using namespace notation; // NOLINT +expr mk_annotation_with_pos(parser & p, name const & a, expr const & e, pos_info const & pos) { + expr r = mk_annotation(a, e); + r.set_tag(nulltag); // mk_annotation copies e's tag + return p.save_pos(r, pos); +} + static name * g_no_universe_annotation = nullptr; bool is_sort_wo_universe(expr const & e) { return is_annotation(e, *g_no_universe_annotation); } -expr mk_sort_wo_universe(parser & p, pos_info const & pos, bool is_type) { - expr r = p.save_pos(mk_sort(is_type ? mk_level_one() : mk_level_zero()), pos); - return p.save_pos(mk_annotation(*g_no_universe_annotation, r), pos); +expr mk_sort_wo_universe(parser & p, ast_data & data, bool is_type) { + level l = is_type ? mk_level_one() : mk_level_zero(); + expr r = p.save_pos((scoped_expr_caching(false), mk_sort(l)), data.m_start); + p.set_ast_pexpr(data.m_id, r); + r = mk_annotation_with_pos(p, *g_no_universe_annotation, r, data.m_start); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_Type(parser & p, unsigned, expr const *, pos_info const & pos) { + auto& data = p.new_ast("Type", pos); if (p.curr_is_token(get_llevel_curly_tk())) { p.next(); - level l = mk_succ(p.parse_level()); + ast_id id; level l; + std::tie(id, l) = p.parse_level(); p.check_token_next(get_rcurly_tk(), "invalid Type expression, '}' expected"); - return p.save_pos(mk_sort(l), pos); + l = mk_succ(l); + expr r = p.save_pos((scoped_expr_caching(false), mk_sort(l)), pos); + p.set_ast_pexpr(data.push(id).m_id, r); + return r; } else { - return mk_sort_wo_universe(p, pos, true); + return mk_sort_wo_universe(p, data.push(0), true); } } static expr parse_Sort(parser & p, unsigned, expr const *, pos_info const & pos) { + auto& data = p.new_ast("Sort", pos); if (p.curr_is_token(get_llevel_curly_tk())) { p.next(); - level l = p.parse_level(); + ast_id id; level l; + std::tie(id, l) = p.parse_level(); p.check_token_next(get_rcurly_tk(), "invalid Sort expression, '}' expected"); - return p.save_pos(mk_sort(l), pos); + expr r = p.save_pos((scoped_expr_caching(false), mk_sort(l)), pos); + p.set_ast_pexpr(data.push(id).m_id, r); + return r; } else { - return mk_sort_wo_universe(p, pos, false); + return mk_sort_wo_universe(p, data.push(0), false); } } static expr parse_Type_star(parser & p, unsigned, expr const *, pos_info const & pos) { - return p.save_pos(mk_sort(mk_succ(mk_level_placeholder())), pos); + level l = mk_succ(mk_level_placeholder()); + expr r = p.save_pos((scoped_expr_caching(false), mk_sort(l)), pos); + p.set_ast_pexpr(p.new_ast("Type*", pos).m_id, r); + return r; } static expr parse_Sort_star(parser & p, unsigned, expr const *, pos_info const & pos) { - return p.save_pos(mk_sort(mk_level_placeholder()), pos); + level l = mk_level_placeholder(); + expr r = p.save_pos((scoped_expr_caching(false), mk_sort(l)), pos); + p.set_ast_pexpr(p.new_ast("Sort*", pos).m_id, r); + return r; } static name * g_let_match_name = nullptr; -static expr parse_let(parser & p, pos_info const & pos, bool in_do_block); -static expr parse_do(parser & p, bool has_braces); -static expr parse_let_body(parser & p, pos_info const & pos, bool in_do_block) { +static pair parse_let(parser & p, pos_info const & pos, ast_data & group, bool in_do_block); +static expr parse_do(parser & p, ast_data & group, bool has_braces); +static pair parse_let_body(parser & p, pos_info const & pos, ast_id last, ast_data & group, bool in_do_block) { if (in_do_block) { if (p.curr_is_token(get_in_tk())) { p.next(); - return p.parse_expr(); + ast_id args = p.new_ast("args", pos).push(last).m_id; + expr e = p.parse_expr(); + ast_id let = p.new_ast(get_let_tk(), pos).push(args).push(p.get_id(e)).m_id; + group.push(p.new_ast("eval", pos).push(let).m_id); + return {0, e}; } else { p.check_token_next(get_comma_tk(), "invalid 'do' block 'let' declaration, ',' or 'in' expected"); + group.push(p.new_ast(get_let_tk(), pos).push(last).m_id); if (p.curr_is_token(get_let_tk())) { p.next(); - return parse_let(p, pos, in_do_block); + return parse_let(p, pos, group, in_do_block); } else { - return parse_do(p, false); + return {0, parse_do(p, group, false)}; } } } else { + group.push(last); if (p.curr_is_token(get_comma_tk())) { p.next(); - return parse_let(p, pos, in_do_block); + return parse_let(p, pos, group, in_do_block); } else if (p.curr_is_token(get_in_tk())) { p.next(); - return p.parse_expr(); } else { p.maybe_throw_error({"invalid let declaration, 'in' or ',' expected", p.pos()}); - return p.parse_expr(); } + expr e = p.parse_expr(); + return {p.new_ast(get_let_tk(), pos).push(group.m_id).push(p.get_id(e)).m_id, e}; } } @@ -137,71 +168,91 @@ static expr mk_typed_expr_distrib_choice(parser & p, expr const & type, expr con } } -static expr parse_let(parser & p, pos_info const & pos, bool in_do_block) { +static pair parse_let(parser & p, pos_info const & pos, ast_data & group, bool in_do_block) { parser::local_scope scope1(p); - if (!in_do_block && p.parse_local_notation_decl()) { - return parse_let_body(p, pos, in_do_block); + ast_id nota; + if (!in_do_block && (nota = p.parse_local_notation_decl())) { + return parse_let_body(p, pos, nota, group, in_do_block); } else if (p.curr_is_identifier()) { auto id_pos = p.pos(); - name id = p.check_atomic_id_next("invalid let declaration, atomic identifier expected"); + ast_id id_ast; name id; + std::tie(id_ast, id) = p.check_atomic_id_next("invalid let declaration, atomic identifier expected"); + auto& var = p.new_ast("var", id_pos).push(id_ast); expr type; expr value; if (p.curr_is_token(get_assign_tk())) { p.next(); type = p.save_pos(mk_expr_placeholder(), id_pos); value = p.parse_expr(); + var.push(0).push(0).push(p.get_id(value)); } else if (p.curr_is_token(get_colon_tk())) { p.next(); type = p.parse_expr(); p.check_token_next(get_assign_tk(), "invalid declaration, ':=' expected"); value = p.parse_expr(); + var.push(0).push(p.get_id(type)).push(p.get_id(value)); } else { parser::local_scope scope2(p); buffer ps; unsigned rbp = 0; - auto lenv = p.parse_binders(ps, rbp); + auto& bis = p.new_ast("binders", p.pos()); + auto lenv = p.parse_binders(&bis, ps, rbp); + var.push(bis.m_id); if (p.curr_is_token(get_colon_tk())) { p.next(); type = p.parse_scoped_expr(ps, lenv); + var.push(p.get_id(type)); type = Pi(ps, type, p); } else { type = p.save_pos(mk_expr_placeholder(), id_pos); + var.push(0); } p.check_token_next(get_assign_tk(), "invalid let declaration, ':=' expected"); value = p.parse_scoped_expr(ps, lenv); + var.push(p.get_id(value)); value = Fun(ps, value, p); } expr x = p.save_pos(mk_local(id, type), id_pos); p.add_local_expr(id, x); - expr b = parse_let_body(p, pos, in_do_block); - return p.save_pos(mk_let(id, type, value, abstract_local(b, x)), pos); + auto r = parse_let_body(p, pos, var.m_id, group, in_do_block); + r.second = p.save_pos(mk_let(id, type, value, abstract_local(r.second, x)), pos); + return r; } else { buffer new_locals; + auto& var = p.new_ast("pat", p.pos()); expr lhs = p.parse_pattern(new_locals); p.check_token_next(get_assign_tk(), "invalid let declaration, ':=' expected"); expr value = p.parse_expr(); + var.push(p.get_id(lhs)).push(p.get_id(value)); for (expr const & l : new_locals) p.add_local(l); - expr body = parse_let_body(p, pos, in_do_block); + ast_id ast; expr body; + std::tie(ast, body) = parse_let_body(p, pos, var.m_id, group, in_do_block); match_definition_scope match_scope(p.env()); expr fn = p.save_pos(mk_local(p.next_name(), *g_let_match_name, mk_expr_placeholder(), mk_rec_info(true)), pos); expr eqn = Fun(fn, Fun(new_locals, p.save_pos(mk_equation(p.rec_save_pos(mk_app(fn, lhs), pos), body), pos), p), p); equations_header h = mk_match_header(match_scope.get_name(), match_scope.get_actual_name()); expr eqns = p.save_pos(mk_equations(h, 1, &eqn), pos); - return p.save_pos(mk_app(eqns, value), pos); + return {ast, p.save_pos(mk_app(eqns, value), pos)}; } } static expr parse_let_expr(parser & p, unsigned, expr const *, pos_info const & pos) { bool in_do_block = false; - return parse_let(p, pos, in_do_block); + auto& group = p.new_ast("args", p.pos()); + ast_id id; expr r; + std::tie(id, r) = parse_let(p, pos, group, in_do_block); + p.set_ast_pexpr(id, r); + return r; } static name * g_do_match_name = nullptr; -static std::tuple, expr, expr, optional> parse_do_action(parser & p, buffer & new_locals) { +static std::tuple, expr, expr, optional> +parse_do_action(parser & p, ast_data & parent, buffer & new_locals) { auto lhs_pos = p.pos(); optional lhs = some(p.parse_pattern_or_expr()); + ast_id lhs_id = p.get_id(*lhs); expr type, curr; optional else_case; if (p.curr_is_token(get_colon_tk())) { @@ -219,26 +270,33 @@ static std::tuple, expr, expr, optional> parse_do_action(pa new_locals.push_back(*lhs); p.check_token_next(get_larrow_tk(), "invalid 'do' block, '←' expected"); curr = p.parse_expr(); + parent.push(p.new_ast(get_larrow_tk(), lhs_pos) + .push(lhs_id).push(p.get_id(type)).push(p.get_id(curr)).push(0).m_id); } else if (p.curr_is_token(get_larrow_tk())) { p.next(); type = p.save_pos(mk_expr_placeholder(), lhs_pos); bool skip_main_fn = false; lhs = p.patexpr_to_pattern(*lhs, skip_main_fn, new_locals); curr = p.parse_expr(); + ast_id else_id = 0; if (p.curr_is_token(get_bar_tk())) { p.next(); else_case = p.parse_expr(); + else_id = p.get_id(*else_case); } + parent.push(p.new_ast(get_larrow_tk(), lhs_pos) + .push(lhs_id).push(0).push(p.get_id(curr)).push(else_id).m_id); } else { curr = p.patexpr_to_expr(*lhs); type = p.save_pos(mk_expr_placeholder(), lhs_pos); lhs = none_expr(); + parent.push(p.new_ast("eval", lhs_pos).push(lhs_id).m_id); } return std::make_tuple(lhs, type, curr, else_case); } static expr mk_bind_fn(parser & p) { - return mk_no_info(p.id_to_expr("bind", pos_info {}, /* resolve_only */ true)); + return mk_no_info(p.id_to_expr("bind", p.new_ast("bind", {}), /* resolve_only */ true)); } static name * g_do_failure_eq = nullptr; @@ -256,7 +314,7 @@ bool is_do_failure_eq(expr const & e) { return is_annotation(equation_rhs(it), *g_do_failure_eq); } -static expr parse_do(parser & p, bool has_braces) { +static expr parse_do(parser & p, ast_data & group, bool has_braces) { parser::local_scope scope(p); buffer es; buffer ps; @@ -268,7 +326,7 @@ static expr parse_do(parser & p, bool has_braces) { auto pos = p.pos(); p.next(); bool in_do_block = true; - es.push_back(parse_let(p, pos, in_do_block)); + es.push_back(parse_let(p, pos, group, in_do_block).second); break; } else { auto lhs_pos = p.pos(); @@ -276,7 +334,7 @@ static expr parse_do(parser & p, bool has_braces) { buffer new_locals; optional lhs, else_case; expr type, curr; - std::tie(lhs, type, curr, else_case) = parse_do_action(p, new_locals); + std::tie(lhs, type, curr, else_case) = parse_do_action(p, group, new_locals); es.push_back(curr); if (p.curr_is_token(get_comma_tk())) { p.next(); @@ -358,114 +416,142 @@ static expr parse_do(parser & p, bool has_braces) { return r; } -static expr parse_do_expr(parser & p, unsigned, expr const *, pos_info const &) { +static expr parse_do_expr(parser & p, unsigned, expr const *, pos_info const & pos) { bool has_braces = false; + auto& data = p.new_ast("do", pos); if (p.curr_is_token(get_lcurly_tk())) { has_braces = true; + data.m_value = "{"; p.next(); } - return parse_do(p, has_braces); + expr r = parse_do(p, data, has_braces); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_unit(parser & p, unsigned, expr const *, pos_info const & pos) { - return p.save_pos(mk_constant(get_unit_star_name()), pos); + auto e = p.save_pos(mk_constant(get_unit_star_name()), pos); + p.set_ast_pexpr(p.new_ast("()", pos).m_id, e); + return e; } -static expr parse_proof(parser & p); - -static expr parse_proof(parser & p) { +static pair parse_proof(parser & p) { + expr e; if (p.curr_is_token(get_from_tk())) { + auto& data = p.new_ast(get_from_tk(), p.pos()); // parse: 'from' expr p.next(); - return p.parse_expr(); + e = p.parse_expr(); + data.push(p.get_id(e)); + return {data.m_id, e}; } else if (p.curr_is_token(get_begin_tk())) { auto pos = p.pos(); - return parse_begin_end_expr(p, pos); + e = parse_begin_end_expr(p, pos); } else if (p.curr_is_token(get_lcurly_tk())) { auto pos = p.pos(); - return parse_curly_begin_end_expr(p, pos); + e = parse_curly_begin_end_expr(p, pos); } else if (p.curr_is_token(get_by_tk())) { auto pos = p.pos(); - return parse_by(p, 0, nullptr, pos); + e = parse_by(p, 0, nullptr, pos); } else { - return p.parser_error_or_expr({"invalid expression, 'by', 'begin', '{', or 'from' expected", p.pos()}); + e = p.parser_error_or_expr({"invalid expression, 'by', 'begin', '{', or 'from' expected", p.pos()}); } + return {p.get_id(e), e}; } static expr parse_have(parser & p, unsigned, expr const *, pos_info const & pos) { - auto id_pos = p.pos(); name id; expr prop; + auto& data = p.new_ast("have", pos); if (p.curr_is_identifier()) { id = p.get_name_val(); + auto& id_ast = p.new_ast("ident", p.pos(), id); p.next(); if (p.curr_is_token(get_colon_tk())) { p.next(); prop = p.parse_expr(); + data.push(id_ast.m_id).push(p.get_id(prop)); } else { - expr left = p.id_to_expr(id, id_pos); + expr left = p.id_to_expr(id, id_ast); id = get_this_tk(); unsigned rbp = 0; while (rbp < p.curr_lbp()) { left = p.parse_led(left); } prop = left; + data.push(0).push(p.get_id(prop)); } } else { id = get_this_tk(); prop = p.parse_expr(); + data.push(0).push(p.get_id(prop)); } expr proof; if (p.curr_is_token(get_assign_tk())) { + auto& ast = p.new_ast(get_assign_tk(), p.pos()); p.next(); proof = p.parse_expr(); + data.push(ast.push(p.get_id(proof)).m_id); } else { p.check_token_next(get_comma_tk(), "invalid 'have' declaration, ',' expected"); - proof = parse_proof(p); + ast_id id; + std::tie(id, proof) = parse_proof(p); + data.push(id); } p.check_token_next(get_comma_tk(), "invalid 'have' declaration, ',' expected"); parser::local_scope scope(p); expr l = p.save_pos(mk_local(id, prop), pos); p.add_local(l); expr body = p.parse_expr(); + data.push(p.get_id(body)); body = abstract(body, l); if (get_parser_checkpoint_have(p.get_options())) body = mk_checkpoint_annotation(body); expr r = p.save_pos(mk_have_annotation(p.save_pos(mk_lambda(id, prop, body), pos)), pos); - return p.mk_app(r, proof, pos); + r = p.mk_app(r, proof, pos); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_show(parser & p, unsigned, expr const *, pos_info const & pos) { expr prop = p.parse_expr(); p.check_token_next(get_comma_tk(), "invalid 'show' declaration, ',' expected"); - expr proof = parse_proof(p); + ast_id id; expr proof; + std::tie(id, proof) = parse_proof(p); expr b = p.save_pos(mk_lambda(get_this_tk(), prop, Var(0)), pos); expr r = p.mk_app(b, proof, pos); - return p.save_pos(mk_show_annotation(r), pos); + r = p.save_pos(mk_show_annotation(r), pos); + p.set_ast_pexpr(p.new_ast("show", pos).push(p.get_id(prop)).push(id).m_id, r); + return r; } static expr parse_suffices(parser & p, unsigned, expr const *, pos_info const & pos) { auto prop_pos = p.pos(); name id; expr from; + auto& data = p.new_ast("suffices", pos); if (p.curr_is_identifier()) { id = p.get_name_val(); + auto& id_ast = p.new_ast("ident", prop_pos, id); p.next(); if (p.curr_is_token(get_colon_tk())) { p.next(); from = p.parse_expr(); + data.push(id_ast.m_id).push(p.get_id(from)); } else { - expr left = p.id_to_expr(id, prop_pos); + expr left = p.id_to_expr(id, id_ast); id = get_this_tk(); unsigned rbp = 0; while (rbp < p.curr_lbp()) { left = p.parse_led(left); } from = left; + data.push(0).push(p.get_id(from)); } } else { id = get_this_tk(); from = p.parse_expr(); + data.push(0).push(p.get_id(from)); } expr local = p.save_pos(mk_local(id, from), prop_pos); p.check_token_next(get_comma_tk(), "invalid 'suffices' declaration, ',' expected"); @@ -473,29 +559,35 @@ static expr parse_suffices(parser & p, unsigned, expr const *, pos_info const & { parser::local_scope scope(p); p.add_local(local); - body = parse_proof(p); + ast_id id; + std::tie(id, body) = parse_proof(p); + data.push(id); } expr proof = p.save_pos(Fun(local, body, p), pos); p.check_token_next(get_comma_tk(), "invalid 'suffices' declaration, ',' expected"); expr rest = p.parse_expr(); + data.push(p.get_id(rest)); expr r = p.mk_app(proof, rest, pos); - return p.save_pos(mk_suffices_annotation(r), pos); + r = p.save_pos(mk_suffices_annotation(r), pos); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr * g_not = nullptr; static unsigned g_then_else_prec = 0; -static expr parse_ite(parser & p, expr const & c, pos_info const & pos) { +static expr parse_ite(parser & p, ast_data & parent, expr const & c, pos_info const & pos) { if (!p.env().find(get_ite_name())) throw parser_error("invalid use of 'if-then-else' expression, environment does not contain 'ite' definition", pos); p.check_token_next(get_then_tk(), "invalid 'if-then-else' expression, 'then' expected"); expr t = p.parse_expr(g_then_else_prec); p.check_token_next(get_else_tk(), "invalid 'if-then-else' expression, 'else' expected"); expr e = p.parse_expr(g_then_else_prec); + parent.push(p.get_id(t)).push(p.get_id(e)); return p.save_pos(mk_app(mk_constant(get_ite_name()), c, t, e), pos); } -static expr parse_dite(parser & p, name const & H_name, expr const & c, pos_info const & pos) { +static expr parse_dite(parser & p, ast_data & parent, name const & H_name, expr const & c, pos_info const & pos) { p.check_token_next(get_then_tk(), "invalid 'if-then-else' expression, 'then' expected"); expr t, e; { @@ -503,7 +595,9 @@ static expr parse_dite(parser & p, name const & H_name, expr const & c, pos_info expr H = mk_local(H_name, c); p.add_local(H); auto pos = p.pos(); - t = p.save_pos(Fun(H, p.parse_expr(g_then_else_prec), p), pos); + t = p.parse_expr(g_then_else_prec); + parent.push(p.get_id(t)); + t = p.save_pos(Fun(H, t, p), pos); } p.check_token_next(get_else_tk(), "invalid 'if-then-else' expression, 'else' expected"); { @@ -511,30 +605,34 @@ static expr parse_dite(parser & p, name const & H_name, expr const & c, pos_info expr H = mk_local(H_name, mk_app(*g_not, c)); p.add_local(H); auto pos = p.pos(); - e = p.save_pos(Fun(H, p.parse_expr(g_then_else_prec), p), pos); + e = p.parse_expr(g_then_else_prec); + parent.push(p.get_id(e)); + e = p.save_pos(Fun(H, e, p), pos); } return p.save_pos(mk_app(p.save_pos(mk_constant(get_dite_name()), pos), c, t, e), pos); } static expr parse_if_then_else(parser & p, unsigned, expr const *, pos_info const & pos) { - pair, expr> ie = p.parse_qualified_expr(); - if (ie.first) - return parse_dite(p, *ie.first, ie.second, pos); - else - return parse_ite(p, ie.second, pos); + ast_id id; optional hyp; expr cond; + std::tie(id, hyp, cond) = p.parse_qualified_expr(); + auto& data = p.new_ast("if", pos).push(id).push(p.get_id(cond)); + auto r = hyp ? parse_dite(p, data, *hyp, cond, pos) : parse_ite(p, data, cond, pos); + p.set_ast_pexpr(data.m_id, r); + return r; } -static expr parse_calc_expr(parser & p, unsigned, expr const *, pos_info const &) { - return parse_calc(p); +static expr parse_calc_expr(parser & p, unsigned, expr const *, pos_info const & pos) { + return parse_calc(p, pos); } static expr parse_explicit_core(parser & p, pos_info const & pos, bool partial) { + auto sym = partial ? "@@" : "@"; if (!p.curr_is_identifier()) - return p.parser_error_or_expr({sstream() << "invalid '" << (partial ? "@@" : "@") << "', identifier expected", p.pos()}); + return p.parser_error_or_expr({sstream() << "invalid '" << sym << "', identifier expected", p.pos()}); expr fn = p.parse_id(/* allow_field_notation */ false); if (is_choice(fn)) { sstream s; - s << "invalid '" << (partial ? "@@" : "@") << "', function is overloaded, use fully qualified names (overloads: "; + s << "invalid '" << sym << "', function is overloaded, use fully qualified names (overloads: "; for (unsigned i = 0; i < get_num_choices(fn); i++) { if (i > 0) s << ", "; expr const & c = get_choice(fn, i); @@ -548,12 +646,15 @@ static expr parse_explicit_core(parser & p, pos_info const & pos, bool partial) s << ")"; return p.parser_error_or_expr({s, pos}); } else if (!is_as_atomic(fn) && !is_constant(fn) && !is_local(fn)) { - return p.parser_error_or_expr({sstream() << "invalid '" << (partial ? "@@" : "@") << "', function must be a constant or variable", pos}); + return p.parser_error_or_expr({sstream() << "invalid '" << sym << "', function must be a constant or variable", pos}); } + expr r; if (partial) - return p.save_pos(mk_partial_explicit(fn), pos); + r = p.save_pos(mk_partial_explicit(fn), pos); else - return p.save_pos(mk_explicit(fn), pos); + r = p.save_pos(mk_explicit(fn), pos); + p.set_ast_pexpr(p.new_ast(sym, pos).push(p.get_id(fn)).m_id, r); + return r; } static expr parse_explicit_expr(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -565,11 +666,15 @@ static expr parse_partial_explicit_expr(parser & p, unsigned, expr const *, pos_ } static expr parse_sorry(parser & p, unsigned, expr const *, pos_info const & pos) { - return p.mk_sorry(pos); + auto r = p.mk_sorry(pos); + p.set_ast_pexpr(p.new_ast("sorry", pos).m_id, r); + return r; } static expr parse_pattern(parser & p, unsigned, expr const * args, pos_info const & pos) { - return p.save_pos(mk_pattern_hint(args[0]), pos); + auto r = p.save_pos(mk_pattern_hint(args[0]), pos); + p.set_ast_pexpr(p.new_ast("(:", pos).push(p.get_id(args[0])).m_id, r); + return r; } static expr parse_lazy_quoted_pexpr(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -578,13 +683,21 @@ static expr parse_lazy_quoted_pexpr(parser & p, unsigned, expr const *, pos_info parser::quote_scope scope1(p, true); restore_decl_meta_scope scope2; expr e = p.parse_expr(); + auto& data = p.new_ast("```()", pos); if (p.curr_is_token(get_colon_tk())) { p.next(); expr t = p.parse_expr(); + ast_id id = p.new_ast(get_colon_tk(), p.pos_of(e)).push(p.get_id(e)).push(p.get_id(t)).m_id; e = mk_typed_expr_distrib_choice(p, t, e, pos); + p.set_ast_pexpr(id, e); + data.push(id); + } else { + data.push(p.get_id(e)); } p.check_token_next(get_rparen_tk(), "invalid quoted expression, `)` expected"); - return p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ false), pos); + auto r = p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ false), pos); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_quoted_pexpr(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -593,18 +706,27 @@ static expr parse_quoted_pexpr(parser & p, unsigned, expr const *, pos_info cons parser::quote_scope scope1(p, true, id_behavior::ErrorIfUndef); restore_decl_meta_scope scope2; expr e = p.parse_expr(); + auto& data = p.new_ast("``()", pos); if (p.curr_is_token(get_colon_tk())) { p.next(); expr t = p.parse_expr(); + ast_id id = p.new_ast(get_colon_tk(), p.pos_of(e)).push(p.get_id(e)).push(p.get_id(t)).m_id; e = mk_typed_expr_distrib_choice(p, t, e, pos); + p.set_ast_pexpr(id, e); + data.push(id); + } else { + data.push(p.get_id(e)); } p.check_token_next(get_rparen_tk(), "invalid quoted expression, `)` expected"); - return p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ true), pos); + auto r = p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ true), pos); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_quoted_expr(parser & p, unsigned, expr const *, pos_info const & pos) { if (p.in_quote()) return p.parser_error_or_expr({"invalid nested quoted expression", pos}); + auto& data = p.new_ast("`()", pos); expr e; { parser::quote_scope scope1(p, true, id_behavior::ErrorIfUndef); @@ -613,11 +735,18 @@ static expr parse_quoted_expr(parser & p, unsigned, expr const *, pos_info const if (p.curr_is_token(get_colon_tk())) { p.next(); expr t = p.parse_expr(); + ast_id id = p.new_ast(get_colon_tk(), p.pos_of(e)).push(p.get_id(e)).push(p.get_id(t)).m_id; e = mk_typed_expr_distrib_choice(p, t, e, pos); + p.set_ast_pexpr(id, e); + data.push(id); + } else { + data.push(p.get_id(e)); } p.check_token_next(get_rparen_tk(), "invalid quoted expression, `)` expected"); } - return p.save_pos(mk_unelaborated_expr_quote(e), pos); + auto r = p.save_pos(mk_unelaborated_expr_quote(e), pos); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_antiquote_expr(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -625,17 +754,22 @@ static expr parse_antiquote_expr(parser & p, unsigned, expr const *, pos_info co return p.parser_error_or_expr({"invalid antiquotation, occurs outside of quoted expressions", pos}); parser::quote_scope scope(p, false); expr e = p.parse_expr(get_max_prec()); - return p.save_pos(mk_antiquote(e), pos); + ast_id id = p.new_ast("%%", pos).push(p.get_id(e)).m_id; + e = p.save_pos(mk_antiquote(e), pos); + p.set_ast_pexpr(id, e); + return e; } static expr parse_quoted_name(parser & p, unsigned, expr const *, pos_info const & pos) { bool resolve = false; name id; + auto& data = p.new_ast("`", pos); if (p.curr_is_token(get_placeholder_tk())) { p.next(); id = "_"; } else { if (p.curr_is_token(get_backtick_tk())) { + data.m_type = "``"; p.next(); resolve = true; } @@ -645,13 +779,14 @@ static expr parse_quoted_name(parser & p, unsigned, expr const *, pos_info const id = p.get_token_info().token(); p.next(); } else { - id = p.check_id_next("invalid quoted name, identifier expected"); + id = p.check_id_next("invalid quoted name, identifier expected").second; } } + data.m_value = id; if (resolve) { parser::error_if_undef_scope scope(p); bool resolve_only = true; - expr e = p.id_to_expr(id, pos, resolve_only); + expr e = p.id_to_expr(id, p.new_ast("ident", pos, id), resolve_only); if (is_constant(e)) { id = const_name(e); } else if (is_local(e)) { @@ -670,7 +805,9 @@ static expr parse_quoted_name(parser & p, unsigned, expr const *, pos_info const } lean_assert(id.is_string()); expr e = quote(id); - return p.rec_save_pos(e, pos); + e = p.rec_save_pos(e, pos); + p.set_ast_pexpr(data.m_id, e); + return e; } static name * g_anonymous_constructor = nullptr; @@ -686,8 +823,10 @@ expr const & get_anonymous_constructor_arg(expr const & e) { static expr parse_constructor_core(parser & p, pos_info const & pos) { buffer args; + auto& data = p.new_ast(get_langle_tk(), pos); while (!p.curr_is_token(get_rangle_tk())) { args.push_back(p.parse_expr()); + data.push(p.get_id(args.back())); if (p.curr_is_token(get_comma_tk())) { p.next(); } else { @@ -696,19 +835,21 @@ static expr parse_constructor_core(parser & p, pos_info const & pos) { } p.check_token_next(get_rangle_tk(), "invalid constructor, `⟩` expected"); expr fn = p.save_pos(mk_expr_placeholder(), pos); - return p.save_pos(mk_anonymous_constructor(p.save_pos(mk_app(fn, args), pos)), pos); + expr e = p.save_pos(mk_anonymous_constructor(p.save_pos(mk_app(fn, args), pos)), pos); + p.set_ast_pexpr(data.m_id, e); + return e; } static expr parse_constructor(parser & p, unsigned, expr const *, pos_info const & pos) { return parse_constructor_core(p, pos); } -static expr parse_lambda_core(parser & p, pos_info const & pos); +static expr parse_lambda_core(parser & p, ast_data & fun, ast_data & bis, pos_info const & pos); -static expr parse_lambda_binder(parser & p, pos_info const & pos) { +static expr parse_lambda_binder(parser & p, ast_data & fun, ast_data & bis, pos_info const & pos) { parser::local_scope scope1(p); buffer locals; - auto new_env = p.parse_binders(locals, 0); + auto new_env = p.parse_binders(&bis, locals, 0); for (expr const & local : locals) p.add_local(local); parser::local_scope scope2(p, new_env); @@ -716,11 +857,13 @@ static expr parse_lambda_binder(parser & p, pos_info const & pos) { if (p.curr_is_token(get_comma_tk())) { p.next(); body = p.parse_expr(); + fun.push(p.get_id(body)); } else if (p.curr_is_token(get_langle_tk())) { - body = parse_lambda_core(p, pos); + body = parse_lambda_core(p, fun, bis, pos); } else { p.maybe_throw_error({"invalid lambda expression, ',' or '⟨' expected", p.pos()}); body = p.parse_expr(); + fun.push(p.get_id(body)); } bool use_cache = false; return p.rec_save_pos(Fun(locals, body, use_cache), pos); @@ -728,21 +871,23 @@ static expr parse_lambda_binder(parser & p, pos_info const & pos) { static name * g_lambda_match_name = nullptr; -static expr parse_lambda_constructor(parser & p, pos_info const & ini_pos) { +static expr parse_lambda_constructor(parser & p, ast_data & fun, ast_data & bis, pos_info const & ini_pos) { lean_assert(p.curr_is_token(get_langle_tk())); parser::local_scope scope(p); auto pos = p.pos(); p.next(); buffer locals; expr pattern = p.parse_pattern([&](parser & p) { return parse_constructor_core(p, pos); }, locals); + bis.push(p.get_id(pattern)); for (expr const & local : locals) p.add_local(local); expr body; if (p.curr_is_token(get_comma_tk())) { p.next(); body = p.parse_expr(); + fun.push(p.get_id(body)); } else { - body = parse_lambda_core(p, ini_pos); + body = parse_lambda_core(p, fun, bis, ini_pos); } match_definition_scope match_scope(p.env()); expr fn = p.save_pos(mk_local(p.next_name(), *g_lambda_match_name, mk_expr_placeholder(), mk_rec_info(true)), pos); @@ -753,32 +898,46 @@ static expr parse_lambda_constructor(parser & p, pos_info const & ini_pos) { return p.rec_save_pos(Fun(x, mk_app(mk_equations(h, 1, &eqn), x), use_cache), pos); } -static expr parse_lambda_core(parser & p, pos_info const & pos) { +static expr parse_lambda_core(parser & p, ast_data & fun, ast_data & bis, pos_info const & pos) { if (p.curr_is_token(get_langle_tk())) { - return parse_lambda_constructor(p, pos); + return parse_lambda_constructor(p, fun, bis, pos); } else { - return parse_lambda_binder(p, pos); + return parse_lambda_binder(p, fun, bis, pos); } } static expr parse_lambda(parser & p, unsigned, expr const *, pos_info const & pos) { - return parse_lambda_core(p, pos); + auto& bis = p.new_ast("binders", p.pos()); + auto& fun = p.new_ast("fun", pos).push(bis.m_id); + expr r = parse_lambda_core(p, fun, bis, pos); + p.set_ast_pexpr(fun.m_id, r); + return r; } static expr parse_assume(parser & p, unsigned, expr const *, pos_info const & pos) { + ast_id id; expr r; if (p.curr_is_token(get_colon_tk())) { // anonymous `assume` p.next(); expr prop = p.parse_expr(); + auto& binder = p.new_ast(name("binder").append_after(0U), p.pos()) + .push(0).push(0).push(p.get_id(prop)); + auto& bis = p.new_ast("binders", p.pos()).push(binder.m_id); p.check_token_next(get_comma_tk(), "invalid 'assume', ',' expected"); parser::local_scope scope(p); expr l = p.save_pos(mk_local(get_this_tk(), prop), pos); p.add_local(l); expr body = p.parse_expr(); - return p.save_pos(Fun(l, body, p), pos); + id = p.new_ast("assume", pos).push(bis.m_id).push(p.get_id(body)).m_id; + r = p.save_pos(Fun(l, body, p), pos); } else { - return parse_lambda_core(p, pos); + auto& bis = p.new_ast("binders", p.pos()); + auto& fun = p.new_ast("assume", pos).push(bis.m_id); + id = fun.m_id; + r = parse_lambda_core(p, fun, bis, pos); } + p.set_ast_pexpr(id, r); + return r; } static void consume_rparen(parser & p) { @@ -829,13 +988,16 @@ static expr parse_infix_paren(parser & p, list const & accs args[0] = mk_local(p.next_name(), "_x", mk_expr_placeholder(), binder_info()); vars.push_back(args[0]); p.next(); + auto& data = p.new_ast(*g_infix_function, pos).push(0); if (p.curr_is_token(get_rparen_tk())) { + data.push(0); p.next(); args[1] = mk_local(p.next_name(), "_y", mk_expr_placeholder(), binder_info()); vars.push_back(args[1]); } else { fixed_second_arg = true; args[1] = p.parse_expr(); + data.push(p.get_id(args[1])); consume_rparen(p); } buffer cs; @@ -851,7 +1013,18 @@ static expr parse_infix_paren(parser & p, list const & accs cs.push_back(r); } } - return p.save_pos(mk_choice(cs.size(), cs.data()), pos); + expr r = p.save_pos(mk_choice(cs.size(), cs.data()), pos); + lean_assert(data.m_children.size() > 0); + if (length(accs) > 1) { + auto& choice = p.new_ast("choice", pos); + data.m_children[0] = choice.m_id; + for (auto& a : accs) + choice.push(p.new_ast(get_notation_tk(), pos, a.get_name()).m_id); + } else { + data.m_children[0] = p.new_ast(get_notation_tk(), pos, head(accs).get_name()).m_id; + } + p.set_ast_pexpr(data.m_id, r); + return r; } expr parse_lparen(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -873,14 +1046,21 @@ expr parse_lparen(parser & p, unsigned, expr const *, pos_info const & pos) { r = p.save_pos(mk_app(p.save_pos(mk_constant(get_prod_mk_name()), pos), args[i], r), pos); } + auto& data = p.new_ast("tuple", pos); + for (expr& e : args) data.push(p.get_id(e)); + p.set_ast_pexpr(data.m_id, r); return r; } else if (p.curr_is_token(get_colon_tk())) { p.next(); expr t = p.parse_expr(); consume_rparen(p); - return mk_typed_expr_distrib_choice(p, t, e, pos); + ast_id id = p.new_ast(get_colon_tk(), p.pos_of(e)).push(p.get_id(e)).push(p.get_id(t)).m_id; + e = mk_typed_expr_distrib_choice(p, t, e, pos); + p.set_ast_pexpr(id, e); + return e; } else { consume_rparen(p); + p.set_ast_pexpr(p.new_ast("(", pos).push(p.get_id(e)).m_id, e); return e; } } @@ -899,12 +1079,25 @@ static expr parse_lambda_cons(parser & p, unsigned, expr const *, pos_info const args[1] = mk_local(p.next_name(), "_y", mk_expr_placeholder(), binder_info()); vars.push_back(args[1]); buffer cs; - for (notation::accepting const & acc : head(ts).second.is_accepting()) { + auto& accs = head(ts).second.is_accepting(); + for (notation::accepting const & acc : accs) { expr r = p.copy_with_new_pos(acc.get_expr(), pos); r = p.save_pos(mk_infix_function(Fun(vars, instantiate_rev(r, 2, args), p)), pos); cs.push_back(r); } - return p.save_pos(mk_choice(cs.size(), cs.data()), pos); + expr r = p.save_pos(mk_choice(cs.size(), cs.data()), pos); + auto& data = p.new_ast(*g_infix_function, pos); + if (length(accs) > 1) { + auto& choice = p.new_ast("choice", pos); + data.push(choice.m_id); + for (auto& a : accs) + choice.push(p.new_ast(get_notation_tk(), pos, a.get_name()).m_id); + } else { + data.push(p.new_ast(get_notation_tk(), pos, head(accs).get_name()).m_id); + } + data.push(0); + p.set_ast_pexpr(data.m_id, r); + return r; } static expr parse_inaccessible(parser & p, unsigned, expr const *, pos_info const & pos) { @@ -914,25 +1107,26 @@ static expr parse_inaccessible(parser & p, unsigned, expr const *, pos_info cons return e; } p.check_token_next(get_rparen_tk(), "invalid inaccesible pattern, `)` expected"); - return p.save_pos(mk_inaccessible(e), pos); + e = p.save_pos(mk_inaccessible(e), pos); + p.set_ast_pexpr(p.new_ast(".(", pos).push(p.get_id(e)).m_id, e); + return e; } static expr parse_atomic_inaccessible(parser & p, unsigned, expr const *, pos_info const & pos) { if (!p.in_pattern()) { return p.parser_error_or_expr({"inaccesible pattern notation `._` can only be used in patterns", pos}); } - return p.save_pos(mk_inaccessible(p.save_pos(mk_expr_placeholder(), pos)), pos); + expr e = p.save_pos(mk_expr_placeholder(), pos); + ast_id id = p.new_ast("_", pos).m_id; + p.set_ast_pexpr(id, e); + e = p.save_pos(mk_inaccessible(e), pos); + p.set_ast_pexpr(p.new_ast(".(", pos).push(id).m_id, e); + return e; } static name * g_begin_hole = nullptr; static name * g_end_hole = nullptr; -expr mk_annotation_with_pos(parser & p, name const & a, expr const & e, pos_info const & pos) { - expr r = mk_annotation(a, e); - r.set_tag(nulltag); // mk_annotation copies e's tag - return p.save_pos(r, pos); -} - expr mk_hole(parser & p, expr const & e, pos_info const & begin_pos, pos_info const & end_pos) { return mk_annotation_with_pos(p, *g_begin_hole, mk_annotation_with_pos(p, *g_end_hole, e, end_pos), begin_pos); } @@ -959,6 +1153,7 @@ expr update_hole_args(expr const & e, expr const & new_args) { static expr parse_hole(parser & p, unsigned, expr const *, pos_info const & begin_pos) { buffer ps; + auto& data = p.new_ast("{!", begin_pos); while (!p.curr_is_token(get_rcurlybang_tk())) { expr e; if (p.in_quote()) { @@ -967,6 +1162,7 @@ static expr parse_hole(parser & p, unsigned, expr const *, pos_info const & begi parser::quote_scope scope(p, false); e = p.parse_expr(); } + data.push(p.get_id(e)); ps.push_back(copy_tag(e, mk_pexpr_quote(e))); if (!p.curr_is_token(get_comma_tk())) break; @@ -976,6 +1172,7 @@ static expr parse_hole(parser & p, unsigned, expr const *, pos_info const & begi p.check_token_next(get_rcurlybang_tk(), "invalid hole, `!}` expected"); end_pos.second += 2; expr r = mk_hole(p, mk_lean_list(ps), begin_pos, end_pos); + p.set_ast_pexpr(data.m_id, r); return r; } @@ -996,18 +1193,24 @@ static expr mk_bin_tree(parser & p, buffer const & args, unsigned start, u static expr parse_bin_tree(parser & p, unsigned, expr const *, pos_info const & pos) { buffer es; + auto& data = p.new_ast("#[", pos); while (!p.curr_is_token(get_rbracket_tk())) { - es.push_back(p.parse_expr()); + expr e = p.parse_expr(); + es.push_back(e); + data.push(p.get_id(e)); if (!p.curr_is_token(get_comma_tk())) break; p.next(); } p.check_token_next(get_rbracket_tk(), "invalid `#[...]`, `]` expected"); + expr r; if (es.empty()) { - return p.save_pos(mk_constant(get_bin_tree_empty_name()), pos); + r = p.save_pos(mk_constant(get_bin_tree_empty_name()), pos); } else { - return mk_bin_tree(p, es, 0, es.size(), pos); + r = mk_bin_tree(p, es, 0, es.size(), pos); } + p.set_ast_pexpr(data.m_id, r); + return r; } parse_table init_nud_table() { @@ -1058,11 +1261,17 @@ parse_table init_nud_table() { static expr parse_field(parser & p, unsigned, expr const * args, pos_info const & pos) { try { if (p.curr_is_numeral()) { - unsigned fidx = p.parse_small_nat(); - return p.save_pos(mk_field_notation(args[0], fidx), pos); + ast_id id; unsigned fidx; + std::tie(id, fidx) = p.parse_small_nat(); + expr r = p.save_pos(mk_field_notation(args[0], fidx), pos); + p.set_ast_pexpr(p.new_ast("^.", pos).push(p.get_id(args[0])).push(id).m_id, r); + return r; } else { - name field = p.check_id_next("identifier or numeral expected"); - return p.save_pos(mk_field_notation(args[0], field), pos); + ast_id id; name field; + std::tie(id, field) = p.check_id_next("identifier or numeral expected"); + expr r = p.save_pos(mk_field_notation(args[0], field), pos); + p.set_ast_pexpr(p.new_ast("^.", pos).push(p.get_id(args[0])).push(id).m_id, r); + return r; } } catch (break_at_pos_exception & ex) { if (!p.get_complete()) { diff --git a/src/frontends/lean/calc.cpp b/src/frontends/lean/calc.cpp index da84a1ce1a..f3a7a86b0b 100644 --- a/src/frontends/lean/calc.cpp +++ b/src/frontends/lean/calc.cpp @@ -75,10 +75,11 @@ static expr mk_op_fn(parser & p, name const & op, unsigned num_placeholders, pos return r; } -static calc_step parse_calc_proof(parser & p, calc_pred const & pred) { +static calc_step parse_calc_proof(parser & p, ast_data & parent, calc_pred const & pred) { p.check_token_next(get_colon_tk(), "invalid 'calc' expression, ':' expected"); auto pos = p.pos(); expr pr = p.parse_expr(); + parent.push(p.get_id(pr)); return calc_step(pred, p.save_pos(mk_calc_annotation(pr), pos)); } @@ -116,9 +117,10 @@ static expr mk_implies(parser & p, expr const & lhs, expr const & rhs, pos_info return p.mk_app(p.mk_app(p.save_pos(mk_constant(get_implies_name()), pos), lhs, pos), rhs, pos); } -static expr parse_pred(parser & p) { +static expr parse_pred(parser & p, ast_data & parent) { auto pos = p.pos(); expr pred = p.parse_expr(); + parent.push(p.get_id(pred)); if (is_arrow(pred)) return mk_implies(p, binding_domain(pred), binding_body(pred), pos); else @@ -127,20 +129,27 @@ static expr parse_pred(parser & p) { static expr parse_next_pred(parser & p, expr const & dummy) { auto pos = p.pos(); + ast_id id = p.new_ast(get_ellipsis_tk(), pos).m_id; if (p.curr_is_token(get_arrow_tk())) { p.next(); expr rhs = p.parse_expr(); - return mk_implies(p, dummy, rhs, pos); + expr r = mk_implies(p, dummy, rhs, pos); + p.set_ast_pexpr(p.new_ast(get_arrow_tk(), pos).push(id).push(p.get_id(rhs)).m_id, r); + return r; } else { + p.set_ast_pexpr(id, dummy); return p.parse_led(dummy); } } -expr parse_calc(parser & p) { +expr parse_calc(parser & p, pos_info const & calc_pos) { + auto& data = p.new_ast(*g_calc_name, calc_pos); auto pos = p.pos(); - expr first_pred_expr = parse_pred(p); + auto& ast = p.new_ast("step", pos); + data.push(ast.m_id); + expr first_pred_expr = parse_pred(p, ast); calc_pred pred = decode_expr(first_pred_expr, pos); - calc_step step = parse_calc_proof(p, pred); + calc_step step = parse_calc_proof(p, ast, pred); bool single = true; // true if calc has only one step expr dummy; @@ -149,21 +158,26 @@ expr parse_calc(parser & p) { pos = p.pos(); p.next(); expr new_pred_expr = parse_next_pred(p, dummy); + auto& ast2 = p.new_ast("step", pos).push(p.get_id(new_pred_expr)); + data.push(ast2.m_id); try { calc_pred new_pred = decode_expr(new_pred_expr, pos); new_pred = calc_pred(pred_op(new_pred), pred_rhs(pred), pred_rhs(new_pred)); - calc_step new_step = parse_calc_proof(p, new_pred); + calc_step new_step = parse_calc_proof(p, ast2, new_pred); step = join(p, step, new_step, pos); } catch (parser_error & ex) { p.maybe_throw_error(std::move(ex)); } } + expr r; if (single) { - return p.save_pos(mk_typed_expr(first_pred_expr, step_proof(step)), pos); + r = p.save_pos(mk_typed_expr(first_pred_expr, step_proof(step)), pos); } else { - return step_proof(step); + r = step_proof(step); } + p.set_ast_pexpr(data.m_id, r); + return r; } void initialize_calc() { diff --git a/src/frontends/lean/calc.h b/src/frontends/lean/calc.h index bbb6aba473..ab8e522c20 100644 --- a/src/frontends/lean/calc.h +++ b/src/frontends/lean/calc.h @@ -8,7 +8,7 @@ Author: Leonardo de Moura #include "frontends/lean/cmd_table.h" namespace lean { class parser; -expr parse_calc(parser & p); +expr parse_calc(parser & p, pos_info const & pos); bool is_calc_annotation(expr const & e); void initialize_calc(); void finalize_calc(); diff --git a/src/frontends/lean/cmd_table.h b/src/frontends/lean/cmd_table.h index 0d1c92bcfa..98241d6774 100644 --- a/src/frontends/lean/cmd_table.h +++ b/src/frontends/lean/cmd_table.h @@ -16,6 +16,7 @@ class parser; struct cmd_meta { decl_attributes m_attrs; + ast_id m_modifiers_id = 0; decl_modifiers m_modifiers; optional m_doc_string; cmd_meta() {} @@ -25,7 +26,7 @@ struct cmd_meta { m_attrs(attrs), m_modifiers(mods), m_doc_string(doc) {} }; -typedef std::function command_fn; +typedef std::function command_fn; template struct cmd_info_tmpl { @@ -36,15 +37,15 @@ struct cmd_info_tmpl { public: cmd_info_tmpl(name const & n, char const * d, F const & fn, bool skip_token = true): m_name(n), m_descr(d), m_fn(fn), m_skip_token(skip_token) {} - cmd_info_tmpl(name const & n, char const * d, std::function const & fn, bool skip_token = true): - cmd_info_tmpl(n, d, [=](parser & p, cmd_meta const & meta) { + cmd_info_tmpl(name const & n, char const * d, std::function const & fn, bool skip_token = true): + cmd_info_tmpl(n, d, [=](parser & p, ast_id & cmd_id, cmd_meta const & meta) { if (meta.m_modifiers) throw exception("command does not accept modifiers"); if (meta.m_attrs) throw exception("command does not accept attributes"); if (meta.m_doc_string) throw exception("command does not accept doc string"); - return fn(p); + return fn(p, cmd_id); }, skip_token) {} cmd_info_tmpl() {} name const & get_name() const { return m_name; } diff --git a/src/frontends/lean/decl_attributes.cpp b/src/frontends/lean/decl_attributes.cpp index 599d859406..8a7035e121 100644 --- a/src/frontends/lean/decl_attributes.cpp +++ b/src/frontends/lean/decl_attributes.cpp @@ -5,6 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #include "util/sexpr/option_declarations.h" +#include "util/task_builder.h" #include "library/attribute_manager.h" #include "library/constants.h" #include "library/class.h" @@ -26,15 +27,17 @@ unsigned get_default_priority(options const & opts) { } // ========================================== -void decl_attributes::parse_core(parser & p, bool compact) { +ast_id decl_attributes::parse_core(parser & p, bool compact) { + auto& attrs = p.new_ast("attrs", p.pos()); while (true) { auto pos = p.pos(); - bool deleted = p.curr_is_token_or_id(get_sub_tk()); + ast_id deleted = p.curr_is_token_or_id(get_sub_tk()) ? p.new_ast(get_sub_tk(), pos).m_id : 0; if (deleted) { if (m_persistent) throw parser_error("cannot remove attribute globally (solution: use 'local attribute')", pos); p.next(); } + p.check_break_before(break_at_pos_exception::token_context::attribute); name id; if (p.curr_is_command()) { @@ -42,16 +45,22 @@ void decl_attributes::parse_core(parser & p, bool compact) { p.next(); } else { id = p.check_id_next("invalid attribute declaration, identifier expected", - break_at_pos_exception::token_context::attribute); + break_at_pos_exception::token_context::attribute).second; } if (id == "priority") { + auto& attr_ast = p.new_ast(id, pos); + attrs.push(attr_ast.m_id); if (deleted) throw parser_error("cannot remove priority attribute", pos); auto pos = p.pos(); expr pre_val = p.parse_expr(); - pre_val = mk_typed_expr(mk_constant(get_nat_name()), pre_val); + pre_val = mk_typed_expr(mk_constant(get_nat_name()), pre_val, pre_val.get_tag()); + ast_id id = p.get_id(pre_val); + p.set_ast_pexpr(id, pre_val); + attr_ast.push(id); expr nat = mk_constant(get_nat_name()); expr val = p.elaborate("_attribute", list(), pre_val).first; + p.get_ast(id).m_expr.emplace(mk_pure_task(std::move(val))); vm_obj prio = eval_closed_expr(p.env(), p.get_options(), "_attribute", nat, val, pos); if (optional _prio = try_to_unsigned(prio)) { m_prio = _prio; @@ -59,6 +68,8 @@ void decl_attributes::parse_core(parser & p, bool compact) { throw parser_error("invalid 'priority', argument does not evaluate to a (small) numeral", pos); } } else { + auto& attr_ast = p.new_ast("attr", pos, id).push(deleted); + attrs.push(attr_ast.m_id); if (!is_attribute(p.env(), id)) throw parser_error(sstream() << "unknown attribute [" << id << "]", pos); @@ -73,7 +84,11 @@ void decl_attributes::parse_core(parser & p, bool compact) { } } } - auto data = deleted ? attr_data_ptr() : attr.parse_data(p); + attr_data_ptr data; + if (deleted) + attr_ast.push(0); + else + data = attr.parse_data(p, attr_ast); m_entries = cons({&attr, data}, m_entries); if (id == "parsing_only") m_parsing_only = true; @@ -90,17 +105,18 @@ void decl_attributes::parse_core(parser & p, bool compact) { break; } } + return attrs.m_id; } -void decl_attributes::parse(parser & p) { +ast_id decl_attributes::parse(parser & p) { if (!p.curr_is_token(get_lbracket_tk())) - return; + return 0; p.next(); - parse_core(p, false); + return parse_core(p, false); } -void decl_attributes::parse_compact(parser & p) { - parse_core(p, true); +ast_id decl_attributes::parse_compact(parser & p) { + return parse_core(p, true); } void decl_attributes::set_attribute(environment const & env, name const & attr_name) { diff --git a/src/frontends/lean/decl_attributes.h b/src/frontends/lean/decl_attributes.h index d8de11cebe..716cf9ef1b 100644 --- a/src/frontends/lean/decl_attributes.h +++ b/src/frontends/lean/decl_attributes.h @@ -24,14 +24,14 @@ class decl_attributes { bool m_parsing_only; list m_entries; optional m_prio; - void parse_core(parser & p, bool compact); + ast_id parse_core(parser & p, bool compact); public: decl_attributes(bool persistent = true): m_persistent(persistent), m_parsing_only(false) {} void set_attribute(environment const & env, name const & attr_name); /* attributes: zero-or-more [ ... ] */ - void parse(parser & p); + ast_id parse(parser & p); /* Parse attributes after `@[` ... ] */ - void parse_compact(parser & p); + ast_id parse_compact(parser & p); environment apply(environment env, io_state const & ios, name const & d) const; list const & get_entries() const { return m_entries; } bool is_parsing_only() const { return m_parsing_only; } diff --git a/src/frontends/lean/decl_cmds.cpp b/src/frontends/lean/decl_cmds.cpp index c60a1e5743..970f19ef99 100644 --- a/src/frontends/lean/decl_cmds.cpp +++ b/src/frontends/lean/decl_cmds.cpp @@ -48,35 +48,42 @@ static environment declare_universe(parser & p, environment env, name const & n, return env; } -static environment universes_cmd_core(parser & p, bool local) { +static environment universes_cmd_core(parser & p, ast_data & parent, bool local) { if (!p.curr_is_identifier()) throw parser_error("invalid 'universes' command, identifier expected", p.pos()); environment env = p.env(); while (p.curr_is_identifier()) { name n = p.get_name_val(); + parent.push(p.new_ast("ident", p.pos(), n).m_id); p.next(); env = declare_universe(p, env, n, local); } return env; } -static environment universe_cmd(parser & p) { +static environment universe_cmd(parser & p, ast_id & cmd_id) { + auto& data = p.get_ast(cmd_id); if (p.curr_is_token(get_variables_tk())) { + data.m_type = "universe_variables"; p.next(); - return universes_cmd_core(p, true); + return universes_cmd_core(p, data, true); } else { bool local = false; if (p.curr_is_token(get_variable_tk())) { + data.m_type = "universe_variable"; p.next(); local = true; } - name n = p.check_decl_id_next("invalid 'universe' command, identifier expected"); + ast_id id; name n; + std::tie(id, n) = p.check_decl_id_next("invalid 'universe' command, identifier expected"); + data.push(id); return declare_universe(p, p.env(), n, local); } } -static environment universes_cmd(parser & p) { - return universes_cmd_core(p, false); +static environment universes_cmd(parser & p, ast_id & cmd_id) { + auto& data = p.get_ast(cmd_id); + return universes_cmd_core(p, data, false); } enum class variable_kind { Constant, Parameter, Variable, Axiom }; @@ -212,26 +219,36 @@ class var_decl_scope { } }; -static environment variable_cmd_core(parser & p, variable_kind k, cmd_meta const & meta) { +static environment variable_cmd_core(parser & p, variable_kind k, ast_id cmd_id, cmd_meta const & meta) { check_variable_kind(p, k); auto pos = p.pos(); module::scope_pos_info scope_pos(pos); + auto& parent = p.get_ast(cmd_id).push(meta.m_modifiers_id); optional bi; - if (k == variable_kind::Parameter || k == variable_kind::Variable) + ast_data * bi_ast = nullptr; + ast_data * vars = nullptr; + bool is_decl = k == variable_kind::Constant || k == variable_kind::Axiom; + if (!is_decl) { bi = parse_binder_info(p, k); + vars = &p.new_ast("vars", p.pos()); + bi_ast = &p.new_ast(name("binder").append_after(bi ? bi->hash() : 0), pos).push(vars->m_id); + parent.push(bi_ast->m_id); + } optional scope1; name n; expr type; buffer ls_buffer; - if (bi && bi->is_inst_implicit() && (k == variable_kind::Parameter || k == variable_kind::Variable)) { + if (bi && bi->is_inst_implicit() && !is_decl) { var_decl_scope var_scope(p, meta.m_modifiers); /* instance implicit */ if (p.curr_is_identifier()) { auto n_pos = p.pos(); n = p.get_name_val(); + auto& n_ast = p.new_ast("ident", n_pos, n); p.next(); if (p.curr_is_token(get_colon_tk())) { /* simple decl: variable [decA : decidable A] */ + vars->push(n_ast.m_id); p.next(); type = p.parse_expr(); } else if (p.curr_is_token(get_rbracket_tk())) { @@ -243,40 +260,50 @@ static environment variable_cmd_core(parser & p, variable_kind k, cmd_meta const auto it = p.get_local(n); if (it && is_local(*it)) { // annotation update: variable [decA] + vars->push(n_ast.m_id); p.parse_close_binder_info(bi); + bi_ast->push(0).push(0); update_local_binder_info(p, k, n, bi, pos); return p.env(); } else { // parameter-less anonymous local instance : variable [io.interface] - type = p.id_to_expr(n, n_pos); + type = p.id_to_expr(n, n_ast); n = p.mk_anonymous_inst_name(); + bi_ast->m_children[0] = 0; } } else { /* anonymous : variable [decidable A] */ - expr left = p.id_to_expr(n, n_pos); + expr left = p.id_to_expr(n, n_ast); n = p.mk_anonymous_inst_name(); unsigned rbp = 0; while (rbp < p.curr_lbp()) { left = p.parse_led(left); } type = left; + bi_ast->m_children[0] = 0; } } else { /* anonymous : variable [forall x y, decidable (x = y)] */ n = p.mk_anonymous_inst_name(); type = p.parse_expr(); + bi_ast->m_children[0] = 0; } + bi_ast->push(0).push(p.get_id(type)); } else { var_decl_scope var_scope(p, meta.m_modifiers); /* non instance implicit cases */ - if (p.curr_is_token(get_lcurly_tk()) && (k == variable_kind::Parameter || k == variable_kind::Variable)) - throw parser_error("invalid declaration, only constants/axioms can be universe polymorphic", p.pos()); - if (k == variable_kind::Constant || k == variable_kind::Axiom) + if (is_decl) { scope1.emplace(p); - parse_univ_params(p, ls_buffer); - n = p.check_decl_id_next("invalid declaration, identifier expected"); + parent.push(parse_univ_params(p, ls_buffer)); + } else if (p.curr_is_token(get_lcurly_tk())) { + throw parser_error("invalid declaration, only constants/axioms can be universe polymorphic", p.pos()); + } + ast_id n_id; + std::tie(n_id, n) = p.check_decl_id_next("invalid declaration, identifier expected"); + (is_decl ? &parent : vars)->push(n_id); if (!p.curr_is_token(get_colon_tk())) { - if (!curr_is_binder_annotation(p) && (k == variable_kind::Parameter || k == variable_kind::Variable)) { + if (!curr_is_binder_annotation(p) && !is_decl) { + bi_ast->push(0).push(0); p.parse_close_binder_info(bi); update_local_binder_info(p, k, n, bi, pos); return p.env(); @@ -284,14 +311,17 @@ static environment variable_cmd_core(parser & p, variable_kind k, cmd_meta const buffer ps; unsigned rbp = 0; bool allow_default = true; - auto lenv = p.parse_binders(ps, rbp, allow_default); + auto& data = p.new_ast("binders", p.pos()); + auto lenv = p.parse_binders(&data, ps, rbp, allow_default); p.check_token_next(get_colon_tk(), "invalid declaration, ':' expected"); type = p.parse_scoped_expr(ps, lenv); + (is_decl ? &parent : bi_ast)->push(data.m_id).push(p.get_id(type)); type = Pi(ps, type, p); } } else { p.next(); type = p.parse_expr(); + (is_decl ? &parent : bi_ast)->push(0).push(p.get_id(type)); } } p.parse_close_binder_info(bi); @@ -306,23 +336,23 @@ static environment variable_cmd_core(parser & p, variable_kind k, cmd_meta const level_param_names new_ls; list ctx = p.locals_to_context(); std::tie(type, new_ls) = p.elaborate_type("_variable", ctx, type, false); - if (k == variable_kind::Variable || k == variable_kind::Parameter) + if (!is_decl) update_local_levels(p, new_ls, k == variable_kind::Variable); return declare_var(p, p.env(), n, append(ls, new_ls), type, k, bi, pos, meta); } -static environment variable_cmd(parser & p, cmd_meta const & meta) { - return variable_cmd_core(p, variable_kind::Variable, meta); +static environment variable_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return variable_cmd_core(p, variable_kind::Variable, cmd_id, meta); } -static environment axiom_cmd(parser & p, cmd_meta const & meta) { +static environment axiom_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { if (meta.m_modifiers.m_is_meta) throw exception("invalid 'meta' modifier for axiom"); - return variable_cmd_core(p, variable_kind::Axiom, meta); + return variable_cmd_core(p, variable_kind::Axiom, cmd_id, meta); } -static environment constant_cmd(parser & p, cmd_meta const & meta) { - return variable_cmd_core(p, variable_kind::Constant, meta); +static environment constant_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return variable_cmd_core(p, variable_kind::Constant, cmd_id, meta); } -static environment parameter_cmd(parser & p, cmd_meta const & meta) { - return variable_cmd_core(p, variable_kind::Parameter, meta); +static environment parameter_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return variable_cmd_core(p, variable_kind::Parameter, cmd_id, meta); } /* @@ -344,38 +374,45 @@ static void ensure_no_match_in_variables_cmd(pos_info const & pos) { } } -static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta const & meta) { - check_variable_kind(p, k); +static environment variables_cmd_core(parser & p, variable_kind k, ast_data & parent, cmd_meta const & meta) { auto pos = p.pos(); module::scope_pos_info scope_pos(pos); declaration_info_scope d_scope(p, decl_cmd_kind::Var, meta.m_modifiers); optional bi = parse_binder_info(p, k); + auto& ast = p.new_ast(name("binder").append_after(bi ? bi->hash() : 0), pos); + parent.push(ast.m_id); + auto& vars = p.new_ast("vars", pos); + ast.push(vars.m_id).push(0); buffer ids; optional scope1; expr type; if (bi && bi->is_inst_implicit() && (k == variable_kind::Parameter || k == variable_kind::Variable)) { /* instance implicit */ if (p.curr_is_identifier()) { - auto id_pos = p.pos(); name id = p.get_name_val(); + auto& id_ast = p.new_ast("ident", p.pos(), id); p.next(); if (p.curr_is_token(get_colon_tk())) { /* simple decl: variables [decA : decidable A] */ + vars.push(id_ast.m_id); p.next(); ids.push_back(id); type = p.parse_expr(); + ast.push(p.get_id(type)); ensure_no_match_in_variables_cmd(pos); } else if (p.curr_is_token(get_rbracket_tk())) { /* annotation update: variables [decA] */ + vars.push(id_ast.m_id); p.parse_close_binder_info(bi); + ast.push(0); update_local_binder_info(p, k, id, bi, pos); if (curr_is_binder_annotation(p)) - return variables_cmd_core(p, k, meta); + return variables_cmd_core(p, k, parent, meta); else return p.env(); } else { /* anonymous : variables [decidable A] */ - expr left = p.id_to_expr(id, id_pos); + expr left = p.id_to_expr(id, id_ast); id = p.mk_anonymous_inst_name(); unsigned rbp = 0; while (rbp < p.curr_lbp()) { @@ -383,6 +420,9 @@ static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta cons } ids.push_back(id); type = left; + lean_assert(ast.m_children.size() > 0); + ast.m_children[0] = 0; + ast.push(p.get_id(type)); ensure_no_match_in_variables_cmd(pos); } } else { @@ -390,12 +430,16 @@ static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta cons name id = p.mk_anonymous_inst_name(); ids.push_back(id); type = p.parse_expr(); + lean_assert(ast.m_children.size() > 0); + ast.m_children[0] = 0; + ast.push(p.get_id(type)); ensure_no_match_in_variables_cmd(pos); } } else { /* non instance implicit cases */ while (p.curr_is_identifier()) { name id = p.get_name_val(); + vars.push(p.new_ast("ident", p.pos(), id).m_id); p.next(); ids.push_back(id); } @@ -406,11 +450,12 @@ static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta cons /* example: variables (A) */ if (k == variable_kind::Parameter || k == variable_kind::Variable) { p.parse_close_binder_info(bi); + ast.push(0); for (name const & id : ids) { update_local_binder_info(p, k, id, bi, pos); } if (curr_is_binder_annotation(p)) - return variables_cmd_core(p, k, meta); + return variables_cmd_core(p, k, parent, meta); else return p.env(); } else { @@ -420,6 +465,7 @@ static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta cons if (k == variable_kind::Constant || k == variable_kind::Axiom) scope1.emplace(p); type = p.parse_expr(); + ast.push(p.get_id(type)); ensure_no_match_in_variables_cmd(pos); } p.parse_close_binder_info(bi); @@ -443,39 +489,45 @@ static environment variables_cmd_core(parser & p, variable_kind k, cmd_meta cons // We must do that to be able to process // constants (A : Type) (a : A) parser::local_scope scope2(p, env); - return variables_cmd_core(p, k, meta); + return variables_cmd_core(p, k, parent, meta); } else { - return variables_cmd_core(p, k, meta); + return variables_cmd_core(p, k, parent, meta); } } return env; } -static environment variables_cmd(parser & p, cmd_meta const & meta) { - return variables_cmd_core(p, variable_kind::Variable, meta); +static environment gen_variables_cmd(parser & p, variable_kind k, ast_id cmd_id, cmd_meta const & meta) { + check_variable_kind(p, k); + auto& parent = p.get_ast(cmd_id).push(meta.m_modifiers_id); + return variables_cmd_core(p, k, parent, meta); +} + +static environment variables_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return gen_variables_cmd(p, variable_kind::Variable, cmd_id, meta); } -static environment parameters_cmd(parser & p, cmd_meta const & meta) { - return variables_cmd_core(p, variable_kind::Parameter, meta); +static environment parameters_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return gen_variables_cmd(p, variable_kind::Parameter, cmd_id, meta); } -static environment constants_cmd(parser & p, cmd_meta const & meta) { - return variables_cmd_core(p, variable_kind::Constant, meta); +static environment constants_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return gen_variables_cmd(p, variable_kind::Constant, cmd_id, meta); } -static environment axioms_cmd(parser & p, cmd_meta const & meta) { - return variables_cmd_core(p, variable_kind::Axiom, meta); +static environment axioms_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return gen_variables_cmd(p, variable_kind::Axiom, cmd_id, meta); } -static environment definition_cmd(parser & p, cmd_meta const & meta) { - return definition_cmd_core(p, decl_cmd_kind::Definition, meta); +static environment definition_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return definition_cmd_core(p, decl_cmd_kind::Definition, cmd_id, meta); } -static environment theorem_cmd(parser & p, cmd_meta const & meta) { - return definition_cmd_core(p, decl_cmd_kind::Theorem, meta); +static environment theorem_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return definition_cmd_core(p, decl_cmd_kind::Theorem, cmd_id, meta); } -static environment abbreviation_cmd(parser & p, cmd_meta const & meta) { - return definition_cmd_core(p, decl_cmd_kind::Abbreviation, meta); +static environment abbreviation_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return definition_cmd_core(p, decl_cmd_kind::Abbreviation, cmd_id, meta); } -static environment example_cmd(parser & p, cmd_meta const & meta) { - return definition_cmd_core(p, decl_cmd_kind::Example, meta); +static environment example_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return definition_cmd_core(p, decl_cmd_kind::Example, cmd_id, meta); } -static environment instance_cmd(parser & p, cmd_meta const & _meta) { +static environment instance_cmd(parser & p, ast_id & cmd_id, cmd_meta const & _meta) { auto meta = _meta; if (meta.m_modifiers.m_is_private) throw exception("invalid 'private' modifier for instance command"); @@ -484,22 +536,32 @@ static environment instance_cmd(parser & p, cmd_meta const & _meta) { if (meta.m_modifiers.m_is_mutual) throw exception("invalid 'mutual' modifier for instance command"); meta.m_modifiers.m_is_protected = true; - return definition_cmd_core(p, decl_cmd_kind::Instance, meta); + return definition_cmd_core(p, decl_cmd_kind::Instance, cmd_id, meta); } -static environment modifiers_cmd(parser & p, cmd_meta const & _meta) { +static environment modifiers_cmd(parser & p, ast_id & cmd_id, cmd_meta const & _meta) { auto meta = _meta; + auto tk = cmd_id; + auto& mods = p.new_modifiers(meta); if (p.curr_is_token(get_private_tk())) { + mods.push(tk); + tk = 0; meta.m_modifiers.m_is_private = true; p.next(); } else if (p.curr_is_token(get_protected_tk())) { + mods.push(tk); + tk = 0; meta.m_modifiers.m_is_protected = true; p.next(); } if (p.curr_is_token(get_noncomputable_tk())) { + if (!tk) tk = p.new_ast(get_noncomputable_tk(), p.pos()).m_id; + mods.push(tk); + tk = 0; p.next(); if (!meta.m_attrs && !meta.m_modifiers && p.curr_is_token_or_id(get_theory_tk())) { + cmd_id = p.new_ast(get_theory_tk(), p.pos()).push(mods.m_id).m_id; // `noncomputable theory` p.next(); p.set_ignore_noncomputable(); @@ -509,11 +571,17 @@ static environment modifiers_cmd(parser & p, cmd_meta const & _meta) { } } if (p.curr_is_token(get_meta_tk())) { + if (!tk) tk = p.new_ast(get_meta_tk(), p.pos()).m_id; + mods.push(tk); + tk = 0; meta.m_modifiers.m_is_meta = true; p.next(); } if (p.curr_is_token(get_mutual_tk())) { + if (!tk) tk = p.new_ast(get_mutual_tk(), p.pos()).m_id; + mods.push(tk); + tk = 0; meta.m_modifiers.m_is_mutual = true; p.next(); } @@ -525,21 +593,31 @@ static environment modifiers_cmd(parser & p, cmd_meta const & _meta) { if (p.curr_is_token(get_attribute_tk()) || p.curr_is_token("@[")) { throw parser_error("unexpected attributes declaration", p.pos()); } - p.parse_command(meta); + cmd_id = p.parse_command(meta); return p.env(); } -static environment attribute_cmd_core(parser & p, bool persistent, cmd_meta const & meta) { +static environment attribute_cmd_core(parser & p, ast_id & cmd_id, bool persistent, cmd_meta const & meta) { buffer ds; decl_attributes attributes(persistent); - attributes.parse(p); + auto& attr = p.get_ast(cmd_id); + attr.m_type = get_attribute_tk(); + attr.push(persistent ? 0 : p.new_ast("local", attr.m_start).m_id).push(attributes.parse(p)); // 'attribute [attr] definition ...' if (p.curr_is_command()) { - return modifiers_cmd(p, {attributes, meta.m_modifiers, meta.m_doc_string}); + auto meta2 = meta; + meta2.m_attrs = attributes; + p.new_modifiers(meta2).push(attr.m_id); + cmd_id = p.new_ast(p.get_token_info().value(), p.pos()).m_id; + return modifiers_cmd(p, cmd_id, meta2); } + attr.m_children.insert(attr.m_children.begin(), meta.m_modifiers_id); + cmd_id = attr.m_id; do { auto pos = p.pos(); - name d = p.check_constant_next("invalid 'attribute' command, constant expected"); + ast_id id; name d; + std::tie(id, d) = p.check_constant_next("invalid 'attribute' command, constant expected"); + attr.push(id); ds.push_back(d); if (get_global_info_manager()) get_global_info_manager()->add_const_info(p.env(), pos, d); @@ -552,27 +630,33 @@ static environment attribute_cmd_core(parser & p, bool persistent, cmd_meta cons return env; } -static environment attribute_cmd(parser & p, cmd_meta const & meta) { - return attribute_cmd_core(p, true, meta); +static environment attribute_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return attribute_cmd_core(p, cmd_id, true, meta); } -environment local_attribute_cmd(parser & p, cmd_meta const & meta) { - return attribute_cmd_core(p, false, meta); +environment local_attribute_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { + return attribute_cmd_core(p, cmd_id, false, meta); } -static environment compact_attribute_cmd(parser & p, cmd_meta const & meta) { +static environment compact_attribute_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { bool persistent = true; decl_attributes attributes(persistent); - attributes.parse_compact(p); - return modifiers_cmd(p, {attributes, meta.m_modifiers, meta.m_doc_string}); + p.get_ast(cmd_id).push(attributes.parse_compact(p)); + auto meta2 = meta; + meta2.m_attrs = attributes; + p.new_modifiers(meta2).push(cmd_id); + cmd_id = p.new_ast(p.get_token_info().value(), p.pos()).m_id; + return modifiers_cmd(p, cmd_id, meta2); } -static environment include_cmd_core(parser & p, bool include) { +static environment include_cmd_core(parser & p, ast_id cmd_id, bool include) { if (!p.curr_is_identifier()) throw parser_error(sstream() << "invalid include/omit command, identifier expected", p.pos()); + auto& data = p.get_ast(cmd_id); while (p.curr_is_identifier()) { auto pos = p.pos(); name n = p.get_name_val(); + data.push(p.new_ast("ident", p.pos(), n).m_id); p.next(); if (!p.get_local(n)) throw parser_error(sstream() << "invalid include/omit command, '" << n << "' is not a parameter/variable", pos); @@ -589,12 +673,12 @@ static environment include_cmd_core(parser & p, bool include) { return p.env(); } -static environment include_cmd(parser & p) { - return include_cmd_core(p, true); +static environment include_cmd(parser & p, ast_id & cmd_id) { + return include_cmd_core(p, cmd_id, true); } -static environment omit_cmd(parser & p) { - return include_cmd_core(p, false); +static environment omit_cmd(parser & p, ast_id & cmd_id) { + return include_cmd_core(p, cmd_id, false); } void register_decl_cmds(cmd_table & r) { diff --git a/src/frontends/lean/decl_cmds.h b/src/frontends/lean/decl_cmds.h index 85375c7345..455d0e0ac0 100644 --- a/src/frontends/lean/decl_cmds.h +++ b/src/frontends/lean/decl_cmds.h @@ -14,7 +14,7 @@ class parser; Store the result in \c ps. Return true when levels were provided. */ -bool parse_univ_params(parser & p, buffer & ps); +ast_id parse_univ_params(parser & p, buffer & ps); /** \brief Add universe levels from \c found_ls to \c ls_buffer (only the levels that do not already occur in \c ls_buffer are added). @@ -23,7 +23,7 @@ bool parse_univ_params(parser & p, buffer & ps); void update_univ_parameters(buffer & ls_buffer, name_set const & found_ls, parser const & p); /** \brief Parse a local attribute command */ -environment local_attribute_cmd(parser & p, cmd_meta const & meta); +environment local_attribute_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta); void register_decl_cmds(cmd_table & r); void initialize_decl_cmds(); diff --git a/src/frontends/lean/decl_util.cpp b/src/frontends/lean/decl_util.cpp index ea78f07085..db4e994c00 100644 --- a/src/frontends/lean/decl_util.cpp +++ b/src/frontends/lean/decl_util.cpp @@ -31,20 +31,23 @@ Author: Leonardo de Moura #include "frontends/lean/builtin_exprs.h" namespace lean { -bool parse_univ_params(parser & p, buffer & lp_names) { +ast_id parse_univ_params(parser & p, buffer & lp_names) { if (p.curr_is_token(get_lcurly_tk())) { + auto& data = p.new_ast("levels", p.pos()); p.next(); while (!p.curr_is_token(get_rcurly_tk())) { auto pos0 = p.pos(); - name l = p.check_atomic_id_next("invalid declaration, identifier expected"); + ast_id id; name l; + std::tie(id, l) = p.check_atomic_id_next("invalid declaration, identifier expected"); + data.push(id); lp_names.push_back(l); p.add_local_level(l, mk_param_univ(l)); if (p.pos() == pos0) break; } p.next(); - return true; + return data.m_id; } else { - return false; + return 0; } } @@ -173,7 +176,7 @@ optional heuristic_inst_name(name const & ns, expr const & type) { } } -expr parse_single_header(parser & p, declaration_name_scope & scope, +expr parse_single_header(parser & p, ast_data & parent, declaration_name_scope & scope, buffer & lp_names, buffer & params, bool is_example, bool is_instance) { auto c_pos = p.pos(); @@ -181,23 +184,31 @@ expr parse_single_header(parser & p, declaration_name_scope & scope, if (is_example) { c_name = "_example"; scope.set_name(c_name); + parent.push(0).push(0); } else { - if (!is_instance) - parse_univ_params(p, lp_names); + parent.push(is_instance ? 0 : parse_univ_params(p, lp_names)); if (!is_instance || p.curr_is_identifier()) { - c_name = p.check_decl_id_next("invalid declaration, identifier expected"); + ast_id id; + std::tie(id, c_name) = p.check_decl_id_next("invalid declaration, identifier expected"); + parent.push(id); scope.set_name(c_name); + } else { + parent.push(0); } } - p.parse_optional_binders(params, /* allow_default */ true, /* explicit delimiters */ true); + auto& bis = p.new_ast("binders", p.pos()); + p.parse_optional_binders(&bis, params, /* allow_default */ true, /* explicit delimiters */ true); + parent.push(bis.m_children.empty() ? 0 : bis.m_id); for (expr const & param : params) p.add_local(param); expr type; if (p.curr_is_token(get_colon_tk())) { p.next(); type = p.parse_expr(); + parent.push(p.get_id(type)); } else { type = p.save_pos(mk_expr_placeholder(), c_pos); + parent.push(0); } if (is_instance && c_name.is_anonymous()) { if (used_match_idx()) @@ -247,11 +258,14 @@ expr parse_single_header(dummy_def_parser & p, declaration_name_scope & scope, b return mk_local(c_name, type); } -void parse_mutual_header(parser & p, buffer & lp_names, buffer & cs, buffer & params) { - parse_univ_params(p, lp_names); +void parse_mutual_header(parser & p, ast_data & parent, buffer & lp_names, buffer & cs, buffer & params) { + auto& names = p.new_ast("mutuals", p.pos()); + parent.push(parse_univ_params(p, lp_names)).push(names.m_id); while (true) { auto c_pos = p.pos(); - name c_name = p.check_decl_id_next("invalid mutual declaration, identifier expected"); + ast_id c_id; name c_name; + std::tie(c_id, c_name) = p.check_decl_id_next("invalid mutual declaration, identifier expected"); + names.push(c_id); cs.push_back(p.save_pos(mk_local(c_name, mk_expr_placeholder()), c_pos)); if (!p.curr_is_token(get_comma_tk())) break; @@ -260,19 +274,23 @@ void parse_mutual_header(parser & p, buffer & lp_names, buffer & c if (cs.size() < 2) { throw parser_error("invalid mutual declaration, must provide more than one identifier (separated by commas)", p.pos()); } - p.parse_optional_binders(params, /* allow_default */ true, /* explicit delimiters */ true); + auto& bis = p.new_ast("binders", p.pos()); + p.parse_optional_binders(&bis, params, /* allow_default */ true, /* explicit delimiters */ true); + parent.push(bis.m_children.empty() ? 0 : bis.m_id); for (expr const & param : params) p.add_local(param); for (expr const & c : cs) p.add_local(c); } -pair parse_inner_header(parser & p, name const & c_expected) { +pair parse_inner_header(parser & p, ast_data & parent, name const & c_expected) { decl_attributes attrs; p.check_token_next(get_with_tk(), "invalid mutual declaration, 'with' expected"); - attrs.parse(p); + parent.push(attrs.parse(p)); auto id_pos = p.pos(); - name n = p.check_decl_id_next("invalid mutual declaration, identifier expected"); + ast_id n_id; name n; + std::tie(n_id, n) = p.check_decl_id_next("invalid mutual declaration, identifier expected"); + parent.push(n_id); if (c_expected != n) throw parser_error(sstream() << "invalid mutual declaration, '" << c_expected << "' expected", id_pos); @@ -280,7 +298,9 @@ pair parse_inner_header(parser & p, name const & c_expect before invoking this procedure. */ declaration_name_scope scope(n); p.check_token_next(get_colon_tk(), "invalid mutual declaration, ':' expected"); - return mk_pair(p.parse_expr(), attrs); + expr e = p.parse_expr(); + parent.push(p.get_id(e)); + return {e, attrs}; } /** \brief Version of collect_locals(expr const & e, collected_locals & ls) that ignores local constants occurring in diff --git a/src/frontends/lean/decl_util.h b/src/frontends/lean/decl_util.h index e1713ac6dd..1c1b4fb2bc 100644 --- a/src/frontends/lean/decl_util.h +++ b/src/frontends/lean/decl_util.h @@ -128,7 +128,7 @@ bool used_match_idx(); {u_1 ... u_k} The universe parameters are automatically added to the parser scope. */ -bool parse_univ_params(parser & p, buffer & lp_names); +ast_id parse_univ_params(parser & p, buffer & lp_names); /** \brief Parse a declaration header of the form @@ -140,7 +140,7 @@ bool parse_univ_params(parser & p, buffer & lp_names); Both lp_names and params are added to the parser scope. \remark Caller is responsible for using: parser::local_scope scope2(p, env); */ -expr parse_single_header(parser & p, declaration_name_scope & s, buffer & lp_names, buffer & params, +expr parse_single_header(parser & p, ast_data & parent, declaration_name_scope & s, buffer & lp_names, buffer & params, bool is_example = false, bool is_instance = false); expr parse_single_header(dummy_def_parser & p, declaration_name_scope & s, buffer & lp_names, buffer & params, @@ -158,14 +158,14 @@ expr parse_single_header(dummy_def_parser & p, declaration_name_scope & s, buffe \remark Caller is responsible for adding expressions encoding the c_names to the parser scope. \remark Caller is responsible for using: parser::local_scope scope2(p, env); */ -void parse_mutual_header(parser & p, buffer & lp_names, buffer & cs, buffer & params); +void parse_mutual_header(parser & p, ast_data & parent, buffer & lp_names, buffer & cs, buffer & params); /** \brief Parse the header for one of the declarations in a mutually recursive declaration. It has the form with id : type The result is (type, attrs). */ -pair parse_inner_header(parser & p, name const & c_expected); +pair parse_inner_header(parser & p, ast_data & parent, name const & c_expected); /** \brief Add section/namespace parameters (and universes) used by params and all_exprs. We also add parameters included using the command 'include'. diff --git a/src/frontends/lean/definition_cmds.cpp b/src/frontends/lean/definition_cmds.cpp index d17e2f0cb0..8ce1f87254 100644 --- a/src/frontends/lean/definition_cmds.cpp +++ b/src/frontends/lean/definition_cmds.cpp @@ -50,37 +50,50 @@ environment ensure_decl_namespaces(environment const & env, name const & full_n) return add_namespace(env, full_n.get_prefix()); } -expr parse_equation_lhs(parser & p, expr const & fn, buffer & locals) { +expr parse_equation_lhs(parser & p, ast_data & parent, expr const & fn, buffer & locals) { auto lhs_pos = p.pos(); buffer lhs_args; + auto& data = p.new_ast("lhs", lhs_pos); + parent.push(data.m_id); lhs_args.push_back(p.parse_pattern_or_expr(get_max_prec())); + data.push(p.get_id(lhs_args.back())); while (!p.curr_is_token(get_assign_tk())) { auto pos0 = p.pos(); lhs_args.push_back(p.parse_pattern_or_expr(get_max_prec())); + data.push(p.get_id(lhs_args.back())); if (p.pos() == pos0) break; } expr lhs = p.mk_app(p.save_pos(mk_explicit(fn), lhs_pos), lhs_args, lhs_pos); bool skip_main_fn = true; - return p.patexpr_to_pattern(lhs, skip_main_fn, locals); + lhs = p.patexpr_to_pattern(lhs, skip_main_fn, locals); + p.set_ast_pexpr(data.m_id, lhs); + return lhs; } -expr parse_equation(parser & p, expr const & fn) { +expr parse_equation(parser & p, ast_data & parent, expr const & fn) { + auto& data = p.new_ast("eqn", p.pos()); + parent.push(data.m_id); p.check_token_next(get_bar_tk(), "invalid equation, '|' expected"); buffer locals; - expr lhs = parse_equation_lhs(p, fn, locals); + expr lhs = parse_equation_lhs(p, data, fn, locals); auto assign_pos = p.pos(); p.check_token_next(get_assign_tk(), "invalid equation, ':=' expected"); expr rhs = p.parse_scoped_expr(locals); + data.push(p.get_id(rhs)); return Fun(locals, p.save_pos(mk_equation(lhs, rhs), assign_pos), p); } -optional parse_using_well_founded(parser & p) { +optional parse_using_well_founded(parser & p, ast_data & parent) { if (p.curr_is_token(get_using_well_founded_tk())) { + auto& data = p.new_ast(get_using_well_founded_tk(), p.pos()); parser::local_scope _(p); p.clear_expr_locals(); p.next(); - return some_expr(p.parse_expr(get_max_prec())); + expr e = p.parse_expr(get_max_prec()); + data.push(p.get_id(e)); + return some_expr(e); } else { + parent.push(0); return none_expr(); } } @@ -118,30 +131,36 @@ void check_valid_end_of_equations(parser & p) { } } -static expr parse_mutual_definition(parser & p, buffer & lp_names, buffer & fns, buffer & prv_names, +static expr parse_mutual_definition(parser & p, ast_data & parent, buffer & lp_names, buffer & fns, buffer & prv_names, buffer & params) { parser::local_scope scope1(p); auto header_pos = p.pos(); buffer pre_fns; - parse_mutual_header(p, lp_names, pre_fns, params); + parse_mutual_header(p, parent, lp_names, pre_fns, params); buffer eqns; buffer full_names; buffer full_actual_names; + auto& fn_asts = p.new_ast("bodies", p.pos()); + parent.push(fn_asts.m_id); for (expr const & pre_fn : pre_fns) { // TODO(leo, dhs): make use of attributes - expr fn_type = parse_inner_header(p, mlocal_pp_name(pre_fn)).first; + auto& fn_ast = p.new_ast("body", p.pos()); + fn_asts.push(fn_ast.m_id); + expr fn_type = parse_inner_header(p, fn_ast, mlocal_pp_name(pre_fn)).first; declaration_name_scope scope2(mlocal_pp_name(pre_fn)); declaration_name_scope scope3("_main"); full_names.push_back(scope3.get_name()); full_actual_names.push_back(scope3.get_actual_name()); prv_names.push_back(scope2.get_actual_name()); + auto& eqn_asts = p.new_ast("eqns", p.pos()); + fn_ast.push(eqn_asts.m_id); if (p.curr_is_token(get_period_tk())) { auto period_pos = p.pos(); p.next(); eqns.push_back(p.save_pos(mk_no_equation(), period_pos)); } else { while (p.curr_is_token(get_bar_tk())) { - eqns.push_back(parse_equation(p, pre_fn)); + eqns.push_back(parse_equation(p, eqn_asts, pre_fn)); } check_valid_end_of_equations(p); } @@ -150,7 +169,7 @@ static expr parse_mutual_definition(parser & p, buffer & lp_names, buffer< } if (p.curr_is_token(get_with_tk())) p.maybe_throw_error({"unexpected 'with' clause", p.pos()}); - optional wf_tacs = parse_using_well_founded(p); + optional wf_tacs = parse_using_well_founded(p, parent); for (expr & eq : eqns) { eq = replace_locals_preserving_pos_info(eq, pre_fns, fns); } @@ -458,7 +477,7 @@ static environment copy_equation_lemmas(environment const & env, name const & d_ return copy_equation_lemmas(env, d_names); } -static environment mutual_definition_cmd_core(parser & p, decl_cmd_kind kind, cmd_meta const & meta) { +static environment mutual_definition_cmd_core(parser & p, decl_cmd_kind kind, ast_data & parent, cmd_meta const & meta) { buffer lp_names; buffer fns, params; declaration_info_scope scope(p, kind, meta.m_modifiers); @@ -468,7 +487,7 @@ static environment mutual_definition_cmd_core(parser & p, decl_cmd_kind kind, cm environment env = p.env(); private_name_scope prv_scope(meta.m_modifiers.m_is_private, env); buffer prv_names; - expr val = parse_mutual_definition(p, lp_names, fns, prv_names, params); + expr val = parse_mutual_definition(p, parent, lp_names, fns, prv_names, params); // skip elaboration of definitions during reparsing if (p.get_break_at_pos()) @@ -521,14 +540,16 @@ static environment mutual_definition_cmd_core(parser & p, decl_cmd_kind kind, cm Note that mlocal_pp_name(fn) and actual_name are different for scoped/private declarations. */ std::tuple -parser::parse_definition(buffer & lp_names, buffer & params, +parser::parse_definition(ast_data * _parent, buffer & lp_names, buffer & params, bool is_example, bool is_instance, bool is_meta, bool is_abbrev) { parser & p = *this; + lean_always_assert(_parent); + auto& parent = *_parent; parser::local_scope scope1(p); auto header_pos = p.pos(); time_task _("parsing", p.mk_message(header_pos, INFORMATION), p.get_options()); declaration_name_scope scope2; - expr fn = parse_single_header(p, scope2, lp_names, params, is_example, is_instance); + expr fn = parse_single_header(p, parent, scope2, lp_names, params, is_example, is_instance); expr val; if (p.curr_is_token(get_assign_tk())) { p.next(); @@ -537,6 +558,7 @@ parser::parse_definition(buffer & lp_names, buffer & params, fn = mk_local(mlocal_name(fn), mlocal_pp_name(fn), mlocal_type(fn), mk_rec_info(true)); p.add_local(fn); val = p.parse_expr(); + parent.push(p.get_id(val)); /* add fake equation */ expr eqn = copy_tag(val, mk_equation(fn, val)); buffer eqns; @@ -544,10 +566,13 @@ parser::parse_definition(buffer & lp_names, buffer & params, val = mk_equations(p, fn, scope2.get_name(), scope2.get_actual_name(), eqns, {}, header_pos); } else { val = p.parse_expr(); + parent.push(p.get_id(val)); } } else if (p.curr_is_token(get_bar_tk()) || p.curr_is_token(get_period_tk())) { if (is_abbrev) throw exception("invalid abbreviation, abbreviations should not be defined using pattern matching"); + auto& eqn_asts = p.new_ast("eqns", p.pos()); + parent.push(eqn_asts.m_id); declaration_name_scope scope2("_main"); fn = mk_local(mlocal_name(fn), mlocal_pp_name(fn), mlocal_type(fn), mk_rec_info(true)); p.add_local(fn); @@ -558,11 +583,11 @@ parser::parse_definition(buffer & lp_names, buffer & params, eqns.push_back(p.save_pos(mk_no_equation(), period_pos)); } else { while (p.curr_is_token(get_bar_tk())) { - eqns.push_back(parse_equation(p, fn)); + eqns.push_back(parse_equation(p, eqn_asts, fn)); } check_valid_end_of_equations(p); } - optional wf_tacs = parse_using_well_founded(p); + optional wf_tacs = parse_using_well_founded(p, parent); val = mk_equations(p, fn, scope2.get_name(), scope2.get_actual_name(), eqns, wf_tacs, header_pos); } else { val = p.parser_error_or_expr({"invalid definition, '|' or ':=' expected", p.pos()}); @@ -572,7 +597,7 @@ parser::parse_definition(buffer & lp_names, buffer & params, } std::tuple -dummy_def_parser::parse_definition(buffer & lp_names, buffer & params, +dummy_def_parser::parse_definition(ast_data *, buffer & lp_names, buffer & params, bool is_example, bool is_instance, bool, bool) { dummy_def_parser & p = *this; parser::local_scope scope1(p); @@ -761,7 +786,7 @@ static bool is_rfl_preexpr(expr const & e) { return is_constant(e, get_rfl_name()); } -environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, cmd_meta meta) { +environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, ast_data * parent, cmd_meta meta) { buffer lp_names; buffer params; expr fn, val; @@ -782,7 +807,7 @@ environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, cmd_ meta.m_attrs.set_attribute(env, "reducible"); } name prv_name; - std::tie(fn, val, prv_name) = p.parse_definition(lp_names, params, is_example, is_instance, meta.m_modifiers.m_is_meta, is_abbrev); + std::tie(fn, val, prv_name) = p.parse_definition(parent, lp_names, params, is_example, is_instance, meta.m_modifiers.m_is_meta, is_abbrev); auto begin_pos = p.cmd_pos(); auto end_pos = p.pos(); @@ -910,10 +935,14 @@ environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, cmd_ } } -environment definition_cmd_core(parser & p, decl_cmd_kind kind, cmd_meta const & meta) { - if (meta.m_modifiers.m_is_mutual) - return mutual_definition_cmd_core(p, kind, meta); - else - return single_definition_cmd_core(p, kind, meta); +environment definition_cmd_core(parser & p, decl_cmd_kind kind, ast_id cmd_id, cmd_meta const & meta) { + auto& data = p.get_ast(cmd_id).push(meta.m_modifiers_id); + if (meta.m_modifiers.m_is_mutual) { + data.push(p.new_ast("mutual", p.pos()).m_id); + return mutual_definition_cmd_core(p, kind, data, meta); + } else { + data.push(0); + return single_definition_cmd_core(p, kind, &data, meta); + } } } diff --git a/src/frontends/lean/definition_cmds.h b/src/frontends/lean/definition_cmds.h index b9bd320aab..3627578305 100644 --- a/src/frontends/lean/definition_cmds.h +++ b/src/frontends/lean/definition_cmds.h @@ -9,9 +9,9 @@ Author: Leonardo de Moura #include "frontends/lean/decl_attributes.h" #include "frontends/lean/decl_util.h" namespace lean { -environment definition_cmd_core(parser & p, decl_cmd_kind k, cmd_meta const & meta); +environment definition_cmd_core(parser & p, decl_cmd_kind k, ast_id cmd_id, cmd_meta const & meta); -environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, cmd_meta meta); +environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, ast_data * parent, cmd_meta meta); environment ensure_decl_namespaces(environment const & env, name const & full_n); } diff --git a/src/frontends/lean/elaborator.cpp b/src/frontends/lean/elaborator.cpp index 1b4b2b18b5..f7cf8850c2 100644 --- a/src/frontends/lean/elaborator.cpp +++ b/src/frontends/lean/elaborator.cpp @@ -4302,7 +4302,7 @@ static vm_obj environment_add_defn_eqns(vm_obj const &_env, vm_obj const &_opts, to_buffer_expr(cfield(o, 0), pat); p.m_val.push_back(std::pair, expr>(std::move(pat), mk_as_is(abstract(to_expr(cfield(o, 1)), sig)))); } - return mk_vm_exceptional_success(to_obj(single_definition_cmd_core(p, kind, meta))); + return mk_vm_exceptional_success(to_obj(single_definition_cmd_core(p, kind, nullptr, meta))); } catch (throwable & ex) { return mk_vm_exceptional_exception(ex); } diff --git a/src/frontends/lean/inductive_cmds.cpp b/src/frontends/lean/inductive_cmds.cpp index ccf8663ff3..8988eee4e9 100644 --- a/src/frontends/lean/inductive_cmds.cpp +++ b/src/frontends/lean/inductive_cmds.cpp @@ -88,10 +88,11 @@ static level subtract_from_max(level const & l, unsigned offset) { class inductive_cmd_fn { parser & m_p; + ast_data & m_cmd; environment m_env; cmd_meta m_meta_info; buffer m_mut_attrs; - type_context_old m_ctx; + type_context_old m_ctx; buffer m_lp_names; pos_info m_pos; name_map m_implicit_infer_map; @@ -361,20 +362,23 @@ class inductive_cmd_fn { } } - void parse_intro_rules(bool has_params, expr const & ind, buffer & intro_rules, + ast_id parse_intro_rules(bool has_params, expr const & ind, buffer & intro_rules, buffer> & intro_rule_docs, bool prepend_ns) { + auto& asts = m_p.new_ast("intros", m_p.pos()); + // If the next token is neither `|` nor a doc block, then the inductive type has no constructors if (!m_p.curr_is_token(get_bar_tk()) && (m_p.curr() != token_kind::DocBlock)) { - return; + return asts.m_id; } while (true) { + ast_id doc_id = 0; optional doc {}; if (m_p.curr() == token_kind::DocBlock) { // If the next token is a doc block, it applies to the following constructor, *if one exists*. // Otherwise, it might be part of the next command. if (m_p.ahead_is_token(get_bar_tk()) || m_p.ahead_is_token(get_comma_tk())) { - doc = m_p.parse_doc_block(); + std::tie(doc_id, doc) = m_p.parse_doc_block(); } else { break; } @@ -386,13 +390,20 @@ class inductive_cmd_fn { m_p.next(); m_pos = m_p.pos(); - name ir_name = mlocal_name(ind) + m_p.check_atomic_id_next("invalid introduction rule, atomic identifier expected"); + auto& ast = m_p.new_ast("intro", m_pos); + asts.push(ast.m_id); + ast_id ir_ast; name ir_name; + std::tie(ir_ast, ir_name) = m_p.check_atomic_id_next("invalid introduction rule, atomic identifier expected"); + ir_name = mlocal_name(ind) + ir_name; if (prepend_ns) ir_name = get_namespace(m_env) + ir_name; parser::local_scope S(m_p); buffer params; + ast_id kind_id = 0; implicit_infer_kind kind = implicit_infer_kind::RelaxedImplicit; - m_p.parse_optional_binders(params, kind); + auto& bis = m_p.new_ast("binders", m_p.pos()); + m_p.parse_optional_binders(&bis, params, kind_id, kind); + ast.push(doc_id).push(ir_ast).push(kind_id).push(bis.m_children.empty() ? 0 : bis.m_id); m_implicit_infer_map.insert(ir_name, kind); for (expr const & param : params) m_p.add_local(param); @@ -400,14 +411,17 @@ class inductive_cmd_fn { if (has_params || m_p.curr_is_token(get_colon_tk())) { m_p.check_token_next(get_colon_tk(), "invalid introduction rule, ':' expected"); ir_type = m_p.parse_expr(); + ast.push(m_p.get_id(ir_type)); } else { ir_type = ind; + ast.push(0); } ir_type = Pi(params, ir_type, m_p); intro_rules.push_back(mk_local(ir_name, ir_type)); intro_rule_docs.push_back(doc); lean_trace(name({"inductive", "parse"}), tout() << ir_name << " : " << ir_type << "\n";); } + return asts.m_id; } /** \brief Add a namespace for each inductive datatype */ @@ -593,7 +607,7 @@ class inductive_cmd_fn { m_pos = m_p.pos(); declaration_name_scope nscope; - expr ind = parse_single_header(m_p, nscope, m_lp_names, params); + expr ind = parse_single_header(m_p, m_cmd, nscope, m_lp_names, params); m_explicit_levels = !m_lp_names.empty(); m_mut_attrs.push_back({}); @@ -603,9 +617,8 @@ class inductive_cmd_fn { tout() << mlocal_name(ind) << " : " << mlocal_type(ind) << "\n";); m_p.add_local(ind); - m_p.parse_local_notation_decl(); - - parse_intro_rules(!params.empty(), ind, intro_rules, intro_rule_docs, false); + m_cmd.push(m_p.parse_local_notation_decl()) + .push(parse_intro_rules(!params.empty(), ind, intro_rules, intro_rule_docs, false)); buffer ind_intro_rules; ind_intro_rules.push_back(ind); @@ -626,21 +639,24 @@ class inductive_cmd_fn { parser::local_scope scope(m_p); buffer pre_inds; - parse_mutual_header(m_p, m_lp_names, pre_inds, params); + parse_mutual_header(m_p, m_cmd, m_lp_names, pre_inds, params); m_explicit_levels = !m_lp_names.empty(); - m_p.parse_local_notation_decl(); + auto& ind_asts = m_p.new_ast("inds", m_p.pos()); + m_cmd.push(m_p.parse_local_notation_decl()).push(ind_asts.m_id); for (expr const & pre_ind : pre_inds) { m_pos = m_p.pos(); + auto& ind_ast = m_p.new_ast("ind", m_pos); + ind_asts.push(ind_ast.m_id); expr ind_type; decl_attributes attrs; - std::tie(ind_type, attrs) = parse_inner_header(m_p, mlocal_pp_name(pre_ind)); + std::tie(ind_type, attrs) = parse_inner_header(m_p, ind_ast, mlocal_pp_name(pre_ind)); check_attrs(attrs); m_mut_attrs.push_back(attrs); lean_trace(name({"inductive", "parse"}), tout() << mlocal_name(pre_ind) << " : " << ind_type << "\n";); intro_rules.emplace_back(); intro_rule_docs.emplace_back(); expr ind = mk_local(resolve_decl_name(m_p.env(), mlocal_name(pre_ind)), ind_type); - parse_intro_rules(!params.empty(), ind, intro_rules.back(), intro_rule_docs.back(), false); + ind_ast.push(parse_intro_rules(!params.empty(), ind, intro_rules.back(), intro_rule_docs.back(), false)); inds.push_back(ind); } @@ -672,8 +688,8 @@ class inductive_cmd_fn { throw_error("invalid 'protected' modifier for inductive type"); } public: - inductive_cmd_fn(parser & p, cmd_meta const & meta): - m_p(p), m_env(p.env()), m_meta_info(meta), m_ctx(p.env()) { + inductive_cmd_fn(parser & p, ast_id cmd_id, cmd_meta const & meta): + m_p(p), m_cmd(p.get_ast(cmd_id)), m_env(p.env()), m_meta_info(meta), m_ctx(p.env()) { m_u_meta = m_ctx.mk_univ_metavar_decl(); check_attrs(m_meta_info.m_attrs); check_modifiers(); @@ -728,9 +744,12 @@ class inductive_cmd_fn { buffer > intro_rules; // intro_rule_docs do not have to go through elaboration, so we don't need a temporary here. + auto& data = m_cmd.push(m_meta_info.m_modifiers_id); if (m_meta_info.m_modifiers.m_is_mutual) { + data.push(m_p.new_ast("mutual", m_pos).m_id); parse_mutual_inductive(params, inds, intro_rules, result.m_intro_rule_docs); } else { + data.push(0); intro_rules.emplace_back(); result.m_intro_rule_docs.emplace_back(); inds.push_back(parse_inductive(params, intro_rules.back(), result.m_intro_rule_docs.back())); @@ -763,17 +782,17 @@ class inductive_cmd_fn { } }; -inductive_decl parse_inductive_decl(parser & p, cmd_meta const & meta) { +inductive_decl parse_inductive_decl(parser & p, ast_id cmd_id, cmd_meta const & meta) { auto pos = p.pos(); module::scope_pos_info scope_pos(pos); - return inductive_cmd_fn(p, meta).parse_and_elaborate(); + return inductive_cmd_fn(p, cmd_id, meta).parse_and_elaborate(); } -environment inductive_cmd(parser & p, cmd_meta const & meta) { +environment inductive_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { p.next(); auto pos = p.pos(); module::scope_pos_info scope_pos(pos); - return inductive_cmd_fn(p, meta).inductive_cmd(); + return inductive_cmd_fn(p, cmd_id, meta).inductive_cmd(); } void register_inductive_cmds(cmd_table & r) { diff --git a/src/frontends/lean/inductive_cmds.h b/src/frontends/lean/inductive_cmds.h index 6309c2ddf4..76c6e4d963 100644 --- a/src/frontends/lean/inductive_cmds.h +++ b/src/frontends/lean/inductive_cmds.h @@ -20,8 +20,8 @@ struct inductive_decl { buffer m_decls; }; -inductive_decl parse_inductive_decl(parser & p, cmd_meta const & meta); -environment inductive_cmd(parser & p, cmd_meta const & meta); +inductive_decl parse_inductive_decl(parser & p, ast_id cmd_id, cmd_meta const & meta); +environment inductive_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta); void register_inductive_cmds(cmd_table & r); void initialize_inductive_cmds(); diff --git a/src/frontends/lean/json.cpp b/src/frontends/lean/json.cpp index e9dd376620..44a76e6505 100644 --- a/src/frontends/lean/json.cpp +++ b/src/frontends/lean/json.cpp @@ -105,6 +105,8 @@ json serialize_decl(name const & d, environment const & env, options const & o) } json json_of_name(name const & n0) { + if (n0.is_atomic() && n0.is_string()) return n0.get_string(); + json j = json::array(); name n = n0; diff --git a/src/frontends/lean/json.h b/src/frontends/lean/json.h index de9d993cef..a92db0d61d 100644 --- a/src/frontends/lean/json.h +++ b/src/frontends/lean/json.h @@ -8,18 +8,18 @@ Author: Gabriel Ebner #pragma once #include "library/messages.h" #include "kernel/environment.h" -#include "util/json.hpp" +#include "util/lean_json.h" namespace lean { -using json = nlohmann::json; - json json_of_severity(message_severity sev); json json_of_message(message const & msg); json json_of_name(name const &); +json to_json(name const &); + void add_source_info(environment const & env, name const & d, json & record); json serialize_decl(name const & short_name, name const & long_name, environment const & env, options const & o); json serialize_decl(name const & d, environment const & env, options const & o); diff --git a/src/frontends/lean/match_expr.cpp b/src/frontends/lean/match_expr.cpp index f7bdf3e97b..b1c0aa3272 100644 --- a/src/frontends/lean/match_expr.cpp +++ b/src/frontends/lean/match_expr.cpp @@ -25,40 +25,55 @@ expr parse_match(parser & p, unsigned, expr const *, pos_info const & pos) { equations_header header = mk_match_header(scope2.get_name(), scope2.get_actual_name()); buffer eqns; buffer ts; + auto& args = p.new_ast("args", p.pos()); + auto& data = p.new_ast("match", pos).push(args.m_id); try { ts.push_back(p.parse_expr()); + args.push(p.get_id(ts.back())); while (p.curr_is_token(get_comma_tk())) { p.next(); ts.push_back(p.parse_expr()); + args.push(p.get_id(ts.back())); } expr fn; /* Parse optional type */ if (p.curr_is_token(get_colon_tk())) { p.next(); expr type = p.parse_expr(); + data.push(p.get_id(type)); fn = mk_local(p.next_name(), *g_match_name, type, mk_rec_info(true)); } else { + data.push(0); fn = mk_local(p.next_name(), *g_match_name, mk_expr_placeholder(), mk_rec_info(true)); } p.check_token_next(get_with_tk(), "invalid 'match' expression, 'with' expected"); + auto& eqn_asts = p.new_ast("eqns", p.pos()); + data.push(eqn_asts.m_id); if (p.curr_is_token(get_end_tk())) { /* Empty match */ p.next(); eqns.push_back(Fun(fn, mk_no_equation())); expr f = p.save_pos(mk_equations(header, eqns.size(), eqns.data()), pos); - return p.mk_app(f, ts, pos); + expr r = p.mk_app(f, ts, pos); + p.set_ast_pexpr(data.m_id, r); + return r; } if (is_eqn_prefix(p)) p.next(); // optional '|' in the first case while (true) { auto lhs_pos = p.pos(); buffer lhs_args; + auto& lhs_ast = p.new_ast("lhs", lhs_pos); + auto& eqn_ast = p.new_ast("eqn", lhs_pos).push(lhs_ast.m_id); + eqn_asts.push(eqn_ast.m_id); lhs_args.push_back(p.parse_pattern_or_expr()); + lhs_ast.push(p.get_id(lhs_args.back())); while (p.curr_is_token(get_comma_tk())) { p.next(); lhs_args.push_back(p.parse_pattern_or_expr()); + lhs_ast.push(p.get_id(lhs_args.back())); } expr lhs = p.mk_app(fn, lhs_args, lhs_pos); buffer locals; @@ -71,6 +86,7 @@ expr parse_match(parser & p, unsigned, expr const *, pos_info const & pos) { for (expr const & local : locals) p.add_local(local); expr rhs = p.parse_expr(); + eqn_ast.push(p.get_id(rhs)); eqns.push_back(Fun(fn, Fun(locals, p.save_pos(mk_equation(lhs, rhs), assign_pos), p), p)); } if (!is_eqn_prefix(p)) @@ -83,7 +99,9 @@ expr parse_match(parser & p, unsigned, expr const *, pos_info const & pos) { } p.check_token_next(get_end_tk(), "invalid 'match' expression, 'end' expected"); expr f = p.save_pos(mk_equations(header, eqns.size(), eqns.data()), pos); - return p.mk_app(f, ts, pos); + expr r = p.mk_app(f, ts, pos); + p.set_ast_pexpr(data.m_id, r); + return r; } void initialize_match_expr() { diff --git a/src/frontends/lean/module_parser.cpp b/src/frontends/lean/module_parser.cpp index e95d0e556d..43550e60e3 100644 --- a/src/frontends/lean/module_parser.cpp +++ b/src/frontends/lean/module_parser.cpp @@ -77,7 +77,7 @@ module_parser::parse_next_command_like(optional> const & depe report_message(ex); self->m_parser.sync_command(); } catch (throwable & ex) { - self->m_parser.mk_message(self->m_parser.m_last_cmd_pos, ERROR).set_exception(ex).report(); + self->m_parser.mk_message(self->m_parser.cmd_pos(), ERROR).set_exception(ex).report(); self->m_parser.sync_command(); } catch (interrupt_parser) { // this exception is thrown by the exit command diff --git a/src/frontends/lean/module_parser.h b/src/frontends/lean/module_parser.h index 9acd995b0f..d18c1a7be7 100644 --- a/src/frontends/lean/module_parser.h +++ b/src/frontends/lean/module_parser.h @@ -48,6 +48,7 @@ class module_parser : public std::enable_shared_from_this { void use_separate_tasks(bool separate_tasks) { m_separate_tasks = separate_tasks; } void save_info(bool save) { m_save_info = save; } + parser & get_parser() { return m_parser; } void break_at_pos(pos_info const & pos, bool complete); pair> diff --git a/src/frontends/lean/notation_cmd.cpp b/src/frontends/lean/notation_cmd.cpp index 478db5681f..3908f640c5 100644 --- a/src/frontends/lean/notation_cmd.cpp +++ b/src/frontends/lean/notation_cmd.cpp @@ -24,8 +24,9 @@ Author: Leonardo de Moura #include "frontends/lean/decl_attributes.h" namespace lean { -static std::string parse_symbol(parser & p, char const * msg) { +static std::string parse_symbol(parser & p, ast_data & parent, char const * msg) { name n; + name type = p.curr_is_quoted_symbol() ? "quoted" : "ident"; if (p.curr_is_identifier() || p.curr_is_quoted_symbol()) { n = p.get_name_val(); } else if (p.curr_is_keyword()) { @@ -33,11 +34,13 @@ static std::string parse_symbol(parser & p, char const * msg) { } else { throw parser_error(msg, p.pos()); } + auto s = n.to_string_unescaped(); + parent.push(p.new_ast(type, p.pos(), s).m_id); p.next(); - return n.to_string_unescaped(); + return s; } -static unsigned parse_precedence_core(parser & p) { +static pair parse_precedence_core(parser & p) { auto pos = p.pos(); if (p.curr_is_numeral()) { return p.parse_small_nat(); @@ -47,28 +50,31 @@ static unsigned parse_precedence_core(parser & p) { env = open_prec_aliases(env); parser::local_scope scope(p, env); expr pre_val = p.parse_expr(get_max_prec()); + auto expr_id = p.new_ast("expr", pos).push(p.get_id(pre_val)).m_id; expr nat = mk_constant(get_nat_name()); pre_val = mk_typed_expr(nat, pre_val); expr val = p.elaborate("notation", list(), pre_val).first; vm_obj p = eval_closed_expr(env, opts, "_precedence", nat, val, pos); if (optional _p = try_to_unsigned(p)) { - return *_p; + return {expr_id, *_p}; } else { throw parser_error("invalid 'precedence', argument does not evaluate to a small numeral", pos); } } } -static optional parse_optional_precedence(parser & p) { +static pair> parse_optional_precedence(parser & p) { if (p.curr_is_token(get_colon_tk())) { p.next(); - return some(parse_precedence_core(p)); + pair> r; + std::tie(r.first, r.second) = parse_precedence_core(p); + return r; } else { - return optional(); + return {0, optional()}; } } -static unsigned parse_precedence(parser & p) { +static pair parse_precedence(parser & p) { return parse_precedence_core(p); } @@ -105,12 +111,12 @@ static optional get_precedence(environment const & env, char const * t return get_expr_precedence(get_token_table(env), tk); } -static auto parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, notation_entry_group grp, bool parse_only, +static auto parse_mixfix_notation(parser & p, ast_data & parent, mixfix_kind k, bool overload, notation_entry_group grp, bool parse_only, unsigned priority) -> pair> { bool explicit_pp = p.curr_is_quoted_symbol(); pos_info tk_pos = p.pos(); - std::string pp_tk = parse_symbol(p, "invalid notation declaration, quoted symbol or identifier expected"); + std::string pp_tk = parse_symbol(p, parent, "invalid notation declaration, quoted symbol or identifier expected"); std::string tk = utf8_trim(pp_tk); char const * tks = tk.c_str(); check_not_forbidden(tks); @@ -140,11 +146,13 @@ static auto parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, nota } pos_info prec_pos; + ast_id prec_id = 0; if (p.curr_is_token(get_colon_tk())) { p.next(); prec_pos = p.pos(); - prec = parse_precedence(p); + std::tie(prec_id, prec) = parse_precedence(p); } + parent.push(prec_id); if (prec && k == mixfix_kind::infixr && *prec == 0) throw parser_error("invalid infixr declaration, precedence must be greater than zero", prec_pos); @@ -205,6 +213,7 @@ static auto parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, nota if (grp == notation_entry_group::Reserve) { // reserve notation commands do not have a denotation + parent.push(0); expr dummy = mk_Prop(); if (p.curr_is_token(get_assign_tk())) throw parser_error("invalid reserve notation, found `:=`", p.pos()); @@ -227,6 +236,7 @@ static auto parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, nota auto f_pos = p.pos(); expr f = p.parse_expr(); check_notation_expr(f, f_pos); + parent.push(p.get_id(f)); switch (k) { case mixfix_kind::infixl: #if defined(__GNUC__) && !defined(__CLANG__) @@ -248,35 +258,42 @@ static auto parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, nota lean_unreachable(); // LCOV_EXCL_LINE } -static notation_entry parse_mixfix_notation(parser & p, mixfix_kind k, bool overload, notation_entry_group grp, +static notation_entry parse_mixfix_notation(parser & p, ast_data & data, mixfix_kind k, bool overload, notation_entry_group grp, buffer & new_tokens, bool parse_only, unsigned priority) { - auto nt = parse_mixfix_notation(p, k, overload, grp, parse_only, priority); + auto nt = parse_mixfix_notation(p, data, k, overload, grp, parse_only, priority); if (nt.second) new_tokens.push_back(*nt.second); return nt.first; } -static name parse_quoted_symbol_or_token(parser & p, buffer & new_tokens, bool & used_default) { +static name parse_quoted_symbol_or_token(parser & p, ast_data & data, buffer & new_tokens, bool & used_default) { used_default = false; + auto& sym = p.new_ast("sym", p.pos()); + data.push(sym.m_id); if (p.curr_is_quoted_symbol()) { environment const & env = p.env(); auto pp_tk = p.get_name_val(); + sym.push(p.new_ast("quoted", p.pos(), pp_tk).m_id); auto tks = utf8_trim(pp_tk.to_string_unescaped()); auto tkcs = tks.c_str(); check_not_forbidden(tkcs); p.next(); + ast_id prec_id = 0; if (p.curr_is_token(get_colon_tk())) { p.next(); - unsigned prec = parse_precedence(p); + unsigned prec; + std::tie(prec_id, prec) = parse_precedence(p); new_tokens.push_back(mk_token_entry(tkcs, prec)); } else if (!get_precedence(env, tkcs)) { new_tokens.push_back(mk_token_entry(tkcs, LEAN_DEFAULT_PRECEDENCE)); used_default = true; } + sym.push(prec_id); return pp_tk; } else if (p.curr_is_keyword()) { auto tk = p.get_token_info().token(); check_not_forbidden(tk.to_string_unescaped().c_str()); + sym.push(p.new_ast("ident", p.pos(), tk).m_id).push(0); p.next(); return tk; } else { @@ -284,27 +301,30 @@ static name parse_quoted_symbol_or_token(parser & p, buffer & new_t } } -static name parse_quoted_symbol_or_token(parser & p, buffer & new_tokens) { +static name parse_quoted_symbol_or_token(parser & p, ast_data & data, buffer & new_tokens) { bool dummy; - return parse_quoted_symbol_or_token(p, new_tokens, dummy); + return parse_quoted_symbol_or_token(p, data, new_tokens, dummy); } -static expr parse_notation_expr(parser & p, buffer const & locals) { +static pair parse_notation_expr(parser & p, buffer const & locals) { auto pos = p.pos(); expr r = p.parse_expr(); + ast_id id = p.get_id(r); r = abstract(r, locals.size(), locals.data()); check_notation_expr(r, pos); - return r; + return {id, r}; } -static void parse_notation_local(parser & p, buffer & locals) { +static ast_id parse_notation_local(parser & p, buffer & locals) { if (p.curr_is_identifier()) { name n = p.get_name_val(); + ast_id id = p.new_ast("ident", p.pos(), n).m_id; p.next(); expr local_type = mk_Prop(); // type used in notation local declarations, it is irrelevant expr l = mk_local(n, local_type); // remark: the type doesn't matter p.add_local(l); locals.push_back(l); + return id; } else { throw parser_error("invalid notation declaration, identifier expected", p.pos()); } @@ -323,55 +343,76 @@ static unsigned get_precedence(environment const & env, buffer cons return 0; } -static action parse_action(parser & p, name const & prev_token, unsigned default_prec, +static action parse_action(parser & p, ast_data & parent, name const & prev_token, unsigned default_prec, buffer & locals, buffer & new_tokens) { if (p.curr_is_token(get_colon_tk())) { p.next(); if (p.curr_is_numeral() || p.curr_is_token_or_id(get_max_tk())) { - unsigned prec = parse_precedence(p); + ast_id id; unsigned prec; + std::tie(id, prec) = parse_precedence(p); + parent.push(id); return mk_expr_action(prec); } else if (p.curr_is_token_or_id(get_prev_tk())) { + parent.push(p.new_ast(get_prev_tk(), p.pos()).m_id); p.next(); return mk_expr_action(get_precedence(p.env(), new_tokens, prev_token)); } else if (p.curr_is_token_or_id(get_scoped_tk())) { + parent.push(p.new_ast(get_scoped_tk(), p.pos()).push(0).push(0).m_id); p.next(); return mk_scoped_expr_action(mk_var(0)); } else { p.check_token_next(get_lparen_tk(), "invalid notation declaration, '(', numeral or 'scoped' expected"); if (p.curr_is_token_or_id(get_foldl_tk()) || p.curr_is_token_or_id(get_foldr_tk())) { bool is_fold_right = p.curr_is_token_or_id(get_foldr_tk()); + auto& fold = p.new_ast(is_fold_right ? get_foldr_tk() : get_foldl_tk(), p.pos()); + parent.push(fold.m_id); p.next(); - auto prec = parse_optional_precedence(p); - name sep = parse_quoted_symbol_or_token(p, new_tokens); + ast_id id; optional prec; + std::tie(id, prec) = parse_optional_precedence(p); + fold.push(id); + name sep = parse_quoted_symbol_or_token(p, fold, new_tokens); + auto& sc = p.new_ast("scope", p.pos()); + fold.push(sc.m_id); expr rec; { parser::local_scope scope(p); p.check_token_next(get_lparen_tk(), "invalid fold notation argument, '(' expected"); - parse_notation_local(p, locals); - parse_notation_local(p, locals); + sc.push(parse_notation_local(p, locals)); + sc.push(parse_notation_local(p, locals)); p.check_token_next(get_comma_tk(), "invalid fold notation argument, ',' expected"); - rec = parse_notation_expr(p, locals); + std::tie(id, rec) = parse_notation_expr(p, locals); + sc.push(id); p.check_token_next(get_rparen_tk(), "invalid fold notation argument, ')' expected"); locals.pop_back(); locals.pop_back(); } - optional ini; + optional ini; id = 0; if (!p.curr_is_token(get_rparen_tk()) && !p.curr_is_quoted_symbol()) - ini = parse_notation_expr(p, locals); + std::tie(id, ini) = parse_notation_expr(p, locals); + fold.push(id); optional terminator; - if (!p.curr_is_token(get_rparen_tk())) - terminator = parse_quoted_symbol_or_token(p, new_tokens); + if (p.curr_is_token(get_rparen_tk())) + fold.push(0); + else + terminator = parse_quoted_symbol_or_token(p, fold, new_tokens); p.check_token_next(get_rparen_tk(), "invalid fold notation argument, ')' expected"); return mk_exprs_action(sep, rec, ini, terminator, is_fold_right, prec ? *prec : 0); } else if (p.curr_is_token_or_id(get_scoped_tk())) { + auto& sc = p.new_ast(get_scoped_tk(), p.pos()); + parent.push(sc.m_id); p.next(); - auto prec = parse_optional_precedence(p); + ast_id id; optional prec; + std::tie(id, prec) = parse_optional_precedence(p); + sc.push(id); + auto& sc1 = p.new_ast("scope", p.pos()); + sc.push(sc1.m_id); expr rec; { parser::local_scope scope(p); - parse_notation_local(p, locals); + sc1.push(parse_notation_local(p, locals)); p.check_token_next(get_comma_tk(), "invalid scoped notation argument, ',' expected"); - rec = parse_notation_expr(p, locals); + std::tie(id, rec) = parse_notation_expr(p, locals); + sc1.push(id); locals.pop_back(); } p.check_token_next(get_rparen_tk(), "invalid scoped notation argument, ')' expected"); @@ -381,6 +422,7 @@ static action parse_action(parser & p, name const & prev_token, unsigned default } } } else { + parent.push(0); return mk_expr_action(default_prec); } } @@ -425,30 +467,38 @@ static list> find_next(optional const return list>(); } -static unsigned parse_binders_rbp(parser & p) { +static pair parse_binders_rbp(parser & p) { if (p.curr_is_token(get_colon_tk())) { p.next(); return parse_precedence(p); } else { - return 0; + return {0, 0}; } } -static transition parse_transition(parser & p, optional const & pt, name const & tk, +static transition parse_transition(parser & p, ast_data & data, optional const & pt, name const & tk, buffer & locals, buffer & new_tokens, name const & pp_tk) { if (p.curr_is_token_or_id(get_binder_tk())) { + auto& b = p.new_ast(get_binder_tk(), p.pos()); p.next(); - unsigned rbp = parse_binders_rbp(p); + ast_id id; unsigned rbp; + std::tie(id, rbp) = parse_binders_rbp(p); + data.push(b.push(id).m_id); return transition(tk, mk_binder_action(rbp), pp_tk); } else if (p.curr_is_token_or_id(get_binders_tk())) { + auto& b = p.new_ast(get_binders_tk(), p.pos()); p.next(); - unsigned rbp = parse_binders_rbp(p); + ast_id id; unsigned rbp; + std::tie(id, rbp) = parse_binders_rbp(p); + data.push(b.push(id).m_id); return transition(tk, mk_binders_action(rbp), pp_tk); } else if (p.curr_is_identifier()) { unsigned default_prec = get_default_prec(pt, tk); name n = p.get_name_val(); + auto& v = p.new_ast("var", p.pos()).push(p.new_ast("ident", p.pos(), n).m_id); + data.push(v.m_id); p.next(); - action a = parse_action(p, tk, default_prec, locals, new_tokens); + action a = parse_action(p, v, tk, default_prec, locals, new_tokens); expr local_type = mk_Prop(); // type used in notation local declarations, it is irrelevant expr l = mk_local(n, local_type); p.add_local(l); @@ -463,7 +513,7 @@ static transition parse_transition(parser & p, optional const & pt, } } -static notation_entry parse_notation_core(parser & p, bool overload, notation_entry_group grp, buffer & new_tokens, bool parse_only, +static notation_entry parse_notation_core(parser & p, ast_data & parent, bool overload, notation_entry_group grp, buffer & new_tokens, bool parse_only, unsigned priority) { buffer locals; buffer ts; @@ -471,17 +521,22 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en bool is_nud = true; optional pt; optional reserved_pt; + auto& args = p.new_ast("args", p.pos()); + parent.push(args.m_id); if (p.curr_is_numeral()) { lean_assert(p.curr_is_numeral()); mpz num = p.get_num_val().get_numerator(); + args.push(p.new_ast("nat", p.pos(), num.to_string()).m_id); p.next(); p.check_token_next(get_assign_tk(), "invalid numeral notation, `:=` expected"); auto e_pos = p.pos(); expr e = p.parse_expr(); + parent.push(p.get_id(e)); check_notation_expr(e, e_pos); return notation_entry(num, e, overload, parse_only); } else if (p.curr_is_identifier()) { - parse_notation_local(p, locals); + ast_id id = parse_notation_local(p, locals); + args.push(p.new_ast("var", p.ast_pos(id)).push(id).push(0).m_id); is_nud = false; pt = get_led_table(p.env()); if (grp != notation_entry_group::Reserve) @@ -494,7 +549,7 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en while ((grp != notation_entry_group::Reserve && !p.curr_is_token(get_assign_tk())) || (grp == notation_entry_group::Reserve && !p.curr_is_command() && !p.curr_is_eof())) { bool used_default = false; - name pp_tk = parse_quoted_symbol_or_token(p, new_tokens, used_default).to_string_unescaped(); + name pp_tk = parse_quoted_symbol_or_token(p, args, new_tokens, used_default).to_string_unescaped(); name tk = utf8_trim(pp_tk.to_string_unescaped()); if (auto at = find_next(reserved_pt, tk)) { // Remark: we are ignoring multiple actions in the reserved notation table @@ -507,47 +562,58 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en case notation::action_kind::Skip: if (!p.curr_is_quoted_symbol() && !p.curr_is_keyword() && !p.curr_is_token(get_assign_tk())) { if (g_allow_local && !p.curr_is_token_or_id(get_binders_tk())) { - ts.push_back(parse_transition(p, pt, tk, locals, new_tokens, pp_tk)); + ts.push_back(parse_transition(p, args, pt, tk, locals, new_tokens, pp_tk)); break; } + auto pos = p.pos(); p.check_token_or_id_next(get_binders_tk(), "invalid notation declaration, quoted-symbol, keyword or `:=` expected " "(declaration prefix matches reserved notation)"); + args.push(p.new_ast(get_binders_tk(), pos).push(0).m_id); } ts.push_back(transition(tk, a, pp_tk)); break; - case notation::action_kind::Binder: + case notation::action_kind::Binder: { if (g_allow_local && !p.curr_is_token_or_id(get_binder_tk())) { - ts.push_back(parse_transition(p, pt, tk, locals, new_tokens, pp_tk)); + ts.push_back(parse_transition(p, args, pt, tk, locals, new_tokens, pp_tk)); break; } + auto pos = p.pos(); p.check_token_or_id_next(get_binder_tk(), "invalid notation declaration, 'binder' expected " "(declaration prefix matches reserved notation)"); + args.push(p.new_ast(get_binder_tk(), pos).push(0).m_id); ts.push_back(transition(tk, a, pp_tk)); break; - case notation::action_kind::Binders: + } + case notation::action_kind::Binders: { if (g_allow_local && !p.curr_is_token_or_id(get_binders_tk())) { - ts.push_back(parse_transition(p, pt, tk, locals, new_tokens, pp_tk)); + ts.push_back(parse_transition(p, args, pt, tk, locals, new_tokens, pp_tk)); break; } + auto pos = p.pos(); p.check_token_or_id_next(get_binders_tk(), "invalid notation declaration, 'binders' expected " "(declaration prefix matches reserved notation)"); + args.push(p.new_ast(get_binders_tk(), pos).push(0).m_id); ts.push_back(transition(tk, a, pp_tk)); break; + } case notation::action_kind::Expr: case notation::action_kind::Exprs: case notation::action_kind::ScopedExpr: - case notation::action_kind::Ext: { + case notation::action_kind::Ext: { if (g_allow_local && !p.curr_is_identifier()) { - ts.push_back(parse_transition(p, pt, tk, locals, new_tokens, pp_tk)); + ts.push_back(parse_transition(p, args, pt, tk, locals, new_tokens, pp_tk)); break; } - name n = p.check_id_next("invalid notation declaration, identifier expected " - "(declaration prefix matches reserved notation)"); + ast_id id; name n; + std::tie(id, n) = p.check_id_next("invalid notation declaration, identifier expected " + "(declaration prefix matches reserved notation)"); + auto& v = p.new_ast("var", p.pos()).push(id); + args.push(v.m_id); if (p.curr_is_token(get_colon_tk())) { if (g_allow_local) { unsigned default_prec = get_default_prec(pt, tk); - action a = parse_action(p, tk, default_prec, locals, new_tokens); + action a = parse_action(p, v, tk, default_prec, locals, new_tokens); expr local_type = mk_Prop(); // type used in notation local declarations, it is irrelevant expr l = mk_local(n, local_type); p.add_local(l); @@ -559,6 +625,7 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en "(declaration prefix matches reserved notation)", p.pos()); } } else { + v.push(0); expr local_type = mk_Prop(); // type used in notation local declarations, it is irrelevant expr l = mk_local(n, local_type); p.add_local(l); @@ -569,7 +636,7 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en }} } else { reserved_pt = optional(); - ts.push_back(parse_transition(p, pt, tk, locals, new_tokens, pp_tk)); + ts.push_back(parse_transition(p, args, pt, tk, locals, new_tokens, pp_tk)); } // for leading tokens where binding power was not set, we set it to max if (is_nud && used_default && ts.size() == 1) { @@ -582,6 +649,7 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en if (grp == notation_entry_group::Reserve) { // reserve notation commands do not have a denotation lean_assert(p.curr_is_command() || p.curr_is_eof()); + parent.push(0); expr dummy = mk_Prop(); // any expression without free variables will do n = dummy; } else { @@ -589,7 +657,9 @@ static notation_entry parse_notation_core(parser & p, bool overload, notation_en p.next(); if (ts.empty()) throw parser_error("invalid notation declaration, empty notation is not allowed", p.pos()); - n = parse_notation_expr(p, locals); + ast_id id; + std::tie(id, n) = parse_notation_expr(p, locals); + parent.push(id); } return notation_entry(is_nud, to_list(ts.begin(), ts.end()), n, overload, priority, grp, parse_only); } @@ -600,35 +670,35 @@ bool curr_is_notation_decl(parser & p) { p.curr_is_token(get_postfix_tk()) || p.curr_is_token(get_prefix_tk()) || p.curr_is_token(get_notation_tk()); } -static notation_entry parse_notation(parser & p, bool overload, notation_entry_group grp, buffer & new_tokens, +static notation_entry parse_notation(parser & p, ast_data & parent, bool overload, notation_entry_group grp, buffer & new_tokens, bool allow_local) { bool parse_only = false; unsigned priority = LEAN_DEFAULT_NOTATION_PRIORITY; flet set_allow_local(g_allow_local, allow_local); if (p.curr_is_token(get_infix_tk()) || p.curr_is_token(get_infixl_tk())) { p.next(); - return parse_mixfix_notation(p, mixfix_kind::infixl, overload, grp, new_tokens, parse_only, priority); + return parse_mixfix_notation(p, parent, mixfix_kind::infixl, overload, grp, new_tokens, parse_only, priority); } else if (p.curr_is_token(get_infixr_tk())) { p.next(); - return parse_mixfix_notation(p, mixfix_kind::infixr, overload, grp, new_tokens, parse_only, priority); + return parse_mixfix_notation(p, parent, mixfix_kind::infixr, overload, grp, new_tokens, parse_only, priority); } else if (p.curr_is_token(get_postfix_tk())) { p.next(); - return parse_mixfix_notation(p, mixfix_kind::postfix, overload, grp, new_tokens, parse_only, priority); + return parse_mixfix_notation(p, parent, mixfix_kind::postfix, overload, grp, new_tokens, parse_only, priority); } else if (p.curr_is_token(get_prefix_tk())) { p.next(); - return parse_mixfix_notation(p, mixfix_kind::prefix, overload, grp, new_tokens, parse_only, priority); + return parse_mixfix_notation(p, parent, mixfix_kind::prefix, overload, grp, new_tokens, parse_only, priority); } else if (p.curr_is_token(get_notation_tk())) { p.next(); - return parse_notation_core(p, overload, grp, new_tokens, parse_only, priority); + return parse_notation_core(p, parent, overload, grp, new_tokens, parse_only, priority); } else { throw parser_error("invalid notation, 'infix', 'infixl', 'infixr', 'prefix', " "'postfix' or 'notation' expected", p.pos()); } } -notation_entry parse_notation(parser & p, bool overload, buffer & new_tokens, bool allow_local) { +notation_entry parse_notation(parser & p, ast_data & parent, bool overload, buffer & new_tokens, bool allow_local) { notation_entry_group grp = notation_entry_group::Main; - return parse_notation(p, overload, grp, new_tokens, allow_local); + return parse_notation(p, parent, overload, grp, new_tokens, allow_local); } static char g_reserved_chars[] = {',', 0}; @@ -679,10 +749,10 @@ struct notation_modifiers { bool m_parse_only; unsigned m_priority; notation_modifiers():m_parse_only(false), m_priority(LEAN_DEFAULT_NOTATION_PRIORITY) {} - void parse(parser & p) { + void parse(parser & p, ast_data & parent) { auto pos = p.pos(); decl_attributes attrs; - attrs.parse(p); + parent.push(attrs.parse(p)); for (auto const & entry : attrs.get_entries()) { if (entry.m_attr->get_name() == "parsing_only") m_parse_only = true; @@ -695,24 +765,26 @@ struct notation_modifiers { } }; -static environment notation_cmd_core(parser & p, bool overload, notation_entry_group grp, bool persistent) { +static environment notation_cmd_core(parser & p, ast_data & parent, ast_id local, bool overload, notation_entry_group grp, bool persistent) { notation_modifiers mods; - mods.parse(p); + parent.push(local); + mods.parse(p, parent); flet set_allow_local(g_allow_local, !persistent); environment env = p.env(); buffer new_tokens; - auto ne = parse_notation_core(p, overload, grp, new_tokens, mods.m_parse_only, mods.m_priority); + auto ne = parse_notation_core(p, parent, overload, grp, new_tokens, mods.m_parse_only, mods.m_priority); for (auto const & te : new_tokens) env = add_user_token(env, te, persistent); env = add_notation(env, ne, persistent); return env; } -static environment mixfix_cmd(parser & p, mixfix_kind k, bool overload, notation_entry_group grp, bool persistent) { +static environment mixfix_cmd(parser & p, ast_data & parent, ast_id local, mixfix_kind k, bool overload, notation_entry_group grp, bool persistent) { notation_modifiers mods; - mods.parse(p); + parent.push(local); + mods.parse(p, parent); flet set_allow_local(g_allow_local, !persistent); - auto nt = parse_mixfix_notation(p, k, overload, grp, mods.m_parse_only, mods.m_priority); + auto nt = parse_mixfix_notation(p, parent, k, overload, grp, mods.m_parse_only, mods.m_priority); environment env = p.env(); if (nt.second) env = add_user_token(env, *nt.second, persistent); @@ -720,87 +792,92 @@ static environment mixfix_cmd(parser & p, mixfix_kind k, bool overload, notation return env; } -static environment notation_cmd(parser & p) { +static environment notation_cmd(parser & p, ast_id & cmd_id) { bool overload = true; notation_entry_group grp = notation_entry_group::Main; bool persistent = true; - return notation_cmd_core(p, overload, grp, persistent); + return notation_cmd_core(p, p.get_ast(cmd_id), 0, overload, grp, persistent); } -static environment infixl_cmd(parser & p) { +static environment infixl_cmd(parser & p, ast_id & cmd_id) { bool overload = true; notation_entry_group grp = notation_entry_group::Main; bool persistent = true; - return mixfix_cmd(p, mixfix_kind::infixl, overload, grp, persistent); + return mixfix_cmd(p, p.get_ast(cmd_id), 0, mixfix_kind::infixl, overload, grp, persistent); } -static environment infixr_cmd(parser & p) { +static environment infixr_cmd(parser & p, ast_id & cmd_id) { bool overload = true; notation_entry_group grp = notation_entry_group::Main; bool persistent = true; - return mixfix_cmd(p, mixfix_kind::infixr, overload, grp, persistent); + return mixfix_cmd(p, p.get_ast(cmd_id), 0, mixfix_kind::infixr, overload, grp, persistent); } -static environment postfix_cmd(parser & p) { +static environment postfix_cmd(parser & p, ast_id & cmd_id) { bool overload = true; notation_entry_group grp = notation_entry_group::Main; bool persistent = true; - return mixfix_cmd(p, mixfix_kind::postfix, overload, grp, persistent); + return mixfix_cmd(p, p.get_ast(cmd_id), 0, mixfix_kind::postfix, overload, grp, persistent); } -static environment prefix_cmd(parser & p) { +static environment prefix_cmd(parser & p, ast_id & cmd_id) { bool overload = true; notation_entry_group grp = notation_entry_group::Main; bool persistent = true; - return mixfix_cmd(p, mixfix_kind::prefix, overload, grp, persistent); + return mixfix_cmd(p, p.get_ast(cmd_id), 0, mixfix_kind::prefix, overload, grp, persistent); } // auxiliary procedure used by local_notation_cmd and reserve_cmd -static environment dispatch_notation_cmd(parser & p, bool overload, notation_entry_group grp, bool persistent) { +static environment dispatch_notation_cmd(parser & p, ast_data & parent, bool overload, notation_entry_group grp, bool persistent) { + ast_id id = p.new_ast(parent.m_type, p.pos()).m_id; + parent.m_type = p.get_token_info().value(); if (p.curr_is_token(get_notation_tk())) { p.next(); - return notation_cmd_core(p, overload, grp, persistent); + return notation_cmd_core(p, parent, id, overload, grp, persistent); } else if (p.curr_is_token(get_infixl_tk())) { p.next(); - return mixfix_cmd(p, mixfix_kind::infixl, overload, grp, persistent); + return mixfix_cmd(p, parent, id, mixfix_kind::infixl, overload, grp, persistent); } else if (p.curr_is_token(get_infix_tk())) { p.next(); - return mixfix_cmd(p, mixfix_kind::infixl, overload, grp, persistent); + return mixfix_cmd(p, parent, id, mixfix_kind::infixl, overload, grp, persistent); } else if (p.curr_is_token(get_infixr_tk())) { p.next(); - return mixfix_cmd(p, mixfix_kind::infixr, overload, grp, persistent); + return mixfix_cmd(p, parent, id, mixfix_kind::infixr, overload, grp, persistent); } else if (p.curr_is_token(get_prefix_tk())) { p.next(); - return mixfix_cmd(p, mixfix_kind::prefix, overload, grp, persistent); + return mixfix_cmd(p, parent, id, mixfix_kind::prefix, overload, grp, persistent); } else if (p.curr_is_token(get_postfix_tk())) { p.next(); - return mixfix_cmd(p, mixfix_kind::postfix, overload, grp, persistent); + return mixfix_cmd(p, parent, id, mixfix_kind::postfix, overload, grp, persistent); } else { throw parser_error("invalid local/reserve notation, 'infix', 'infixl', 'infixr', 'prefix', " "'postfix' or 'notation' expected", p.pos()); } } -environment local_notation_cmd(parser & p) { +environment local_notation_cmd(parser & p, ast_id & cmd_id) { parser::in_notation_ctx ctx(p); bool overload = false; // REMARK: local notation override global one notation_entry_group grp = notation_entry_group::Main; bool persistent = false; - return dispatch_notation_cmd(p, overload, grp, persistent); + return dispatch_notation_cmd(p, p.get_ast(cmd_id), overload, grp, persistent); } -static environment reserve_cmd(parser & p) { +static environment reserve_cmd(parser & p, ast_id & cmd_id) { parser::in_notation_ctx ctx(p); bool overload = false; notation_entry_group grp = notation_entry_group::Reserve; bool persistent = true; - return dispatch_notation_cmd(p, overload, grp, persistent); + return dispatch_notation_cmd(p, p.get_ast(cmd_id), overload, grp, persistent); } -static environment precedence_cmd(parser & p) { - std::string tk = parse_symbol(p, "invalid precedence declaration, quoted symbol or identifier expected"); +static environment precedence_cmd(parser & p, ast_id & cmd_id) { + auto& data = p.get_ast(cmd_id); + std::string tk = parse_symbol(p, data, "invalid precedence declaration, quoted symbol or identifier expected"); p.check_token_next(get_colon_tk(), "invalid precedence declaration, ':' expected"); - unsigned prec = parse_precedence(p); + ast_id id; unsigned prec; + std::tie(id, prec) = parse_precedence(p); + data.push(id); return add_user_token(p.env(), tk.c_str(), prec); } diff --git a/src/frontends/lean/notation_cmd.h b/src/frontends/lean/notation_cmd.h index 3552037591..e8a2121ee9 100644 --- a/src/frontends/lean/notation_cmd.h +++ b/src/frontends/lean/notation_cmd.h @@ -15,10 +15,10 @@ bool curr_is_notation_decl(parser & p); /** \brief Parse a notation declaration, throws an error if the current token is not a "notation declaration". If allow_local is true, then notation may contain reference to local constants. */ -notation_entry parse_notation(parser & p, bool overload, buffer & new_tokens, bool allow_local); +notation_entry parse_notation(parser & p, ast_data & data, bool overload, buffer & new_tokens, bool allow_local); /** \brief Parse local notation */ -environment local_notation_cmd(parser & p); +environment local_notation_cmd(parser & p, ast_id & cmd_id); void register_notation_cmds(cmd_table & r); diff --git a/src/frontends/lean/parse_table.cpp b/src/frontends/lean/parse_table.cpp index d5fbeef17b..977e94c22e 100644 --- a/src/frontends/lean/parse_table.cpp +++ b/src/frontends/lean/parse_table.cpp @@ -369,13 +369,13 @@ static void validate_transitions(bool nud, unsigned num, transition const * ts, throw exception("invalid notation declaration, expression template has more free variables than arguments"); } -static list insert(list const & l, unsigned priority, list const & p, expr const & a) { +static list insert(list const & l, unsigned priority, list const & p, name const & n, expr const & a) { if (!l) { - return to_list(accepting(priority, p, a)); + return to_list(accepting(priority, p, n, a)); } else if (priority <= head(l).get_prio()) { - return cons(accepting(priority, p, a), l); + return cons(accepting(priority, p, n, a), l); } else { - return cons(head(l), insert(tail(l), priority, p, a)); + return cons(head(l), insert(tail(l), priority, p, n, a)); } } @@ -387,18 +387,18 @@ static bool contains_equivalent_action(list> const return false; } -parse_table parse_table::add_core(unsigned num, transition const * ts, expr const & a, +parse_table parse_table::add_core(unsigned num, transition const * ts, name const & n, expr const & a, unsigned priority, bool overload, buffer & post_buffer) const { parse_table r(new cell(*m_ptr)); if (num == 0) { list postponed = to_list(post_buffer); if (!overload) { - r.m_ptr->m_accept = to_list(accepting(priority, postponed, a)); + r.m_ptr->m_accept = to_list(accepting(priority, postponed, n, a)); } else { auto new_accept = filter(r.m_ptr->m_accept, [&](accepting const & p) { return p.get_expr() != a || p.get_postponed() != postponed; }); - r.m_ptr->m_accept = insert(new_accept, priority, postponed, a); + r.m_ptr->m_accept = insert(new_accept, priority, postponed, n, a); } } else { list> const * it = r.m_ptr->m_children.find(ts->get_token()); @@ -413,7 +413,7 @@ parse_table parse_table::add_core(unsigned num, transition const * ts, expr cons to_buffer(*it, tmp); for (pair & p : tmp) { if (p.first.get_action().is_equivalent(ts_act)) { - p.second = p.second.add_core(num-1, ts+1, a, priority, overload, post_buffer); + p.second = p.second.add_core(num-1, ts+1, n, a, priority, overload, post_buffer); break; } } @@ -423,11 +423,11 @@ parse_table parse_table::add_core(unsigned num, transition const * ts, expr cons new_lst = filter(*it, [&](pair const & p) { return p.first.get_action().is_compatible(ts_act); }); - parse_table new_child = parse_table().add_core(num-1, ts+1, a, priority, overload, post_buffer); + parse_table new_child = parse_table().add_core(num-1, ts+1, n, a, priority, overload, post_buffer); new_lst = cons(mk_pair(*ts, new_child), new_lst); } } else { - parse_table new_child = parse_table().add_core(num-1, ts+1, a, priority, overload, post_buffer); + parse_table new_child = parse_table().add_core(num-1, ts+1, n, a, priority, overload, post_buffer); new_lst = to_list(mk_pair(*ts, new_child)); } r.m_ptr->m_children.insert(ts->get_token(), new_lst); @@ -497,11 +497,11 @@ optional get_head_index(unsigned num, transition const * ts, expr co return optional(); } -parse_table parse_table::add(unsigned num, transition const * ts, expr const & a, unsigned priority, bool overload) const { +parse_table parse_table::add(unsigned num, transition const * ts, name const & n, expr const & a, unsigned priority, bool overload) const { buffer postponed; expr new_a = annotate_macro_subterms(a); validate_transitions(is_nud(), num, ts, new_a); - return add_core(num, ts, new_a, priority, overload, postponed); + return add_core(num, ts, n, new_a, priority, overload, postponed); } void parse_table::for_each(buffer & ts, @@ -530,7 +530,7 @@ parse_table parse_table::merge(parse_table const & s, bool overload) const { parse_table r(*this); s.for_each([&](unsigned num, transition const * ts, list const & accept) { for (accepting const & acc : accept) { - r = r.add(num, ts, acc.get_expr(), acc.get_prio(), overload); + r = r.add(num, ts, acc.get_name(), acc.get_expr(), acc.get_prio(), overload); } }); return r; diff --git a/src/frontends/lean/parse_table.h b/src/frontends/lean/parse_table.h index ddb3dd6afc..efb11e1029 100644 --- a/src/frontends/lean/parse_table.h +++ b/src/frontends/lean/parse_table.h @@ -134,12 +134,14 @@ transition replace(transition const & t, std::function const class accepting { unsigned m_prio; list m_postponed; // exprs and scoped_expr actions + name m_name; // notation name expr m_expr; // resulting expression public: - accepting(unsigned prio, list const & post, expr const & e): - m_prio(prio), m_postponed(post), m_expr(e) {} + accepting(unsigned prio, list const & post, name const & n, expr const & e): + m_prio(prio), m_postponed(post), m_name(n), m_expr(e) {} unsigned get_prio() const { return m_prio; } list const & get_postponed() const { return m_postponed; } + name const & get_name() const { return m_name; } expr const & get_expr() const { return m_expr; } }; @@ -154,7 +156,7 @@ class parse_table { struct cell; cell * m_ptr; explicit parse_table(cell * c); - parse_table add_core(unsigned num, transition const * ts, expr const & a, unsigned priority, bool overload, buffer & postponed) const; + parse_table add_core(unsigned num, transition const * ts, name const & n, expr const & a, unsigned priority, bool overload, buffer & postponed) const; void for_each(buffer & ts, std::function const &)> const & fn) const; public: @@ -166,11 +168,15 @@ class parse_table { parse_table & operator=(parse_table&& n); bool is_nud() const; - parse_table add(unsigned num, transition const * ts, expr const & a, unsigned priority, bool overload) const; + parse_table add(unsigned num, transition const * ts, name const & n, expr const & a, unsigned priority, bool overload) const; + parse_table add(std::initializer_list const & ts, name const & n, expr const & a) const { + return add(ts.size(), ts.begin(), n, a, LEAN_DEFAULT_NOTATION_PRIORITY, true); + } parse_table add(std::initializer_list const & ts, expr const & a) const { - return add(ts.size(), ts.begin(), a, LEAN_DEFAULT_NOTATION_PRIORITY, true); + lean_assert(ts.size() != 0); + return add(ts, ts.begin()->get_token(), a); } - parse_table merge(parse_table const & s, bool overload) const;\ + parse_table merge(parse_table const & s, bool overload) const; list> find(name const & tk) const; list const & is_accepting() const; void for_each(std::function const &)> const & fn) const; diff --git a/src/frontends/lean/parser.cpp b/src/frontends/lean/parser.cpp index 5ecb587c9f..2ffaa953a6 100644 --- a/src/frontends/lean/parser.cpp +++ b/src/frontends/lean/parser.cpp @@ -15,6 +15,7 @@ Author: Leonardo de Moura #include "util/sstream.h" #include "util/flet.h" #include "util/sexpr/option_declarations.h" +#include "util/task_builder.h" #include "kernel/for_each_fn.h" #include "kernel/replace_fn.h" #include "kernel/find_fn.h" @@ -161,7 +162,7 @@ parser::parser(environment const & env, io_state const & ios, m_import_fn(import_fn), m_file_name(file_name), m_scanner(strm, m_file_name.c_str()), - m_imports_parsed(false) { + m_ast{new ast_data(0, {}, {}), new ast_data(AST_TOP_ID, {}, "file")} { m_next_inst_idx = 1; m_ignore_noncomputable = false; m_in_quote = false; @@ -174,6 +175,7 @@ parser::parser(environment const & env, io_state const & ios, } parser::~parser() { + for (auto p : m_ast) delete p; } void parser::check_break_at_pos(break_at_pos_exception::token_context ctxt) { @@ -274,6 +276,37 @@ tag parser::get_tag(expr e) { return t; } +ast_data & parser::new_ast(name type, pos_info start, name value) { + ast_id id = m_ast.size(); + m_ast.push_back(new ast_data{id, start, type, value}); + return *m_ast.back(); +} + +void parser::set_ast_pexpr(ast_id id, expr const & e) { + m_ast[id]->m_pexpr = e; + auto t = get_tag(e); + // if (!m_tag_ast_table.contains(t)) + m_tag_ast_table.insert(t, id); +} + +void parser::set_ast_expr(ast_id id, expr e) { + if (id) m_ast[id]->m_expr.emplace(mk_pure_task(std::move(e))); +} + +ast_id parser::get_id(expr const & e) const { + tag t = e.get_tag(); + if (t == nulltag) return 0; + auto id = m_tag_ast_table.find(t); + return id ? *id : 0; +} + +ast_data & parser::new_modifiers(cmd_meta & meta) { + if (meta.m_modifiers_id) return get_ast(meta.m_modifiers_id); + auto& mods = new_ast("modifiers", pos()); + meta.m_modifiers_id = mods.m_id; + return mods; +} + name parser::mk_anonymous_inst_name() { name n = ::lean::mk_anonymous_inst_name(m_next_inst_idx); m_next_inst_idx++; @@ -394,17 +427,17 @@ void parser::check_token_or_id_next(name const & tk, char const * msg) { next(); } -name parser::check_id_next(char const * msg, break_at_pos_exception::token_context ctxt) { +pair parser::check_id_next(char const * msg, break_at_pos_exception::token_context ctxt) { // initiate empty completion even if following token is not an identifier if (get_complete()) check_break_before(ctxt); - name r; + pair r; if (!curr_is_identifier()) { auto _ = no_error_recovery_scope_if(curr_is_command()); maybe_throw_error({msg, pos()}); - return "_"; + return {0, "_"}; } else { - r = get_name_val(); + r.first = new_ast("ident", pos(), r.second = get_name_val()).m_id; } try { next(); @@ -423,26 +456,26 @@ void parser::check_not_internal(name const & id, pos_info const & p) { p}); } -name parser::check_decl_id_next(char const * msg, break_at_pos_exception::token_context ctxt) { +pair parser::check_decl_id_next(char const * msg, break_at_pos_exception::token_context ctxt) { auto p = pos(); - name id = check_id_next(msg, ctxt); - check_not_internal(id, p); - return id; + auto r = check_id_next(msg, ctxt); + check_not_internal(r.second, p); + return r; } -name parser::check_atomic_id_next(char const * msg) { +pair parser::check_atomic_id_next(char const * msg) { auto p = pos(); - name id = check_id_next(msg); - if (!id.is_atomic()) + auto r = check_id_next(msg); + if (!r.second.is_atomic()) maybe_throw_error({msg, p}); - return id; + return r; } -name parser::check_atomic_decl_id_next(char const * msg) { +pair parser::check_atomic_decl_id_next(char const * msg) { auto p = pos(); - name id = check_atomic_id_next(msg); - check_not_internal(id, p); - return id; + auto r = check_atomic_id_next(msg); + check_not_internal(r.second, p); + return r; } expr parser::mk_app(expr fn, expr arg, pos_info const & p) { @@ -467,6 +500,13 @@ expr parser::mk_app(std::initializer_list const & args, pos_info const & p return r; } +static expr mk_app_ast(parser & p, expr left, expr right) { + pos_info pos = p.pos_of(left); + expr r = p.mk_app(left, right, pos); + p.set_ast_pexpr(p.new_ast("app", pos).push(p.get_id(left)).push(p.get_id(right)).m_id, r); + return r; +} + parser_scope parser_info::mk_parser_scope(optional const & opts) { return parser_scope(opts, m_level_variables, m_variables, m_include_vars, m_next_inst_idx, m_has_params, @@ -679,31 +719,35 @@ unsigned parser::get_small_nat() { return val.get_unsigned_int(); } -std::string parser::parse_string_lit() { - std::string v = get_str_val(); +pair parser::parse_string_lit() { + pair r; + r.second = get_str_val(); + r.first = new_ast("string", pos(), r.second).m_id; next(); - return v; + return r; } -unsigned parser::parse_small_nat() { - unsigned r = 0; +pair parser::parse_small_nat() { + pair r = {0, 0}; if (!curr_is_numeral()) { auto _ = no_error_recovery_scope_if(curr_is_command()); maybe_throw_error({"(small) natural number expected", pos()}); } else { - r = get_small_nat(); + r.second = get_small_nat(); + r.first = new_ast("nat", pos(), std::to_string(r.second)).m_id; } next(); return r; } -double parser::parse_double() { - double r = 0; +pair parser::parse_double() { + pair r = {0, 0}; if (curr() != token_kind::Decimal) { auto _ = no_error_recovery_scope_if(curr_is_command()); maybe_throw_error({"decimal value expected", pos()}); } else { - r = get_num_val().get_double(); + r.second = get_num_val().get_double(); + r.first = new_ast("double", pos(), std::to_string(r.second)).m_id; } next(); return r; @@ -724,12 +768,15 @@ unsigned parser::curr_level_lbp() const { return 0; } -level parser::parse_max_imax(bool is_max) { +pair parser::parse_max_imax(bool is_max) { auto p = pos(); + auto& data = new_ast(is_max ? get_max_tk() : get_imax_tk(), p); next(); buffer lvls; while (curr_is_identifier() || curr_is_numeral() || curr_is_token(get_lparen_tk())) { - lvls.push_back(parse_level(get_max_prec())); + auto r = parse_level(get_max_prec()); + data.push(r.first); + lvls.push_back(r.second); } if (lvls.size() < 2) { return parser_error_or_level( @@ -744,35 +791,39 @@ level parser::parse_max_imax(bool is_max) { else r = mk_imax(lvls[i], r); } - return r; + return {data.m_id, r}; } -level parser::parse_level_id() { +pair parser::parse_level_id() { auto p = pos(); name id = get_name_val(); + auto a = new_ast("param", p, id).m_id; next(); if (auto it = m_local_level_decls.find(id)) - return *it; + return {a, *it}; return parser_error_or_level({sstream() << "unknown universe '" << id << "'", p}); } -level parser::parse_level_nud() { +pair parser::parse_level_nud() { if (curr_is_token_or_id(get_max_tk())) { return parse_max_imax(true); } else if (curr_is_token_or_id(get_imax_tk())) { return parse_max_imax(false); } else if (curr_is_token_or_id(get_placeholder_tk())) { + auto a = new_ast(get_placeholder_tk(), pos()).m_id; next(); - return mk_level_placeholder(); + return {a, mk_level_placeholder()}; } else if (curr_is_token(get_lparen_tk())) { + auto p = pos(); next(); - level l = parse_level(); + auto r = parse_level(); check_token_next(get_rparen_tk(), "invalid level expression, ')' expected"); - return l; + return {new_ast(get_lparen_tk(), p).push(r.first).m_id, r.second}; } else if (curr_is_numeral()) { - unsigned k = parse_small_nat(); - return lift(level(), k); + ast_id ak; unsigned k; + std::tie(ak, k) = parse_small_nat(); + return {ak, lift(level(), k)}; } else if (curr_is_identifier()) { return parse_level_id(); } else { @@ -780,13 +831,15 @@ level parser::parse_level_nud() { } } -level parser::parse_level_led(level left) { +pair parser::parse_level_led(pair left) { auto p = pos(); if (curr_is_token(get_add_tk())) { next(); if (curr_is_numeral()) { - unsigned k = parse_small_nat(); - return lift(left, k); + ast_id ak; unsigned k; + std::tie(ak, k) = parse_small_nat(); + ast_id id = new_ast(get_add_tk(), ast_pos(left.first)).push(left.first).push(ak).m_id; + return {id, lift(left.second, k)}; } else { return parser_error_or_level( {"invalid level expression, right hand side of '+' " @@ -797,8 +850,8 @@ level parser::parse_level_led(level left) { } } -level parser::parse_level(unsigned rbp) { - level left = parse_level_nud(); +pair parser::parse_level(unsigned rbp) { + pair left = parse_level_nud(); while (rbp < curr_level_lbp()) { left = parse_level_led(left); } @@ -814,6 +867,7 @@ pair parser::elaborate(name const & decl_name, check_unassigned, m_error_recovery, freeze_instances); expr new_e = r.first; new_e = adapter.translate_from(new_e); + set_ast_expr(get_id(e), new_e); return mk_pair(new_e, r.second); } @@ -934,21 +988,28 @@ void parser::parse_close_binder_info(optional const & bi) { /** \brief Parse ID ':' expr, where the expression represents the type of the identifier. */ expr parser::parse_binder_core(binder_info const & bi, unsigned rbp) { auto p = pos(); - name id; + auto& vars = new_ast("vars", p); + auto& ast = new_ast(name("binder").append_after(bi.hash()), p).push(vars.m_id).push(0); + pair id; if (curr_is_token(get_placeholder_tk())) { - id = "_x"; + id = {new_ast(get_placeholder_tk(), p, "_x").m_id, "_x"}; next(); } else { id = check_atomic_id_next("invalid binder, atomic identifier expected"); } + vars.push(id.first); expr type; if (curr_is_token(get_colon_tk())) { next(); type = parse_expr(rbp); + ast.push(get_id(type)); } else { type = save_pos(mk_expr_placeholder(), p); + ast.push(0); } - return save_pos(mk_local(id, type, bi), p); + expr r = save_pos(mk_local(id.second, type, bi), p); + set_ast_pexpr(ast.m_id, r); + return r; } expr parser::parse_binder(unsigned rbp) { @@ -972,27 +1033,28 @@ expr parser::parse_binder(unsigned rbp) { This method return true if the next token is an infix operator, and populates r with the locals above. */ -bool parser::parse_binder_collection(buffer> const & names, binder_info const & bi, buffer & r) { +ast_id parser::parse_binder_collection(buffer> const & names, binder_info const & bi, buffer & r) { if (!curr_is_keyword()) - return false; + return 0; name tk = get_token_info().value(); list> trans_list = led().find(tk); if (length(trans_list) != 1) - return false; + return 0; pair const & p = head(trans_list); list const & acc_lst = p.second.is_accepting(); if (length(acc_lst) != 1) - return false; // no overloading + return 0; // no overloading notation::accepting const & acc = head(acc_lst); lean_assert(!acc.get_postponed()); expr pred = acc.get_expr(); auto k = p.first.get_action().kind(); if (k == notation::action_kind::Skip || k == notation::action_kind::Ext) - return false; + return 0; unsigned rbp = p.first.get_action().rbp(); next(); // consume tk expr S = parse_expr(rbp); + ast_id id = new_ast("collection", names[0].first, acc.get_name()).push(get_id(S)).m_id; unsigned old_sz = r.size(); /* Add (ID_1 ... ID_n : _) to r */ for (auto p : names) { @@ -1012,43 +1074,66 @@ bool parser::parse_binder_collection(buffer> const & names, r.push_back(local); i++; } - return true; + return id; } /** \brief Parse ID ... ID ':' expr, where the expression represents the type of the identifiers. */ -void parser::parse_binder_block(buffer & r, binder_info const & bi, unsigned rbp, bool allow_default) { +ast_id parser::parse_binder_block(buffer & r, binder_info const & bi, unsigned rbp, bool allow_default) { buffer> names; + auto& vars = new_ast("vars", pos()); + auto& ast = new_ast(name("binder").append_after(bi.hash()), pos()).push(vars.m_id).push(0); while (curr_is_identifier() || curr_is_token(get_placeholder_tk())) { auto p = pos(); if (curr_is_identifier()) { - names.emplace_back(p, check_atomic_id_next("invalid binder, atomic identifier expected")); + ast_id id; name n; + std::tie(id, n) = check_atomic_id_next("invalid binder, atomic identifier expected"); + vars.push(id); + names.emplace_back(p, n); } else { + vars.push(new_ast(get_placeholder_tk(), p).m_id); names.emplace_back(p, "_x"); next(); } } - if (names.empty()) - return maybe_throw_error({"invalid binder, identifier expected", pos()}); + if (names.empty()) { + maybe_throw_error({"invalid binder, identifier expected", pos()}); + return 0; + } optional type; if (curr_is_token(get_colon_tk())) { next(); type = parse_expr(rbp); + ast.push(get_id(*type)); if (allow_default && curr_is_token(get_assign_tk())) { + auto& ast1 = new_ast(get_assign_tk(), pos()); + ast.push(ast1.m_id); next(); expr val = parse_expr(rbp); + ast1.push(get_id(val)); type = mk_opt_param(*type, val); } else if (allow_default && curr_is_token(get_period_tk())) { - type = parse_auto_param(*this, *type); + auto& ast1 = new_ast(get_period_tk(), pos()); + ast.push(ast1.m_id); + ast_id id; + std::tie(id, type) = parse_auto_param(*this, *type); + ast1.push(id); + } + } else { + ast.push(0); + if (allow_default && curr_is_token(get_assign_tk())) { + auto& ast1 = new_ast(get_assign_tk(), pos()); + ast.push(ast1.m_id); + next(); + expr val = parse_expr(rbp); + ast1.push(get_id(val)); + type = mk_opt_param(copy_tag(val, mk_expr_placeholder()), val); + } else if (auto id = parse_binder_collection(names, bi, r)) { + ast.push(id); + return ast.m_id; } - } else if (allow_default && curr_is_token(get_assign_tk())) { - next(); - expr val = parse_expr(rbp); - type = mk_opt_param(copy_tag(val, mk_expr_placeholder()), val); - } else if (parse_binder_collection(names, bi, r)) { - return; } for (auto p : names) { expr arg_type = type ? *type : save_pos(mk_expr_placeholder(), p.first); @@ -1056,21 +1141,25 @@ void parser::parse_binder_block(buffer & r, binder_info const & bi, unsign add_local(local); r.push_back(local); } + return ast.m_id; } -expr parser::parse_inst_implicit_decl() { +pair parser::parse_inst_implicit_decl() { binder_info bi = mk_inst_implicit_binder_info(); auto id_pos = pos(); + auto& ast = new_ast(name("binder").append_after(bi.hash()), id_pos); name id; expr type; + ast_id var = 0; if (curr_is_identifier()) { id = get_name_val(); next(); if (curr_is_token(get_colon_tk())) { next(); type = parse_expr(); + var = new_ast("vars", id_pos).push(new_ast("ident", id_pos, id).m_id).m_id; } else { - expr left = id_to_expr(id, id_pos); + expr left = id_to_expr(id, new_ast("ident", id_pos, id)); id = mk_anonymous_inst_name(); unsigned rbp = 0; while (rbp < curr_lbp()) { @@ -1082,18 +1171,20 @@ expr parser::parse_inst_implicit_decl() { id = mk_anonymous_inst_name(); type = parse_expr(); } + ast.push(var).push(0).push(get_id(type)); expr local = save_pos(mk_local(id, type, bi), id_pos); add_local(local); - return local; + return {ast.m_id, local}; } -void parser::parse_inst_implicit_decl(buffer & r) { - expr local = parse_inst_implicit_decl(); - r.push_back(local); +void parser::parse_inst_implicit_decl(ast_data * parent, buffer & r) { + auto p = parse_inst_implicit_decl(); + if (parent) parent->push(p.first); + r.push_back(p.second); } -void parser::parse_binders_core(buffer & r, parse_binders_config & cfg) { +void parser::parse_binders_core(ast_data * parent, buffer & r, parse_binders_config & cfg) { bool first = true; while (true) { if (curr_is_identifier() || curr_is_token(get_placeholder_tk())) { @@ -1104,7 +1195,8 @@ void parser::parse_binders_core(buffer & r, parse_binders_config & cfg) { /* We only allow the default parameter value syntax for declarations with surrounded by () */ bool new_allow_default = false; - parse_binder_block(r, binder_info(), cfg.m_rbp, new_allow_default); + ast_id id = parse_binder_block(r, binder_info(), cfg.m_rbp, new_allow_default); + if (parent) parent->push(id); cfg.m_last_block_delimited = false; } else { /* We only allow the default parameter value syntax for declarations with @@ -1115,31 +1207,40 @@ void parser::parse_binders_core(buffer & r, parse_binders_config & cfg) { if (first && cfg.m_infer_kind != nullptr) { /* Parse {}, [], or () prefix */ if (bi->is_implicit() && curr_is_token(get_rcurly_tk())) { + ast_id id = new_ast("{}", pos()).m_id; next(); *cfg.m_infer_kind = implicit_infer_kind::RelaxedImplicit; + *cfg.m_infer_kind_id = id; first = false; continue; } else if (is_explicit(*bi) && curr_is_token(get_rparen_tk())) { + ast_id id = new_ast("()", pos()).m_id; next(); *cfg.m_infer_kind = implicit_infer_kind::None; + *cfg.m_infer_kind_id = id; first = false; continue; } else if (bi->is_inst_implicit() && curr_is_token(get_rbracket_tk())) { + ast_id id = new_ast("[]", pos()).m_id; next(); *cfg.m_infer_kind = implicit_infer_kind::Implicit; + *cfg.m_infer_kind_id = id; first = false; continue; } else { *cfg.m_infer_kind = implicit_infer_kind::RelaxedImplicit; + *cfg.m_infer_kind_id = 0; } } unsigned rbp = 0; cfg.m_last_block_delimited = true; if (bi->is_inst_implicit()) { - parse_inst_implicit_decl(r); + parse_inst_implicit_decl(parent, r); } else { - if (cfg.m_simple_only || !parse_local_notation_decl(cfg.m_nentries)) - parse_binder_block(r, *bi, rbp, new_allow_default); + ast_id id = 0; + if (!cfg.m_simple_only) id = parse_local_notation_decl(cfg.m_nentries); + if (!id) id = parse_binder_block(r, *bi, rbp, new_allow_default); + if (parent) parent->push(id); } parse_close_binder_info(bi); } else { @@ -1150,30 +1251,31 @@ void parser::parse_binders_core(buffer & r, parse_binders_config & cfg) { } } -local_environment parser::parse_binders(buffer & r, parse_binders_config & cfg) { +local_environment parser::parse_binders(ast_data * parent, buffer & r, parse_binders_config & cfg) { flet save1(m_env, m_env); // save environment flet save2(m_local_decls, m_local_decls); unsigned old_sz = r.size(); - parse_binders_core(r, cfg); + parse_binders_core(parent, r, cfg); if (!cfg.m_allow_empty && old_sz == r.size()) throw_invalid_open_binder(pos()); return local_environment(m_env); } -bool parser::parse_local_notation_decl(buffer * nentries) { +ast_id parser::parse_local_notation_decl(buffer * nentries) { if (curr_is_notation_decl(*this)) { parser::in_notation_ctx ctx(*this); buffer new_tokens; bool overload = false; bool allow_local = true; - auto ne = ::lean::parse_notation(*this, overload, new_tokens, allow_local); + auto& data = new_ast(get_token_info().value(), pos()); + auto ne = ::lean::parse_notation(*this, data, overload, new_tokens, allow_local); for (auto const & te : new_tokens) m_env = add_token(m_env, te); if (nentries) nentries->push_back(ne); m_env = add_notation(m_env, ne); - return true; + return data.m_id; } else { - return false; + return 0; } } @@ -1336,6 +1438,7 @@ expr parser::parse_notation(parse_table t, expr * left) { buffer> nargs; // nary args buffer ps; buffer> scoped_info; // size of ps and binder_pos for scoped_exprs + buffer ast_ids; // Invariants: // args.size() == kinds.size() if not left // args.size() == kinds.size()+1 if left @@ -1344,8 +1447,10 @@ expr parser::parse_notation(parse_table t, expr * left) { bool has_Exprs = false; local_environment lenv(m_env); pos_info binder_pos; - if (left) + if (left) { args.push_back(*left); + ast_ids.push_back(get_id(args.back())); + } while (true) { if (curr() != token_kind::Keyword) break; @@ -1378,6 +1483,7 @@ expr parser::parse_notation(parse_table t, expr * left) { break; case notation::action_kind::Expr: args.push_back(parse_expr(a.rbp())); + ast_ids.push_back(get_id(args.back())); kinds.push_back(a.kind()); break; case notation::action_kind::Exprs: { @@ -1406,26 +1512,35 @@ expr parser::parse_notation(parse_table t, expr * left) { has_Exprs = true; args.push_back(expr()); // placeholder kinds.push_back(a.kind()); + auto& data = new_ast("exprs", pos()); + for (auto& e : r_args) data.push(get_id(e)); + ast_ids.push_back(data.m_id); nargs.push_back(to_list(r_args)); break; } case notation::action_kind::Binder: binder_pos = pos(); ps.push_back(parse_binder(a.rbp())); + ast_ids.push_back(get_id(ps.back())); break; - case notation::action_kind::Binders: + case notation::action_kind::Binders: { binder_pos = pos(); - lenv = parse_binders(ps, a.rbp()); + auto& data = new_ast("binders", pos()); + ast_ids.push_back(data.m_id); + lenv = parse_binders(&data, ps, a.rbp()); break; + } case notation::action_kind::ScopedExpr: { expr r = parse_scoped_expr(ps, lenv, a.rbp()); args.push_back(r); + ast_ids.push_back(get_id(r)); kinds.push_back(a.kind()); scoped_info.push_back(mk_pair(ps.size(), binder_pos)); break; } case notation::action_kind::Ext: args.push_back(a.get_parse_fn()(*this, args.size(), args.data(), p)); + ast_ids.push_back(get_id(args.back())); kinds.push_back(a.kind()); break; } @@ -1442,8 +1557,7 @@ expr parser::parse_notation(parse_table t, expr * left) { } return parser_error_or_expr({msg, pos()}); } - lean_assert(left || args.size() == kinds.size()); - lean_assert(!left || args.size() == kinds.size() + 1); + lean_assert(args.size() == kinds.size() + (left ? 1 : 0)); /* IF there are multiple choices and Exprs was not used, we create a lambda for each choice. In this case, we copy args to actual_args and store local constants in args. */ @@ -1484,6 +1598,17 @@ expr parser::parse_notation(parse_table t, expr * left) { if (create_lambdas) { r = rec_save_pos(::lean::mk_app(r, actual_args), p); } + ast_data* data; + if (length(as) > 1) { + auto& choices = new_ast("choices", p); + data = &new_ast("choice", p).push(choices.m_id); + for (auto& a : as) + choices.push(new_ast(get_notation_tk(), p, a.get_name()).m_id); + } else { + data = &new_ast(get_notation_tk(), p, head(as).get_name()); + } + for (auto id : ast_ids) data->push(id); + set_ast_pexpr(data->m_id, r); return r; } @@ -1495,27 +1620,31 @@ expr parser::parse_inaccessible() { auto p = pos(); next(); expr t = parse_expr(get_max_prec()); - return save_pos(mk_inaccessible(t), p); + t = save_pos(mk_inaccessible(t), p); + set_ast_pexpr(new_ast(".(", p).push(get_id(t)).m_id, t); + return t; } expr parser::parse_placeholder() { auto p = pos(); next(); - return save_pos(mk_explicit_expr_placeholder(), p); + expr r = save_pos(mk_explicit_expr_placeholder(), p); + set_ast_pexpr(new_ast("_", p).m_id, r); + return r; } -expr parser::parse_anonymous_var_pattern() { - auto p = pos(); - next(); - expr t = mk_local(next_name(), "_x", mk_expr_placeholder(), binder_info()); - return save_pos(t, p); -} +// expr parser::parse_anonymous_var_pattern() { +// auto p = pos(); +// next(); +// expr t = mk_local(next_name(), "_x", mk_expr_placeholder(), binder_info()); +// return save_pos(t, p); +// } expr parser::parse_led_notation(expr left) { if (led().find(get_token_info().value())) { return parse_notation(led(), &left); } else { - return mk_app(left, parse_expr(get_max_prec()), pos_of(left)); + return mk_app_ast(*this, left, parse_expr(get_max_prec())); } } @@ -1628,7 +1757,7 @@ struct to_pattern_fn { void collect_new_local(expr const & e) { name const & n = mlocal_pp_name(e); bool resolve_only = true; - expr new_e = m_parser.id_to_expr(n, m_parser.pos_of(e), resolve_only); + expr new_e = m_parser.id_to_expr(n, m_parser.expr_ast(e), resolve_only); if (is_as_atomic(new_e)) { new_e = get_app_fn(get_as_atomic_arg(new_e)); if (is_explicit(new_e)) @@ -1929,7 +2058,7 @@ class patexpr_to_expr_fn : public replace_visitor { } virtual expr visit_local(expr const & e) override { - return m_p.id_to_expr(mlocal_pp_name(e), m_p.pos_of(e), true, true, m_locals); + return m_p.id_to_expr(mlocal_pp_name(e), m_p.expr_ast(e), true, true, m_locals); } public: @@ -1983,17 +2112,52 @@ optional parser::resolve_local(name const & id, pos_info const & p, list const & extra_locals) { +static expr mk_constant(parser & p, const name & id, const levels & ls, const ast_data & id_data, ast_id levels_id) { + expr r = mk_constant(id, ls); + ast_id c = p.new_ast("const", id_data.m_start).push(id_data.m_id).push(levels_id).m_id; + p.set_ast_pexpr(c, r); + return r; +} + +static expr mk_field_notation_ast(parser & p, const expr & s, const expr & r, ast_id field_id) { + ast_id sid = p.get_id(s); + ast_id ast = p.new_ast("field", p.ast_pos(sid)).push(sid).push(field_id).m_id; + p.set_ast_pexpr(ast, r); + return r; +} + +static expr mk_field_notation_compact(parser & p, const expr & s, const pos_info & field_pos, const char * field) { + ast_id field_id = p.new_ast("ident", field_pos, field).m_id; + expr r = p.save_pos(mk_field_notation_compact(s, field), field_pos); + return mk_field_notation_ast(p, s, r, field_id); +} + +static expr mk_field_notation(parser & p, const expr & s, const pos_info & field_pos, const name & field) { + ast_id field_id = p.new_ast("ident", field_pos, field).m_id; + expr r = p.save_pos(mk_field_notation(s, field), field_pos); + return mk_field_notation_ast(p, s, r, field_id); +} + +static expr mk_field_notation(parser & p, const expr & s, const pos_info & field_pos, ast_id field_id, unsigned fidx) { + expr r = p.save_pos(mk_field_notation(s, fidx), field_pos); + return mk_field_notation_ast(p, s, r, field_id); +} + +expr parser::id_to_expr(name const & id, ast_data & id_data, + bool resolve_only, bool allow_field_notation, list const & extra_locals) { buffer lvl_buffer; levels ls; - bool explicit_levels = false; + ast_id explicit_levels = 0; if (!resolve_only && curr_is_token(get_llevel_curly_tk())) { next(); - explicit_levels = true; + auto& ast = new_ast("levels", pos()); + explicit_levels = ast.m_id; while (!curr_is_token(get_rcurly_tk())) { auto _ = backtracking_scope(); try { - lvl_buffer.push_back(parse_level()); + auto l = parse_level(); + ast.push(l.first); + lvl_buffer.push_back(l.second); } catch (backtracking_exception) {} if (!consumed_input()) break; } @@ -2009,29 +2173,37 @@ expr parser::id_to_expr(name const & id, pos_info const & p, bool resolve_only, } }; + auto& p = id_data.m_start; if (!explicit_levels && m_id_behavior == id_behavior::AllLocal) { - return save_pos(mk_local(id, save_pos(mk_expr_placeholder(), p)), p); + expr r = save_pos(mk_local(id, save_pos(mk_expr_placeholder(), p)), p); + set_ast_pexpr(id_data.m_id, r); + return r; } if (auto r = resolve_local(id, p, extra_locals, allow_field_notation)) { check_no_levels(ls, p); + set_ast_pexpr(id_data.m_id, *r); return *r; } if (!explicit_levels && m_id_behavior == id_behavior::AssumeLocalIfNotLocal) { - return save_pos(mk_local(id, save_pos(mk_expr_placeholder(), p)), p); + expr r = save_pos(mk_local(id, save_pos(mk_expr_placeholder(), p)), p); + set_ast_pexpr(id_data.m_id, r); + return r; } if (auto ref = get_local_ref(m_env, id)) { check_no_levels(ls, p); - return copy_with_new_pos(*ref, p); + expr r = copy_with_new_pos(*ref, p); + set_ast_pexpr(id_data.m_id, r); + return r; } for (name const & ns : get_namespaces(m_env)) { auto new_id = ns + id; if (!ns.is_anonymous() && m_env.find(new_id) && (!id.is_atomic() || !is_protected(m_env, new_id))) { - return save_pos(mk_constant(new_id, ls), p); + return save_pos(mk_constant(*this, new_id, ls, id_data, explicit_levels), p); } } @@ -2039,7 +2211,7 @@ expr parser::id_to_expr(name const & id, pos_info const & p, bool resolve_only, name new_id = id; new_id = remove_root_prefix(new_id); if (m_env.find(new_id)) { - return save_pos(mk_constant(new_id, ls), p); + return save_pos(mk_constant(*this, new_id, ls, id_data, explicit_levels), p); } } @@ -2047,7 +2219,7 @@ expr parser::id_to_expr(name const & id, pos_info const & p, bool resolve_only, // globals if (m_env.find(id)) - r = save_pos(mk_constant(id, ls), p); + r = save_pos(mk_constant(*this, id, ls, id_data, explicit_levels), p); // aliases auto as = get_expr_aliases(m_env, id); if (!is_nil(as)) { @@ -2058,20 +2230,25 @@ expr parser::id_to_expr(name const & id, pos_info const & p, bool resolve_only, new_as.push_back(copy_with_new_pos(mk_constant(e, ls), p)); } r = save_pos(mk_choice(new_as.size(), new_as.data()), p); + ast_id c = new_ast(new_as.size() == 1 ? "const" : "choice_const", id_data.m_start) + .push(id_data.m_id).push(explicit_levels).m_id; + set_ast_pexpr(c, *r); } if (!r) { if (m_id_behavior == id_behavior::AssumeLocalIfUndef) { expr local = mk_local(id, save_pos(mk_expr_placeholder(), p)); r = save_pos(local, p); + set_ast_pexpr(id_data.m_id, *r); } } if (!r && allow_field_notation && !id.is_atomic() && id.is_string()) { try { auto _ = no_error_recovery_scope(); - expr s = id_to_expr(id.get_prefix(), p, resolve_only, allow_field_notation, extra_locals); + name n = id_data.m_value = id.get_prefix(); + expr s = id_to_expr(n, id_data, resolve_only, allow_field_notation, extra_locals); auto field_pos = p; - field_pos.second += id.get_prefix().utf8_size(); - r = save_pos(mk_field_notation_compact(s, id.get_string()), field_pos); + field_pos.second += n.utf8_size(); + r = mk_field_notation_compact(*this, s, field_pos, id.get_string()); } catch (exception &) {} } if (!r) { @@ -2140,22 +2317,24 @@ name parser::to_constant(name const & id, char const * msg, pos_info const & p) return head(to_constants(id, msg, p)); } -name parser::check_constant_next(char const * msg) { - auto p = pos(); - name id = check_id_next(msg); +pair parser::check_constant_next(char const * msg) { + auto p = pos(); + auto r = check_id_next(msg); try { - return to_constant(id, msg, p); + r.second = to_constant(r.second, msg, p); } catch (parser_error & e) { maybe_throw_error(std::move(e)); - return id; } + return r; } expr parser::parse_id(bool allow_field_notation) { - auto p = pos(); + auto p = pos(); lean_assert(curr_is_identifier()); - name id = check_id_next("", break_at_pos_exception::token_context::expr); - expr e = id_to_expr(id, p, /* resolve_only */ false, allow_field_notation); + ast_id aid; name id; + std::tie(aid, id) = check_id_next("", break_at_pos_exception::token_context::expr); + expr e = id_to_expr(id, get_ast(aid), /* resolve_only */ false, allow_field_notation); + set_ast_pexpr(aid, e); if (is_constant(e) && get_global_info_manager()) { get_global_info_manager()->add_const_info(m_env, p, const_name(e)); } @@ -2169,18 +2348,21 @@ expr parser::parse_numeral_expr(bool user_notation) { list vals; if (user_notation) vals = get_mpz_notation(m_env, n); + expr r; if (!vals) { - return save_pos(mk_prenum(n), p); + r = save_pos(mk_prenum(n), p); } else { buffer cs; cs.push_back(save_pos(mk_prenum(n), p)); for (expr const & c : vals) cs.push_back(copy_with_new_pos(c, p)); if (cs.size() == 1) - return cs[0]; + r = cs[0]; else - return save_pos(mk_choice(cs.size(), cs.data()), p); + r = save_pos(mk_choice(cs.size(), cs.data()), p); } + set_ast_pexpr(new_ast("nat", p, n.to_string()).m_id, r); + return r; } expr parser::parse_decimal_expr() { @@ -2188,19 +2370,26 @@ expr parser::parse_decimal_expr() { mpq val = get_num_val(); next(); expr num = save_pos(mk_prenum(val.get_numerator()), p); + expr r; if (val.get_denominator() == 1) { - return num; + r = num; } else { expr den = save_pos(mk_prenum(val.get_denominator()), p); expr div = save_pos(mk_constant(get_has_div_div_name()), p); - return save_pos(lean::mk_app(div, num, den), p); + r = save_pos(lean::mk_app(div, num, den), p); } + std::ostringstream out; out << val; + set_ast_pexpr(new_ast("decimal", p, out.str()).m_id, r); + return r; } expr parser::parse_string_expr() { std::string v = get_str_val(); + ast_id id = new_ast("string", pos(), v).m_id; next(); - return from_string(v); + expr r = from_string(v); + set_ast_pexpr(id, r); + return r; } expr parser::parse_char_expr() { @@ -2209,10 +2398,13 @@ expr parser::parse_char_expr() { buffer tmp; utf8_decode(v, tmp); lean_assert(tmp.size() == 1); + ast_id id = new_ast("char", pos(), v).m_id; next(); - return mk_app(save_pos(mk_constant(get_char_of_nat_name()), p), - save_pos(mk_prenum(mpz(tmp[0])), p), - p); + expr r = mk_app(save_pos(mk_constant(get_char_of_nat_name()), p), + save_pos(mk_prenum(mpz(tmp[0])), p), + p); + set_ast_pexpr(id, r); + return r; } expr parser::parse_nud() { @@ -2236,7 +2428,9 @@ expr parser::parse_nud() { if (p.first == id_pos.first && p.second == id_pos.second + id_len) { next(); auto pat = parse_expr(get_max_prec()); - return save_pos(mk_as_pattern(e, pat), id_pos); + ast_id eid = new_ast("at_pat", id_pos).push(get_id(e)).push(get_id(pat)).m_id; + e = save_pos(mk_as_pattern(e, pat), id_pos); + set_ast_pexpr(eid, e); } } return e; @@ -2266,28 +2460,35 @@ bool parser::curr_starts_expr() { expr parser::parse_led(expr left) { if (is_sort_wo_universe(left) && (curr_is_numeral() || curr_is_identifier() || curr_is_token(get_lparen_tk()) || curr_is_token(get_placeholder_tk()))) { - left = get_annotation_arg(left); - level l = parse_level(get_max_prec()); + left = get_annotation_arg(left); + auto& data = expr_ast(left); + lean_always_assert(!data.m_children.empty()); + level l; + std::tie(data.m_children[0], l) = parse_level(get_max_prec()); lean_assert(sort_level(left) == mk_level_one() || sort_level(left) == mk_level_zero()); if (sort_level(left) == mk_level_one()) l = mk_succ(l); - return copy_tag(left, update_sort(left, l)); + expr r = copy_tag(left, update_sort(left, l)); + set_ast_pexpr(data.m_id, r); + return r; } else { switch (curr()) { case token_kind::Keyword: return parse_led_notation(left); case token_kind::FieldName: { - expr r = save_pos(mk_field_notation(left, get_name_val()), pos()); + expr r = mk_field_notation(*this, left, pos(), get_name_val()); next(); return r; } case token_kind::FieldNum: { - expr r = save_pos(mk_field_notation(left, get_small_nat()), pos()); + unsigned fidx = get_small_nat(); + ast_id id = new_ast("nat", pos(), std::to_string(fidx)).m_id; + expr r = mk_field_notation(*this, left, pos(), id, fidx); next(); return r; } default: - return mk_app(left, parse_expr(get_max_prec()), pos_of(left)); + return mk_app_ast(*this, left, parse_expr(get_max_prec())); } } } @@ -2344,31 +2545,31 @@ expr parser::parse_expr(unsigned rbp) { return parse_led_loop(left, rbp); } -pair, expr> parser::parse_id_tk_expr(name const & tk, unsigned rbp) { +std::tuple, expr> parser::parse_id_tk_expr(name const & tk, unsigned rbp) { if (curr_is_identifier()) { - auto id_pos = pos(); name id = get_name_val(); + auto& id_data = new_ast("ident", pos(), id); next(); if (curr_is_token(tk)) { next(); - return mk_pair(optional(id), parse_expr(rbp)); + return {id_data.m_id, optional(id), parse_expr(rbp)}; } else { - expr left = id_to_expr(id, id_pos); + expr left = id_to_expr(id, id_data); while (rbp < curr_lbp()) { left = parse_led(left); } - return mk_pair(optional(), left); + return {0, {}, left}; } } else { - return mk_pair(optional(), parse_expr(rbp)); + return {0, {}, parse_expr(rbp)}; } } -pair, expr> parser::parse_qualified_expr(unsigned rbp) { +std::tuple, expr> parser::parse_qualified_expr(unsigned rbp) { return parse_id_tk_expr(get_colon_tk(), rbp); } -pair, expr> parser::parse_optional_assignment(unsigned rbp) { +std::tuple, expr> parser::parse_optional_assignment(unsigned rbp) { return parse_id_tk_expr(get_assign_tk(), rbp); } @@ -2407,17 +2608,19 @@ class lazy_type_context : public abstract_type_context { virtual name next_name() override { return ctx().next_name(); } }; -void parser::parse_command(cmd_meta const & meta) { +ast_id parser::parse_command(cmd_meta const & meta) { if (curr() != token_kind::CommandKeyword) { auto p = pos(); maybe_throw_error({"expected command", p}); - return; + return 0; } reset_thread_local(); m_last_cmd_pos = pos(); + scoped_expr_caching scope(false); + name cmd_name = get_token_info().value(); - m_cmd_token = get_token_info().token(); if (auto it = cmds().find(cmd_name)) { + auto cmd_id = new_ast(cmd_name, pos()).m_id; lazy_type_context tc(m_env, get_options()); scope_global_ios scope1(m_ios); scope_trace_env scope2(m_env, m_ios.get_options(), tc); @@ -2426,49 +2629,60 @@ void parser::parse_command(cmd_meta const & meta) { in_notation_ctx ctx(*this); if (it->get_skip_token()) next(); - m_env = it->get_fn()(*this, meta); + m_env = it->get_fn()(*this, cmd_id, meta); } else { if (it->get_skip_token()) next(); - m_env = it->get_fn()(*this, meta); + m_env = it->get_fn()(*this, cmd_id, meta); } + return cmd_id; } else { auto p = pos(); next(); maybe_throw_error({sstream() << "unknown command '" << cmd_name << "'", p}); + return 0; } } -std::string parser::parse_doc_block() { +pair parser::parse_doc_block() { auto val = m_scanner.get_str_val(); + auto id = new_ast("doc", m_scanner.get_pos_info(), val).m_id; next(); - return val; + return {id, val}; } -void parser::parse_mod_doc_block() { - m_env = add_module_doc_string(m_env, m_scanner.get_str_val(), m_scanner.get_pos_info() ); +ast_id parser::parse_mod_doc_block() { + auto doc = m_scanner.get_str_val(); + auto pos = m_scanner.get_pos_info(); + auto id = new_ast("mdoc", pos, doc).m_id; + m_env = add_module_doc_string(m_env, doc, pos); next(); + return id; } #if defined(__GNUC__) && !defined(__CLANG__) #pragma GCC diagnostic ignored "-Wmaybe-uninitialized" #endif -bool parser::parse_imports(unsigned & fingerprint, std::vector & imports) { +bool parser::parse_imports(unsigned & fingerprint, ast_data * parent, std::vector & imports) { init_scanner(); scanner::field_notation_scope scope(m_scanner, false); m_last_cmd_pos = pos(); - bool prelude = false; + ast_id prelude = 0; if (curr_is_token(get_prelude_tk())) { + prelude = new_ast("prelude", pos()).m_id; next(); - prelude = true; } if (!prelude) { module_name m("init"); imports.push_back(m); } + auto& imports_cmd = new_ast("imports", pos()); + if (parent) parent->push(prelude).push(imports_cmd.m_id); while (curr_is_token(get_import_tk())) { + auto& import_cmd = new_ast("import", pos()); m_last_cmd_pos = pos(); + imports_cmd.push(import_cmd.m_id); next(); while (true) { pos_info p = pos(); @@ -2501,10 +2715,14 @@ bool parser::parse_imports(unsigned & fingerprint, std::vector & im if (k_init) { fingerprint = hash(fingerprint, h); } + auto& ast_mod = new_ast("module", p); + import_cmd.push(ast_mod.m_id); if (k_init) { + ast_mod.m_value = f.append_after(k); module_name m(f, k); imports.push_back(m); } else { + ast_mod.m_value = f; module_name m(f); imports.push_back(m); } @@ -2536,7 +2754,7 @@ bool parser::parse_imports(unsigned & fingerprint, std::vector & im return false; } -void parser::process_imports() { +void parser::process_imports(ast_data * parent) { unsigned fingerprint = 0; std::vector imports; @@ -2544,7 +2762,7 @@ void parser::process_imports() { auto begin_pos = pos(); bool needs_to_scan_again = false; try { - needs_to_scan_again = parse_imports(fingerprint, imports); + needs_to_scan_again = parse_imports(fingerprint, parent, imports); } catch (parser_exception &) { exception_during_scanning = std::current_exception(); } @@ -2594,6 +2812,8 @@ void parser::process_imports() { m_env = activate_export_decls(m_env, {}); // explicitly activate exports in root namespace m_env = replay_export_decls_core(m_env, m_ios); m_imports_parsed = true; + m_commands = new_ast("commands", pos()).m_id; + if (parent) parent->push(m_commands); if (needs_to_scan_again) { next(); @@ -2605,10 +2825,10 @@ void parser::process_imports() { void parser::get_imports(std::vector & imports) { scope_pos_info_provider scope1(*this); unsigned fingerprint; - parse_imports(fingerprint, imports); + parse_imports(fingerprint, nullptr, imports); } -bool parser::parse_command_like() { +bool parser::parse_command_like(ast_data * parent) { init_scanner(); m_error_since_last_cmd = false; @@ -2618,23 +2838,32 @@ bool parser::parse_command_like() { check_interrupted(); - if (!m_imports_parsed) { - process_imports(); + if (!imports_parsed()) { + process_imports(&get_ast(AST_TOP_ID)); return false; } + if (!parent) { + if (!m_commands) m_commands = new_ast("commands", pos()).m_id; + parent = &get_ast(m_commands); + } module::scope_pos_info scope2(pos()); + ast_id cmd_id = 0; switch (curr()) { case token_kind::CommandKeyword: - parse_command(cmd_meta()); + cmd_id = parse_command(cmd_meta()); updt_options(); break; - case token_kind::DocBlock: - parse_command(cmd_meta({}, {}, some(parse_doc_block()))); + case token_kind::DocBlock: { + auto r = parse_doc_block(); + cmd_meta meta({}, {}, some(std::move(r.second))); + new_modifiers(meta).push(r.first); + cmd_id = parse_command(meta); break; + } case token_kind::ModDocBlock: - parse_mod_doc_block(); + cmd_id = parse_mod_doc_block(); break; case token_kind::Eof: if (has_open_scopes(m_env)) { @@ -2651,6 +2880,7 @@ bool parser::parse_command_like() { default: throw parser_error("command expected", pos()); } + if (cmd_id) parent->push(cmd_id); return false; } @@ -2729,6 +2959,9 @@ void parser::from_snapshot(snapshot const & s) { m_ignore_noncomputable = s.m_noncomputable_theory; m_parser_scope_stack = s.m_parser_scope_stack; m_next_inst_idx = s.m_next_inst_idx; + m_ast_invalid = true; // invalidate AST because we don't snapshot it + m_commands = 0; + m_vm_parser_ast_id = 0; } void dummy_def_parser::maybe_throw_error(parser_error && err) { @@ -2755,9 +2988,9 @@ expr parser::parser_error_or_expr(parser_error && err) { return mk_sorry(err_pos, true); } -level parser::parser_error_or_level(parser_error && err) { +pair parser::parser_error_or_level(parser_error && err) { maybe_throw_error(std::move(err)); - return mk_level_placeholder(); + return {0, mk_level_placeholder()}; } bool parse_commands(environment & env, io_state & ios, char const * fname) { diff --git a/src/frontends/lean/parser.h b/src/frontends/lean/parser.h index fcd967579e..3d1935507a 100644 --- a/src/frontends/lean/parser.h +++ b/src/frontends/lean/parser.h @@ -35,6 +35,8 @@ class metavar_context; class local_context_adapter; class scope_message_context; +#define AST_TOP_ID 1 + class parser_info { protected: environment m_env; @@ -125,10 +127,10 @@ class parser_info { parser_scope mk_parser_scope(optional const & opts = optional()); void restore_parser_scope(parser_scope const & s); - virtual std::tuple parse_definition(buffer & lp_names, buffer & params, + virtual std::tuple parse_definition(ast_data * parent, buffer & lp_names, buffer & params, bool is_example, bool is_instance, bool, bool) = 0; - friend environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, cmd_meta meta); + friend environment single_definition_cmd_core(parser_info & p, decl_cmd_kind kind, ast_data * parent, cmd_meta meta); public: void updt_options(); options get_options() const { return m_ios.get_options(); } @@ -149,6 +151,8 @@ class parser_info { bool has_error_recovery() const { return m_error_recovery; } }; +typedef rb_map tag_ast_table; + class parser : public abstract_parser, public parser_info { name_generator m_ngen; bool m_use_exceptions; @@ -156,12 +160,17 @@ class parser : public abstract_parser, public parser_info { std::string m_file_name; scanner m_scanner; token_kind m_curr; - bool m_imports_parsed; + bool m_imports_parsed = false; bool m_scanner_inited = false; parser_scope_stack m_quote_stack; pos_info m_last_cmd_pos; unsigned m_next_tag_idx; + ast_id m_vm_parser_ast_id = 0; pos_info_table m_pos_table; + tag_ast_table m_tag_ast_table; + std::vector m_ast; + bool m_ast_invalid = false; + ast_id m_commands = 0; // By default, when the parser finds a unknown identifier, it signs an error. // When the following flag is true, it creates a constant. // This flag is when we are trying to parse mutually recursive declarations. @@ -175,22 +184,19 @@ class parser : public abstract_parser, public parser_info { pos_info m_last_recovered_error_pos {0, 0}; optional m_backtracking_pos; - // curr command token - name m_cmd_token; - void sync_command(); tag get_tag(expr e); unsigned curr_level_lbp() const; - level parse_max_imax(bool is_max); - level parse_level_id(); - level parse_level_nud(); - level parse_level_led(level left); + pair parse_max_imax(bool is_max); + pair parse_level_id(); + pair parse_level_nud(); + pair parse_level_led(pair left); - void parse_mod_doc_block(); + ast_id parse_mod_doc_block(); - void process_imports(); + void process_imports(ast_data * parent); void process_postponed(buffer const & args, bool is_left, buffer const & kinds, buffer> const & nargs, buffer const & ps, buffer> const & scoped_info, list const & postponed, pos_info const & p, buffer & new_args); @@ -199,18 +205,18 @@ class parser : public abstract_parser, public parser_info { expr parse_led_notation(expr left); expr parse_inaccessible(); expr parse_placeholder(); - expr parse_anonymous_var_pattern(); + // expr parse_anonymous_var_pattern(); expr parse_nud(); bool curr_starts_expr(); expr parse_numeral_expr(bool user_notation = true); expr parse_decimal_expr(); expr parse_string_expr(); expr parse_char_expr(); - expr parse_inst_implicit_decl(); - void parse_inst_implicit_decl(buffer & r); + pair parse_inst_implicit_decl(); + void parse_inst_implicit_decl(ast_data * parent, buffer & r); expr parse_binder_core(binder_info const & bi, unsigned rbp); - bool parse_binder_collection(buffer> const & names, binder_info const & bi, buffer & r); - void parse_binder_block(buffer & r, binder_info const & bi, unsigned rbp, bool allow_default); + ast_id parse_binder_collection(buffer> const & names, binder_info const & bi, buffer & r); + ast_id parse_binder_block(buffer & r, binder_info const & bi, unsigned rbp, bool allow_default); struct parse_binders_config { /* (input) If m_allow_empty is true, then parse_binders will succeed even if not binder is parsed. */ bool m_allow_empty{false}; @@ -259,21 +265,23 @@ class parser : public abstract_parser, public parser_info { The default is `RelaxedImplicit` */ implicit_infer_kind * m_infer_kind{nullptr}; + /* (output) The AST ID associated to the m_infer_kind */ + ast_id * m_infer_kind_id{nullptr}; /* (output) It is set to true if the last binder is surrounded with some kind of bracket (e.g., '()', '{}', '[]'). */ bool m_last_block_delimited{false}; /* (output) If m_nentries != nullptr, then local notation declarations are stored here */ buffer * m_nentries{nullptr}; }; - void parse_binders_core(buffer & r, parse_binders_config & cfg); - local_environment parse_binders(buffer & r, parse_binders_config & cfg); - bool parse_local_notation_decl(buffer * entries); + void parse_binders_core(ast_data * parent, buffer & r, parse_binders_config & cfg); + local_environment parse_binders(ast_data * parent, buffer & r, parse_binders_config & cfg); + ast_id parse_local_notation_decl(buffer * entries); - pair, expr> parse_id_tk_expr(name const & tk, unsigned rbp); + std::tuple, expr> parse_id_tk_expr(name const & tk, unsigned rbp); - friend environment section_cmd(parser & p); - friend environment namespace_cmd(parser & p); - friend environment end_scoped_cmd(parser & p); + friend environment section_cmd(parser & p, ast_id &); + friend environment namespace_cmd(parser & p, ast_id &); + friend environment end_scoped_cmd(parser & p, ast_id &); std::shared_ptr mk_snapshot(); @@ -300,6 +308,21 @@ class parser : public abstract_parser, public parser_info { void init_scanner(); + ast_data & new_ast(name type, pos_info start, name value = {}); + void set_ast_pexpr(ast_id id, expr const & e); + bool is_ast_invalid() { return m_ast_invalid; } + void set_ast_expr(ast_id id, expr e); + ast_data & get_ast(ast_id id) { return *m_ast[id]; } + ast_id get_id(expr const & e) const; + ast_data & expr_ast(expr const & e) { return get_ast(get_id(e)); } + const char * ast_string(ast_id id) const { return m_ast[id]->m_value.get_string(); } + pos_info ast_pos(ast_id id) const { return m_ast[id]->m_start; } + void push_from_vm_parse(ast_id id) { if (m_vm_parser_ast_id) get_ast(m_vm_parser_ast_id).push(id); } + ast_id get_vm_parse_parent() { return m_vm_parser_ast_id; } + void set_vm_parse_parent(ast_id id) { m_vm_parser_ast_id = id; } + ast_data & new_modifiers(cmd_meta & meta); + friend void export_ast(parser & p); + void from_snapshot(snapshot const & snap); name next_name() { return m_ngen.next(); } @@ -352,7 +375,6 @@ class parser : public abstract_parser, public parser_info { pos_info pos_of(expr const & e, pos_info default_pos) const; pos_info pos_of(expr const & e) const { return pos_of(e, pos()); } pos_info cmd_pos() const override { return m_last_cmd_pos; } - name const & get_cmd_token() const { return m_cmd_token; } parser_pos_provider get_parser_pos_provider(pos_info const & some_pos) const override { return parser_pos_provider(m_pos_table, m_file_name, some_pos, m_next_tag_idx); @@ -404,20 +426,20 @@ class parser : public abstract_parser, public parser_info { void check_token_or_id_next(name const & tk, char const * msg); /** \brief Check if the current token is an identifier, if it is return it and move to next token, otherwise throw an exception. */ - name check_id_next(char const * msg, break_at_pos_exception::token_context ctxt = + pair check_id_next(char const * msg, break_at_pos_exception::token_context ctxt = break_at_pos_exception::token_context::none); void check_not_internal(name const & n, pos_info const & pos); /** \brief Similar to check_id_next, but also ensures the identifier is *not* an internal/reserved name. */ - name check_decl_id_next(char const * msg, break_at_pos_exception::token_context ctxt = + pair check_decl_id_next(char const * msg, break_at_pos_exception::token_context ctxt = break_at_pos_exception::token_context::none); /** \brief Check if the current token is an atomic identifier, if it is, return it and move to next token, otherwise throw an exception. */ - name check_atomic_id_next(char const * msg); - name check_atomic_decl_id_next(char const * msg); + pair check_atomic_id_next(char const * msg); + pair check_atomic_decl_id_next(char const * msg); list to_constants(name const & id, char const * msg, pos_info const & p) const; name to_constant(name const & id, char const * msg, pos_info const & p); /** \brief Check if the current token is a constant, if it is, return it and move to next token, otherwise throw an exception. */ - name check_constant_next(char const * msg); + pair check_constant_next(char const * msg); mpq const & get_num_val() const { return m_scanner.get_num_val(); } name const & get_name_val() const { return m_scanner.get_name_val(); } @@ -426,59 +448,60 @@ class parser : public abstract_parser, public parser_info { std::string const & get_stream_name() const { return m_scanner.get_stream_name(); } unsigned get_small_nat(); - virtual unsigned parse_small_nat() override final; - virtual std::string parse_string_lit() override final; - double parse_double(); + virtual pair parse_small_nat() override final; + virtual pair parse_string_lit() override final; + pair parse_double(); /** \brief Parses a documentation block (`/-- TEXT -/`). For example, `/-- Doc\ndoc -/` returns " Doc\ndoc ". */ - std::string parse_doc_block(); + pair parse_doc_block(); - bool parse_local_notation_decl() { return parse_local_notation_decl(nullptr); } + ast_id parse_local_notation_decl() { return parse_local_notation_decl(nullptr); } - level parse_level(unsigned rbp = 0); + pair parse_level(unsigned rbp = 0); expr parse_binder(unsigned rbp); - local_environment parse_binders(buffer & r, bool & last_block_delimited) { + local_environment parse_binders(ast_data * parent, buffer & r, bool & last_block_delimited) { parse_binders_config cfg; - auto new_env = parse_binders(r, cfg); + auto new_env = parse_binders(parent, r, cfg); last_block_delimited = cfg.m_last_block_delimited; return new_env; } - local_environment parse_binders(buffer & r, unsigned rbp, bool allow_default = false) { + local_environment parse_binders(ast_data * parent, buffer & r, unsigned rbp, bool allow_default = false) { parse_binders_config cfg; cfg.m_rbp = rbp; cfg.m_allow_default = allow_default; - return parse_binders(r, cfg); + return parse_binders(parent, r, cfg); } - void parse_simple_binders(buffer & r, unsigned rbp) { + void parse_simple_binders(ast_data * parent, buffer & r, unsigned rbp) { parse_binders_config cfg; cfg.m_simple_only = true; cfg.m_rbp = rbp; - parse_binders(r, cfg); + parse_binders(parent, r, cfg); } - local_environment parse_optional_binders(buffer & r, bool allow_default = false, bool explicit_delimiters = false) { + local_environment parse_optional_binders(ast_data * parent, buffer & r, bool allow_default = false, bool explicit_delimiters = false) { parse_binders_config cfg; cfg.m_allow_empty = true; cfg.m_allow_default = allow_default; cfg.m_explicit_delimiters = explicit_delimiters; - return parse_binders(r, cfg); + return parse_binders(parent, r, cfg); } - local_environment parse_optional_binders(buffer & r, implicit_infer_kind & kind) { + local_environment parse_optional_binders(ast_data * parent, buffer & r, ast_id & kind_id, implicit_infer_kind & kind) { parse_binders_config cfg; cfg.m_allow_empty = true; cfg.m_infer_kind = &kind; - return parse_binders(r, cfg); + cfg.m_infer_kind_id = &kind_id; + return parse_binders(parent, r, cfg); } - local_environment parse_binders(buffer & r, buffer & nentries) { + local_environment parse_binders(ast_data * parent, buffer & r, buffer & nentries) { parse_binders_config cfg; cfg.m_nentries = &nentries; - return parse_binders(r, cfg); + return parse_binders(parent, r, cfg); } optional parse_optional_binder_info(bool simple_only = false); @@ -488,7 +511,8 @@ class parser : public abstract_parser, public parser_info { void parse_close_binder_info(binder_info const & bi) { return parse_close_binder_info(optional(bi)); } /** \brief Convert an identifier into an expression (constant or local constant) based on the current scope */ - expr id_to_expr(name const & id, pos_info const & p, bool resolve_only = false, bool allow_field_notation = true, + expr id_to_expr(name const & id, ast_data & data, + bool resolve_only = false, bool allow_field_notation = true, list const & extra_locals = list()); /** Always parses an expression. Returns a synthetic sorry even if no input is consumed. */ @@ -501,10 +525,10 @@ class parser : public abstract_parser, public parser_info { /** \brief Parse an (optionally) qualified expression. If the input is of the form : , then return the pair (some(id), expr). Otherwise, parse the next expression and return (none, expr). */ - pair, expr> parse_qualified_expr(unsigned rbp = 0); + std::tuple, expr> parse_qualified_expr(unsigned rbp = 0); /** \brief If the input is of the form := , then return the pair (some(id), expr). Otherwise, parse the next expression and return (none, expr). */ - pair, expr> parse_optional_assignment(unsigned rbp = 0); + std::tuple, expr> parse_optional_assignment(unsigned rbp = 0); /** \brief Parse a pattern or expression */ expr parse_pattern_or_expr(unsigned rbp = 0); @@ -537,9 +561,9 @@ class parser : public abstract_parser, public parser_info { expr parse_scoped_expr(buffer const & ps, unsigned rbp = 0) { return parse_scoped_expr(ps.size(), ps.data(), rbp); } expr parse_expr_with_env(local_environment const & lenv, unsigned rbp = 0); - bool parse_command_like(); - void parse_command(cmd_meta const & meta); - bool parse_imports(unsigned & fingerprint, std::vector &); + bool parse_command_like(ast_data * parent = nullptr); + ast_id parse_command(cmd_meta const & meta); + bool parse_imports(unsigned & fingerprint, ast_data * parent, std::vector &); bool imports_parsed() const { return m_imports_parsed; } struct quote_scope { @@ -555,7 +579,7 @@ class parser : public abstract_parser, public parser_info { bool in_quote() const { return m_in_quote; } void maybe_throw_error(parser_error && err) override; - level parser_error_or_level(parser_error && err); + pair parser_error_or_level(parser_error && err); expr parser_error_or_expr(parser_error && err); void throw_invalid_open_binder(pos_info const & pos); @@ -610,7 +634,7 @@ class parser : public abstract_parser, public parser_info { virtual char const * get_file_name() const override; protected: - virtual std::tuple parse_definition(buffer & lp_names, buffer & params, + virtual std::tuple parse_definition(ast_data * parent, buffer & lp_names, buffer & params, bool is_example, bool is_instance, bool, bool) override; }; @@ -662,7 +686,7 @@ class dummy_def_parser : public parser_info { optional get_well_founded_tac() const { return m_wf_tac; } protected: - virtual std::tuple parse_definition(buffer & lp_names, buffer & params, + virtual std::tuple parse_definition(ast_data * parent, buffer & lp_names, buffer & params, bool is_example, bool is_instance, bool, bool) override; }; diff --git a/src/frontends/lean/parser_config.cpp b/src/frontends/lean/parser_config.cpp index bf09f5b7cb..b3ecb37021 100644 --- a/src/frontends/lean/parser_config.cpp +++ b/src/frontends/lean/parser_config.cpp @@ -38,16 +38,15 @@ std::string notation_entry_group_to_string (notation_entry_group const & g) { notation_entry replace(notation_entry const & e, std::function const & f) { if (e.is_numeral()) - return notation_entry(e.get_num(), f(e.get_expr()), e.overload(), e.parse_only()); + return notation_entry(e.get_num(), f(e.get_expr()), e.overload(), e.parse_only(), e.get_name()); else return notation_entry(e.is_nud(), map(e.get_transitions(), [&](transition const & t) { return notation::replace(t, f); }), - f(e.get_expr()), e.overload(), e.priority(), e.group(), e.parse_only()); + f(e.get_expr()), e.overload(), e.priority(), e.group(), e.parse_only(), e.get_name()); } -notation_entry::notation_entry():m_kind(notation_entry_kind::NuD) {} notation_entry::notation_entry(notation_entry const & e): - m_kind(e.m_kind), m_expr(e.m_expr), m_overload(e.m_overload), + m_kind(e.m_kind), m_name(e.m_name), m_expr(e.m_expr), m_overload(e.m_overload), m_safe_ascii(e.m_safe_ascii), m_group(e.m_group), m_parse_only(e.m_parse_only), m_priority(e.m_priority) { if (is_numeral()) @@ -57,19 +56,22 @@ notation_entry::notation_entry(notation_entry const & e): } notation_entry::notation_entry(bool is_nud, list const & ts, expr const & e, bool overload, - unsigned priority, notation_entry_group g, bool parse_only): + unsigned priority, notation_entry_group g, bool parse_only, name const & n): m_kind(is_nud ? notation_entry_kind::NuD : notation_entry_kind::LeD), - m_expr(e), m_overload(overload), m_group(g), m_parse_only(parse_only), m_priority(priority) { + m_name(n), m_expr(e), m_overload(overload), m_group(g), m_parse_only(parse_only), m_priority(priority) { new (&m_transitions) list(ts); m_safe_ascii = std::all_of(ts.begin(), ts.end(), [](transition const & t) { return t.is_safe_ascii(); }); + if (n.is_anonymous()) m_name = heuristic_name(); } notation_entry::notation_entry(notation_entry const & e, bool overload): notation_entry(e) { m_overload = overload; } -notation_entry::notation_entry(mpz const & val, expr const & e, bool overload, bool parse_only): - m_kind(notation_entry_kind::Numeral), m_expr(e), m_overload(overload), m_safe_ascii(true), m_group(notation_entry_group::Main), m_parse_only(parse_only) { +notation_entry::notation_entry(mpz const & val, expr const & e, bool overload, bool parse_only, name const & n): + m_kind(notation_entry_kind::Numeral), m_name(n), m_expr(e), m_overload(overload), + m_safe_ascii(true), m_group(notation_entry_group::Main), m_parse_only(parse_only) { new (&m_num) mpz(val); + if (n.is_anonymous()) m_name = heuristic_name(); } notation_entry::~notation_entry() { @@ -89,6 +91,29 @@ bool operator==(notation_entry const & e1, notation_entry const & e2) { return e1.get_transitions() == e2.get_transitions(); } +name notation_entry::heuristic_name() const { + std::ostringstream out; + out << "expr"; + if (m_kind == notation_entry_kind::Numeral) { + out << m_num; + } else { + if (m_kind == notation_entry_kind::LeD) out << " "; + for (auto& t : m_transitions) { + t.get_token().display(out, false); + auto& a = t.get_action(); + switch (a.kind()) { + case action_kind::Skip: break; + case action_kind::Exprs: + a.get_sep().display(out << " ", false); + if (auto term = a.get_terminator()) term->display(out, false); + break; + default: out << " "; break; + } + } + } + return out.str(); +} + struct token_state { token_table m_table; token_state() { m_table = mk_default_token_table(); } @@ -286,7 +311,7 @@ struct notation_config { if (auto idx = get_head_index(ts.size(), ts.data(), e.get_expr())) updt_inv_map(s, *idx, e); parse_table & nud = s.nud(e.group()); - nud = nud.add(ts.size(), ts.data(), e.get_expr(), e.priority(), e.overload()); + nud = nud.add(ts.size(), ts.data(), e.get_name(), e.get_expr(), e.priority(), e.overload()); break; } case notation_entry_kind::LeD: { @@ -294,7 +319,7 @@ struct notation_config { if (auto idx = get_head_index(ts.size(), ts.data(), e.get_expr())) updt_inv_map(s, *idx, e); parse_table & led = s.led(e.group()); - led = led.add(ts.size(), ts.data(), e.get_expr(), e.priority(), e.overload()); + led = led.add(ts.size(), ts.data(), e.get_name(), e.get_expr(), e.priority(), e.overload()); break; } case notation_entry_kind::Numeral: diff --git a/src/frontends/lean/parser_config.h b/src/frontends/lean/parser_config.h index 44136b0c1a..5fdca7e8a0 100644 --- a/src/frontends/lean/parser_config.h +++ b/src/frontends/lean/parser_config.h @@ -33,18 +33,19 @@ class notation_entry { list m_transitions; mpz m_num; }; + name m_name; expr m_expr; bool m_overload; bool m_safe_ascii; notation_entry_group m_group; bool m_parse_only; unsigned m_priority; + name heuristic_name() const; public: - notation_entry(); notation_entry(notation_entry const & e); notation_entry(bool is_nud, list const & ts, expr const & e, bool overload, unsigned priority, - notation_entry_group g, bool parse_only); - notation_entry(mpz const & val, expr const & e, bool overload, bool parse_only); + notation_entry_group g, bool parse_only, name const & n = {}); + notation_entry(mpz const & val, expr const & e, bool overload, bool parse_only, name const & n = {}); notation_entry(notation_entry const & e, bool overload); ~notation_entry(); notation_entry_kind kind() const { return m_kind; } @@ -52,6 +53,7 @@ class notation_entry { bool is_nud() const { return m_kind == notation_entry_kind::NuD; } list const & get_transitions() const { lean_assert(!is_numeral()); return m_transitions; } mpz const & get_num() const { lean_assert(is_numeral()); return m_num; } + name const & get_name() const { return m_name; } expr const & get_expr() const { return m_expr; } bool overload() const { return m_overload; } bool is_safe_ascii() const { return m_safe_ascii; } diff --git a/src/frontends/lean/prenum.cpp b/src/frontends/lean/prenum.cpp index f062acaace..eabffa36e3 100644 --- a/src/frontends/lean/prenum.cpp +++ b/src/frontends/lean/prenum.cpp @@ -7,6 +7,7 @@ Author: Leonardo de Moura #include #include "library/kernel_serializer.h" #include "util/numerics/mpz.h" +#include "frontends/lean/json.h" namespace lean { static name * g_prenum_name = nullptr; @@ -26,6 +27,11 @@ class prenum_macro_definition_cell : public macro_definition_cell { s.write_string(*g_prenum_opcode); s << m_value; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["value"] = m_value.to_string(); + } +#endif virtual bool operator==(macro_definition_cell const & other) const { if (auto other_ptr = dynamic_cast(&other)) { return m_value == other_ptr->m_value; diff --git a/src/frontends/lean/print_cmd.cpp b/src/frontends/lean/print_cmd.cpp index 7c7581ceea..4d89cdf874 100644 --- a/src/frontends/lean/print_cmd.cpp +++ b/src/frontends/lean/print_cmd.cpp @@ -78,14 +78,17 @@ struct print_axioms_deps { } }; -static void print_axioms(parser & p, message_builder & out) { +static void print_axioms(parser & p, ast_data & parent, message_builder & out) { if (p.curr_is_identifier()) { - name c = p.check_constant_next("invalid '#print axioms', constant expected"); + ast_id id; name c; + std::tie(id, c) = p.check_constant_next("invalid '#print axioms', constant expected"); + parent.push(id); auto env = p.env(); type_context_old tc(env, p.get_options()); auto new_out = io_state_stream(env, p.ios(), tc, out.get_text_stream().get_channel()); print_axioms_deps(env, new_out)(c); } else { + parent.push(0); bool has_axioms = false; p.env().for_each_declaration([&](declaration const & d) { name const & n = d.get_name(); @@ -99,8 +102,10 @@ static void print_axioms(parser & p, message_builder & out) { } } -static void print_prefix(parser & p, message_builder & out) { - name prefix = p.check_id_next("invalid '#print prefix' command, identifier expected"); +static void print_prefix(parser & p, ast_data & parent, message_builder & out) { + ast_id id; name prefix; + std::tie(id, prefix) = p.check_id_next("invalid '#print prefix' command, identifier expected"); + parent.push(id); buffer to_print; p.env().for_each_declaration([&](declaration const & d) { if (is_prefix_of(prefix, d.get_name())) { @@ -158,10 +163,11 @@ static bool print_parse_table(parser const & p, message_builder & rep, parse_tab return found; } -static void print_notation(parser & p, message_builder & out) { +static void print_notation(parser & p, ast_data & parent, message_builder & out) { buffer tokens; while (p.curr_is_keyword()) { tokens.push_back(p.get_token_info().token()); + parent.push(p.new_ast("token", p.pos(), tokens.back()).m_id); p.next(); } bool found = false; @@ -297,8 +303,10 @@ static void print_inductive(parser const & p, message_builder & out, name const } } -static void print_recursor_info(parser & p, message_builder & out) { - name c = p.check_constant_next("invalid '#print [recursor]', constant expected"); +static void print_recursor_info(parser & p, ast_data & parent, message_builder & out) { + ast_id id; name c; + std::tie(id, c) = p.check_constant_next("invalid '#print [recursor]', constant expected"); + parent.push(id); recursor_info info = get_recursor_info(p.env(), c); out << "recursor information\n" << " num. parameters: " << info.get_num_params() << "\n" @@ -448,19 +456,22 @@ bool print_token_info(parser const & p, message_builder & out, name const & tk) return found; } -void print_polymorphic(parser & p, message_builder & out) { +void print_polymorphic(parser & p, ast_data & data, message_builder & out) { auto pos = p.pos(); // notation if (p.curr_is_keyword()) { name tk = p.get_token_info().token(); if (print_token_info(p, out, tk)) { + data.push(p.new_ast("token", pos, tk).m_id); p.next(); return; } } - name id = p.check_id_next("invalid #print command", break_at_pos_exception::token_context::expr); + ast_id aid; name id; + std::tie(aid, id) = p.check_id_next("invalid #print command", break_at_pos_exception::token_context::expr); + data.push(aid); bool show_value = true; print_id_info(p, out, id, show_value, pos); } @@ -469,14 +480,18 @@ static void print_unification_hints(parser & p, message_builder & out) { out << pp_unification_hints(get_unification_hints(p.env()), out.get_formatter()); } -static void print_simp_rules(parser & p, message_builder & out) { - name attr = p.check_id_next("invalid '#print [simp]' command, identifier expected"); +static void print_simp_rules(parser & p, ast_data & parent, message_builder & out) { + ast_id id; name attr; + std::tie(id, attr) = p.check_id_next("invalid '#print [simp]' command, identifier expected"); + parent.push(id); simp_lemmas slss = get_simp_lemmas(p.env(), attr); out << slss.pp_simp(out.get_formatter()); } -static void print_congr_rules(parser & p, message_builder & out) { - name attr = p.check_id_next("invalid '#print [congr]' command, identifier expected"); +static void print_congr_rules(parser & p, ast_data & parent, message_builder & out) { + ast_id id; name attr; + std::tie(id, attr) = p.check_id_next("invalid '#print [congr]' command, identifier expected"); + parent.push(id); simp_lemmas slss = get_simp_lemmas(p.env(), attr); out << slss.pp_congr(out.get_formatter()); } @@ -516,36 +531,46 @@ static void print_attribute(parser & p, message_builder & out, attribute const & } } -environment print_cmd(parser & p) { +environment print_cmd(parser & p, ast_id & cmd_id) { // Fallbacks are handled via exceptions. auto _ = p.no_error_recovery_scope(); + auto& data = p.get_ast(cmd_id); - auto out = p.mk_message(p.cmd_pos(), INFORMATION); + auto out = p.mk_message(data.m_start, INFORMATION); out.set_caption("print result"); auto env = p.env(); if (p.curr() == token_kind::String) { out << p.get_str_val() << endl; + data.push(p.new_ast("string", p.pos(), p.get_str_val()).m_id); p.next(); } else if (p.curr_is_token_or_id(get_raw_tk())) { + data.push(p.new_ast(get_raw_tk(), p.pos()).m_id); p.next(); auto _ = p.error_recovery_scope(true); expr e = p.parse_expr(); + data.push(p.get_id(e)); options opts = out.get_text_stream().get_options(); opts = opts.update(get_pp_notation_name(), false); out.get_text_stream().update_options(opts) << e << endl; } else if (p.curr_is_token_or_id(get_options_tk())) { + data.push(p.new_ast(get_options_tk(), p.pos()).m_id); p.next(); out << p.ios().get_options() << endl; } else if (p.curr_is_token_or_id(get_trust_tk())) { + data.push(p.new_ast(get_trust_tk(), p.pos()).m_id); p.next(); out << "trust level: " << p.env().trust_lvl() << endl; } else if (p.curr_is_token_or_id(get_key_equivalences_tk())) { + data.push(p.new_ast(get_key_equivalences_tk(), p.pos()).m_id); p.next(); print_key_equivalences(p, out); } else if (p.curr_is_token_or_id(get_definition_tk())) { + data.push(p.new_ast(get_definition_tk(), p.pos()).m_id); p.next(); auto pos = p.pos(); - name id = p.check_id_next("invalid '#print definition', constant expected"); + ast_id id_ast; name id; + std::tie(id_ast, id) = p.check_id_next("invalid '#print definition', constant expected"); + data.push(id_ast); list cs = p.to_constants(id, "invalid '#print definition', constant expected", pos); bool first = true; for (name const & c : cs) { @@ -565,12 +590,16 @@ environment print_cmd(parser & p) { } } } else if (p.curr_is_token_or_id(get_instances_tk())) { + data.push(p.new_ast(get_instances_tk(), p.pos()).m_id); p.next(); - name c = p.check_constant_next("invalid '#print instances', constant expected"); + ast_id cid; name c; + std::tie(cid, c) = p.check_constant_next("invalid '#print instances', constant expected"); + data.push(cid); for (name const & i : get_class_instances(env, c)) { out << i << " : " << env.get(i).get_type() << endl; } } else if (p.curr_is_token_or_id(get_classes_tk())) { + data.push(p.new_ast(get_classes_tk(), p.pos()).m_id); p.next(); buffer classes; get_classes(env, classes); @@ -579,6 +608,7 @@ environment print_cmd(parser & p) { out << c << " : " << env.get(c).get_type() << endl; } } else if (p.curr_is_token_or_id(get_attributes_tk())) { + data.push(p.new_ast(get_attributes_tk(), p.pos()).m_id); p.next(); buffer attrs; get_attributes(p.env(), attrs); @@ -589,49 +619,64 @@ environment print_cmd(parser & p) { out << "[" << attr->get_name() << "] " << attr->get_description() << endl; } } else if (p.curr_is_token_or_id(get_prefix_tk())) { + data.push(p.new_ast(get_prefix_tk(), p.pos()).m_id); p.next(); - print_prefix(p, out); + print_prefix(p, data, out); } else if (p.curr_is_token_or_id(get_aliases_tk())) { + data.push(p.new_ast(get_aliases_tk(), p.pos()).m_id); p.next(); print_aliases(p, out); } else if (p.curr_is_token_or_id(get_axioms_tk())) { + data.push(p.new_ast(get_axioms_tk(), p.pos()).m_id); p.next(); - print_axioms(p, out); + print_axioms(p, data, out); } else if (p.curr_is_token_or_id(get_fields_tk())) { + data.push(p.new_ast(get_fields_tk(), p.pos()).m_id); p.next(); auto pos = p.pos(); - name S = p.check_constant_next("invalid '#print fields' command, constant expected"); + ast_id id; name S; + std::tie(id, S) = p.check_constant_next("invalid '#print fields' command, constant expected"); + data.push(id); print_fields(p, out, S, pos); } else if (p.curr_is_token_or_id(get_notation_tk())) { + data.push(p.new_ast(get_notation_tk(), p.pos()).m_id); p.next(); - print_notation(p, out); + print_notation(p, data, out); } else if (p.curr_is_token_or_id(get_inductive_tk())) { + data.push(p.new_ast(get_inductive_tk(), p.pos()).m_id); p.next(); auto pos = p.pos(); - name c = p.check_constant_next("invalid '#print inductive', constant expected"); + ast_id id; name c; + std::tie(id, c) = p.check_constant_next("invalid '#print inductive', constant expected"); + data.push(id); print_inductive(p, out, c, pos); } else if (p.curr_is_token(get_lbracket_tk())) { + data.push(p.new_ast(get_attribute_tk(), p.pos()).m_id); p.next(); auto pos = p.pos(); - auto name = p.check_id_next("invalid attribute declaration, identifier expected"); + auto name = p.check_id_next("invalid attribute declaration, identifier expected").second; p.check_token_next(get_rbracket_tk(), "invalid '#print []', ']' expected"); - if (name == "recursor") { - print_recursor_info(p, out); + data.push(p.new_ast(name, pos).m_id); + print_recursor_info(p, data, out); } else if (name == "unify") { + data.push(p.new_ast(name, pos).m_id); print_unification_hints(p, out); } else if (name == "simp") { - print_simp_rules(p, out); + data.push(p.new_ast(name, pos).m_id); + print_simp_rules(p, data, out); } else if (name == "congr") { - print_congr_rules(p, out); + data.push(p.new_ast(name, pos).m_id); + print_congr_rules(p, data, out); } else { if (!is_attribute(p.env(), name)) throw parser_error(sstream() << "unknown attribute [" << name << "]", pos); + data.push(p.new_ast("attr", pos, name).m_id); auto const & attr = get_attribute(p.env(), name); print_attribute(p, out, attr); } } else { - print_polymorphic(p, out); + print_polymorphic(p, data, out); } out.set_end_pos(p.pos()); out.report(); diff --git a/src/frontends/lean/print_cmd.h b/src/frontends/lean/print_cmd.h index c85e0414c8..3812672891 100644 --- a/src/frontends/lean/print_cmd.h +++ b/src/frontends/lean/print_cmd.h @@ -8,5 +8,5 @@ Author: Leonardo de Moura #include "kernel/environment.h" namespace lean { class parser; -environment print_cmd(parser & p); +environment print_cmd(parser & p, ast_id &); } diff --git a/src/frontends/lean/structure_cmd.cpp b/src/frontends/lean/structure_cmd.cpp index 9ef36f012c..d85198c756 100644 --- a/src/frontends/lean/structure_cmd.cpp +++ b/src/frontends/lean/structure_cmd.cpp @@ -266,9 +266,10 @@ struct structure_cmd_fn { }; parser & m_p; + ast_data & m_cmd; cmd_meta m_meta_info; environment m_env; - type_context_old m_ctx; + type_context_old m_ctx; name m_name; name m_given_name; pos_info m_name_pos; @@ -294,8 +295,9 @@ struct structure_cmd_fn { buffer m_ctx_locals; // context local constants for creating aliases unsigned m_prio; - structure_cmd_fn(parser & p, cmd_meta const & meta): + structure_cmd_fn(parser & p, ast_id cmd_id, cmd_meta const & meta): m_p(p), + m_cmd(p.get_ast(cmd_id)), m_meta_info(meta), m_env(p.env()), m_ctx(p.env()) { @@ -314,13 +316,16 @@ struct structure_cmd_fn { void parse_decl_name() { m_name_pos = m_p.pos(); buffer ls_buffer; - if (parse_univ_params(m_p, ls_buffer)) { + auto id = parse_univ_params(m_p, ls_buffer); + m_cmd.push(id); + if (id) { m_explicit_universe_params = true; m_level_names.append(ls_buffer); } else { m_explicit_universe_params = false; } - m_given_name = m_p.check_decl_id_next("invalid 'structure', identifier expected"); + std::tie(id, m_given_name) = m_p.check_decl_id_next("invalid 'structure', identifier expected"); + m_cmd.push(id); if (is_private()) { std::tie(m_env, m_private_prefix) = mk_private_prefix(m_env); m_name = m_private_prefix + m_given_name; @@ -332,12 +337,14 @@ struct structure_cmd_fn { /** \brief Parse structure parameters */ void parse_params() { + auto& bis = m_p.new_ast("binders", m_p.pos()); if (!m_p.curr_is_token(get_extends_tk()) && !m_p.curr_is_token(get_assign_tk()) && !m_p.curr_is_token(get_colon_tk())) { unsigned rbp = 0; bool allow_default = true; - m_p.parse_binders(m_params, rbp, allow_default); + m_p.parse_binders(&bis, m_params, rbp, allow_default); } + m_cmd.push(bis.m_children.empty() ? 0 : bis.m_id); for (expr const & l : m_params) m_p.add_local(l); } @@ -377,18 +384,24 @@ struct structure_cmd_fn { /** \brief Parse optional extends clause */ void parse_extends() { + ast_id extends_id = 0; if (m_p.curr_is_token(get_extends_tk())) { + auto& exts = m_p.new_ast(get_extends_tk(), m_p.pos()); + extends_id = exts.m_id; m_p.next(); while (true) { auto pos = m_p.pos(); - bool is_private_parent = false; + auto& ast = m_p.new_ast("parent", pos); + exts.push(ast.m_id); + ast_id is_private_parent = 0; if (m_p.curr_is_token(get_private_tk())) { + is_private_parent = m_p.new_ast(get_private_tk(), pos).m_id; m_p.next(); - is_private_parent = true; } - pair, expr> qparent = m_p.parse_qualified_expr(); - m_parent_refs.push_back(qparent.first); - expr const & parent = qparent.second; + ast_id q_id; optional q; expr parent; + std::tie(q_id, q, parent) = m_p.parse_qualified_expr(); + ast.push(is_private_parent).push(q_id).push(m_p.get_id(parent)); + m_parent_refs.push_back(q); m_parents.push_back(parent); m_private_parents.push_back(is_private_parent); name const & parent_name = check_parent(parent); @@ -401,13 +414,17 @@ struct structure_cmd_fn { throw parser_error("invalid 'structure' extends, parent structure seems to be ill-formed", pos); intro_type = binding_body(intro_type); } + ast_id renames_id = 0; m_renames.push_back(rename_vector()); if (!m_subobjects && m_p.curr_is_token(get_renaming_tk())) { + auto& ast1 = m_p.new_ast(get_renaming_tk(), m_p.pos()); + renames_id = ast1.m_id; m_p.next(); rename_vector & v = m_renames.back(); while (m_p.curr_is_identifier()) { auto from_pos = m_p.pos(); name from_id = m_p.get_name_val(); + ast_id from_ast = m_p.new_ast("ident", from_pos, from_id).m_id; if (std::find_if(v.begin(), v.end(), [&](pair const & p) { return p.first == from_id; }) != v.end()) throw parser_error(sstream() << "invalid 'structure' renaming, a rename from '" << @@ -415,17 +432,21 @@ struct structure_cmd_fn { check_from_rename(parent_name, intro_type, from_id, from_pos); m_p.next(); m_p.check_token_next(get_arrow_tk(), "invalid 'structure' renaming, '->' expected"); - name to_id = m_p.check_id_next("invalid 'structure' renaming, identifier expected"); + ast_id to_ast; name to_id; + std::tie(to_ast, to_id) = m_p.check_id_next("invalid 'structure' renaming, identifier expected"); if (from_id == to_id) throw parser_error(sstream() << "invalid 'structure' renaming, redundant rename", from_pos); + ast1.push(m_p.new_ast(get_arrow_tk(), from_pos).push(from_ast).push(to_ast).m_id); v.emplace_back(from_id, to_id); } } + ast.push(renames_id); if (!m_p.curr_is_token(get_comma_tk())) break; m_p.next(); } } + m_cmd.push(extends_id); } /** \brief Parse resultant universe */ @@ -434,6 +455,7 @@ struct structure_cmd_fn { if (m_p.curr_is_token(get_colon_tk())) { m_p.next(); m_type = m_p.parse_expr(); + m_cmd.push(m_p.get_id(m_type)); while (is_annotation(m_type)) m_type = get_annotation_arg(m_type); if (!is_sort(m_type)) @@ -457,6 +479,7 @@ struct structure_cmd_fn { } else { m_infer_result_universe = true; m_type = m_p.save_pos(mk_sort(mk_level_placeholder()), pos); + m_cmd.push(0); } } @@ -795,12 +818,16 @@ struct structure_cmd_fn { } } - void parse_field_block(binder_info const & bi) { + ast_id parse_field_block(binder_info const & bi) { buffer> names; auto start_pos = m_p.pos(); + auto& vars = m_p.new_ast("vars", start_pos); + auto& ast = m_p.new_ast(name("field").append_after(bi.hash()), start_pos).push(vars.m_id); while (m_p.curr_is_identifier()) { auto p = m_p.pos(); - auto n = m_p.check_atomic_id_next("invalid field, atomic identifier expected"); + ast_id id; name n; + std::tie(id, n) = m_p.check_atomic_id_next("invalid field, atomic identifier expected"); + vars.push(id); if (is_internal_subobject_field(n)) throw parser_error(sstream() << "invalid field name '" << n << "', identifiers starting with '_' are reserved to the system", p); names.emplace_back(p, n); @@ -810,28 +837,44 @@ struct structure_cmd_fn { expr type; optional default_value; + ast_id& kind_id = ast.push(0).m_children.back(); implicit_infer_kind kind = implicit_infer_kind::RelaxedImplicit; { parser::local_scope scope(m_p); buffer params; if (names.size() == 1) { - m_p.parse_optional_binders(params, kind); + auto& bis = m_p.new_ast("binders", m_p.pos()); + m_p.parse_optional_binders(&bis, params, kind_id, kind); + ast.push(bis.m_children.empty() ? 0 : bis.m_id); for (expr const & param : params) m_p.add_local(param); + } else { + ast.push(0); } if (m_p.curr_is_token(get_assign_tk())) { + auto& ast1 = m_p.new_ast(get_assign_tk(), m_p.pos()); + ast.push(0).push(ast1.m_id); type = m_p.save_pos(mk_expr_placeholder(), m_p.pos()); m_p.next(); default_value = m_p.parse_expr(); + ast1.push(m_p.get_id(*default_value)); } else { m_p.check_token_next(get_colon_tk(), "invalid field, ':' expected"); type = m_p.parse_expr(); + ast.push(m_p.get_id(type)); if (m_p.curr_is_token(get_assign_tk())) { + auto& ast1 = m_p.new_ast(get_assign_tk(), m_p.pos()); + ast.push(ast1.m_id); m_p.next(); default_value = m_p.parse_expr(); + ast1.push(m_p.get_id(*default_value)); } else if (m_p.curr_is_token(get_period_tk())) { - type = parse_auto_param(m_p, type); + auto& ast1 = m_p.new_ast(get_period_tk(), m_p.pos()); + ast.push(ast1.m_id); + ast_id id; + std::tie(id, type) = parse_auto_param(m_p, type); + ast1.push(id); } } type = Pi(params, type); @@ -839,6 +882,7 @@ struct structure_cmd_fn { *default_value = Fun(params, *default_value); } + if (default_value && !is_explicit(bi)) { throw parser_error("invalid field, it is not explicit, but it has a default value", start_pos); } @@ -871,16 +915,19 @@ struct structure_cmd_fn { m_fields.emplace_back(local, default_value, field_kind::new_field, kind); } } + return ast.m_id; } /** \brief Parse new fields declared in this structure */ void parse_new_fields() { parser::local_scope scope(m_p); add_locals(); + auto& fields = m_p.new_ast("fields", m_p.pos()); + m_cmd.push(fields.m_id); while (!m_p.curr_is_command_like()) { if (auto bi = m_p.parse_optional_binder_info()) { - if (!m_p.parse_local_notation_decl()) - parse_field_block(*bi); + ast_id nota = m_p.parse_local_notation_decl(); + fields.push(nota ? nota : parse_field_block(*bi)); m_p.parse_close_binder_info(*bi); } else { break; @@ -1301,6 +1348,7 @@ struct structure_cmd_fn { } environment operator()() { + m_cmd.push(m_meta_info.m_modifiers_id); process_header(); module::scope_pos_info scope(m_name_pos); if (m_p.curr_is_token(get_assign_tk())) { @@ -1308,11 +1356,14 @@ struct structure_cmd_fn { m_mk_pos = m_p.pos(); if (m_p.curr_is_token(get_lparen_tk()) || m_p.curr_is_token(get_lcurly_tk()) || m_p.curr_is_token(get_lbracket_tk())) { + m_cmd.push(0); m_mk_short = LEAN_DEFAULT_STRUCTURE_INTRO; m_mk_infer = implicit_infer_kind::RelaxedImplicit; } else { - m_mk_short = m_p.check_atomic_id_next("invalid 'structure', atomic identifier expected"); - m_mk_infer = parse_implicit_infer_modifier(m_p); + ast_id id1, id2; + std::tie(id1, m_mk_short) = m_p.check_atomic_id_next("invalid 'structure', atomic identifier expected"); + std::tie(id2, m_mk_infer) = parse_implicit_infer_modifier(m_p); + m_cmd.push(m_p.new_ast("mk", m_mk_pos).push(id1).push(id2).m_id); if (!m_p.curr_is_command_like()) m_p.check_token_next(get_dcolon_tk(), "invalid 'structure', '::' expected"); } @@ -1323,6 +1374,7 @@ struct structure_cmd_fn { m_mk_short = LEAN_DEFAULT_STRUCTURE_INTRO; m_mk_infer = implicit_infer_kind::RelaxedImplicit; m_mk = m_name + m_mk_short; + m_cmd.push(0).push(0); process_empty_new_fields(); } infer_resultant_universe(); @@ -1345,20 +1397,21 @@ struct structure_cmd_fn { } }; -environment structure_cmd(parser & p, cmd_meta const & meta) { +environment structure_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta) { p.next(); - return structure_cmd_fn(p, meta)(); + return structure_cmd_fn(p, cmd_id, meta)(); } -environment class_cmd(parser & p, cmd_meta const & _meta) { +environment class_cmd(parser & p, ast_id & cmd_id, cmd_meta const & _meta) { auto meta = _meta; meta.m_attrs.set_persistent(true); meta.m_attrs.set_attribute(p.env(), "class"); p.next(); if (p.curr_is_token(get_inductive_tk())) { - return inductive_cmd(p, meta); + p.get_ast(cmd_id).m_type = "class_inductive"; + return inductive_cmd(p, cmd_id, meta); } else { - return structure_cmd_fn(p, meta)(); + return structure_cmd_fn(p, cmd_id, meta)(); } } diff --git a/src/frontends/lean/structure_cmd.h b/src/frontends/lean/structure_cmd.h index 644e7e244f..946da1a89b 100644 --- a/src/frontends/lean/structure_cmd.h +++ b/src/frontends/lean/structure_cmd.h @@ -9,8 +9,8 @@ Author: Leonardo de Moura #include "frontends/lean/decl_util.h" #include "frontends/lean/cmd_table.h" namespace lean { -environment structure_cmd(parser & p, cmd_meta const & meta); -environment class_cmd(parser & p, cmd_meta const & meta); +environment structure_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta); +environment class_cmd(parser & p, ast_id & cmd_id, cmd_meta const & meta); buffer get_structure_fields(environment const & env, name const & S); void register_structure_cmd(cmd_table & r); /** \brief Return true iff \c S is a structure created with the structure command */ diff --git a/src/frontends/lean/structure_instance.cpp b/src/frontends/lean/structure_instance.cpp index 4e482129e9..af9d562590 100644 --- a/src/frontends/lean/structure_instance.cpp +++ b/src/frontends/lean/structure_instance.cpp @@ -7,6 +7,7 @@ Author: Leonardo de Moura #include #include "util/sstream.h" #include "library/kernel_serializer.h" +#include "frontends/lean/json.h" #include "frontends/lean/parser.h" #include "frontends/lean/tokens.h" #include "frontends/lean/structure_instance.h" @@ -42,6 +43,14 @@ class structure_instance_macro_cell : public macro_definition_cell { s << *g_structure_instance_opcode << m_struct << m_catchall; write_list(s, m_fields); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["struct"] = json_of_name(m_struct); + j["catchall"] = m_catchall; + auto& jl = j["fields"] = json::array(); + for (auto& f : m_fields) jl += json_of_name(f); + } +#endif name const & get_struct() const { return m_struct; } bool get_catchall() const { return m_catchall; } list const & get_field_names() const { return m_fields; } diff --git a/src/frontends/lean/tactic_notation.cpp b/src/frontends/lean/tactic_notation.cpp index ada0721b19..b883705abc 100644 --- a/src/frontends/lean/tactic_notation.cpp +++ b/src/frontends/lean/tactic_notation.cpp @@ -118,7 +118,7 @@ static bool is_curr_exact_shortcut(parser & p) { return p.curr_is_token(get_calc_tk()); } -static optional is_interactive_tactic(parser & p, name const & tac_class) { +static optional is_interactive_tactic(parser & p, ast_data & id_ast, name const & tac_class) { name id; switch (p.curr()) { case token_kind::Identifier: @@ -130,6 +130,7 @@ static optional is_interactive_tactic(parser & p, name const & tac_class) default: return {}; } + id_ast.m_value = id; id = get_interactive_tactic_full_name(tac_class, id); if (p.env().find(id)) return optional(id); @@ -141,13 +142,15 @@ static expr parse_begin_end_block(parser & p, pos_info const & start_pos, name c expr parse_nested_interactive_tactic(parser & p, name const & tac_class, bool use_istep) { auto pos = p.pos(); + expr r; if (p.curr_is_token(get_lcurly_tk())) { - return parse_begin_end_block(p, pos, get_rcurly_tk(), tac_class, use_istep); + r = parse_begin_end_block(p, pos, get_rcurly_tk(), tac_class, use_istep); } else if (p.curr_is_token(get_begin_tk())) { - return parse_begin_end_block(p, pos, get_end_tk(), tac_class, use_istep); + r = parse_begin_end_block(p, pos, get_end_tk(), tac_class, use_istep); } else { - return p.parser_error_or_expr({"invalid nested auto-quote tactic, '{' or 'begin' expected", pos}); + r = p.parser_error_or_expr({"invalid nested auto-quote tactic, '{' or 'begin' expected", pos}); } + return r; } static optional is_itactic(expr const & type) { @@ -166,7 +169,7 @@ static optional is_itactic(expr const & type) { return optional(pre.get_prefix()); } -static expr parse_interactive_tactic(parser & p, name const & decl_name, name const & tac_class, bool use_istep) { +static expr parse_interactive_tactic(parser & p, name const & decl_name, ast_data & parent, name const & tac_class, bool use_istep) { auto pos = p.pos(); expr type = p.env().get(decl_name).get_type(); buffer args; @@ -179,9 +182,11 @@ static expr parse_interactive_tactic(parser & p, name const & decl_name, name co expr arg_type = binding_domain(type); if (is_app_of(arg_type, get_interactive_parse_name())) { parser::quote_scope scope(p, true); - args.push_back(parse_interactive_param(p, arg_type)); + args.push_back(parse_interactive_param(p, parent, arg_type)); } else if (auto new_tac_class = is_itactic(arg_type)) { - args.push_back(parse_nested_interactive_tactic(p, *new_tac_class, use_istep)); + expr e = parse_nested_interactive_tactic(p, *new_tac_class, use_istep); + parent.push(p.get_id(e)); + args.push_back(e); } else { break; } @@ -192,6 +197,7 @@ static expr parse_interactive_tactic(parser & p, name const & decl_name, name co p.check_break_before(); auto pos = p.pos(); args.push_back(p.parse_expr(get_max_prec())); + parent.push(p.new_ast("expr", pos).push(p.get_id(args.back())).m_id); if (p.pos() == pos) { // parse_expr does not necessarily consume input if there is a syntax error break; @@ -239,12 +245,12 @@ struct parse_tactic_fn { return m_p.save_pos(mk_app(mk_constant(get_has_orelse_orelse_name()), tac1, tac2), pos); } - expr parse_qexpr(unsigned rbp = 0) { + pair parse_qexpr(unsigned rbp = 0) { auto p = m_p.pos(); parser::quote_scope scope1(m_p, true); restore_decl_meta_scope scope2; expr e = m_p.parse_expr(rbp); - return m_p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ false), p); + return {m_p.get_id(e), m_p.save_pos(mk_pexpr_quote_and_substs(e, /* is_strict */ false), p)}; } expr parse_elem_core(bool save_info) { @@ -257,11 +263,13 @@ struct parse_tactic_fn { e.m_token_info.m_param = m_tac_class; throw; } - expr r; + ast_id id; expr r; auto pos = m_p.pos(); - if (auto dname = is_interactive_tactic(m_p, m_tac_class)) { + auto& id_ast = m_p.new_ast("tactic", pos); + if (auto dname = is_interactive_tactic(m_p, id_ast, m_tac_class)) { try { - r = parse_interactive_tactic(m_p, *dname, m_tac_class, m_use_istep); + id = id_ast.m_id; + r = parse_interactive_tactic(m_p, *dname, id_ast, m_tac_class, m_use_istep); } catch (break_at_pos_exception & e) { if (!m_p.get_complete() && (!e.m_token_info.m_token.size() || @@ -274,17 +282,20 @@ struct parse_tactic_fn { throw; } } else if (is_curr_exact_shortcut(m_p)) { - expr arg = parse_qexpr(); + auto& data = m_p.new_ast("exact_shortcut", pos, m_p.get_token_info().value()); + ast_id arg_ast; expr arg; + std::tie(arg_ast, arg) = parse_qexpr(); + id = data.push(arg_ast).m_id; r = m_p.mk_app(m_p.save_pos(mk_constant(get_interactive_tactic_full_name(m_tac_class, "refine")), pos), arg, pos); if (m_use_istep) r = mk_tactic_istep(m_p, r, pos, pos, m_tac_class); } else { r = m_p.parse_expr(); + id = m_p.new_ast("(", pos).push(m_p.get_id(r)).m_id; if (m_use_istep) r = mk_tactic_istep(m_p, r, pos, pos, m_tac_class); } - if (save_info) - return concat(mk_tactic_save_info(m_p, pos, m_tac_class), r, pos); - else - return r; + if (save_info) r = concat(mk_tactic_save_info(m_p, pos, m_tac_class), r, pos); + m_p.set_ast_pexpr(id, r); + return r; } expr parse_block(pos_info const & pos, name const & end_tk) { @@ -297,23 +308,26 @@ struct parse_tactic_fn { auto pos = m_p.pos(); name const & end_tk = m_p.curr_is_token(get_begin_tk()) ? get_end_tk() : get_rcurly_tk(); expr next_tac = parse_block(pos, end_tk); + ast_id id = m_p.get_id(next_tac); auto end_pos = m_p.pos_of(next_tac); if (use_solve1) { - next_tac = mk_tactic_solve1(m_p, next_tac, pos, end_pos, m_tac_class, m_use_istep && save_info); + next_tac = mk_tactic_solve1(m_p, next_tac, pos, end_pos, m_tac_class, m_use_istep && save_info); } if (save_info) { expr info_tac = mk_tactic_save_info(m_p, pos, m_tac_class); - return concat(info_tac, next_tac, pos); - } else { - return next_tac; + next_tac = concat(info_tac, next_tac, pos); } + m_p.set_ast_pexpr(id, next_tac); + return next_tac; } else if (m_p.curr_is_token(get_lbracket_tk())) { auto pos = m_p.pos(); m_p.next(); buffer args; + auto& data = m_p.new_ast("[", pos); if (!m_p.curr_is_token(get_rbracket_tk())) { while (true) { args.push_back(parse_elem(save_info, false)); + data.push(m_p.get_id(args.back())); if (!m_p.curr_is_token(get_comma_tk())) break; m_p.next(); @@ -323,21 +337,24 @@ struct parse_tactic_fn { expr r = mk_lean_list(m_p, args, pos); expr type = mk_app(mk_constant(get_list_name()), mk_tactic_unit(m_tac_class)); r = m_p.save_pos(mk_typed_expr(type, r), pos); + m_p.set_ast_pexpr(data.m_id, r); return r; } else { if (m_p.curr_is_token(get_by_tk())) { // `by tac` ~> `solve1 { tac }` + auto& data = m_p.new_ast("by", m_p.pos()); m_p.next(); auto pos = m_p.pos(); expr tac = operator()(save_info); + data.push(m_p.get_id(tac)); auto end_pos = m_p.pos_of(tac); tac = mk_tactic_solve1(m_p, tac, pos, end_pos, m_tac_class, m_use_istep && save_info); if (save_info) { expr info_tac = mk_tactic_save_info(m_p, pos, m_tac_class); - return concat(info_tac, tac, pos); - } else { - return tac; + tac = concat(info_tac, tac, pos); } + m_p.set_ast_pexpr(data.m_id, tac); + return tac; } else { return parse_elem_core(save_info); } @@ -347,11 +364,14 @@ struct parse_tactic_fn { expr parse_orelse(expr left) { auto start_pos = m_p.pos(); expr r = left; + auto& data = m_p.new_ast("<|>", m_p.pos_of(r)).push(m_p.get_id(r)); while (m_p.curr_is_token(get_orelse_tk())) { m_p.next(); expr curr = parse_elem(false); + data.push(m_p.get_id(curr)); r = orelse(r, curr, start_pos); } + m_p.set_ast_pexpr(data.m_id, r); return r; } @@ -359,6 +379,7 @@ struct parse_tactic_fn { auto start_pos = m_p.pos(); optional pos; expr r = left; + auto& data = m_p.new_ast(";", m_p.pos_of(r)).push(m_p.get_id(r)); try { while (m_p.curr_is_token(get_semicolon_tk())) { m_p.next(); @@ -366,6 +387,7 @@ struct parse_tactic_fn { expr curr = parse_elem(save_info); if (m_p.curr_is_token(get_orelse_tk())) curr = parse_orelse(curr); + data.push(m_p.get_id(curr)); r = andthen(r, curr, start_pos); } } catch (break_at_pos_exception & ex) { @@ -373,6 +395,7 @@ struct parse_tactic_fn { ex.report_goal_pos(*pos); throw; } + m_p.set_ast_pexpr(data.m_id, r); return r; } @@ -411,21 +434,22 @@ static optional is_tactic_class(environment const & env, name const & n) { } } -static name parse_tactic_class(parser & p, name tac_class) { +static pair parse_tactic_class(parser & p, name tac_class) { if (p.curr_is_token(get_lbracket_tk())) { p.next(); auto id_pos = p.pos(); - name id = p.check_id_next("invalid 'begin [...] ... end' block, identifier expected"); + ast_id id_ast; name id; + std::tie(id_ast, id) = p.check_id_next("invalid 'begin [...] ... end' block, identifier expected"); auto new_class = is_tactic_class(p.env(), id); if (!new_class) { p.maybe_throw_error({sstream() << "invalid 'begin [" << id << "] ...end' block, " << "'" << id << "' is not a valid tactic class", id_pos}); - return tac_class; + return {id_ast, tac_class}; } p.check_token_next(get_rbracket_tk(), "invalid 'begin [...] ... end block', ']' expected"); - return *new_class; + return {id_ast, *new_class}; } else { - return tac_class; + return {0, tac_class}; } } @@ -493,7 +517,7 @@ struct parse_begin_end_block_fn { return ::lean::parse_tactic(m_p, m_tac_class, m_use_istep); } - expr operator()(pos_info const & start_pos, name const & end_token) { + expr operator()(pos_info const & pos, name const & end_token) { auto sync = [&] { while (!m_p.curr_is_token(get_comma_tk()) && !m_p.curr_is_token(end_token) && !m_p.curr_is_token(get_semicolon_tk()) && !m_p.curr_is_token(get_orelse_tk())) { @@ -505,10 +529,12 @@ struct parse_begin_end_block_fn { m_p.maybe_throw_error({"sync", m_p.pos()}); }; + auto& group = m_p.new_ast(m_p.get_token_info().value(), pos); m_p.next(); name new_tac_class = m_tac_class; + ast_id tac_class_id = 0; if (m_tac_class == get_tactic_name()) - new_tac_class = parse_tactic_class(m_p, m_tac_class); + std::tie(tac_class_id, new_tac_class) = parse_tactic_class(m_p, m_tac_class); optional cfg; bool is_ext_tactic_class = m_tac_class == get_tactic_name() && new_tac_class != get_tactic_name(); if (is_ext_tactic_class && m_p.curr_is_token(get_with_tk())) { @@ -516,13 +542,15 @@ struct parse_begin_end_block_fn { cfg = m_p.parse_expr(); m_p.check_token_next(get_comma_tk(), "invalid begin [...] with cfg, ... end block, ',' expected"); } + group.push(tac_class_id).push(cfg ? m_p.get_id(*cfg) : 0); m_tac_class = new_tac_class; buffer to_concat; - to_concat.push_back(mk_tactic_save_info(m_p, start_pos, m_tac_class)); + to_concat.push_back(mk_tactic_save_info(m_p, group.m_start, m_tac_class)); while (!m_p.curr_is_token(end_token)) { pos_info pos = m_p.pos(); try { to_concat.push_back(parse_tactic()); + group.push(m_p.get_id(to_concat.back())); if (!m_p.curr_is_token(end_token)) { m_p.without_break_at_pos([&]() { if (!m_p.check_token_next(get_comma_tk(), "invalid 'begin-end' expression, ',' expected")) { @@ -539,7 +567,7 @@ struct parse_begin_end_block_fn { if (m_p.pos() == pos) break; } auto end_pos = m_p.pos(); - expr r = concat(to_concat, start_pos); + expr r = concat(to_concat, group.m_start); r = concat(r, mk_tactic_save_info(m_p, end_pos, m_tac_class), end_pos); try { m_p.next(); @@ -547,18 +575,20 @@ struct parse_begin_end_block_fn { ex.report_goal_pos(end_pos); throw; } - if (!is_ext_tactic_class && m_tac_class != get_tactic_name()) { - return r; - } else if (cfg) { - return copy_tag(r, mk_tactic_execute_with(r, *cfg, m_tac_class)); - } else { - return copy_tag(r, mk_tactic_execute(r, m_tac_class)); + if (is_ext_tactic_class || m_tac_class == get_tactic_name()) { + if (cfg) { + r = copy_tag(r, mk_tactic_execute_with(r, *cfg, m_tac_class)); + } else { + r = copy_tag(r, mk_tactic_execute(r, m_tac_class)); + } } + m_p.set_ast_pexpr(group.m_id, r); + return r; } }; -static expr parse_begin_end_block(parser & p, pos_info const & start_pos, name const & end_token, name tac_class, bool use_istep) { - return parse_begin_end_block_fn(p, tac_class, use_istep)(start_pos, end_token); +static expr parse_begin_end_block(parser & p, pos_info const & pos, name const & end_token, name tac_class, bool use_istep) { + return parse_begin_end_block_fn(p, tac_class, use_istep)(pos, end_token); } expr parse_begin_end_expr_core(parser & p, pos_info const & pos, name const & end_token) { @@ -591,10 +621,13 @@ expr parse_by(parser & p, unsigned, expr const *, pos_info const & pos) { try { bool use_istep = true; expr tac = parse_tactic(p, get_tactic_name(), use_istep); + ast_id id = p.new_ast(get_by_tk(), pos).push(p.get_id(tac)).m_id; tac = mk_tactic_execute(tac, get_tactic_name()); expr type = mk_tactic_unit(get_tactic_name()); expr r = p.save_pos(mk_typed_expr(type, tac), tac_pos); - return p.save_pos(mk_by(r), pos); + r = p.save_pos(mk_by(r), pos); + p.set_ast_pexpr(id, r); + return r; } catch (break_at_pos_exception & ex) { ex.report_goal_pos(tac_pos); throw ex; @@ -630,14 +663,17 @@ expr parse_interactive_tactic_block(parser & p, unsigned, expr const *, pos_info name const & tac_class = get_tactic_name(); bool use_istep = false; expr r = parse_tactic(p, tac_class, use_istep); + auto& data = p.new_ast("`[", pos).push(p.get_id(r)); erase_quoted_terms_pos_info(p, r); while (p.curr_is_token(get_comma_tk())) { p.next(); expr next = parse_tactic(p, tac_class, use_istep); + data.push(p.get_id(next)); erase_quoted_terms_pos_info(p, next); r = p.mk_app({p.save_pos(mk_constant(get_has_bind_and_then_name()), pos), r, next}, pos); } p.check_token_next(get_rbracket_tk(), "invalid auto-quote tactic block, ']' expected"); + p.set_ast_pexpr(data.m_id, r); return r; } diff --git a/src/frontends/lean/type_util.cpp b/src/frontends/lean/type_util.cpp index 24544983dd..07214c67e6 100644 --- a/src/frontends/lean/type_util.cpp +++ b/src/frontends/lean/type_util.cpp @@ -32,21 +32,24 @@ environment add_alias(parser & p, environment env, bool composite, return add_alias(p, env, id, full_id, ctx_levels, ctx_params); } -implicit_infer_kind parse_implicit_infer_modifier(parser & p) { +pair parse_implicit_infer_modifier(parser & p) { if (p.curr_is_token(get_lcurly_tk())) { + ast_id id = p.new_ast("{}", p.pos()).m_id; p.next(); p.check_token_next(get_rcurly_tk(), "invalid introduction rule, '}' expected"); - return implicit_infer_kind::RelaxedImplicit; + return {id, implicit_infer_kind::RelaxedImplicit}; } else if (p.curr_is_token(get_lparen_tk())) { + ast_id id = p.new_ast("()", p.pos()).m_id; p.next(); p.check_token_next(get_rparen_tk(), "invalid introduction rule, ')' expected"); - return implicit_infer_kind::None; + return {id, implicit_infer_kind::None}; } else if (p.curr_is_token(get_lbracket_tk())) { + ast_id id = p.new_ast("[]", p.pos()).m_id; p.next(); p.check_token_next(get_rbracket_tk(), "invalid introduction rule, ']' expected"); - return implicit_infer_kind::Implicit; + return {id, implicit_infer_kind::Implicit}; } else { - return implicit_infer_kind::RelaxedImplicit; + return {0, implicit_infer_kind::RelaxedImplicit}; } } } diff --git a/src/frontends/lean/type_util.h b/src/frontends/lean/type_util.h index b5f04e3a77..50c13cd40a 100644 --- a/src/frontends/lean/type_util.h +++ b/src/frontends/lean/type_util.h @@ -6,6 +6,7 @@ Author: Leonardo de Moura */ #pragma once #include "library/util.h" +#include "library/abstract_parser.h" namespace lean { class parser; @@ -28,5 +29,5 @@ environment add_alias(parser & p, environment env, name const & id, name const & environment add_alias(parser & p, environment env, bool composite, name const & full_id, levels const & ctx_levels, buffer const & ctx_params); -implicit_infer_kind parse_implicit_infer_modifier(parser & p); +pair parse_implicit_infer_modifier(parser & p); } diff --git a/src/frontends/lean/user_command.cpp b/src/frontends/lean/user_command.cpp index b5d5256caf..177bef153a 100644 --- a/src/frontends/lean/user_command.cpp +++ b/src/frontends/lean/user_command.cpp @@ -51,20 +51,29 @@ static environment add_user_command(environment const & env, name const & d) { expr dummy = mk_true(); - auto run = [=](parser & p, cmd_meta const & meta) { + auto run = [=](parser & p, ast_id & cmd_id, cmd_meta const & meta) { + auto& data = p.get_ast(cmd_id); + data.m_value = data.m_type; + data.m_type = "user_command"; + expr parser = mk_constant(d); // we don't want to reflect `meta` into an expr, so abstract the parameter and pass it as a vm_obj arg if (takes_meta_infos) { parser = mk_app(parser, mk_var(0)); + data.push(meta.m_modifiers_id); + } else { + data.push(0); } // `parse (tk c)` arg parser = mk_app(parser, mk_constant(get_unit_star_name())); for (expr t = type; is_pi(t); t = binding_body(t)) { expr arg_type = binding_domain(t); if (is_app_of(arg_type, get_interactive_parse_name())) { - parser = mk_app(parser, parse_interactive_param(p, arg_type)); + parser = mk_app(parser, parse_interactive_param(p, data, arg_type)); } else { + auto pos = p.pos(); expr e = p.parse_expr(get_max_prec()); + data.push(p.new_ast("expr", pos).push(p.get_id(e)).m_id); if (!closed(e) || has_local(e)) { throw elaborator_exception(e, "invalid argument to user-defined command, must be closed term"); } @@ -84,7 +93,7 @@ static environment add_user_command(environment const & env, name const & d) { if (takes_meta_infos) { return add_command(env, tk, cmd_info(tk, "description", run)); } else { - return add_command(env, tk, cmd_info(tk, "description", [=](parser & p) { return run(p, {}); })); + return add_command(env, tk, cmd_info(tk, "description", [=](parser & p, ast_id & cmd_id) { return run(p, cmd_id, {}); })); } } diff --git a/src/frontends/lean/user_notation.cpp b/src/frontends/lean/user_notation.cpp index 6f1ab37e89..154e0150ba 100644 --- a/src/frontends/lean/user_notation.cpp +++ b/src/frontends/lean/user_notation.cpp @@ -67,12 +67,15 @@ static environment add_user_notation(environment const & env, name const & d, un parser = mk_app(parser, mk_pexpr_quote(args[0])); // `parse (tk c)` arg parser = mk_app(parser, mk_constant(get_unit_star_name())); + auto& data = p.new_ast("user_notation", pos, d); for (expr t = type; is_pi(t); t = binding_body(t)) { expr arg_type = binding_domain(t); if (is_app_of(arg_type, get_interactive_parse_name())) { - parser = mk_app(parser, parse_interactive_param(p, arg_type)); + parser = mk_app(parser, parse_interactive_param(p, data, arg_type)); } else { + auto pos = p.pos(); expr e = p.parse_expr(get_max_prec()); + data.push(p.new_ast("expr", pos).push(p.get_id(e)).m_id); if (!closed(e) || has_local(e)) { throw elaborator_exception(e, "invalid argument to user-defined notation, must be closed term"); } @@ -81,7 +84,13 @@ static environment add_user_notation(environment const & env, name const & d, un } parser = p.elaborate("_user_notation", {}, parser).first; try { - return to_expr(run_parser(p, parser)); + ast_id id; vm_obj obj; + std::tie(id, obj) = run_parser(p, parser); + data.push(id); + expr r = to_expr(obj); + r.set_tag(nulltag); + p.set_ast_pexpr(data.m_id, r); + return r; } catch (formatted_exception const & ex) { if (ex.get_pos() && *ex.get_pos() >= pos) { throw; diff --git a/src/frontends/lean/util.cpp b/src/frontends/lean/util.cpp index fe8e07d4f7..eada00e2f2 100644 --- a/src/frontends/lean/util.cpp +++ b/src/frontends/lean/util.cpp @@ -35,6 +35,7 @@ Author: Leonardo de Moura #include "frontends/lean/parser.h" #include "frontends/lean/tokens.h" #include "frontends/lean/decl_util.h" // TODO(Leo): remove +#include "frontends/lean/json.h" #include "frontends/lean/prenum.h" #ifndef LEAN_DEFAULT_AUTO_PARAM_CHECK_EXISTS @@ -254,8 +255,9 @@ level mk_result_level(buffer const & r_lvls) { } } -std::tuple parse_local_expr(parser & p, name const & decl_name, metavar_context & mctx, bool relaxed) { +std::tuple parse_local_expr(parser & p, ast_data & parent, name const & decl_name, metavar_context & mctx, bool relaxed) { expr e = p.parse_expr(); + parent.push(p.get_id(e)); bool check_unassigend = !relaxed; expr new_e; level_param_names ls; std::tie(new_e, ls) = p.elaborate(decl_name, mctx, e, check_unassigend); @@ -263,9 +265,9 @@ std::tuple parse_local_expr(parser & p, name const & de return std::make_tuple(new_e, new_ls); } -std::tuple parse_local_expr(parser & p, name const & decl_name, bool relaxed) { +std::tuple parse_local_expr(parser & p, ast_data & parent, name const & decl_name, bool relaxed) { metavar_context mctx; - return parse_local_expr(p, decl_name, mctx, relaxed); + return parse_local_expr(p, parent, decl_name, mctx, relaxed); } optional is_uniquely_aliased(environment const & env, name const & n) { @@ -310,9 +312,11 @@ char const * close_binder_string(binder_info const & bi, bool unicode) { else return ")"; } -pair parse_option_name(parser & p, char const * error_msg) { +pair parse_option_name(parser & p, ast_data & parent, char const * error_msg) { auto id_pos = p.pos(); - name id = p.check_id_next(error_msg, break_at_pos_exception::token_context::option); + ast_id id_ast; name id; + std::tie(id_ast, id) = p.check_id_next(error_msg, break_at_pos_exception::token_context::option); + parent.push(id_ast); option_declarations decls = get_option_declarations(); auto it = decls.find(id); if (!it) { @@ -368,17 +372,18 @@ static bool is_tactic_unit(environment const & env, expr const & c) { return tc.is_def_eq(tc.infer(c), mk_tactic_unit()); } -expr parse_auto_param(parser & p, expr const & type) { +pair parse_auto_param(parser & p, expr const & type) { p.next(); auto pos = p.pos(); - name tac_id = p.check_decl_id_next("invalid auto_param, identifier expected"); + ast_id id; name tac_id; + std::tie(id, tac_id) = p.check_decl_id_next("invalid auto_param, identifier expected"); if (get_auto_param_check_exists(p.get_options())) { - expr tac_expr = p.id_to_expr(tac_id, pos, true); + expr tac_expr = p.id_to_expr(tac_id, p.get_ast(id), true); if (!is_tactic_unit(p.env(), tac_expr)) throw parser_error(sstream() << "invalid auto_param, '" << tac_id << "' must have type (tactic unit)", pos); - return mk_auto_param(type, const_name(tac_expr)); + return {id, mk_auto_param(type, const_name(tac_expr))}; } else { - return mk_auto_param(type, tac_id); + return {id, mk_auto_param(type, tac_id)}; } } @@ -416,6 +421,12 @@ class field_notation_macro_cell : public macro_definition_cell { virtual expr check_type(expr const &, abstract_type_context &, bool) const override { throw_pn_ex(); } virtual optional expand(expr const &, abstract_type_context &) const override { throw_pn_ex(); } virtual void write(serializer & s) const override { s << *g_field_notation_opcode << m_field << m_field_idx; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["field"] = json_of_name(m_field); + j["idx"] = m_field_idx; + } +#endif bool is_anonymous() const { return m_field.is_anonymous(); } name const & get_field_name() const { lean_assert(!is_anonymous()); return m_field; } unsigned get_field_idx() const { lean_assert(is_anonymous()); return m_field_idx; } diff --git a/src/frontends/lean/util.h b/src/frontends/lean/util.h index 92c1acae55..d9b1896e55 100644 --- a/src/frontends/lean/util.h +++ b/src/frontends/lean/util.h @@ -8,6 +8,7 @@ Author: Leonardo de Moura #include "kernel/expr.h" #include "kernel/expr_sets.h" #include "kernel/type_checker.h" +#include "library/abstract_parser.h" #include "library/util.h" #include "library/locals.h" #include "library/vm/vm.h" @@ -82,7 +83,7 @@ expr Pi_as_is(expr const & local, expr const & e); level mk_result_level(buffer const & r_lvls); /** \brief Auxiliary function for check/eval/find_decl */ -std::tuple parse_local_expr(parser & p, name const & decl_name, bool relaxed = true); +std::tuple parse_local_expr(parser & p, ast_data & parent, name const & decl_name, bool relaxed = true); optional is_uniquely_aliased(environment const & env, name const & n); @@ -99,7 +100,7 @@ char const * open_binder_string(binder_info const & bi, bool unicode); char const * close_binder_string(binder_info const & bi, bool unicode); /** \brief Parse option name */ -pair parse_option_name(parser & p, char const * error_msg); +pair parse_option_name(parser & p, ast_data & parent, char const * error_msg); expr quote(unsigned u); expr quote(char const * str); @@ -110,7 +111,7 @@ bool is_no_info(expr const & e); expr mk_opt_param(expr const & t, expr const & val); expr mk_auto_param(expr const & t, name const & tac_name); -expr parse_auto_param(parser & p, expr const & type); +pair parse_auto_param(parser & p, expr const & type); /* Add frozen annotation around constants and local constants occurring in \c e. This annotation is used to prevent lean from resolving the names again diff --git a/src/kernel/expr.h b/src/kernel/expr.h index 6affa7b0b4..1c9f6bcefb 100644 --- a/src/kernel/expr.h +++ b/src/kernel/expr.h @@ -16,6 +16,7 @@ Author: Leonardo de Moura #include "util/name.h" #include "util/hash.h" #include "util/buffer.h" +#include "util/lean_json.h" #include "util/list_fn.h" #include "util/optional.h" #include "util/serializer.h" @@ -27,6 +28,7 @@ Author: Leonardo de Moura namespace lean { class abstract_type_context; class tlean_exporter; +class abstract_ast_exporter; // Tags are used by frontends to mark expressions. They are automatically propagated by // procedures such as update_app, update_binder, etc. @@ -369,6 +371,9 @@ class macro_definition_cell { virtual bool can_textualize() const { return false; } virtual void textualize(tlean_exporter &) const { throw exception("macro::textualize not implemented by default"); } typedef std::function reader; +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const = 0; +#endif }; /** \brief Smart pointer for macro definitions */ diff --git a/src/library/CMakeLists.txt b/src/library/CMakeLists.txt index ba9bac42f1..12f4b25937 100644 --- a/src/library/CMakeLists.txt +++ b/src/library/CMakeLists.txt @@ -22,7 +22,7 @@ add_library(library OBJECT deep_copy.cpp expr_lt.cpp io_state.cpp documentation.cpp check.cpp parray.cpp process.cpp pipe.cpp handle.cpp profiling.cpp time_task.cpp abstract_context_cache.cpp context_cache.cpp unique_id.cpp persistent_context_cache.cpp elab_context.cpp - expr_address.cpp tlean_exporter.cpp + expr_address.cpp tlean_exporter.cpp ast_exporter.cpp ) if(EMSCRIPTEN) add_dependencies(library gmp) diff --git a/src/library/abstract_parser.h b/src/library/abstract_parser.h index 93d301f48a..71b4c74752 100644 --- a/src/library/abstract_parser.h +++ b/src/library/abstract_parser.h @@ -6,6 +6,7 @@ Author: Sebastian Ullrich */ #pragma once #include +#include #include "util/name_map.h" #include "util/exception_with_pos.h" #include "kernel/pos_info_provider.h" @@ -22,12 +23,33 @@ struct parser_error : public exception_with_pos { virtual void rethrow() const override { throw *this; } }; +typedef unsigned ast_id; + +struct ast_data { + ast_id m_id; + pos_info m_start; + optional m_end; + name m_type; + std::vector m_children; + name m_value; + optional m_pexpr; + optional> m_expr; + + ast_data(ast_id id, pos_info start, name type, name value = {}): + m_id(id), m_start(start), m_type(type), m_value(value) {} + + ast_data & push(ast_id id) { + lean_assert(m_id != 0); + m_children.push_back(id); + return *this; + } +}; + /** \brief Base class for frontend parsers with basic functions */ class abstract_parser : public pos_info_provider { public: /** \brief Return the current position information */ virtual pos_info pos() const = 0; - /** \brief Return true iff the current token is a keyword (or command keyword) named \c tk */ virtual bool curr_is_token(name const & tk) const = 0; /** \brief Return true iff the current token is a numeral */ @@ -35,7 +57,7 @@ class abstract_parser : public pos_info_provider { /** \brief Read the next token if the current one is not End-of-file. */ virtual void next() = 0; - virtual unsigned parse_small_nat() = 0; - virtual std::string parse_string_lit() = 0; + virtual pair parse_small_nat() = 0; + virtual pair parse_string_lit() = 0; }; } diff --git a/src/library/annotation.cpp b/src/library/annotation.cpp index b6ea4418e1..313ce92445 100644 --- a/src/library/annotation.cpp +++ b/src/library/annotation.cpp @@ -11,6 +11,7 @@ Author: Leonardo de Moura #include "kernel/abstract_type_context.h" #include "library/kernel_serializer.h" #include "library/annotation.h" +#include "frontends/lean/json.h" namespace lean { static name * g_annotation = nullptr; @@ -45,6 +46,11 @@ class annotation_macro_definition_cell : public macro_definition_cell { s.write_string(get_annotation_opcode()); s << m_name; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["name"] = json_of_name(m_name); + } +#endif virtual bool operator==(macro_definition_cell const & other) const override { if (auto other_ptr = dynamic_cast(&other)) { return m_name == other_ptr->m_name; diff --git a/src/library/ast_exporter.cpp b/src/library/ast_exporter.cpp new file mode 100644 index 0000000000..dfa8313d38 --- /dev/null +++ b/src/library/ast_exporter.cpp @@ -0,0 +1,188 @@ +/* +Copyright (c) 2021 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Mario Carneiro +*/ +#include +#include +#include "library/ast_exporter.h" +#include "frontends/lean/json.h" +#include "frontends/lean/parser.h" +#include "util/file_lock.h" +#include "kernel/expr.h" + +namespace lean { + +#ifdef LEAN_JSON + +struct ast_exporter : abstract_ast_exporter { + std::vector const & m_ast; + tag_ast_table const & m_tag_ast_table; + std::vector m_reachable; + json m_levels; + json m_exprs = {nullptr}; + level_map m_level2idx; + std::unordered_map m_expr2idx; + + void mark(ast_id id) { + if (!m_reachable[id]) { + m_reachable[id] = true; + for (ast_id i : m_ast[id]->m_children) mark(i); + } + } + + ast_exporter(std::vector const & ast, tag_ast_table const & tag_ast_table): + m_ast(ast), m_tag_ast_table(tag_ast_table), + m_reachable(ast.size(), false) { + m_levels.push_back("0"); + m_level2idx.emplace(mk_level_zero(), 0); + m_reachable[0] = true; + mark(AST_TOP_ID); + } + + unsigned export_level(const level & l) override { + auto it = m_level2idx.find(l); + if (it != m_level2idx.end()) return it->second; + json j; + switch (l.kind()) { + case level_kind::Zero: lean_unreachable(); + case level_kind::Succ: + j["suc"] = export_level(succ_of(l)); break; + case level_kind::Max: + j["max"] = {export_level(max_lhs(l)), export_level(max_rhs(l))}; break; + case level_kind::IMax: + j["imax"] = {export_level(imax_lhs(l)), export_level(imax_rhs(l))}; break; + case level_kind::Param: + j["param"] = json_of_name(param_id(l)); break; + case level_kind::Meta: + j["mvar"] = json_of_name(meta_id(l)); break; + } + auto r = m_levels.size(); + m_levels.push_back(std::move(j)); + m_level2idx.emplace(l, r); + return r; + } + + unsigned export_expr(const expr & e) override { + auto it = m_expr2idx.find(e.raw()); + if (it != m_expr2idx.end()) return it->second; + json j; + switch (e.kind()) { + case expr_kind::Var: + j["var"] = var_idx(e); break; + case expr_kind::Sort: + j["sort"] = export_level(sort_level(e)); break; + case expr_kind::Meta: + j["mvar"] = { + {"name", json_of_name(mlocal_name(e))}, + {"pp", json_of_name(mlocal_pp_name(e))}, + {"type", export_expr(mlocal_type(e))}}; + break; + case expr_kind::Local: + j["local"] = { + {"name", json_of_name(mlocal_name(e))}, + {"pp", json_of_name(mlocal_pp_name(e))}, + {"bi", local_info(e).hash()}, + {"type", export_expr(mlocal_type(e))}}; + break; + case expr_kind::Constant: { + json jl = json::array(); + for (auto l : const_levels(e)) jl += export_level(l); + j["const"] = {json_of_name(const_name(e)), std::move(jl)}; + break; + } + case expr_kind::Macro: { + auto& jm = j[macro_def(e).get_name().get_string()] = json::object(); + macro_def(e).raw()->write_json(*this, jm); + auto& jargs = jm["args"] = json::array(); + for (unsigned i = 0; i < macro_num_args(e); i++) + jargs += export_expr(macro_arg(e, i)); + break; + } + case expr_kind::Lambda: + case expr_kind::Pi: + j[(e.kind() == expr_kind::Pi) ? "Pi" : "lam"] = { + {"name", json_of_name(binding_name(e))}, + {"bi", binding_info(e).hash()}, + {"dom", export_expr(binding_domain(e))}, + {"body", export_expr(binding_body(e))}}; + break; + case expr_kind::App: + j["app"] = {export_expr(app_fn(e)), export_expr(app_arg(e))}; + break; + case expr_kind::Let: + j["let"] = { + {"name", json_of_name(let_name(e))}, + {"type", export_expr(let_type(e))}, + {"value", export_expr(let_value(e))}, + {"body", export_expr(let_body(e))}}; + break; + } + auto r = m_exprs.size(); + m_exprs.push_back(std::move(j)); + m_expr2idx.emplace(e.raw(), r); + return r; + } + + void write_ast(std::ostream & out) { + json asts = json::array({nullptr}); + for (unsigned i = 1; i < m_ast.size(); i++) { + if (!m_reachable[i] || !m_ast[i]) { + asts.push_back(nullptr); + continue; + } + auto& data = *m_ast[i]; + lean_assert(data.m_type.is_atomic()); + json j { + // {"id", data.m_id}, + {"kind", data.m_type.get_string()}, + {"start", {data.m_start.first, data.m_start.second}}, + }; + if (auto end = data.m_end) j["end"] = {end->first, end->second}; + if (data.m_value) { + j["value"] = json_of_name(data.m_value); + } + if (!data.m_children.empty()) + j["children"] = data.m_children; + if (data.m_pexpr) j["pexpr"] = export_expr(*data.m_pexpr); + if (data.m_expr) { + auto res = (*data.m_expr)->peek(); + j["expr"] = res ? export_expr(*res) : 0; + } + asts.push_back(j); + } + json r = { + {"ast", asts}, + {"file", AST_TOP_ID}, + {"level", std::move(m_levels)}, + {"expr", std::move(m_exprs)}, + }; + out << r; + } +}; + +std::string json_of_lean(std::string const & lean_fn) { + if (lean_fn.size() > 5 && lean_fn.substr(lean_fn.size() - 5) == ".lean") { + return lean_fn.substr(0, lean_fn.size() - 5) + ".ast.json"; + } else { + throw exception(sstream() << "not a .lean file: " << lean_fn); + } +} + +void export_ast(parser & p) { + if (p.m_ast.empty() || p.m_ast_invalid) return; + auto ast_fn = json_of_lean(p.m_file_name); + std::cerr << "exporting " << ast_fn << std::endl; + exclusive_file_lock output_lock(ast_fn); + std::ofstream out(ast_fn); + ast_exporter(p.m_ast, p.m_tag_ast_table).write_ast(out); + out.close(); + if (!out) throw exception(sstream() << "failed to write ast file: " << ast_fn); +} + +#else +void export_ast(parser &) {} +#endif // LEAN_JSON + +} diff --git a/src/library/ast_exporter.h b/src/library/ast_exporter.h new file mode 100644 index 0000000000..c0d05169f8 --- /dev/null +++ b/src/library/ast_exporter.h @@ -0,0 +1,21 @@ +/* +Copyright (c) 2021 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Mario Carneiro +*/ +#pragma once + +#include "frontends/lean/parser.h" + +namespace lean { + +class abstract_ast_exporter { +public: + virtual unsigned export_level(const level & l) = 0; + virtual unsigned export_expr(const expr & l) = 0; +}; + +void export_ast(parser & p); + +} diff --git a/src/library/attribute_manager.cpp b/src/library/attribute_manager.cpp index e7bc0a80c0..1bbb7cd436 100644 --- a/src/library/attribute_manager.cpp +++ b/src/library/attribute_manager.cpp @@ -12,11 +12,27 @@ Author: Leonardo de Moura #include "util/sstream.h" #include "library/attribute_manager.h" #include "library/scoped_ext.h" +#include "frontends/lean/parser.h" namespace lean { template class typed_attribute; template class typed_attribute; +ast_id key_value_data::parse(abstract_parser & p) { + std::cout << "in extern parser" << std::endl; + lean_assert(dynamic_cast(&p)); + auto& p2 = *static_cast(&p); + auto& data = p2.new_ast("key_value", p2.pos()); + auto n = p.parse_string_lit(); + auto l = p.parse_string_lit(); + data.push(n.first).push(l.first); + std::cout << "link symbol: " << n.second << std::endl; + std::cout << "library symbol: " << l.second << std::endl; + this->m_symbol = n.second; + this->m_library = l.second; + return data.m_id; +} + static name_map * g_system_attributes = nullptr; static user_attribute_ext * g_user_attribute_ext = nullptr; static attr_data_ptr * g_default_attr_data_ptr = nullptr; @@ -253,20 +269,26 @@ priority_queue attribute::get_instances_by_prio(environmen return q; } -attr_data_ptr attribute::parse_data(abstract_parser &) const { +attr_data_ptr attribute::parse_data(abstract_parser &, ast_data & parent) const { + parent.push(0); return get_default_attr_data(); } -void indices_attribute_data::parse(abstract_parser & p) { +ast_id indices_attribute_data::parse(abstract_parser & p) { buffer vs; + lean_assert(dynamic_cast(&p)); + auto& p2 = *static_cast(&p); + auto& data = p2.new_ast("indices", p2.pos()); while (p.curr_is_numeral()) { auto pos = p.pos(); - unsigned v = p.parse_small_nat(); - if (v == 0) + auto v = p.parse_small_nat(); + data.push(v.first); + if (v.second == 0) throw parser_error("invalid attribute parameter, value must be positive", pos); - vs.push_back(v - 1); + vs.push_back(v.second - 1); } m_idxs = to_list(vs); + return data.m_id; } void register_incompatible(char const * attr1, char const * attr2) { diff --git a/src/library/attribute_manager.h b/src/library/attribute_manager.h index 00b1d5b896..ec5d59f338 100644 --- a/src/library/attribute_manager.h +++ b/src/library/attribute_manager.h @@ -22,7 +22,7 @@ struct attr_data { virtual unsigned hash() const { return 0; } - virtual void parse(abstract_parser &) {} + virtual ast_id parse(abstract_parser &) { return 0; } virtual void print(std::ostream &) {} virtual void textualize(tlean_exporter &) const {} virtual ~attr_data() {} @@ -69,7 +69,7 @@ class attribute { bool get_persistent(environment const &, name const &) const; virtual void get_instances(environment const &, buffer &) const; priority_queue get_instances_by_prio(environment const &) const; - virtual attr_data_ptr parse_data(abstract_parser &) const; + virtual attr_data_ptr parse_data(abstract_parser & p, ast_data & parent) const; virtual environment unset(environment env, io_state const & ios, name const & n, bool persistent) const; virtual unsigned get_fingerprint(environment const & env) const; @@ -136,9 +136,9 @@ class typed_attribute : public attribute { typed_attribute(name const & id, char const * descr, after_set_proc after_set = {}, before_unset_proc before_unset = {}) : attribute(id, descr, after_set, before_unset) {} - virtual attr_data_ptr parse_data(abstract_parser & p) const override { + virtual attr_data_ptr parse_data(abstract_parser & p, ast_data & parent) const override { auto data = new Data; - data->parse(p); + parent.push(data->parse(p)); return attr_data_ptr(data); } @@ -212,7 +212,7 @@ struct indices_attribute_data : public attr_data { void read(deserializer & d) { m_idxs = read_list(d); } - void parse(abstract_parser & p) override; + ast_id parse(abstract_parser & p) override; virtual void print(std::ostream & out) override { for (auto p : m_idxs) { out << " " << p + 1; @@ -247,15 +247,7 @@ struct key_value_data : public attr_data { d >> m_library; } - void parse(abstract_parser & p) override { - std::cout << "in extern parser" << std::endl; - std::string n = p.parse_string_lit(); - std::string l = p.parse_string_lit(); - std::cout << "link symbol: " << n << std::endl; - std::cout << "library symbol: " << l << std::endl; - this->m_symbol = n; - this->m_library = l; - } + ast_id parse(abstract_parser & p) override; virtual void print(std::ostream & out) override { out << "external"; diff --git a/src/library/choice.cpp b/src/library/choice.cpp index cf0fe69f1d..4f6f9ad8e2 100644 --- a/src/library/choice.cpp +++ b/src/library/choice.cpp @@ -31,6 +31,9 @@ class choice_macro_cell : public macro_definition_cell { // we should be able to write choice expressions because of notation declarations s.write_string(*g_choice_opcode); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif }; static macro_definition * g_choice = nullptr; diff --git a/src/library/compiler/nat_value.cpp b/src/library/compiler/nat_value.cpp index 17d3e7cb2f..73cfb88e6d 100644 --- a/src/library/compiler/nat_value.cpp +++ b/src/library/compiler/nat_value.cpp @@ -12,6 +12,7 @@ Author: Leonardo de Moura #include "library/util.h" #include "library/kernel_serializer.h" #include "library/replace_visitor_with_tc.h" +#include "frontends/lean/json.h" namespace lean { static name * g_nat_macro = nullptr; @@ -65,6 +66,11 @@ class nat_value_macro : public macro_definition_cell { virtual void write(serializer & s) const override { s << *g_nat_opcode << m_value; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["value"] = m_value.to_string(); + } +#endif mpz const & get_value() const { return m_value; } }; diff --git a/src/library/compiler/rec_fn_macro.cpp b/src/library/compiler/rec_fn_macro.cpp index 59bba6f50c..a44b93102d 100644 --- a/src/library/compiler/rec_fn_macro.cpp +++ b/src/library/compiler/rec_fn_macro.cpp @@ -9,6 +9,7 @@ Author: Leonardo de Moura #include "library/kernel_serializer.h" #include "library/compiler/rec_fn_macro.h" #include "kernel/abstract_type_context.h" +#include "frontends/lean/json.h" namespace lean { static name * g_rec_fn_macro_id = nullptr; @@ -51,6 +52,12 @@ class rec_fn_macro_definition_cell : public macro_definition_cell { s << m_name; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["name"] = json_of_name(m_name); + } +#endif + virtual bool operator==(macro_definition_cell const & other) const override { if (auto other_ptr = dynamic_cast(&other)) { return m_name == other_ptr->m_name; diff --git a/src/library/constructions/projection.cpp b/src/library/constructions/projection.cpp index 90d83452aa..4b276795a4 100644 --- a/src/library/constructions/projection.cpp +++ b/src/library/constructions/projection.cpp @@ -12,6 +12,7 @@ Author: Leonardo de Moura #include "kernel/instantiate.h" #include "kernel/kernel_exception.h" #include "kernel/inductive/inductive.h" +#include "library/ast_exporter.h" #include "library/reducible.h" #include "library/projection.h" #include "library/module.h" @@ -22,6 +23,7 @@ Author: Leonardo de Moura #include "library/class.h" #include "library/constructions/projection.h" #include "library/constructions/util.h" +#include "frontends/lean/json.h" namespace lean { [[ noreturn ]] static void throw_ill_formed(name const & n) { @@ -138,6 +140,19 @@ class projection_macro_definition_cell : public macro_definition_cell { x.out() << "#PROJ_MACRO " << I_name << " " << constructor_name << " " << proj_name << " " << m_idx; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter & exp, json & j) const override { + j["I"] = json_of_name(m_I_name); + j["constr"] = json_of_name(m_constructor_name); + j["proj"] = json_of_name(m_proj_name); + j["idx"] = m_idx; + auto& jl = j["params"] = json::array(); + for (auto& p : m_ps) jl += json_of_name(p); + j["type"] = exp.export_expr(m_type); + j["val"] = exp.export_expr(m_val); + } +#endif + virtual bool operator==(macro_definition_cell const & other) const { if (auto other_ptr = dynamic_cast(&other)) { return m_proj_name == other_ptr->m_proj_name; diff --git a/src/library/delayed_abstraction.cpp b/src/library/delayed_abstraction.cpp index 9cd266a297..506f40e95f 100644 --- a/src/library/delayed_abstraction.cpp +++ b/src/library/delayed_abstraction.cpp @@ -10,6 +10,7 @@ Author: Leonardo de Moura #include "kernel/abstract_type_context.h" #include "library/replace_visitor.h" #include "library/metavar_context.h" +#include "frontends/lean/json.h" namespace lean { static name * g_delayed_abstraction_macro = nullptr; @@ -55,6 +56,12 @@ class delayed_abstraction_macro : public macro_definition_cell { return length(m_value); } virtual void write(serializer &) const override { lean_unreachable(); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + auto& x = j["value"] = json::array(); + for (auto& n : m_value) x += json_of_name(n); + } +#endif list const & get_names() const { return m_value; } }; diff --git a/src/library/equations_compiler/equations.cpp b/src/library/equations_compiler/equations.cpp index 5c4e95adb7..bf96456556 100644 --- a/src/library/equations_compiler/equations.cpp +++ b/src/library/equations_compiler/equations.cpp @@ -27,6 +27,7 @@ Author: Leonardo de Moura #include "library/normalize.h" #include "library/pp_options.h" #include "library/equations_compiler/equations.h" +#include "frontends/lean/json.h" namespace lean { static name * g_equations_name = nullptr; @@ -49,6 +50,9 @@ class as_pattern_macro_cell : public macro_definition_cell { virtual expr check_type(expr const &, abstract_type_context &, bool) const override { throw_asp_ex(); } virtual optional expand(expr const &, abstract_type_context &) const override { throw_asp_ex(); } virtual void write(serializer & s) const override { s << *g_as_pattern_opcode; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif }; bool operator==(equations_header const & h1, equations_header const & h2) { @@ -80,6 +84,24 @@ class equations_macro_cell : public macro_definition_cell { write_list(s, m_header.m_fn_names); write_list(s, m_header.m_fn_actual_names); } + +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["num_fns"] = m_header.m_num_fns; + j["is_private"] = m_header.m_is_private; + j["is_meta"] = m_header.m_is_meta; + j["is_noncomputable"] = m_header.m_is_noncomputable; + j["is_lemma"] = m_header.m_is_lemma; + j["aux_lemmas"] = m_header.m_aux_lemmas; + j["prev_errors"] = m_header.m_prev_errors; + j["gen_code"] = m_header.m_gen_code; + auto& names = j["fn_names"] = json::array(); + for (auto& n : m_header.m_fn_names) names += json_of_name(n); + auto& actual_names = j["fn_actual_names"] = json::array(); + for (auto& n : m_header.m_fn_actual_names) actual_names += json_of_name(n); + } +#endif + virtual bool operator==(macro_definition_cell const & other) const override { if (auto other_ptr = dynamic_cast(&other)) { return m_header == other_ptr->m_header; @@ -112,6 +134,11 @@ class equation_macro_cell : public equation_base_macro_cell { s.write_string(*g_equation_opcode); s.write_bool(m_ignore_if_unused); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["ignore_if_unused"] = m_ignore_if_unused; + } +#endif virtual bool operator==(macro_definition_cell const & other) const override { if (auto other_ptr = dynamic_cast(&other)) { return m_ignore_if_unused == other_ptr->m_ignore_if_unused; @@ -127,6 +154,9 @@ class no_equation_macro_cell : public equation_base_macro_cell { public: virtual name get_name() const override { return *g_no_equation_name; } virtual void write(serializer & s) const override { s.write_string(*g_no_equation_opcode); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif }; static macro_definition * g_as_pattern = nullptr; @@ -282,6 +312,9 @@ class equations_result_macro_cell : public macro_definition_cell { return some_expr(macro_arg(m, 0)); } virtual void write(serializer & s) const { s << *g_equations_result_opcode; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif }; static macro_definition * g_equations_result = nullptr; @@ -298,7 +331,7 @@ void initialize_equations() { g_equations_name = new name("equations"); g_equation_name = new name("equation"); g_no_equation_name = new name("no_equation"); - g_inaccessible_name = new name("innaccessible"); + g_inaccessible_name = new name("inaccessible"); g_equations_result_name = new name("equations_result"); g_as_pattern_name = new name("as_pattern"); g_equation = new macro_definition(new equation_macro_cell(false)); diff --git a/src/library/module_mgr.cpp b/src/library/module_mgr.cpp index 84d74127b6..24f23b0423 100644 --- a/src/library/module_mgr.cpp +++ b/src/library/module_mgr.cpp @@ -17,6 +17,7 @@ Author: Gabriel Ebner #include "frontends/lean/pp.h" #include "frontends/lean/parser.h" #include "library/library_task_builder.h" +#include "library/ast_exporter.h" namespace lean { @@ -279,12 +280,16 @@ void module_mgr::build_lean(std::shared_ptr const & mod, name_set c mod->m_snapshots = snapshots; } + // make a copy of the mod_parser only if exporting AST + auto mod_parser_for_export = m_export_ast ? mod_parser_fn : nullptr; + auto initial_env = m_initial_env; unsigned src_hash = mod->m_src_hash; unsigned trans_hash = mod->m_trans_hash; mod->m_result = map( get_end(snapshots), - [id, initial_env, ldr, src_hash, trans_hash](module_parser_result const & res) { + [id, initial_env, ldr, src_hash, trans_hash, mod_parser_for_export] + (module_parser_result const & res) { module_info::parse_result parse_res; lean_always_assert(res.m_snapshot_at_end); @@ -293,7 +298,7 @@ void module_mgr::build_lean(std::shared_ptr const & mod, name_set c initial_env, [=] { return ldr; }); parse_res.m_opts = res.m_snapshot_at_end->m_options; - + if (mod_parser_for_export) export_ast(mod_parser_for_export->get_parser()); return parse_res; }).build(); diff --git a/src/library/module_mgr.h b/src/library/module_mgr.h index 6ee2e3d34b..7a618b2013 100644 --- a/src/library/module_mgr.h +++ b/src/library/module_mgr.h @@ -51,6 +51,7 @@ struct module_info { struct parse_result { options m_opts; std::shared_ptr m_loaded_module; + std::function m_ast_export_fn; }; task m_result; @@ -95,6 +96,7 @@ class module_mgr { bool m_save_olean = false; bool m_use_old_oleans = false; bool m_report_widgets = true; + bool m_export_ast = false; search_path m_path; environment m_initial_env; @@ -141,6 +143,7 @@ class module_mgr { bool get_use_old_oleans() const { return m_use_old_oleans; } void set_report_widgets(bool report_widgets) { m_report_widgets = report_widgets; } bool get_report_widgets() const { return m_report_widgets; } + void set_export_ast(bool export_ast) { m_export_ast = export_ast; } environment get_initial_env() const { return m_initial_env; } options get_options() const { return m_ios.get_options(); } diff --git a/src/library/quote.cpp b/src/library/quote.cpp index e394315566..b75de5acb2 100644 --- a/src/library/quote.cpp +++ b/src/library/quote.cpp @@ -16,6 +16,7 @@ Author: Leonardo de Moura #include "library/util.h" #include "library/quote.h" #include "library/type_context.h" +#include "library/ast_exporter.h" namespace lean { static std::string * g_expr_quote_opcode = nullptr; @@ -61,6 +62,12 @@ class expr_quote_macro : public macro_definition_cell { } virtual unsigned hash() const override { return m_value.hash(); } virtual void write(serializer & s) const override { s << *g_expr_quote_opcode << m_value << m_reflected; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter & exp, json & j) const override { + j["value"] = exp.export_expr(m_value); + j["reflected"] = m_reflected; + } +#endif expr const & get_value() const { return m_value; } bool const & is_reflected() const { return m_reflected; } }; diff --git a/src/library/scoped_ext.h b/src/library/scoped_ext.h index c5b73116ec..ef8d84ca94 100644 --- a/src/library/scoped_ext.h +++ b/src/library/scoped_ext.h @@ -143,7 +143,6 @@ class scoped_ext : public environment_extension { entry m_entry; - modification() {} modification(entry const & e) : m_entry(e) {} void perform(environment & env) const override { diff --git a/src/library/sorry.cpp b/src/library/sorry.cpp index 58979df71b..9da8283f73 100644 --- a/src/library/sorry.cpp +++ b/src/library/sorry.cpp @@ -12,6 +12,7 @@ Author: Leonardo de Moura #include "library/module.h" #include "kernel/for_each_fn.h" #include "library/kernel_serializer.h" +#include "frontends/lean/json.h" namespace lean { static name * g_sorry_name = nullptr; @@ -50,6 +51,11 @@ class sorry_macro_cell : public macro_definition_cell { s.write_string(*g_sorry_opcode); s.write_bool(m_synthetic); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["synthetic"] = m_synthetic; + } +#endif }; void initialize_sorry() { diff --git a/src/library/string.cpp b/src/library/string.cpp index c6e618b121..a5da2061f7 100644 --- a/src/library/string.cpp +++ b/src/library/string.cpp @@ -14,6 +14,7 @@ Author: Leonardo de Moura #include "library/constants.h" #include "library/num.h" #include "library/trace.h" +#include "frontends/lean/json.h" namespace lean { static name * g_string_macro = nullptr; @@ -109,6 +110,11 @@ class string_macro : public macro_definition_cell { virtual unsigned hash() const { return std::hash()(m_value); } virtual void write(serializer & s) const { s << *g_string_opcode << m_value; } std::string const & get_value() const { return m_value; } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json & j) const override { + j["value"] = m_value; + } +#endif }; expr mk_string_macro(std::string const & v) { diff --git a/src/library/tactic/ac_tactics.cpp b/src/library/tactic/ac_tactics.cpp index ec1cec811c..7202bc80e6 100644 --- a/src/library/tactic/ac_tactics.cpp +++ b/src/library/tactic/ac_tactics.cpp @@ -115,6 +115,10 @@ class ac_app_macro_cell : public macro_definition_cell { s.write_string(*g_ac_app_opcode); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif + virtual bool operator==(macro_definition_cell const & other) const { ac_app_macro_cell const * other_ptr = dynamic_cast(&other); return other_ptr; @@ -625,6 +629,10 @@ class perm_ac_macro_cell : public macro_definition_cell { s.write_string(*g_perm_ac_opcode); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif + virtual bool operator==(macro_definition_cell const & other) const { perm_ac_macro_cell const * other_ptr = dynamic_cast(&other); return other_ptr; diff --git a/src/library/tactic/backward/backward_lemmas.cpp b/src/library/tactic/backward/backward_lemmas.cpp index b72c7020cf..3b55a8522a 100644 --- a/src/library/tactic/backward/backward_lemmas.cpp +++ b/src/library/tactic/backward/backward_lemmas.cpp @@ -19,6 +19,7 @@ Author: Leonardo de Moura #include "library/vm/vm_expr.h" #include "library/tactic/tactic_state.h" #include "library/tactic/backward/backward_lemmas.h" +#include "frontends/lean/parser.h" namespace lean { static optional get_backward_target(type_context_old & ctx, expr type) { @@ -51,11 +52,16 @@ struct intro_attr_data : public attr_data { m_eager = d.read_bool(); } - void parse(abstract_parser & p) override { + ast_id parse(abstract_parser & p) override { + ast_id r = 0; if (p.curr_is_token("!")) { - p.next(); + lean_assert(dynamic_cast(&p)); + auto& p2 = *static_cast(&p); + r = p2.new_ast("!", p2.pos()).m_id; + p2.next(); m_eager = true; } + return r; } virtual void print(std::ostream & out) override { if (m_eager) diff --git a/src/library/tactic/smt/util.cpp b/src/library/tactic/smt/util.cpp index 08c8250b10..c1e23f2c4f 100644 --- a/src/library/tactic/smt/util.cpp +++ b/src/library/tactic/smt/util.cpp @@ -32,6 +32,10 @@ class cc_proof_macro_cell : public macro_definition_cell { lean_unreachable(); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif + virtual bool operator==(macro_definition_cell const & other) const { cc_proof_macro_cell const * other_ptr = dynamic_cast(&other); return other_ptr; diff --git a/src/library/tactic/user_attribute.cpp b/src/library/tactic/user_attribute.cpp index d42f3811cf..3d3863abb7 100644 --- a/src/library/tactic/user_attribute.cpp +++ b/src/library/tactic/user_attribute.cpp @@ -66,12 +66,18 @@ class user_attribute : public typed_attribute { user_attribute(name const & id, name const & d, char const * descr, after_set_proc const & after_set, before_unset_proc const & before_unset) : typed_attribute(id, descr, after_set, before_unset), m_decl(d) {} - attr_data_ptr parse_data(abstract_parser & p) const override final { + attr_data_ptr parse_data(abstract_parser & p, ast_data & parent) const override final { lean_assert(dynamic_cast(&p)); auto & p2 = *static_cast(&p); type_context_old ctx(p2.env(), p2.get_options()); expr parser = mk_app(ctx, get_user_attribute_parse_reflect_name(), 3, mk_constant(m_decl)); - expr param = to_expr(run_parser(p2, parser)); + ast_id id; vm_obj obj; + std::tie(id, obj) = run_parser(p2, parser); + p2.get_ast(id).m_value = m_decl; + expr param = to_expr(obj); + param.set_tag(nulltag); + p2.set_ast_pexpr(id, param); + parent.push(id); return attr_data_ptr(new user_attribute_data(param)); } }; diff --git a/src/library/typed_expr.cpp b/src/library/typed_expr.cpp index fe3fdf2201..4fd2c86730 100644 --- a/src/library/typed_expr.cpp +++ b/src/library/typed_expr.cpp @@ -59,15 +59,18 @@ class typed_expr_macro_definition_cell : public macro_definition_cell { virtual void write(serializer & s) const { s.write_string(get_typed_expr_opcode()); } +#ifdef LEAN_JSON + virtual void write_json(abstract_ast_exporter &, json &) const override {} +#endif }; bool is_typed_expr(expr const & e) { return is_macro(e) && macro_def(e) == *g_typed_expr; } -expr mk_typed_expr(expr const & t, expr const & v) { +expr mk_typed_expr(expr const & t, expr const & v, tag g = nulltag) { expr args[2] = {t, v}; - return mk_macro(*g_typed_expr, 2, args); + return mk_macro(*g_typed_expr, 2, args, g); } expr get_typed_expr_type(expr const & e) { lean_assert(is_typed_expr(e)); return macro_arg(e, 0); } diff --git a/src/library/typed_expr.h b/src/library/typed_expr.h index 78ef36d6df..bfbab1b03e 100644 --- a/src/library/typed_expr.h +++ b/src/library/typed_expr.h @@ -9,7 +9,7 @@ Author: Leonardo de Moura namespace lean { class expr; /** \brief Create an expression definitionally equal to \c e, but it must have type \c t. */ -expr mk_typed_expr(expr const & t, expr const & e); +expr mk_typed_expr(expr const & t, expr const & e, tag g = nulltag); /** \brief Return true iff \c e was created using #mk_typed_expr */ bool is_typed_expr(expr const & e); /** \brief Return the type of a typed expression diff --git a/src/library/vm/vm_override.cpp b/src/library/vm/vm_override.cpp index 871ae0816c..04287b5f62 100644 --- a/src/library/vm/vm_override.cpp +++ b/src/library/vm/vm_override.cpp @@ -32,17 +32,22 @@ struct vm_override_attribute_data : public attr_data { void read(deserializer & d) { m_name = read_name(d); } - void parse(abstract_parser & p) override { + ast_id parse(abstract_parser & p) override { lean_assert(dynamic_cast(&p)); - auto & p2 = *static_cast(&p); + auto& p2 = *static_cast(&p); auto n = p2.check_constant_next("not a constant"); - m_name = n; + auto& data = p2.new_ast("vm_override", p2.pos()); + data.push(n.first); + m_name = n.second; if (p2.curr_is_identifier()) { m_ns = optional(p2.get_name_val()); + data.push(p2.new_ast("ident", p2.pos(), *m_ns).m_id); p2.next(); } else { m_ns = optional(); + data.push(0); } + return data.m_id; } }; bool operator==(vm_override_attribute_data const & d1, vm_override_attribute_data const & d2) { diff --git a/src/library/vm/vm_parser.cpp b/src/library/vm/vm_parser.cpp index b7f498c13b..76f424119f 100644 --- a/src/library/vm/vm_parser.cpp +++ b/src/library/vm/vm_parser.cpp @@ -45,13 +45,16 @@ typedef interaction_monad lean_parser; #define CATCH } catch (break_at_pos_exception const & ex) { throw; }\ catch (exception const & ex) { return lean_parser::mk_exception(ex, s); } -vm_obj run_parser(parser & p, expr const & spec, buffer const & args, bool allow_profiler) { +pair run_parser(parser & p, expr const & spec, buffer const & args, bool allow_profiler) { type_context_old ctx(p.env(), p.get_options()); + auto& data = p.new_ast("parse", p.pos()); + p.set_vm_parse_parent(data.m_id); auto r = lean_parser::evaluator(ctx, p.get_options(), allow_profiler)(spec, args, {&p}); - return lean_parser::get_success_value(r); + p.set_vm_parse_parent(0); + return {data.m_id, lean_parser::get_success_value(r)}; } -expr parse_interactive_param(parser & p, expr const & param_ty) { +expr parse_interactive_param(parser & p, ast_data & parent, expr const & param_ty) { lean_assert(is_app_of(param_ty, get_interactive_parse_name())); buffer param_args; get_app_args(param_ty, param_args); @@ -63,12 +66,16 @@ expr parse_interactive_param(parser & p, expr const & param_ty) { try { expr pr = mk_app({mk_constant(get_lean_parser_reflectable_expr_name()), param_args[0], param_args[1], param_args[2]}); - expr r = to_expr(run_parser(p, pr)); - if (is_app_of(r, get_expr_subst_name())) { - return r; // HACK - } else { - return mk_as_is(r); + ast_id id; vm_obj obj; + std::tie(id, obj) = run_parser(p, pr); + expr r = to_expr(obj); + if (!is_app_of(r, get_expr_subst_name())) { // HACK + r = mk_as_is(r); } + r.set_tag(nulltag); + p.set_ast_pexpr(id, r); + parent.push(id); + return r; } catch (exception & ex) { if (!p.has_error_recovery()) throw; p.mk_message(ERROR).set_exception(ex).report(); @@ -172,7 +179,9 @@ vm_obj vm_parser_ident(vm_obj const & o) { auto const & s = lean_parser::to_state(o); TRY; auto _ = s.m_p->no_error_recovery_scope(); - name ident = s.m_p->check_id_next("identifier expected"); + ast_id id; name ident; + std::tie(id, ident) = s.m_p->check_id_next("identifier expected"); + s.m_p->push_from_vm_parse(id); return lean_parser::mk_success(to_obj(ident), s); CATCH; } @@ -181,7 +190,9 @@ vm_obj vm_parser_small_nat(vm_obj const & o) { auto const & s = lean_parser::to_state(o); TRY; auto _ = s.m_p->no_error_recovery_scope(); - unsigned n = s.m_p->parse_small_nat(); + ast_id id; unsigned n; + std::tie(id, n) = s.m_p->parse_small_nat(); + s.m_p->push_from_vm_parse(id); return lean_parser::mk_success(mk_vm_nat(n), s); CATCH; } @@ -192,6 +203,7 @@ vm_obj vm_parser_tk(vm_obj const & vm_tk, vm_obj const & o) { name tk = to_string(vm_tk); if (!s.m_p->curr_is_token(tk)) throw parser_error(sstream() << "'" << tk << "' expected", s.m_p->pos()); + s.m_p->push_from_vm_parse(s.m_p->new_ast("token", s.m_p->pos(), tk).m_id); s.m_p->next(); return lean_parser::mk_success(s); CATCH; @@ -286,6 +298,7 @@ vm_obj vm_parser_pexpr(vm_obj const & vm_rbp, vm_obj const & vm_pat, vm_obj cons restore_decl_meta_scope scope; auto rbp = to_unsigned(vm_rbp); auto pat = to_bool(vm_pat); + auto pos = s.m_p->pos(); optional e; if (pat) { e = s.m_p->maybe_parse_pattern(rbp); @@ -293,6 +306,8 @@ vm_obj vm_parser_pexpr(vm_obj const & vm_rbp, vm_obj const & vm_pat, vm_obj cons e = s.m_p->maybe_parse_expr(rbp); } if (e) { + auto id = s.m_p->new_ast(pat ? "pat" : "expr", pos).push(s.m_p->get_id(*e)).m_id; + s.m_p->push_from_vm_parse(id); return lean_parser::mk_success(to_obj(*e), s); } else { throw parser_error(sstream() << "expression expected", s.m_p->pos()); @@ -305,8 +320,8 @@ vm_obj vm_parser_itactic_reflected(vm_obj const & o) { TRY; metavar_context mctx; auto _ = s.m_p->no_error_recovery_scope(); - expr e = parse_nested_interactive_tactic(*s.m_p, get_tactic_name(), true); + s.m_p->push_from_vm_parse(s.m_p->get_id(e)); vm_obj r = to_obj(s.m_p->elaborate({}, mctx, e, false).first); r = mk_vm_constructor(0, r, r); return lean_parser::mk_success(r, s); @@ -336,12 +351,17 @@ vm_obj vm_parser_with_input(vm_obj const &, vm_obj const & vm_p, vm_obj const & std::string input = to_string(vm_input); std::istringstream strm(input); vm_obj vm_state; pos_info pos; + auto old_id = s.m_p->get_vm_parse_parent(); + ast_id new_id = s.m_p->new_ast("with_input", s.m_p->pos()).m_id; + s.m_p->push_from_vm_parse(new_id); + s.m_p->set_vm_parse_parent(new_id); auto _ = s.m_p->no_error_recovery_scope(); TRY; std::tie(vm_state, pos) = s.m_p->with_input(strm, [&]() { return invoke(vm_p, o); }); CATCH; + s.m_p->set_vm_parse_parent(old_id); if (lean_parser::is_result_exception(vm_state)) { return vm_state; @@ -368,7 +388,9 @@ vm_obj vm_parser_with_input(vm_obj const &, vm_obj const & vm_p, vm_obj const & static vm_obj vm_parser_command_like(vm_obj const & vm_s) { auto s = lean_parser::to_state(vm_s); TRY; - s.m_p->parse_command_like(); + auto& data = s.m_p->new_ast("command", s.m_p->pos()); + s.m_p->push_from_vm_parse(data.m_id); + s.m_p->parse_command_like(&data); return lean_parser::mk_success(s); CATCH; } @@ -395,7 +417,9 @@ static vm_obj to_obj(inductive_decl const & d) { static vm_obj interactive_inductive_decl_parse(vm_obj const & vm_meta, vm_obj const & vm_s) { auto s = lean_parser::to_state(vm_s); TRY; - auto decl = parse_inductive_decl(*s.m_p, to_cmd_meta(vm_meta)); + auto& data = s.m_p->new_ast("inductive", s.m_p->pos()); + s.m_p->push_from_vm_parse(data.m_id); + auto decl = parse_inductive_decl(*s.m_p, data.m_id, to_cmd_meta(vm_meta)); return lean_parser::mk_success(to_obj(decl), s); CATCH; } @@ -404,7 +428,9 @@ static vm_obj interactive_parse_binders(vm_obj const & vm_rbp, vm_obj const & vm auto s = lean_parser::to_state(vm_s); TRY; buffer binders; - s.m_p->parse_binders(binders, to_unsigned(vm_rbp)); + auto& data = s.m_p->new_ast("binders", s.m_p->pos()); + s.m_p->push_from_vm_parse(data.m_id); + s.m_p->parse_binders(&data, binders, to_unsigned(vm_rbp)); return lean_parser::mk_success(to_obj(binders), s); CATCH; } diff --git a/src/library/vm/vm_parser.h b/src/library/vm/vm_parser.h index ea51a78c66..43b602b48b 100644 --- a/src/library/vm/vm_parser.h +++ b/src/library/vm/vm_parser.h @@ -9,8 +9,8 @@ Author: Sebastian Ullrich #include "frontends/lean/parser.h" namespace lean { -vm_obj run_parser(parser & p, expr const & spec, buffer const & args = {}, bool allow_profiler = false); -expr parse_interactive_param(parser & p, expr const & param_ty); +pair run_parser(parser & p, expr const & spec, buffer const & args = {}, bool allow_profiler = false); +expr parse_interactive_param(parser & p, ast_data & parent, expr const & param_ty); vm_obj to_obj(cmd_meta const & meta); diff --git a/src/shell/lean.cpp b/src/shell/lean.cpp index 3a918e6432..b80acfc0e9 100644 --- a/src/shell/lean.cpp +++ b/src/shell/lean.cpp @@ -210,6 +210,7 @@ static void display_help(std::ostream & out) { std::cout << " --server start lean in server mode\n"; std::cout << " --server=file start lean in server mode, redirecting standard input from the specified file (for debugging)\n"; // std::cout << " --no-widgets turn off reporting on widgets\n"; + std::cout << " --ast export a .ast.json file for every .lean file\n"; #endif std::cout << " --profile display elaboration/type checking time for each definition/theorem\n"; DEBUG_CODE( @@ -247,6 +248,7 @@ static struct option g_long_options[] = { {"path", no_argument, 0, 'p'}, {"server", optional_argument, 0, 'S'}, {"no-widgets", no_argument, 0, 'W'}, + {"ast", no_argument, 0, 'X'}, #endif #if defined(LEAN_MULTI_THREAD) {"tstack", required_argument, 0, 's'}, @@ -431,6 +433,7 @@ int main(int argc, char ** argv) { ::initializer init; bool make_mode = false; bool export_tlean = false; + bool export_ast = false; bool use_old_oleans = false; bool report_widgets = true; bool recursive = false; @@ -541,6 +544,9 @@ int main(int argc, char ** argv) { std::cout << std::setw(2) << out << std::endl; return 0; } + case 'X': + export_ast = true; + break; #endif case 'P': opts = opts.update("profiler", true); @@ -634,6 +640,7 @@ int main(int argc, char ** argv) { fs_module_vfs vfs; module_mgr mod_mgr(&vfs, lt.get_root(), path.get_path(), env, ios); mod_mgr.set_use_old_oleans(use_old_oleans); + mod_mgr.set_export_ast(export_ast); set_global_module_mgr(mod_mgr); if (run_arg) { diff --git a/src/util/lean_json.h b/src/util/lean_json.h new file mode 100644 index 0000000000..fee0233427 --- /dev/null +++ b/src/util/lean_json.h @@ -0,0 +1,20 @@ +/* +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Mario Carneiro +*/ +#pragma once +#ifdef LEAN_JSON +#include "util/json.hpp" + +namespace lean { +using json = nlohmann::json; +} + +#else + +namespace lean { +struct json; +} +#endif From 5df7df88b9afa273578a72881dc81a6d440cbcf3 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Wed, 4 Aug 2021 22:05:50 -0400 Subject: [PATCH 2/6] fix(vm_parser.cpp): restore nested context --- src/frontends/lean/parser.h | 6 +++++- src/library/vm/vm_parser.cpp | 4 ++-- 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/src/frontends/lean/parser.h b/src/frontends/lean/parser.h index 3d1935507a..baf2a4a676 100644 --- a/src/frontends/lean/parser.h +++ b/src/frontends/lean/parser.h @@ -319,7 +319,11 @@ class parser : public abstract_parser, public parser_info { pos_info ast_pos(ast_id id) const { return m_ast[id]->m_start; } void push_from_vm_parse(ast_id id) { if (m_vm_parser_ast_id) get_ast(m_vm_parser_ast_id).push(id); } ast_id get_vm_parse_parent() { return m_vm_parser_ast_id; } - void set_vm_parse_parent(ast_id id) { m_vm_parser_ast_id = id; } + ast_id set_vm_parse_parent(ast_id id) { + auto old = m_vm_parser_ast_id; + m_vm_parser_ast_id = id; + return old; + } ast_data & new_modifiers(cmd_meta & meta); friend void export_ast(parser & p); diff --git a/src/library/vm/vm_parser.cpp b/src/library/vm/vm_parser.cpp index 76f424119f..0d0339b47d 100644 --- a/src/library/vm/vm_parser.cpp +++ b/src/library/vm/vm_parser.cpp @@ -48,9 +48,9 @@ typedef interaction_monad lean_parser; pair run_parser(parser & p, expr const & spec, buffer const & args, bool allow_profiler) { type_context_old ctx(p.env(), p.get_options()); auto& data = p.new_ast("parse", p.pos()); - p.set_vm_parse_parent(data.m_id); + auto old = p.set_vm_parse_parent(data.m_id); auto r = lean_parser::evaluator(ctx, p.get_options(), allow_profiler)(spec, args, {&p}); - p.set_vm_parse_parent(0); + p.set_vm_parse_parent(old); return {data.m_id, lean_parser::get_success_value(r)}; } From 0f8f772d4587d5d6cf796de4a7c4cf4d912b1526 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 5 Aug 2021 22:25:18 -0400 Subject: [PATCH 3/6] feat(vm_parser): with_input size AST --- src/library/vm/vm_parser.cpp | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/library/vm/vm_parser.cpp b/src/library/vm/vm_parser.cpp index 0d0339b47d..58a1c012ab 100644 --- a/src/library/vm/vm_parser.cpp +++ b/src/library/vm/vm_parser.cpp @@ -352,9 +352,9 @@ vm_obj vm_parser_with_input(vm_obj const &, vm_obj const & vm_p, vm_obj const & std::istringstream strm(input); vm_obj vm_state; pos_info pos; auto old_id = s.m_p->get_vm_parse_parent(); - ast_id new_id = s.m_p->new_ast("with_input", s.m_p->pos()).m_id; - s.m_p->push_from_vm_parse(new_id); - s.m_p->set_vm_parse_parent(new_id); + auto& new_data = s.m_p->new_ast("with_input", s.m_p->pos()); + s.m_p->push_from_vm_parse(new_data.m_id); + s.m_p->set_vm_parse_parent(new_data.m_id); auto _ = s.m_p->no_error_recovery_scope(); TRY; std::tie(vm_state, pos) = s.m_p->with_input(strm, [&]() { @@ -381,7 +381,7 @@ vm_obj vm_parser_with_input(vm_obj const &, vm_obj const & vm_p, vm_obj const & } spos += get_utf8_size(input[spos]); } - + new_data.m_value = std::to_string(spos); vm_res = mk_vm_pair(vm_res, to_obj(input.substr(spos))); return lean_parser::mk_success(vm_res, lean_parser::get_success_state(vm_state)); } From fc5f1191859fe1074c5fb00aabace91b16fa2cf2 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 16 Aug 2021 15:10:19 -0400 Subject: [PATCH 4/6] fix(user_command): get AST for command parser --- src/frontends/lean/user_command.cpp | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/frontends/lean/user_command.cpp b/src/frontends/lean/user_command.cpp index 177bef153a..2a66a442a3 100644 --- a/src/frontends/lean/user_command.cpp +++ b/src/frontends/lean/user_command.cpp @@ -86,7 +86,7 @@ static environment add_user_command(environment const & env, name const & d) { args.push_back(to_obj(meta)); } parser = p.elaborate("_user_command", {}, parser).first; - run_parser(p, parser, args, true); + data.push(run_parser(p, parser, args, true).first); return p.env(); }; From 43f569a4ee91b2399efd13cad2d56a9fc5e51371 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 23 Aug 2021 15:32:06 -0400 Subject: [PATCH 5/6] fix(json.cpp): un-reverse json_of_name --- src/frontends/lean/json.cpp | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/frontends/lean/json.cpp b/src/frontends/lean/json.cpp index 44a76e6505..74802cb569 100644 --- a/src/frontends/lean/json.cpp +++ b/src/frontends/lean/json.cpp @@ -107,20 +107,20 @@ json serialize_decl(name const & d, environment const & env, options const & o) json json_of_name(name const & n0) { if (n0.is_atomic() && n0.is_string()) return n0.get_string(); - json j = json::array(); - - name n = n0; - while (!n.is_anonymous()) { + buffer js; + for (name n = n0; !n.is_anonymous();) { if (n.is_numeral()) { - j.push_back(n.get_numeral()); + js.push_back(n.get_numeral()); } else if (n.is_string()) { - j.push_back(n.get_string()); + js.push_back(n.get_string()); } else { - j.push_back(json()); + js.push_back(json()); } n = n.get_prefix(); } + json j = json::array(); + for (auto i = js.size(); i != 0;) j.push_back(js[--i]); return j; } From 8716276a8b169b2f70971c4f6b4c78f0587bc146 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Mon, 23 Aug 2021 15:32:17 -0400 Subject: [PATCH 6/6] fix warning --- src/library/tlean_exporter.cpp | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/library/tlean_exporter.cpp b/src/library/tlean_exporter.cpp index 728e659f1c..fe9fbbfc72 100644 --- a/src/library/tlean_exporter.cpp +++ b/src/library/tlean_exporter.cpp @@ -161,7 +161,7 @@ unsigned tlean_exporter::export_expr_core(expr const & e) { case expr_kind::Macro: if (macro_def(e).can_textualize()) { buffer args; - for (auto i = 0; i < macro_num_args(e); ++i) { + for (unsigned i = 0; i < macro_num_args(e); ++i) { args.push_back(export_expr_core(macro_arg(e, i))); } i = static_cast(m_expr2idx.size());