feat(library/tactic): add 'contradiction' tactic
see issue #500 Remark: this tactic also applies no_confusion to take care of a contradiction
This commit is contained in:
parent
3233008039
commit
9c8a63caec
17 changed files with 225 additions and 21 deletions
|
@ -42,6 +42,7 @@ opaque definition id : tactic := builtin
|
|||
opaque definition beta : tactic := builtin
|
||||
opaque definition info : tactic := builtin
|
||||
opaque definition whnf : tactic := builtin
|
||||
opaque definition contradiction : tactic := builtin
|
||||
opaque definition rotate_left (k : num) := builtin
|
||||
opaque definition rotate_right (k : num) := builtin
|
||||
definition rotate (k : num) := rotate_left k
|
||||
|
|
|
@ -42,6 +42,7 @@ opaque definition id : tactic := builtin
|
|||
opaque definition beta : tactic := builtin
|
||||
opaque definition info : tactic := builtin
|
||||
opaque definition whnf : tactic := builtin
|
||||
opaque definition contradiction : tactic := builtin
|
||||
opaque definition rotate_left (k : num) := builtin
|
||||
opaque definition rotate_right (k : num) := builtin
|
||||
definition rotate (k : num) := rotate_left k
|
||||
|
|
|
@ -135,7 +135,7 @@
|
|||
"apply" "fapply" "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"))
|
||||
"unfold" "change" "check_expr" "contradiction"))
|
||||
word-end)
|
||||
(1 'font-lock-constant-face))
|
||||
;; Types
|
||||
|
|
|
@ -14,6 +14,8 @@ name const * g_bool_tt = nullptr;
|
|||
name const * g_char = nullptr;
|
||||
name const * g_char_mk = nullptr;
|
||||
name const * g_dite = nullptr;
|
||||
name const * g_empty = nullptr;
|
||||
name const * g_empty_rec = nullptr;
|
||||
name const * g_eq = nullptr;
|
||||
name const * g_eq_elim_inv_inv = nullptr;
|
||||
name const * g_eq_intro = nullptr;
|
||||
|
@ -24,6 +26,7 @@ name const * g_eq_symm = nullptr;
|
|||
name const * g_eq_trans = nullptr;
|
||||
name const * g_exists_elim = nullptr;
|
||||
name const * g_false = nullptr;
|
||||
name const * g_false_rec = nullptr;
|
||||
name const * g_heq = nullptr;
|
||||
name const * g_heq_refl = nullptr;
|
||||
name const * g_heq_to_eq = nullptr;
|
||||
|
@ -133,6 +136,8 @@ void initialize_constants() {
|
|||
g_char = new name{"char"};
|
||||
g_char_mk = new name{"char", "mk"};
|
||||
g_dite = new name{"dite"};
|
||||
g_empty = new name{"empty"};
|
||||
g_empty_rec = new name{"empty", "rec"};
|
||||
g_eq = new name{"eq"};
|
||||
g_eq_elim_inv_inv = new name{"eq", "elim_inv_inv"};
|
||||
g_eq_intro = new name{"eq", "intro"};
|
||||
|
@ -143,6 +148,7 @@ void initialize_constants() {
|
|||
g_eq_trans = new name{"eq", "trans"};
|
||||
g_exists_elim = new name{"exists", "elim"};
|
||||
g_false = new name{"false"};
|
||||
g_false_rec = new name{"false", "rec"};
|
||||
g_heq = new name{"heq"};
|
||||
g_heq_refl = new name{"heq", "refl"};
|
||||
g_heq_to_eq = new name{"heq", "to_eq"};
|
||||
|
@ -253,6 +259,8 @@ void finalize_constants() {
|
|||
delete g_char;
|
||||
delete g_char_mk;
|
||||
delete g_dite;
|
||||
delete g_empty;
|
||||
delete g_empty_rec;
|
||||
delete g_eq;
|
||||
delete g_eq_elim_inv_inv;
|
||||
delete g_eq_intro;
|
||||
|
@ -263,6 +271,7 @@ void finalize_constants() {
|
|||
delete g_eq_trans;
|
||||
delete g_exists_elim;
|
||||
delete g_false;
|
||||
delete g_false_rec;
|
||||
delete g_heq;
|
||||
delete g_heq_refl;
|
||||
delete g_heq_to_eq;
|
||||
|
@ -372,6 +381,8 @@ name const & get_bool_tt_name() { return *g_bool_tt; }
|
|||
name const & get_char_name() { return *g_char; }
|
||||
name const & get_char_mk_name() { return *g_char_mk; }
|
||||
name const & get_dite_name() { return *g_dite; }
|
||||
name const & get_empty_name() { return *g_empty; }
|
||||
name const & get_empty_rec_name() { return *g_empty_rec; }
|
||||
name const & get_eq_name() { return *g_eq; }
|
||||
name const & get_eq_elim_inv_inv_name() { return *g_eq_elim_inv_inv; }
|
||||
name const & get_eq_intro_name() { return *g_eq_intro; }
|
||||
|
@ -382,6 +393,7 @@ name const & get_eq_symm_name() { return *g_eq_symm; }
|
|||
name const & get_eq_trans_name() { return *g_eq_trans; }
|
||||
name const & get_exists_elim_name() { return *g_exists_elim; }
|
||||
name const & get_false_name() { return *g_false; }
|
||||
name const & get_false_rec_name() { return *g_false_rec; }
|
||||
name const & get_heq_name() { return *g_heq; }
|
||||
name const & get_heq_refl_name() { return *g_heq_refl; }
|
||||
name const & get_heq_to_eq_name() { return *g_heq_to_eq; }
|
||||
|
|
|
@ -16,6 +16,8 @@ name const & get_bool_tt_name();
|
|||
name const & get_char_name();
|
||||
name const & get_char_mk_name();
|
||||
name const & get_dite_name();
|
||||
name const & get_empty_name();
|
||||
name const & get_empty_rec_name();
|
||||
name const & get_eq_name();
|
||||
name const & get_eq_elim_inv_inv_name();
|
||||
name const & get_eq_intro_name();
|
||||
|
@ -26,6 +28,7 @@ name const & get_eq_symm_name();
|
|||
name const & get_eq_trans_name();
|
||||
name const & get_exists_elim_name();
|
||||
name const & get_false_name();
|
||||
name const & get_false_rec_name();
|
||||
name const & get_heq_name();
|
||||
name const & get_heq_refl_name();
|
||||
name const & get_heq_to_eq_name();
|
||||
|
|
|
@ -9,6 +9,8 @@ bool.tt
|
|||
char
|
||||
char.mk
|
||||
dite
|
||||
empty
|
||||
empty.rec
|
||||
eq
|
||||
eq.elim_inv_inv
|
||||
eq.intro
|
||||
|
@ -19,6 +21,7 @@ eq.symm
|
|||
eq.trans
|
||||
exists.elim
|
||||
false
|
||||
false.rec
|
||||
heq
|
||||
heq.refl
|
||||
heq.to_eq
|
||||
|
|
|
@ -4,6 +4,6 @@ exact_tactic.cpp generalize_tactic.cpp
|
|||
inversion_tactic.cpp whnf_tactic.cpp revert_tactic.cpp
|
||||
assert_tactic.cpp clear_tactic.cpp expr_to_tactic.cpp location.cpp
|
||||
rewrite_tactic.cpp util.cpp class_instance_synth.cpp init_module.cpp
|
||||
change_tactic.cpp check_expr_tactic.cpp let_tactic.cpp)
|
||||
change_tactic.cpp check_expr_tactic.cpp let_tactic.cpp contradiction_tactic.cpp)
|
||||
|
||||
target_link_libraries(tactic ${LEAN_LIBS})
|
||||
|
|
80
src/library/tactic/contradiction_tactic.cpp
Normal file
80
src/library/tactic/contradiction_tactic.cpp
Normal file
|
@ -0,0 +1,80 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#include "kernel/inductive/inductive.h"
|
||||
#include "library/constants.h"
|
||||
#include "library/util.h"
|
||||
#include "library/reducible.h"
|
||||
#include "library/tactic/intros_tactic.h"
|
||||
#include "library/tactic/expr_to_tactic.h"
|
||||
#include "kernel/kernel_exception.h"
|
||||
|
||||
namespace lean {
|
||||
tactic contradiction_tactic() {
|
||||
auto fn = [=](environment const & env, io_state const & ios, proof_state const & s) {
|
||||
goals const & gs = s.get_goals();
|
||||
if (empty(gs)) {
|
||||
throw_no_goal_if_enabled(s);
|
||||
return optional<proof_state>();
|
||||
}
|
||||
goal const & g = head(gs);
|
||||
expr const & t = g.get_type();
|
||||
substitution subst = s.get_subst();
|
||||
name_generator ngen = s.get_ngen();
|
||||
auto tc = mk_type_checker(env, ngen.mk_child(), s.relax_main_opaque());
|
||||
buffer<expr> hyps;
|
||||
g.get_hyps(hyps);
|
||||
for (expr const & h : hyps) {
|
||||
expr h_type = mlocal_type(h);
|
||||
h_type = tc->whnf(h_type).first;
|
||||
expr lhs, rhs;
|
||||
if (is_false(env, h_type)) {
|
||||
assign(subst, g, mk_false_rec(*tc, h, t));
|
||||
return some_proof_state(proof_state(s, tail(gs), subst, ngen));
|
||||
} else if (is_eq(h_type, lhs, rhs)) {
|
||||
lhs = tc->whnf(lhs).first;
|
||||
rhs = tc->whnf(rhs).first;
|
||||
optional<name> lhs_c = is_constructor_app(env, lhs);
|
||||
optional<name> rhs_c = is_constructor_app(env, rhs);
|
||||
if (lhs_c && rhs_c && *lhs_c != *rhs_c) {
|
||||
if (optional<name> I_name = inductive::is_intro_rule(env, *lhs_c)) {
|
||||
name no_confusion(*I_name, "no_confusion");
|
||||
try {
|
||||
expr I = tc->whnf(tc->infer(lhs).first).first;
|
||||
buffer<expr> args;
|
||||
expr I_fn = get_app_args(I, args);
|
||||
if (is_constant(I_fn)) {
|
||||
level t_lvl = sort_level(tc->ensure_type(t).first);
|
||||
expr V = mk_app(mk_app(mk_constant(no_confusion, cons(t_lvl, const_levels(I_fn))), args),
|
||||
t, lhs, rhs, h);
|
||||
if (auto r = lift_down_if_hott(*tc, V)) {
|
||||
tc->check_ignore_undefined_universes(*r);
|
||||
assign(subst, g, *r);
|
||||
return some_proof_state(proof_state(s, tail(gs), subst, ngen));
|
||||
}
|
||||
}
|
||||
} catch (kernel_exception & ex) {
|
||||
regular(env, ios) << ex << "\n";
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
return none_proof_state();
|
||||
};
|
||||
return tactic01(fn);
|
||||
}
|
||||
|
||||
void initialize_contradiction_tactic() {
|
||||
register_tac(name{"tactic", "contradiction"},
|
||||
[](type_checker &, elaborate_fn const &, expr const &, pos_info_provider const *) {
|
||||
list<name> empty;
|
||||
return then(orelse(intros_tactic(empty), id_tactic()), contradiction_tactic());
|
||||
});
|
||||
}
|
||||
void finalize_contradiction_tactic() {
|
||||
}
|
||||
}
|
11
src/library/tactic/contradiction_tactic.h
Normal file
11
src/library/tactic/contradiction_tactic.h
Normal file
|
@ -0,0 +1,11 @@
|
|||
/*
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Author: Leonardo de Moura
|
||||
*/
|
||||
#pragma once
|
||||
namespace lean {
|
||||
void initialize_contradiction_tactic();
|
||||
void finalize_contradiction_tactic();
|
||||
}
|
|
@ -23,6 +23,7 @@ Author: Leonardo de Moura
|
|||
#include "library/tactic/change_tactic.h"
|
||||
#include "library/tactic/check_expr_tactic.h"
|
||||
#include "library/tactic/let_tactic.h"
|
||||
#include "library/tactic/contradiction_tactic.h"
|
||||
|
||||
namespace lean {
|
||||
void initialize_tactic_module() {
|
||||
|
@ -45,9 +46,11 @@ void initialize_tactic_module() {
|
|||
initialize_change_tactic();
|
||||
initialize_check_expr_tactic();
|
||||
initialize_let_tactic();
|
||||
initialize_contradiction_tactic();
|
||||
}
|
||||
|
||||
void finalize_tactic_module() {
|
||||
finalize_contradiction_tactic();
|
||||
finalize_let_tactic();
|
||||
finalize_check_expr_tactic();
|
||||
finalize_change_tactic();
|
||||
|
|
|
@ -11,7 +11,7 @@ Author: Leonardo de Moura
|
|||
#include "library/tactic/expr_to_tactic.h"
|
||||
|
||||
namespace lean {
|
||||
tactic intros_tactic(list<name> _ns, bool relax_main_opaque) {
|
||||
tactic intros_tactic(list<name> _ns) {
|
||||
auto fn = [=](environment const & env, io_state const &, proof_state const & s) {
|
||||
list<name> ns = _ns;
|
||||
goals const & gs = s.get_goals();
|
||||
|
@ -21,7 +21,7 @@ tactic intros_tactic(list<name> _ns, bool relax_main_opaque) {
|
|||
}
|
||||
goal const & g = head(gs);
|
||||
name_generator ngen = s.get_ngen();
|
||||
auto tc = mk_type_checker(env, ngen.mk_child(), relax_main_opaque);
|
||||
auto tc = mk_type_checker(env, ngen.mk_child(), s.relax_main_opaque());
|
||||
expr t = g.get_type();
|
||||
expr m = g.get_meta();
|
||||
bool gen_names = empty(ns);
|
||||
|
|
|
@ -7,7 +7,7 @@ Author: Leonardo de Moura
|
|||
#pragma once
|
||||
#include "library/tactic/tactic.h"
|
||||
namespace lean {
|
||||
tactic intros_tactic(list<name> ns, bool relax_main_opaque = true);
|
||||
tactic intros_tactic(list<name> ns);
|
||||
void initialize_intros_tactic();
|
||||
void finalize_intros_tactic();
|
||||
}
|
||||
|
|
|
@ -507,7 +507,7 @@ class inversion_tac {
|
|||
throw inversion_exception();
|
||||
}
|
||||
|
||||
void throw_lift_down_failure() {
|
||||
[[ noreturn ]] void throw_lift_down_failure() {
|
||||
if (m_throw_tactic_exception)
|
||||
throw tactic_exception("invalid 'cases' tactic, lift.down failed");
|
||||
else
|
||||
|
@ -664,19 +664,10 @@ class inversion_tac {
|
|||
We must apply lift.down to eliminate the auxiliary lift.
|
||||
*/
|
||||
expr lift_down(expr const & v) {
|
||||
if (!m_proof_irrel) {
|
||||
expr v_type = whnf(infer_type(v));
|
||||
if (!is_app(v_type)) {
|
||||
throw_lift_down_failure();
|
||||
}
|
||||
expr const & lift = app_fn(v_type);
|
||||
if (!is_constant(lift) || const_name(lift) != get_lift_name()) {
|
||||
throw_lift_down_failure();
|
||||
}
|
||||
return mk_app(mk_constant(get_lift_down_name(), const_levels(lift)), app_arg(v_type), v);
|
||||
} else {
|
||||
return v;
|
||||
}
|
||||
if (auto r = lift_down_if_hott(m_tc, v))
|
||||
return *r;
|
||||
else
|
||||
throw_lift_down_failure();
|
||||
}
|
||||
|
||||
buffer<expr> m_c_args; // current intro/constructor args that may be renamed by unify_eqs
|
||||
|
|
|
@ -15,6 +15,10 @@ Author: Leonardo de Moura
|
|||
#include "library/constants.h"
|
||||
|
||||
namespace lean {
|
||||
bool is_standard(environment const & env) {
|
||||
return env.prop_proof_irrel() && env.impredicative();
|
||||
}
|
||||
|
||||
bool is_def_app(environment const & env, expr const & e) {
|
||||
if (!is_app(e))
|
||||
return false;
|
||||
|
@ -151,8 +155,8 @@ level get_datatype_level(expr ind_type) {
|
|||
}
|
||||
|
||||
bool is_inductive_predicate(environment const & env, name const & n) {
|
||||
if (!env.impredicative())
|
||||
return false; // environment does not have Prop
|
||||
if (!is_standard(env))
|
||||
return false;
|
||||
if (!inductive::is_inductive_decl(env, n))
|
||||
return false; // n is not inductive datatype
|
||||
return is_zero(get_datatype_level(env.get(n).get_type()));
|
||||
|
@ -240,6 +244,43 @@ expr to_telescope(type_checker & tc, expr type, buffer<expr> & telescope, option
|
|||
return to_telescope(tc, type, telescope, binfo, cs);
|
||||
}
|
||||
|
||||
bool is_false(expr const & e) {
|
||||
return is_constant(e) && const_name(e) == get_false_name();
|
||||
}
|
||||
|
||||
bool is_empty(expr const & e) {
|
||||
return is_constant(e) && const_name(e) == get_empty_name();
|
||||
}
|
||||
|
||||
bool is_false(environment const & env, expr const & e) {
|
||||
return is_standard(env) ? is_false(e) : is_empty(e);
|
||||
}
|
||||
|
||||
expr mk_false_rec(type_checker & tc, expr const & f, expr const & t) {
|
||||
level t_lvl = sort_level(tc.ensure_type(t).first);
|
||||
if (is_standard(tc.env())) {
|
||||
return mk_app(mk_constant(get_false_rec_name(), {t_lvl}), t, f);
|
||||
} else {
|
||||
expr f_type = tc.infer(f).first;
|
||||
level f_lvl = sort_level(tc.ensure_type(f_type).first);
|
||||
return mk_app(mk_constant(get_empty_rec_name(), {t_lvl, f_lvl}), mk_lambda("e", f_type, t), f);
|
||||
}
|
||||
}
|
||||
|
||||
optional<expr> lift_down_if_hott(type_checker & tc, expr const & v) {
|
||||
if (is_standard(tc.env())) {
|
||||
return some_expr(v);
|
||||
} else {
|
||||
expr v_type = tc.whnf(tc.infer(v).first).first;
|
||||
if (!is_app(v_type))
|
||||
return none_expr();
|
||||
expr const & lift = app_fn(v_type);
|
||||
if (!is_constant(lift) || const_name(lift) != get_lift_name())
|
||||
return none_expr();
|
||||
return some_expr(mk_app(mk_constant(get_lift_down_name(), const_levels(lift)), app_arg(v_type), v));
|
||||
}
|
||||
}
|
||||
|
||||
static expr * g_true = nullptr;
|
||||
static expr * g_true_intro = nullptr;
|
||||
static expr * g_and = nullptr;
|
||||
|
@ -387,6 +428,14 @@ bool is_eq(expr const & e) {
|
|||
return is_constant(fn) && const_name(fn) == get_eq_name();
|
||||
}
|
||||
|
||||
bool is_eq(expr const & e, expr & lhs, expr & rhs) {
|
||||
if (!is_eq(e) || !is_app(app_fn(e)))
|
||||
return false;
|
||||
lhs = app_arg(app_fn(e));
|
||||
rhs = app_arg(e);
|
||||
return true;
|
||||
}
|
||||
|
||||
bool is_eq_a_a(expr const & e) {
|
||||
if (!is_eq(e))
|
||||
return false;
|
||||
|
|
|
@ -26,6 +26,9 @@ optional<expr> unfold_app(environment const & env, expr const & e);
|
|||
*/
|
||||
optional<level> dec_level(level const & l);
|
||||
|
||||
/** \brief Return true iff \c env has been configured with an impredicative and proof irrelevant Prop. */
|
||||
bool is_standard(environment const & env);
|
||||
|
||||
bool has_unit_decls(environment const & env);
|
||||
bool has_eq_decls(environment const & env);
|
||||
bool has_heq_decls(environment const & env);
|
||||
|
@ -123,11 +126,20 @@ expr mk_pair(type_checker & tc, expr const & a, expr const & b, bool prop);
|
|||
expr mk_pr1(type_checker & tc, expr const & p, bool prop);
|
||||
expr mk_pr2(type_checker & tc, expr const & p, bool prop);
|
||||
|
||||
|
||||
bool is_false(expr const & e);
|
||||
bool is_empty(expr const & e);
|
||||
/** \brief Return true iff \c e is false (in standard mode) or empty (in HoTT) mode */
|
||||
bool is_false(environment const & env, expr const & e);
|
||||
/** \brief Return an element of type t given an element \c f : false (in standard mode) and empty (in HoTT) mode */
|
||||
expr mk_false_rec(type_checker & tc, expr const & f, expr const & t);
|
||||
|
||||
expr mk_eq(type_checker & tc, expr const & lhs, expr const & rhs);
|
||||
expr mk_refl(type_checker & tc, expr const & a);
|
||||
expr mk_symm(type_checker & tc, expr const & H);
|
||||
bool is_eq_rec(expr const & e);
|
||||
bool is_eq(expr const & e);
|
||||
bool is_eq(expr const & e, expr & lhs, expr & rhs);
|
||||
/** \brief Return true iff \c e is of the form (eq A a a) */
|
||||
bool is_eq_a_a(expr const & e);
|
||||
/** \brief Return true iff \c e is of the form (eq A a a') where \c a and \c a' are definitionally equal */
|
||||
|
@ -137,6 +149,12 @@ bool is_iff(expr const & e);
|
|||
expr mk_iff(expr const & lhs, expr const & rhs);
|
||||
expr mk_iff_refl(expr const & a);
|
||||
|
||||
/** \brief If in HoTT mode, apply lift.down.
|
||||
The no_confusion constructions uses lifts in the proof relevant version (aka HoTT mode).
|
||||
We must apply lift.down to eliminate the auxiliary lift.
|
||||
*/
|
||||
optional<expr> lift_down_if_hott(type_checker & tc, expr const & v);
|
||||
|
||||
/** \brief Create a telescope equality for HoTT library.
|
||||
This procedure assumes eq supports dependent elimination.
|
||||
For HoTT, we can't use heterogeneous equality.
|
||||
|
|
16
tests/lean/hott/contra1.hlean
Normal file
16
tests/lean/hott/contra1.hlean
Normal file
|
@ -0,0 +1,16 @@
|
|||
example (a b : nat) (h : empty) : a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), empty → a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), 0 = 1 → a = b :=
|
||||
by contradiction
|
||||
|
||||
definition id {A : Type} (a : A) := a
|
||||
|
||||
example : ∀ (a b : nat), id empty → a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), id (0 = 1) → a = b :=
|
||||
by contradiction
|
16
tests/lean/run/contra1.lean
Normal file
16
tests/lean/run/contra1.lean
Normal file
|
@ -0,0 +1,16 @@
|
|||
example (a b : nat) (h : false) : a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), false → a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), 0 = 1 → a = b :=
|
||||
by contradiction
|
||||
|
||||
definition id {A : Type} (a : A) := a
|
||||
|
||||
example : ∀ (a b : nat), id false → a = b :=
|
||||
by contradiction
|
||||
|
||||
example : ∀ (a b : nat), id (0 = 1) → a = b :=
|
||||
by contradiction
|
Loading…
Reference in a new issue