fix(library/tactic/apply_tactic): add eapply, and fix issue #361

This commit is contained in:
Leonardo de Moura 2015-05-01 15:07:28 -07:00
parent e948dd239c
commit 9ba8b284a1
19 changed files with 59 additions and 28 deletions

View file

@ -204,7 +204,7 @@ namespace functor
(q : (λ(a b : C) (f : hom a b), hom_of_eq (p b) ∘ F₁ f ∘ inv_of_eq (p a)) 3 to_fun_hom F₂) (c : C) :
ap010 to_fun_ob (functor_eq p q) c = p c :=
begin
cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intro p q,
cases F₂, revert q, eapply (homotopy.rec_on p), clear p, esimp, intro p q,
apply sorry,
--apply (homotopy3.rec_on q), clear q, intro q,
--cases p, --TODO: report: this fails

View file

@ -67,6 +67,7 @@ definition identifier_list := expr_list
definition opt_identifier_list := expr_list
opaque definition apply (e : expr) : tactic := builtin
opaque definition eapply (e : expr) : tactic := builtin
opaque definition fapply (e : expr) : tactic := builtin
opaque definition rename (a b : identifier) : tactic := builtin
opaque definition intro (e : identifier_list) : tactic := builtin

View file

@ -40,7 +40,7 @@ namespace is_trunc
definition is_hprop_is_trunc [instance] (n : trunc_index) :
Π (A : Type), is_hprop (is_trunc n A) :=
begin
apply (trunc_index.rec_on n),
eapply (trunc_index.rec_on n),
{ intro A,
apply is_trunc_is_equiv_closed,
{ apply equiv.to_is_equiv, apply is_contr.sigma_char},

View file

@ -158,7 +158,7 @@ namespace pi
[H : ∀a, is_trunc n (B a)] : is_trunc n (Πa, B a) :=
begin
revert B H,
apply (trunc_index.rec_on n),
eapply (trunc_index.rec_on n),
{intro B H,
fapply is_contr.mk,
intro a, apply center,

View file

@ -365,7 +365,7 @@ namespace sigma
[HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) :=
begin
revert A B HA HB,
apply (trunc_index.rec_on n),
eapply (trunc_index.rec_on n),
intro A B HA HB,
fapply is_trunc.is_trunc_equiv_closed,
apply equiv.symm,

View file

@ -58,7 +58,7 @@ namespace is_trunc
(n : trunc_index) [HA : is_trunc n A] : is_trunc n B :=
begin
revert A B f Hf HA,
apply (trunc_index.rec_on n),
eapply (trunc_index.rec_on n),
{ clear n, intro A B f Hf HA, cases Hf with g ε, fapply is_contr.mk,
{ exact f (center A)},
{ intro b, apply concat,
@ -154,7 +154,7 @@ namespace trunc
fapply equiv.MK,
{ intro p, cases p, apply (trunc.rec_on aa),
intro a, esimp [trunc_eq_type,trunc.rec_on], exact (tr idp)},
{ apply (trunc.rec_on aa'), apply (trunc.rec_on aa),
{ eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
intro a a' x, esimp [trunc_eq_type, trunc.rec_on] at x,
apply (trunc.rec_on x), intro p, exact (ap tr p)},
{
@ -173,14 +173,14 @@ namespace trunc
definition is_trunc_trunc_of_is_trunc [instance] [priority 500] (A : Type)
(n m : trunc_index) [H : is_trunc n A] : is_trunc n (trunc m A) :=
begin
revert A m H, apply (trunc_index.rec_on n),
revert A m H, eapply (trunc_index.rec_on n),
{ clear n, intro A m H, apply is_contr_equiv_closed,
{ apply equiv_trunc, apply (@is_trunc_of_leq _ -2), exact unit.star} },
{ clear n, intro n IH A m H, cases m with m,
{ apply (@is_trunc_of_leq _ -2), exact unit.star},
{ apply is_trunc_succ_intro, intro aa aa',
apply (@trunc.rec_on _ _ _ aa (λy, !is_trunc_succ_of_is_hprop)),
apply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_hprop)),
eapply (@trunc.rec_on _ _ _ aa' (λy, !is_trunc_succ_of_is_hprop)),
intro a a', apply (is_trunc_equiv_closed_rev),
{ apply tr_eq_tr_equiv},
{ exact (IH _ _ _)}}}

View file

@ -181,7 +181,7 @@ private theorem mem_ltype_elems {A : Type} {s : list A} {a : ⟪s⟫}
begin
revert vaeqb h,
-- TODO(Leo): check why 'cases a with va, ma' produces an incorrect proof
apply as_type.cases_on a,
eapply as_type.cases_on a,
intro va ma vaeqb,
rewrite -vaeqb, intro h,
apply mem_cons

View file

@ -259,7 +259,7 @@ theorem zip_unzip : ∀ (l : list (A × B)), zip (pr₁ (unzip l)) (pr₂ (unzip
rewrite unzip_cons,
have r : zip (pr₁ (unzip l)) (pr₂ (unzip l)) = l, from zip_unzip l,
revert r,
apply prod.cases_on (unzip l),
eapply prod.cases_on (unzip l),
intro la lb r,
rewrite -r
end

View file

@ -138,7 +138,7 @@ assume p, gen [a] p rfl
theorem eq_singlenton_of_perm (a b : A) : [a] ~ [b] → a = b :=
assume p,
begin
injection eq_singlenton_of_perm_inv a p with e₁ e₂,
injection eq_singlenton_of_perm_inv a p with e₁,
rewrite e₁
end

View file

@ -67,6 +67,7 @@ definition identifier_list := expr_list
definition opt_identifier_list := expr_list
opaque definition apply (e : expr) : tactic := builtin
opaque definition eapply (e : expr) : tactic := builtin
opaque definition fapply (e : expr) : tactic := builtin
opaque definition rename (a b : identifier) : tactic := builtin
opaque definition intro (e : identifier_list) : tactic := builtin

View file

@ -132,7 +132,7 @@
(,(rx (not (any "\.")) word-start
(group
(or "\\b.*_tac" "Cond" "or_else" "then" "try" "when" "assumption" "eassumption" "rapply"
"apply" "fapply" "rename" "intro" "intros" "all_goals" "fold"
"apply" "fapply" "eapply" "rename" "intro" "intros" "all_goals" "fold"
"generalize" "generalizes" "clear" "clears" "revert" "reverts" "back" "beta" "done" "exact" "rexact"
"refine" "repeat" "whnf" "rotate" "rotate_left" "rotate_right" "inversion" "cases" "rewrite" "esimp"
"unfold" "change" "check_expr" "contradiction" "exfalso" "split" "existsi" "constructor" "left" "right"

View file

@ -68,6 +68,7 @@ name const * g_tactic = nullptr;
name const * g_tactic_all_goals = nullptr;
name const * g_tactic_apply = nullptr;
name const * g_tactic_assert_hypothesis = nullptr;
name const * g_tactic_eapply = nullptr;
name const * g_tactic_fapply = nullptr;
name const * g_tactic_eassumption = nullptr;
name const * g_tactic_and_then = nullptr;
@ -193,6 +194,7 @@ void initialize_constants() {
g_tactic_all_goals = new name{"tactic", "all_goals"};
g_tactic_apply = new name{"tactic", "apply"};
g_tactic_assert_hypothesis = new name{"tactic", "assert_hypothesis"};
g_tactic_eapply = new name{"tactic", "eapply"};
g_tactic_fapply = new name{"tactic", "fapply"};
g_tactic_eassumption = new name{"tactic", "eassumption"};
g_tactic_and_then = new name{"tactic", "and_then"};
@ -319,6 +321,7 @@ void finalize_constants() {
delete g_tactic_all_goals;
delete g_tactic_apply;
delete g_tactic_assert_hypothesis;
delete g_tactic_eapply;
delete g_tactic_fapply;
delete g_tactic_eassumption;
delete g_tactic_and_then;
@ -444,6 +447,7 @@ name const & get_tactic_name() { return *g_tactic; }
name const & get_tactic_all_goals_name() { return *g_tactic_all_goals; }
name const & get_tactic_apply_name() { return *g_tactic_apply; }
name const & get_tactic_assert_hypothesis_name() { return *g_tactic_assert_hypothesis; }
name const & get_tactic_eapply_name() { return *g_tactic_eapply; }
name const & get_tactic_fapply_name() { return *g_tactic_fapply; }
name const & get_tactic_eassumption_name() { return *g_tactic_eassumption; }
name const & get_tactic_and_then_name() { return *g_tactic_and_then; }

View file

@ -70,6 +70,7 @@ name const & get_tactic_name();
name const & get_tactic_all_goals_name();
name const & get_tactic_apply_name();
name const & get_tactic_assert_hypothesis_name();
name const & get_tactic_eapply_name();
name const & get_tactic_fapply_name();
name const & get_tactic_eassumption_name();
name const & get_tactic_and_then_name();

View file

@ -63,6 +63,7 @@ tactic
tactic.all_goals
tactic.apply
tactic.assert_hypothesis
tactic.eapply
tactic.fapply
tactic.eassumption
tactic.and_then

View file

@ -67,9 +67,11 @@ static void remove_redundant_metas(buffer<expr> & metas) {
enum subgoals_action_kind { IgnoreSubgoals, AddRevSubgoals, AddSubgoals, AddAllSubgoals };
enum add_meta_kind { DoNotAdd, AddDiff, AddAll };
static proof_state_seq apply_tactic_core(environment const & env, io_state const & ios, proof_state const & s,
expr const & _e, buffer<constraint> & cs,
bool add_meta, subgoals_action_kind subgoals_action) {
add_meta_kind add_meta, subgoals_action_kind subgoals_action) {
goals const & gs = s.get_goals();
if (empty(gs)) {
throw_no_goal_if_enabled(s);
@ -90,12 +92,17 @@ static proof_state_seq apply_tactic_core(environment const & env, io_state const
local_context ctx;
bool initialized_ctx = false;
unifier_config cfg(ios.get_options());
if (add_meta) {
// unsigned num_t = get_expect_num_args(*tc, t);
if (add_meta != DoNotAdd) {
unsigned num_e_t = get_expect_num_args(*tc, e_t);
// Remark: we used to add (num_e_t - num_t) arguments.
// This would allow us to handle (A -> B) without using intros,
// but it was preventing us from solving other problems
if (add_meta == AddDiff) {
unsigned num_t = get_expect_num_args(*tc, t);
if (num_t <= num_e_t)
num_e_t -= num_t;
else
num_e_t = 0;
} else {
lean_assert(add_meta == AddAll);
}
for (unsigned i = 0; i < num_e_t; i++) {
auto e_t_cs = tc->whnf(e_t);
e_t_cs.second.linearize(cs);
@ -171,7 +178,7 @@ static proof_state_seq apply_tactic_core(environment const & env, io_state const
}
static proof_state_seq apply_tactic_core(environment const & env, io_state const & ios, proof_state const & s, expr const & e,
bool add_meta, subgoals_action_kind subgoals_action) {
add_meta_kind add_meta, subgoals_action_kind subgoals_action) {
buffer<constraint> cs;
return apply_tactic_core(env, ios, s, e, cs, add_meta, subgoals_action);
}
@ -179,7 +186,7 @@ static proof_state_seq apply_tactic_core(environment const & env, io_state const
proof_state_seq apply_tactic_core(environment const & env, io_state const & ios, proof_state const & s, expr const & e, constraint_seq const & cs) {
buffer<constraint> tmp_cs;
cs.linearize(tmp_cs);
return apply_tactic_core(env, ios, s, e, tmp_cs, true, AddSubgoals);
return apply_tactic_core(env, ios, s, e, tmp_cs, AddDiff, AddSubgoals);
}
tactic apply_tactic_core(expr const & e, constraint_seq const & cs) {
@ -201,13 +208,13 @@ tactic eassumption_tactic() {
buffer<expr> hs;
get_app_args(g.get_meta(), hs);
for (expr const & h : hs) {
r = append(r, apply_tactic_core(env, ios, new_s, h, false, IgnoreSubgoals));
r = append(r, apply_tactic_core(env, ios, new_s, h, DoNotAdd, IgnoreSubgoals));
}
return r;
});
}
tactic apply_tactic_core(elaborate_fn const & elab, expr const & e, subgoals_action_kind k) {
tactic apply_tactic_core(elaborate_fn const & elab, expr const & e, add_meta_kind add_meta, subgoals_action_kind k) {
return tactic([=](environment const & env, io_state const & ios, proof_state const & s) {
goals const & gs = s.get_goals();
if (empty(gs)) {
@ -223,16 +230,20 @@ tactic apply_tactic_core(elaborate_fn const & elab, expr const & e, subgoals_act
to_buffer(ecs.second, cs);
to_buffer(s.get_postponed(), cs);
proof_state new_s(s, ngen, constraints());
return apply_tactic_core(env, ios, new_s, new_e, cs, true, k);
return apply_tactic_core(env, ios, new_s, new_e, cs, add_meta, k);
});
}
tactic apply_tactic(elaborate_fn const & elab, expr const & e) {
return apply_tactic_core(elab, e, AddSubgoals);
return apply_tactic_core(elab, e, AddDiff, AddSubgoals);
}
tactic fapply_tactic(elaborate_fn const & elab, expr const & e) {
return apply_tactic_core(elab, e, AddAllSubgoals);
return apply_tactic_core(elab, e, AddDiff, AddAllSubgoals);
}
tactic eapply_tactic(elaborate_fn const & elab, expr const & e) {
return apply_tactic_core(elab, e, AddAll, AddSubgoals);
}
int mk_eassumption_tactic(lua_State * L) { return push_tactic(L, eassumption_tactic()); }
@ -252,6 +263,12 @@ void initialize_apply_tactic() {
return apply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
});
register_tac(get_tactic_eapply_name(),
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
check_tactic_expr(app_arg(e), "invalid 'eapply' tactic, invalid argument");
return eapply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
});
register_tac(get_tactic_fapply_name(),
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
check_tactic_expr(app_arg(e), "invalid 'fapply' tactic, invalid argument");

View file

@ -3,7 +3,7 @@ open eq
definition concat_pV_p {A : Type} {x y z : A} (p : x = z) (q : y = z) : (p ⬝ q⁻¹) ⬝ q = p :=
begin
generalize p,
apply (eq.rec_on q),
eapply (eq.rec_on q),
intro p,
apply (eq.rec_on p),
apply idp

6
tests/lean/run/361.lean Normal file
View file

@ -0,0 +1,6 @@
variables {P Q R : Prop}
theorem foo (H : P → Q → R) (x : P) : Q → R :=
begin
apply H,
exact x
end

View file

@ -12,7 +12,7 @@ theorem foo1 {A : Type} (a b c : A) (P : A → Prop) : P a → a = b → P b :=
begin
intros [Hp, Heq],
revert Hp,
apply (eq.rec_on Heq),
eapply (eq.rec_on Heq),
intro Hpa,
apply Hpa
end

View file

@ -13,7 +13,7 @@ begin
intros [Hp, Heq],
reverts [Heq, Hp],
intro Heq,
apply (eq.rec_on Heq),
eapply (eq.rec_on Heq),
intro Pa,
apply Pa
end