feat(library,hott): remove rapply tactic
This commit is contained in:
parent
88023e9597
commit
ca8943f45b
10 changed files with 15 additions and 34 deletions
|
@ -65,7 +65,6 @@ definition identifier_list := expr_list
|
||||||
definition opt_identifier_list := expr_list
|
definition opt_identifier_list := expr_list
|
||||||
|
|
||||||
opaque definition apply (e : expr) : tactic := builtin
|
opaque definition apply (e : expr) : tactic := builtin
|
||||||
opaque definition rapply (e : expr) : tactic := builtin
|
|
||||||
opaque definition fapply (e : expr) : tactic := builtin
|
opaque definition fapply (e : expr) : tactic := builtin
|
||||||
opaque definition rename (a b : identifier) : tactic := builtin
|
opaque definition rename (a b : identifier) : tactic := builtin
|
||||||
opaque definition intro (e : identifier) : tactic := builtin
|
opaque definition intro (e : identifier) : tactic := builtin
|
||||||
|
|
|
@ -65,7 +65,6 @@ definition identifier_list := expr_list
|
||||||
definition opt_identifier_list := expr_list
|
definition opt_identifier_list := expr_list
|
||||||
|
|
||||||
opaque definition apply (e : expr) : tactic := builtin
|
opaque definition apply (e : expr) : tactic := builtin
|
||||||
opaque definition rapply (e : expr) : tactic := builtin
|
|
||||||
opaque definition fapply (e : expr) : tactic := builtin
|
opaque definition fapply (e : expr) : tactic := builtin
|
||||||
opaque definition rename (a b : identifier) : tactic := builtin
|
opaque definition rename (a b : identifier) : tactic := builtin
|
||||||
opaque definition intro (e : identifier) : tactic := builtin
|
opaque definition intro (e : identifier) : tactic := builtin
|
||||||
|
|
|
@ -63,7 +63,6 @@ name const * g_tactic_all_goals = nullptr;
|
||||||
name const * g_tactic_apply = nullptr;
|
name const * g_tactic_apply = nullptr;
|
||||||
name const * g_tactic_assert_hypothesis = nullptr;
|
name const * g_tactic_assert_hypothesis = nullptr;
|
||||||
name const * g_tactic_fapply = nullptr;
|
name const * g_tactic_fapply = nullptr;
|
||||||
name const * g_tactic_rapply = nullptr;
|
|
||||||
name const * g_tactic_eassumption = nullptr;
|
name const * g_tactic_eassumption = nullptr;
|
||||||
name const * g_tactic_and_then = nullptr;
|
name const * g_tactic_and_then = nullptr;
|
||||||
name const * g_tactic_append = nullptr;
|
name const * g_tactic_append = nullptr;
|
||||||
|
@ -182,7 +181,6 @@ void initialize_constants() {
|
||||||
g_tactic_apply = new name{"tactic", "apply"};
|
g_tactic_apply = new name{"tactic", "apply"};
|
||||||
g_tactic_assert_hypothesis = new name{"tactic", "assert_hypothesis"};
|
g_tactic_assert_hypothesis = new name{"tactic", "assert_hypothesis"};
|
||||||
g_tactic_fapply = new name{"tactic", "fapply"};
|
g_tactic_fapply = new name{"tactic", "fapply"};
|
||||||
g_tactic_rapply = new name{"tactic", "rapply"};
|
|
||||||
g_tactic_eassumption = new name{"tactic", "eassumption"};
|
g_tactic_eassumption = new name{"tactic", "eassumption"};
|
||||||
g_tactic_and_then = new name{"tactic", "and_then"};
|
g_tactic_and_then = new name{"tactic", "and_then"};
|
||||||
g_tactic_append = new name{"tactic", "append"};
|
g_tactic_append = new name{"tactic", "append"};
|
||||||
|
@ -302,7 +300,6 @@ void finalize_constants() {
|
||||||
delete g_tactic_apply;
|
delete g_tactic_apply;
|
||||||
delete g_tactic_assert_hypothesis;
|
delete g_tactic_assert_hypothesis;
|
||||||
delete g_tactic_fapply;
|
delete g_tactic_fapply;
|
||||||
delete g_tactic_rapply;
|
|
||||||
delete g_tactic_eassumption;
|
delete g_tactic_eassumption;
|
||||||
delete g_tactic_and_then;
|
delete g_tactic_and_then;
|
||||||
delete g_tactic_append;
|
delete g_tactic_append;
|
||||||
|
@ -421,7 +418,6 @@ 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_apply_name() { return *g_tactic_apply; }
|
||||||
name const & get_tactic_assert_hypothesis_name() { return *g_tactic_assert_hypothesis; }
|
name const & get_tactic_assert_hypothesis_name() { return *g_tactic_assert_hypothesis; }
|
||||||
name const & get_tactic_fapply_name() { return *g_tactic_fapply; }
|
name const & get_tactic_fapply_name() { return *g_tactic_fapply; }
|
||||||
name const & get_tactic_rapply_name() { return *g_tactic_rapply; }
|
|
||||||
name const & get_tactic_eassumption_name() { return *g_tactic_eassumption; }
|
name const & get_tactic_eassumption_name() { return *g_tactic_eassumption; }
|
||||||
name const & get_tactic_and_then_name() { return *g_tactic_and_then; }
|
name const & get_tactic_and_then_name() { return *g_tactic_and_then; }
|
||||||
name const & get_tactic_append_name() { return *g_tactic_append; }
|
name const & get_tactic_append_name() { return *g_tactic_append; }
|
||||||
|
|
|
@ -65,7 +65,6 @@ name const & get_tactic_all_goals_name();
|
||||||
name const & get_tactic_apply_name();
|
name const & get_tactic_apply_name();
|
||||||
name const & get_tactic_assert_hypothesis_name();
|
name const & get_tactic_assert_hypothesis_name();
|
||||||
name const & get_tactic_fapply_name();
|
name const & get_tactic_fapply_name();
|
||||||
name const & get_tactic_rapply_name();
|
|
||||||
name const & get_tactic_eassumption_name();
|
name const & get_tactic_eassumption_name();
|
||||||
name const & get_tactic_and_then_name();
|
name const & get_tactic_and_then_name();
|
||||||
name const & get_tactic_append_name();
|
name const & get_tactic_append_name();
|
||||||
|
|
|
@ -58,7 +58,6 @@ tactic.all_goals
|
||||||
tactic.apply
|
tactic.apply
|
||||||
tactic.assert_hypothesis
|
tactic.assert_hypothesis
|
||||||
tactic.fapply
|
tactic.fapply
|
||||||
tactic.rapply
|
|
||||||
tactic.eassumption
|
tactic.eassumption
|
||||||
tactic.and_then
|
tactic.and_then
|
||||||
tactic.append
|
tactic.append
|
||||||
|
|
|
@ -225,10 +225,6 @@ tactic fapply_tactic(elaborate_fn const & elab, expr const & e) {
|
||||||
return apply_tactic_core(elab, e, AddAllSubgoals);
|
return apply_tactic_core(elab, e, AddAllSubgoals);
|
||||||
}
|
}
|
||||||
|
|
||||||
tactic rapply_tactic(elaborate_fn const & elab, expr const & e) {
|
|
||||||
return apply_tactic_core(elab, e, AddRevSubgoals);
|
|
||||||
}
|
|
||||||
|
|
||||||
int mk_eassumption_tactic(lua_State * L) { return push_tactic(L, eassumption_tactic()); }
|
int mk_eassumption_tactic(lua_State * L) { return push_tactic(L, eassumption_tactic()); }
|
||||||
void open_apply_tactic(lua_State * L) {
|
void open_apply_tactic(lua_State * L) {
|
||||||
SET_GLOBAL_FUN(mk_eassumption_tactic, "eassumption_tac");
|
SET_GLOBAL_FUN(mk_eassumption_tactic, "eassumption_tac");
|
||||||
|
@ -246,12 +242,6 @@ void initialize_apply_tactic() {
|
||||||
return apply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
|
return apply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
|
||||||
});
|
});
|
||||||
|
|
||||||
register_tac(get_tactic_rapply_name(),
|
|
||||||
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
|
||||||
check_tactic_expr(app_arg(e), "invalid 'rapply' tactic, invalid argument");
|
|
||||||
return rapply_tactic(fn, get_tactic_expr_expr(app_arg(e)));
|
|
||||||
});
|
|
||||||
|
|
||||||
register_tac(get_tactic_fapply_name(),
|
register_tac(get_tactic_fapply_name(),
|
||||||
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
[](type_checker &, elaborate_fn const & fn, expr const & e, pos_info_provider const *) {
|
||||||
check_tactic_expr(app_arg(e), "invalid 'fapply' tactic, invalid argument");
|
check_tactic_expr(app_arg(e), "invalid 'fapply' tactic, invalid argument");
|
||||||
|
|
|
@ -9,7 +9,6 @@ Author: Leonardo de Moura
|
||||||
#include "library/tactic/elaborate.h"
|
#include "library/tactic/elaborate.h"
|
||||||
namespace lean {
|
namespace lean {
|
||||||
tactic apply_tactic(elaborate_fn const & fn, expr const & e);
|
tactic apply_tactic(elaborate_fn const & fn, expr const & e);
|
||||||
tactic rapply_tactic(elaborate_fn const & fn, expr const & e);
|
|
||||||
tactic fapply_tactic(elaborate_fn const & fn, expr const & e);
|
tactic fapply_tactic(elaborate_fn const & fn, expr const & e);
|
||||||
tactic eassumption_tactic();
|
tactic eassumption_tactic();
|
||||||
void open_apply_tactic(lua_State * L);
|
void open_apply_tactic(lua_State * L);
|
||||||
|
|
|
@ -5,7 +5,7 @@ theorem tst1 (a b : Prop) : a → b → b :=
|
||||||
by intro Ha; intro Hb; apply Hb
|
by intro Ha; intro Hb; apply Hb
|
||||||
|
|
||||||
theorem tst2 (a b : Prop) : a → b → a ∧ b :=
|
theorem tst2 (a b : Prop) : a → b → a ∧ b :=
|
||||||
by intro Ha; intro Hb; rapply and.intro; apply Hb; apply Ha
|
by intro Ha; intro Hb; apply and.intro; apply Ha; apply Hb
|
||||||
|
|
||||||
theorem tst3 (a b : Prop) : a → b → a ∧ b :=
|
theorem tst3 (a b : Prop) : a → b → a ∧ b :=
|
||||||
begin
|
begin
|
||||||
|
@ -19,9 +19,9 @@ end
|
||||||
theorem tst4 (a b : Prop) : a → b → a ∧ b :=
|
theorem tst4 (a b : Prop) : a → b → a ∧ b :=
|
||||||
begin
|
begin
|
||||||
intros [Ha, Hb],
|
intros [Ha, Hb],
|
||||||
rapply and.intro,
|
apply and.intro,
|
||||||
|
apply Ha,
|
||||||
apply Hb,
|
apply Hb,
|
||||||
apply Ha
|
|
||||||
end
|
end
|
||||||
|
|
||||||
theorem tst5 (a b : Prop) : a → b → a ∧ b :=
|
theorem tst5 (a b : Prop) : a → b → a ∧ b :=
|
||||||
|
|
|
@ -1,10 +1,10 @@
|
||||||
import logic
|
import logic
|
||||||
|
|
||||||
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a
|
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a
|
||||||
:= by rapply iff.intro;
|
:= by apply iff.intro;
|
||||||
intro Ha;
|
|
||||||
apply (iff.elim_left H Ha);
|
|
||||||
intro Hb;
|
intro Hb;
|
||||||
apply (iff.elim_right H Hb)
|
apply (iff.elim_right H Hb);
|
||||||
|
intro Ha;
|
||||||
|
apply (iff.elim_left H Ha)
|
||||||
|
|
||||||
check tst
|
check tst
|
||||||
|
|
|
@ -3,19 +3,19 @@ import logic
|
||||||
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a
|
theorem tst (a b : Prop) (H : a ↔ b) : b ↔ a
|
||||||
:= have H1 [visible] : a → b, -- We need to mark H1 as fact, otherwise it is not visible by tactics
|
:= have H1 [visible] : a → b, -- We need to mark H1 as fact, otherwise it is not visible by tactics
|
||||||
from iff.elim_left H,
|
from iff.elim_left H,
|
||||||
by rapply iff.intro;
|
by apply iff.intro;
|
||||||
intro Ha;
|
|
||||||
apply (H1 Ha);
|
|
||||||
intro Hb;
|
intro Hb;
|
||||||
apply (iff.elim_right H Hb)
|
apply (iff.elim_right H Hb);
|
||||||
|
intro Ha;
|
||||||
|
apply (H1 Ha)
|
||||||
|
|
||||||
theorem tst2 (a b : Prop) (H : a ↔ b) : b ↔ a
|
theorem tst2 (a b : Prop) (H : a ↔ b) : b ↔ a
|
||||||
:= have H1 [visible] : a → b,
|
:= have H1 [visible] : a → b,
|
||||||
from iff.elim_left H,
|
from iff.elim_left H,
|
||||||
begin
|
begin
|
||||||
rapply iff.intro,
|
apply iff.intro,
|
||||||
intro Ha;
|
|
||||||
apply (H1 Ha),
|
|
||||||
intro Hb;
|
intro Hb;
|
||||||
apply (iff.elim_right H Hb)
|
apply (iff.elim_right H Hb),
|
||||||
|
intro Ha;
|
||||||
|
apply (H1 Ha)
|
||||||
end
|
end
|
||||||
|
|
Loading…
Reference in a new issue