feat(library/tactic/induction_tactic): add support for user-defined recursors that contain parameters that should be synthesized by type class resolution
This commit is contained in:
parent
78ee055de8
commit
3e87f09d78
6 changed files with 106 additions and 59 deletions
|
@ -194,7 +194,7 @@ theorem card_insert_of_not_mem {a : A} {s : finset A} : a ∉ s → card (insert
|
|||
quot.induction_on s
|
||||
(λ (l : nodup_list A) (nainl : a ∉ ⟦l⟧), list.length_insert_of_not_mem nainl)
|
||||
|
||||
protected theorem induction {P : finset A → Prop}
|
||||
protected theorem induction [recursor 6] {P : finset A → Prop}
|
||||
(H1 : P empty)
|
||||
(H2 : ∀⦃s : finset A⦄, ∀{a : A}, a ∉ s → P s → P (insert a s)) :
|
||||
∀s, P s :=
|
||||
|
|
|
@ -75,18 +75,19 @@ section deceqA
|
|||
-- this will eventually be an instance of something more general
|
||||
theorem inter_Union (s : finset B) (t : finset A) (f : A → finset B) :
|
||||
s ∩ (⋃ x ∈ t, f x) = (⋃ x ∈ t, s ∩ f x) :=
|
||||
finset.induction_on t
|
||||
(by rewrite [*Union_empty, inter_empty])
|
||||
(take s' x, assume H : x ∉ s',
|
||||
assume IH,
|
||||
by rewrite [*Union_insert_of_not_mem _ H, inter.distrib_left, IH])
|
||||
begin
|
||||
induction t with s' x H IH,
|
||||
rewrite [*Union_empty, inter_empty],
|
||||
rewrite [*Union_insert_of_not_mem _ H, inter.distrib_left, IH],
|
||||
end
|
||||
|
||||
theorem mem_Union_iff (s : finset A) (f : A → finset B) (b : B) :
|
||||
b ∈ (⋃ x ∈ s, f x) ↔ (∃ x, x ∈ s ∧ b ∈ f x ) :=
|
||||
finset.induction_on s
|
||||
(by rewrite [exists_mem_empty_eq])
|
||||
(take s' a, assume H : a ∉ s', assume IH,
|
||||
by rewrite [Union_insert_of_not_mem _ H, mem_union_eq, IH, exists_mem_insert_eq])
|
||||
begin
|
||||
induction s with s' a H IH,
|
||||
rewrite [exists_mem_empty_eq],
|
||||
rewrite [Union_insert_of_not_mem _ H, mem_union_eq, IH, exists_mem_insert_eq]
|
||||
end
|
||||
|
||||
theorem mem_Union_eq (s : finset A) (f : A → finset B) (b : B) :
|
||||
b ∈ (⋃ x ∈ s, f x) = (∃ x, x ∈ s ∧ b ∈ f x ) :=
|
||||
|
|
|
@ -245,8 +245,12 @@ static void print_recursor_info(parser & p) {
|
|||
<< " dep. elimination: " << info.has_dep_elim() << "\n";
|
||||
if (info.get_num_params() > 0) {
|
||||
out << " parameters pos. at major:";
|
||||
for (unsigned p : info.get_params_pos())
|
||||
out << " " << p+1;
|
||||
for (optional<unsigned> const & p : info.get_params_pos()) {
|
||||
if (p)
|
||||
out << " " << *p+1;
|
||||
else
|
||||
out << " [instance]";
|
||||
}
|
||||
out << "\n";
|
||||
}
|
||||
if (info.get_num_indices() > 0) {
|
||||
|
|
|
@ -14,6 +14,7 @@ Author: Leonardo de Moura
|
|||
#include "library/locals.h"
|
||||
#include "library/tactic/tactic.h"
|
||||
#include "library/tactic/expr_to_tactic.h"
|
||||
#include "library/tactic/class_instance_synth.h"
|
||||
|
||||
namespace lean {
|
||||
class has_rec_opaque_converter : public default_converter {
|
||||
|
@ -38,6 +39,8 @@ class induction_tac {
|
|||
substitution m_subst;
|
||||
bool m_throw_ex;
|
||||
bool m_standard;
|
||||
constraint_seq m_cs;
|
||||
expr m_ref; // reference expression for location information
|
||||
|
||||
void assign(goal const & g, expr const & val) {
|
||||
::lean::assign(m_subst, g, val);
|
||||
|
@ -68,11 +71,28 @@ class induction_tac {
|
|||
<< "' is ill-formed");
|
||||
}
|
||||
|
||||
expr mk_type_class_param(goal const & g, expr const & rec, recursor_info const & rec_info) {
|
||||
expr rec_type = m_tc.whnf(m_tc.infer(rec, m_cs), m_cs);
|
||||
if (!is_pi(rec_type)) {
|
||||
throw_ill_formed_recursor(rec_info);
|
||||
}
|
||||
bool use_local_insts = true;
|
||||
bool is_strict = false;
|
||||
local_context ctx = g.to_local_context();
|
||||
unifier_config cfg(m_ios.get_options());
|
||||
auto mc = mk_class_instance_elaborator(
|
||||
m_env, m_ios, ctx, m_ngen.next(), optional<name>(),
|
||||
use_local_insts, is_strict,
|
||||
some_expr(binding_domain(rec_type)), m_ref.get_tag(), cfg, nullptr);
|
||||
m_cs += mc.second;
|
||||
return mc.first;
|
||||
}
|
||||
|
||||
goals apply_recursor(goal const & g, unsigned num_deps, expr const & h, expr const & h_type,
|
||||
buffer<expr> const & params, buffer<expr> const & indices,
|
||||
buffer<optional<expr>> const & params, buffer<expr> const & indices,
|
||||
recursor_info const & rec_info) {
|
||||
expr const & g_type = g.get_type();
|
||||
level g_lvl = sort_level(m_tc.ensure_type(g_type).first);
|
||||
level g_lvl = sort_level(m_tc.ensure_type(g_type, m_cs));
|
||||
levels rec_levels;
|
||||
expr const & I = get_app_fn(h_type);
|
||||
if (auto lpos = rec_info.get_motive_univ_pos()) {
|
||||
|
@ -95,7 +115,13 @@ class induction_tac {
|
|||
rec_levels = const_levels(I);
|
||||
}
|
||||
expr rec = mk_constant(rec_info.get_name(), rec_levels);
|
||||
rec = mk_app(rec, params);
|
||||
for (optional<expr> const & p : params) {
|
||||
if (p) {
|
||||
rec = mk_app(rec, *p);
|
||||
} else {
|
||||
rec = mk_app(rec, mk_type_class_param(g, rec, rec_info));
|
||||
}
|
||||
}
|
||||
expr motive = g_type;
|
||||
if (rec_info.has_dep_elim())
|
||||
motive = Fun(h, motive);
|
||||
|
@ -111,7 +137,7 @@ class induction_tac {
|
|||
new_hyps.push_back(curr_h);
|
||||
}
|
||||
}
|
||||
expr rec_type = m_tc.whnf(m_tc.infer(rec).first).first;
|
||||
expr rec_type = m_tc.whnf(m_tc.infer(rec, m_cs), m_cs);
|
||||
unsigned curr_pos = params.size() + 1 /* motive */;
|
||||
unsigned first_idx_pos = rec_info.get_first_index_pos();
|
||||
bool consumed_major = false;
|
||||
|
@ -120,13 +146,14 @@ class induction_tac {
|
|||
if (first_idx_pos == curr_pos) {
|
||||
for (expr const & idx : indices) {
|
||||
rec = mk_app(rec, idx);
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), idx)).first;
|
||||
if (!is_pi(rec_type))
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), idx), m_cs);
|
||||
if (!is_pi(rec_type)) {
|
||||
throw_ill_formed_recursor(rec_info);
|
||||
}
|
||||
curr_pos++;
|
||||
}
|
||||
rec = mk_app(rec, h);
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), h)).first;
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), h), m_cs);
|
||||
consumed_major = true;
|
||||
curr_pos++;
|
||||
} else {
|
||||
|
@ -137,7 +164,7 @@ class induction_tac {
|
|||
// introduce minor arguments
|
||||
buffer<expr> minor_args;
|
||||
for (unsigned i = 0; i < arity; i++) {
|
||||
expr arg_type = binding_domain(new_type);
|
||||
expr arg_type = head_beta_reduce(binding_domain(new_type));
|
||||
name arg_name;
|
||||
if (m_ids) {
|
||||
arg_name = head(m_ids);
|
||||
|
@ -155,8 +182,9 @@ class induction_tac {
|
|||
buffer<expr> new_deps;
|
||||
// introduce deps
|
||||
for (unsigned i = 0; i < num_deps; i++) {
|
||||
if (!is_pi(new_type))
|
||||
if (!is_pi(new_type)) {
|
||||
throw_ill_formed_recursor(rec_info);
|
||||
}
|
||||
expr dep_type = binding_domain(new_type);
|
||||
expr new_dep = mk_local(m_ngen.next(), get_unused_name(binding_name(new_type), new_goal_hyps),
|
||||
dep_type, binder_info());
|
||||
|
@ -168,12 +196,13 @@ class induction_tac {
|
|||
goal new_g(new_meta, new_type);
|
||||
new_goals.push_back(new_g);
|
||||
rec = mk_app(rec, Fun(minor_args, Fun(new_deps, new_meta)));
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), new_meta)).first;
|
||||
rec_type = m_tc.whnf(instantiate(binding_body(rec_type), new_meta), m_cs);
|
||||
curr_pos++;
|
||||
}
|
||||
}
|
||||
if (!consumed_major)
|
||||
if (!consumed_major) {
|
||||
throw_ill_formed_recursor(rec_info);
|
||||
}
|
||||
assign(g, rec);
|
||||
return to_list(new_goals);
|
||||
}
|
||||
|
@ -183,11 +212,15 @@ class induction_tac {
|
|||
recursor_info rec_info = get_recursor_info(m_env, rec);
|
||||
buffer<expr> h_type_args;
|
||||
get_app_args(h_type, h_type_args);
|
||||
buffer<expr> params;
|
||||
for (unsigned pos : rec_info.get_params_pos()) {
|
||||
if (pos >= h_type_args.size())
|
||||
buffer<optional<expr>> params;
|
||||
for (optional<unsigned> const & pos : rec_info.get_params_pos()) {
|
||||
if (!pos) {
|
||||
params.push_back(none_expr());
|
||||
} else if (*pos >= h_type_args.size()) {
|
||||
throw tactic_exception("invalid 'induction' tactic, major premise type is ill-formed");
|
||||
params.push_back(h_type_args[pos]);
|
||||
} else {
|
||||
params.push_back(some_expr(h_type_args[*pos]));
|
||||
}
|
||||
}
|
||||
buffer<expr> indices;
|
||||
for (unsigned pos : rec_info.get_indices_pos()) {
|
||||
|
@ -232,9 +265,9 @@ public:
|
|||
induction_tac(environment const & env, io_state const & ios, name_generator const & ngen,
|
||||
type_checker & tc, substitution const & subst, name const & h_name,
|
||||
optional<name> const & rec_name, list<name> const & ids,
|
||||
bool throw_ex):
|
||||
bool throw_ex, expr const & ref):
|
||||
m_env(env), m_ios(ios), m_tc(tc), m_h_name(h_name), m_rec_name(rec_name), m_ids(ids),
|
||||
m_ngen(ngen), m_subst(subst), m_throw_ex(throw_ex) {
|
||||
m_ngen(ngen), m_subst(subst), m_throw_ex(throw_ex), m_ref(ref) {
|
||||
m_standard = is_standard(env);
|
||||
}
|
||||
|
||||
|
@ -242,6 +275,8 @@ public:
|
|||
|
||||
substitution const & get_subst() const { return m_subst; }
|
||||
|
||||
constraint_seq get_new_constraints() const { return m_cs; }
|
||||
|
||||
expr normalize_H_type(expr const & H) {
|
||||
lean_assert(is_local(H));
|
||||
type_checker aux_tc(m_env, m_ngen.mk_child(), std::unique_ptr<converter>(new has_rec_opaque_converter(m_env)));
|
||||
|
@ -275,7 +310,7 @@ public:
|
|||
}
|
||||
};
|
||||
|
||||
tactic induction_tactic(name const & H, optional<name> rec, list<name> const & ids) {
|
||||
tactic induction_tactic(name const & H, optional<name> rec, list<name> const & ids, expr const & ref) {
|
||||
name rec1 = "unknown";
|
||||
if (rec) rec1 = *rec;
|
||||
auto fn = [=](environment const & env, io_state const & ios, proof_state const & ps) -> optional<proof_state> {
|
||||
|
@ -288,10 +323,13 @@ tactic induction_tactic(name const & H, optional<name> rec, list<name> const & i
|
|||
goals tail_gs = tail(gs);
|
||||
name_generator ngen = ps.get_ngen();
|
||||
std::unique_ptr<type_checker> tc = mk_type_checker(env, ngen.mk_child());
|
||||
induction_tac tac(env, ios, ngen, *tc, ps.get_subst(), H, rec, ids, ps.report_failure());
|
||||
induction_tac tac(env, ios, ngen, *tc, ps.get_subst(), H, rec, ids, ps.report_failure(), ref);
|
||||
if (auto res = tac.execute(g)) {
|
||||
proof_state new_s(ps, append(*res, tail_gs), tac.get_subst(), tac.get_ngen());
|
||||
return some_proof_state(new_s);
|
||||
if (solve_constraints(env, ios, new_s, tac.get_new_constraints()))
|
||||
return some_proof_state(new_s);
|
||||
else
|
||||
return none_proof_state();
|
||||
} else {
|
||||
return none_proof_state();
|
||||
}
|
||||
|
@ -316,9 +354,9 @@ void initialize_induction_tactic() {
|
|||
}
|
||||
name const & cname = const_name(rec);
|
||||
if (cname == get_tactic_none_expr_name()) {
|
||||
return induction_tactic(H, optional<name>(), to_list(ids.begin(), ids.end()));
|
||||
return induction_tactic(H, optional<name>(), to_list(ids.begin(), ids.end()), e);
|
||||
} else {
|
||||
return induction_tactic(H, optional<name>(cname), to_list(ids.begin(), ids.end()));
|
||||
return induction_tactic(H, optional<name>(cname), to_list(ids.begin(), ids.end()), e);
|
||||
}
|
||||
});
|
||||
}
|
||||
|
|
|
@ -23,14 +23,14 @@ bool recursor_info::is_minor(unsigned pos) const {
|
|||
}
|
||||
|
||||
recursor_info::recursor_info(name const & r, name const & I, optional<unsigned> const & motive_univ_pos,
|
||||
bool dep_elim, unsigned num_args, unsigned major_pos, list<unsigned> const & params_pos,
|
||||
list<unsigned> const & indices_pos):
|
||||
bool dep_elim, unsigned num_args, unsigned major_pos,
|
||||
list<optional<unsigned>> const & params_pos, list<unsigned> const & indices_pos):
|
||||
m_recursor(r), m_type_name(I), m_motive_univ_pos(motive_univ_pos), m_dep_elim(dep_elim),
|
||||
m_num_args(num_args), m_major_pos(major_pos), m_params_pos(params_pos), m_indices_pos(indices_pos) {}
|
||||
recursor_info::recursor_info() {}
|
||||
|
||||
void recursor_info::write(serializer & s) const {
|
||||
s << m_recursor << m_type_name << m_motive_univ_pos << m_dep_elim << m_major_pos;
|
||||
s << m_recursor << m_type_name << m_motive_univ_pos << m_dep_elim << m_num_args << m_major_pos;
|
||||
write_list(s, m_params_pos);
|
||||
write_list(s, m_indices_pos);
|
||||
}
|
||||
|
@ -38,8 +38,8 @@ void recursor_info::write(serializer & s) const {
|
|||
recursor_info recursor_info::read(deserializer & d) {
|
||||
recursor_info info;
|
||||
d >> info.m_recursor >> info.m_type_name >> info.m_motive_univ_pos >> info.m_dep_elim
|
||||
>> info.m_major_pos;
|
||||
info.m_params_pos = read_list<unsigned>(d);
|
||||
>> info.m_num_args >> info.m_major_pos;
|
||||
info.m_params_pos = read_list<optional<unsigned>>(d);
|
||||
info.m_indices_pos = read_list<unsigned>(d);
|
||||
return info;
|
||||
}
|
||||
|
@ -64,7 +64,8 @@ recursor_info mk_recursor_info(environment const & env, name const & r, optional
|
|||
unsigned num_params = *inductive::get_num_params(env, *I);
|
||||
unsigned num_minors = *inductive::get_num_minor_premises(env, *I);
|
||||
unsigned num_args = num_params + 1 /* motive */ + num_minors + num_indices + 1 /* major */;
|
||||
list<unsigned> params_pos = mk_list_range(0, num_params);
|
||||
list<optional<unsigned>> params_pos = map2<optional<unsigned>>(mk_list_range(0, num_params),
|
||||
[](unsigned i) { return optional<unsigned>(i); });
|
||||
list<unsigned> indices_pos = mk_list_range(num_params, num_params + num_indices);
|
||||
return recursor_info(r, *I, motive_univ_pos, inductive::has_dep_elim(env, *I),
|
||||
num_args, major_pos, params_pos, indices_pos);
|
||||
|
@ -133,7 +134,7 @@ recursor_info mk_recursor_info(environment const & env, name const & r, optional
|
|||
}
|
||||
|
||||
// Store position of the recursor parameters in the major premise.
|
||||
buffer<unsigned> params_pos;
|
||||
buffer<optional<unsigned>> params_pos;
|
||||
for (unsigned i = 0; i < num_params; i++) {
|
||||
expr const & A = tele[i];
|
||||
unsigned j = 0;
|
||||
|
@ -142,10 +143,15 @@ recursor_info mk_recursor_info(environment const & env, name const & r, optional
|
|||
break;
|
||||
}
|
||||
if (j == I_args.size()) {
|
||||
throw exception(sstream() << "invalid user defined recursor, type of the major premise '" << major
|
||||
<< "' does not contain the recursor parameter '" << A << "'");
|
||||
if (local_info(tele[i]).is_inst_implicit()) {
|
||||
params_pos.push_back(optional<unsigned>());
|
||||
} else {
|
||||
throw exception(sstream() << "invalid user defined recursor, type of the major premise '" << major
|
||||
<< "' does not contain the recursor parameter '" << A << "'");
|
||||
}
|
||||
} else {
|
||||
params_pos.push_back(optional<unsigned>(j));
|
||||
}
|
||||
params_pos.push_back(j);
|
||||
}
|
||||
|
||||
// Store position of the recursor indices in the major premise
|
||||
|
|
|
@ -9,22 +9,20 @@ Author: Leonardo de Moura
|
|||
namespace lean {
|
||||
/** \brief Information for a user defined recursor */
|
||||
class recursor_info {
|
||||
name m_recursor;
|
||||
name m_type_name;
|
||||
optional<unsigned> m_motive_univ_pos; // if none, then recursor can only eliminate to Prop
|
||||
bool m_dep_elim;
|
||||
unsigned m_num_args; // total number of arguments
|
||||
unsigned m_major_pos;
|
||||
list<unsigned> m_params_pos; // position of the recursor parameters in the major premise
|
||||
list<unsigned> m_indices_pos; // position of the recursor indices in the major premise
|
||||
name m_recursor;
|
||||
name m_type_name;
|
||||
optional<unsigned> m_motive_univ_pos; // if none, then recursor can only eliminate to Prop
|
||||
bool m_dep_elim;
|
||||
unsigned m_num_args; // total number of arguments
|
||||
unsigned m_major_pos;
|
||||
// if param is <none>, then it should be resolved by type class resolution
|
||||
list<optional<unsigned>> m_params_pos; // position of the recursor parameters in the major premise
|
||||
list<unsigned> m_indices_pos; // position of the recursor indices in the major premise
|
||||
|
||||
public:
|
||||
recursor_info(name const & r, name const & I,
|
||||
optional<unsigned> const & motive_univ_pos,
|
||||
bool dep_elim,
|
||||
unsigned num_args, unsigned major_pos,
|
||||
list<unsigned> const & params_pos,
|
||||
list<unsigned> const & indices_pos);
|
||||
recursor_info(name const & r, name const & I, optional<unsigned> const & motive_univ_pos,
|
||||
bool dep_elim, unsigned num_args, unsigned major_pos,
|
||||
list<optional<unsigned>> const & params_pos, list<unsigned> const & indices_pos);
|
||||
recursor_info();
|
||||
|
||||
name const & get_name() const { return m_recursor; }
|
||||
|
@ -37,7 +35,7 @@ public:
|
|||
unsigned get_major_pos() const { return m_major_pos; }
|
||||
optional<unsigned> const & get_motive_univ_pos() const { return m_motive_univ_pos; }
|
||||
/** \brief Return position of the recursor parameters in the major premise. */
|
||||
list<unsigned> const & get_params_pos() const { return m_params_pos; }
|
||||
list<optional<unsigned>> const & get_params_pos() const { return m_params_pos; }
|
||||
/** \brief Return position of the recursor indices in the major premise. */
|
||||
list<unsigned> const & get_indices_pos() const { return m_indices_pos; }
|
||||
bool has_dep_elim() const { return m_dep_elim; }
|
||||
|
|
Loading…
Reference in a new issue