feat(frontends/lean): add exact_apply
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
3650ffdd4f
commit
8d584e54da
8 changed files with 116 additions and 9 deletions
|
@ -650,6 +650,36 @@ public:
|
|||
return apply(s, r);
|
||||
}
|
||||
|
||||
static format pp_type_mismatch(formatter const & fmt, environment const & env, options const & opts,
|
||||
expr const & expected_type, expr const & given_type) {
|
||||
format r("type mismatch, expected type");
|
||||
r += pp_indent_expr(fmt, env, opts, expected_type);
|
||||
r += compose(line(), format("given type:"));
|
||||
r += pp_indent_expr(fmt, env, opts, given_type);
|
||||
return r;
|
||||
}
|
||||
|
||||
expr operator()(expr const & e, expr const & expected_type) {
|
||||
expr r = visit(e);
|
||||
expr r_type = infer_type(r);
|
||||
environment env = m_env;
|
||||
justification j = mk_justification(e, [=](formatter const & fmt, options const & opts, substitution const & subst) {
|
||||
return pp_type_mismatch(fmt, env, opts,
|
||||
subst.instantiate_metavars_wo_jst(expected_type),
|
||||
subst.instantiate_metavars_wo_jst(r_type));
|
||||
});
|
||||
if (!m_tc.is_def_eq(r_type, expected_type, j)) {
|
||||
throw_kernel_exception(env, e,
|
||||
[=](formatter const & fmt, options const & o) {
|
||||
return pp_type_mismatch(fmt, env, o, expected_type, r_type);
|
||||
});
|
||||
}
|
||||
auto p = solve().pull();
|
||||
lean_assert(p);
|
||||
substitution s = p->first;
|
||||
return apply(s, r);
|
||||
}
|
||||
|
||||
std::pair<expr, expr> operator()(expr const & t, expr const & v, name const & n) {
|
||||
expr r_t = visit(t);
|
||||
expr r_v = visit(v);
|
||||
|
@ -688,4 +718,9 @@ std::pair<expr, expr> elaborate(environment const & env, io_state const & ios, n
|
|||
hint_table const & htable, pos_info_provider * pp) {
|
||||
return elaborator(env, ios, name_generator(g_tmp_prefix), htable, substitution(), list<parameter>(), pp)(t, v, n);
|
||||
}
|
||||
|
||||
expr elaborate(environment const & env, io_state const & ios, expr const & e, expr const & expected_type, name_generator const & ngen,
|
||||
hint_table const & htable, list<parameter> const & ctx, pos_info_provider * pp) {
|
||||
return elaborator(env, ios, ngen, htable, substitution(), ctx, pp)(e, expected_type);
|
||||
}
|
||||
}
|
||||
|
|
|
@ -16,6 +16,8 @@ namespace lean {
|
|||
expr elaborate(environment const & env, io_state const & ios, expr const & e, name_generator const & ngen,
|
||||
hint_table const & htable = hint_table(), substitution const & s = substitution(),
|
||||
list<parameter> const & ctx = list<parameter>(), pos_info_provider * pp = nullptr);
|
||||
expr elaborate(environment const & env, io_state const & ios, expr const & e, expr const & expected_type, name_generator const & ngen,
|
||||
hint_table const & htable = hint_table(), list<parameter> const & ctx = list<parameter>(), pos_info_provider * pp = nullptr);
|
||||
expr elaborate(environment const & env, io_state const & ios, expr const & e, hint_table const & htable = hint_table(),
|
||||
pos_info_provider * pp = nullptr);
|
||||
std::pair<expr, expr> elaborate(environment const & env, io_state const & ios, name const & n, expr const & t, expr const & v,
|
||||
|
|
|
@ -949,25 +949,76 @@ expr parser::abstract(unsigned num_params, parameter const * ps, expr const & e,
|
|||
return r;
|
||||
}
|
||||
|
||||
tactic parser::parse_apply() {
|
||||
auto p = pos();
|
||||
expr e = parse_expr();
|
||||
/** \brief Throw an exception if \c e contains a local constant that is not marked as contextual in \c local_decls */
|
||||
static void check_only_contextual_locals(expr const & e, local_expr_decls const & local_decls, pos_info const & pos) {
|
||||
for_each(e, [&](expr const & e, unsigned) {
|
||||
if (is_local(e)) {
|
||||
auto it = get_local(mlocal_name(e));
|
||||
auto it = local_decls.find(mlocal_name(e));
|
||||
lean_assert(it);
|
||||
if (!it->m_bi.is_contextual())
|
||||
throw parser_error(sstream() << "invalid 'apply' tactic, it references non contextual local '"
|
||||
<< local_pp_name(e) << "'", p);
|
||||
throw parser_error(sstream() << "invalid 'apply' tactic, it references non-contextual local '"
|
||||
<< local_pp_name(e) << "'", pos);
|
||||
}
|
||||
return has_local(e);
|
||||
});
|
||||
}
|
||||
|
||||
|
||||
/** \brief Return a list of all contextual parameters in local_decls */
|
||||
static list<parameter> collect_contextual_parameters(local_expr_decls const & local_decls) {
|
||||
buffer<parameter> tmp;
|
||||
for (parameter const & p : m_local_decls.get_values()) {
|
||||
for (parameter const & p : local_decls.get_values()) {
|
||||
if (p.m_bi.is_contextual() && is_local(p.m_local))
|
||||
tmp.push_back(p);
|
||||
}
|
||||
list<parameter> ctx = to_list(tmp.begin(), tmp.end());
|
||||
return to_list(tmp.begin(), tmp.end());
|
||||
}
|
||||
|
||||
tactic parser::parse_exact_apply() {
|
||||
auto p = pos();
|
||||
expr e = parse_expr();
|
||||
check_only_contextual_locals(e, m_local_decls, p);
|
||||
list<parameter> ctx = collect_contextual_parameters(m_local_decls);
|
||||
return tactic01([=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||
if (empty(s.get_goals()))
|
||||
return none_proof_state();
|
||||
name gname = head(s.get_goals()).first;
|
||||
goal g = head(s.get_goals()).second;
|
||||
goals new_gs = tail(s.get_goals());
|
||||
auto const & s_locals = s.get_init_locals();
|
||||
// Remark: the proof state initial hypotheses (s_locals) are essentially ctx "after elaboration".
|
||||
// So, to process \c e, we must first replace its local constants with (s_locals).
|
||||
// We should also create a new_ctx based on (s_locals), but using the binder info from ctx.
|
||||
name_generator ngen = s.ngen();
|
||||
lean_assert(length(s_locals) == length(ctx));
|
||||
buffer<parameter> new_ctx;
|
||||
buffer<expr> locals;
|
||||
auto it1 = ctx;
|
||||
auto it2 = s_locals;
|
||||
while (!is_nil(it1) && !is_nil(it2)) {
|
||||
auto const & p = head(it1);
|
||||
locals.push_back(p.m_local);
|
||||
new_ctx.push_back(parameter(head(it2), p.m_bi));
|
||||
it1 = tail(it1);
|
||||
it2 = tail(it2);
|
||||
}
|
||||
std::reverse(locals.begin(), locals.end());
|
||||
expr new_e = abstract_locals(e, locals.size(), locals.data());
|
||||
for (auto const & l : s_locals)
|
||||
new_e = instantiate(new_e, l);
|
||||
parser_pos_provider pp(m_pos_table, get_stream_name(), m_last_cmd_pos);
|
||||
new_e = ::lean::elaborate(env, ios, new_e, g.get_conclusion(), ngen.mk_child(),
|
||||
m_hints, to_list(new_ctx.begin(), new_ctx.end()), &pp);
|
||||
proof_builder new_pb = add_proof(s.get_pb(), gname, new_e);
|
||||
return some_proof_state(proof_state(s, new_gs, new_pb, ngen));
|
||||
});
|
||||
}
|
||||
|
||||
tactic parser::parse_apply() {
|
||||
auto p = pos();
|
||||
expr e = parse_expr();
|
||||
check_only_contextual_locals(e, m_local_decls, p);
|
||||
list<parameter> ctx = collect_contextual_parameters(m_local_decls);
|
||||
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||
name_generator ngen = s.ngen();
|
||||
auto const & s_locals = s.get_init_locals();
|
||||
|
@ -1018,7 +1069,7 @@ tactic parser::parse_tactic(unsigned /* rbp */) {
|
|||
}
|
||||
} else if (curr_is_token(g_have) || curr_is_token(g_show) || curr_is_token(g_assume) ||
|
||||
curr_is_token(g_take) || curr_is_token(g_fun)) {
|
||||
return parse_apply();
|
||||
return parse_exact_apply();
|
||||
} else {
|
||||
name id;
|
||||
auto p = pos();
|
||||
|
|
|
@ -121,6 +121,7 @@ class parser {
|
|||
parameter parse_binder_core(binder_info const & bi);
|
||||
void parse_binder_block(buffer<parameter> & r, binder_info const & bi);
|
||||
void parse_binders_core(buffer<parameter> & r);
|
||||
tactic parse_exact_apply();
|
||||
|
||||
friend environment section_cmd(parser & p);
|
||||
friend environment end_scoped_cmd(parser & p);
|
||||
|
|
|
@ -29,6 +29,10 @@ proof_builder add_proofs(proof_builder const & pb, list<std::pair<name, expr>> c
|
|||
});
|
||||
}
|
||||
|
||||
proof_builder add_proof(proof_builder const & pb, name const & goal_name, expr const & pr) {
|
||||
return add_proofs(pb, list<std::pair<name, expr>>(mk_pair(goal_name, pr)));
|
||||
}
|
||||
|
||||
DECL_UDATA(proof_map)
|
||||
|
||||
static int mk_proof_map(lua_State * L) {
|
||||
|
|
|
@ -25,6 +25,7 @@ expr find(proof_map const & m, name const & n);
|
|||
*/
|
||||
typedef std::function<expr(proof_map const &, substitution const &)> proof_builder;
|
||||
proof_builder add_proofs(proof_builder const & pb, list<std::pair<name, expr>> const & prs);
|
||||
proof_builder add_proof(proof_builder const & pb, name const & goal_name, expr const & pr);
|
||||
|
||||
/** \brief Convert a Lua function on position \c idx (on the Lua stack) into a proof_builder_fn */
|
||||
proof_builder to_proof_builder(lua_State * L, int idx);
|
||||
|
|
6
tests/lean/run/tactic10.lean
Normal file
6
tests/lean/run/tactic10.lean
Normal file
|
@ -0,0 +1,6 @@
|
|||
import logic
|
||||
|
||||
theorem tst (a b : Bool) (H : a ↔ b) : b ↔ a
|
||||
:= by (apply iff_intro,
|
||||
assume Hb, iff_mp_right H Hb,
|
||||
assume Ha, iff_mp_left H Ha)
|
7
tests/lean/run/tactic11.lean
Normal file
7
tests/lean/run/tactic11.lean
Normal file
|
@ -0,0 +1,7 @@
|
|||
import logic
|
||||
|
||||
theorem tst (a b : Bool) (H : a ↔ b) : b ↔ a
|
||||
:= have H1 [fact] : a → b, from iff_elim_left H,
|
||||
by (apply iff_intro,
|
||||
assume Hb, iff_mp_right H Hb,
|
||||
assume Ha, H1 Ha)
|
Loading…
Add table
Reference in a new issue