feat(builtin/kernel): add left_comm theorem

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
Leonardo de Moura 2014-01-11 18:08:08 -08:00
parent a1a467a65f
commit 781720a26a
4 changed files with 26 additions and 0 deletions

View file

@ -365,6 +365,20 @@ theorem exists_unfold {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) =
:= boolext (λ H : (∃ x : A, P x), exists_unfold1 a H)
(λ H : (P a (∃ x : A, x ≠ a ∧ P x)), exists_unfold2 a H)
-- Remark: ordered rewriting + assoc + comm + left_comm sorts a term lexicographically
theorem left_comm {A : TypeU} {R : A -> A -> A} (comm : ∀ x y, R x y = R y x) (assoc : ∀ x y z, R (R x y) z = R x (R y z)) :
∀ x y z, R x (R y z) = R y (R x z)
:= λ x y z, calc R x (R y z) = R (R x y) z : symm (assoc x y z)
... = R (R y x) z : { comm x y }
... = R y (R x z) : assoc y x z
theorem and_left_comm (a b c : Bool) : (a ∧ (b ∧ c)) = (b ∧ (a ∧ c))
:= left_comm and_comm and_assoc a b c
theorem or_left_comm (a b c : Bool) : (a (b c)) = (b (a c))
:= left_comm or_comm or_assoc a b c
-- Congruence theorems for contextual simplification
-- Simplify a → b, by first simplifying a to c using the fact that ¬ b is true, and then

Binary file not shown.

View file

@ -100,6 +100,9 @@ MK_CONSTANT(not_exists_elim_fn, name("not_exists_elim"));
MK_CONSTANT(exists_unfold1_fn, name("exists_unfold1"));
MK_CONSTANT(exists_unfold2_fn, name("exists_unfold2"));
MK_CONSTANT(exists_unfold_fn, name("exists_unfold"));
MK_CONSTANT(left_comm_fn, name("left_comm"));
MK_CONSTANT(and_left_comm_fn, name("and_left_comm"));
MK_CONSTANT(or_left_comm_fn, name("or_left_comm"));
MK_CONSTANT(imp_congrr_fn, name("imp_congrr"));
MK_CONSTANT(imp_congrl_fn, name("imp_congrl"));
MK_CONSTANT(imp_congr_fn, name("imp_congr"));

View file

@ -291,6 +291,15 @@ inline expr mk_exists_unfold2_th(expr const & e1, expr const & e2, expr const &
expr mk_exists_unfold_fn();
bool is_exists_unfold_fn(expr const & e);
inline expr mk_exists_unfold_th(expr const & e1, expr const & e2, expr const & e3) { return mk_app({mk_exists_unfold_fn(), e1, e2, e3}); }
expr mk_left_comm_fn();
bool is_left_comm_fn(expr const & e);
inline expr mk_left_comm_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4, expr const & e5, expr const & e6, expr const & e7) { return mk_app({mk_left_comm_fn(), e1, e2, e3, e4, e5, e6, e7}); }
expr mk_and_left_comm_fn();
bool is_and_left_comm_fn(expr const & e);
inline expr mk_and_left_comm_th(expr const & e1, expr const & e2, expr const & e3) { return mk_app({mk_and_left_comm_fn(), e1, e2, e3}); }
expr mk_or_left_comm_fn();
bool is_or_left_comm_fn(expr const & e);
inline expr mk_or_left_comm_th(expr const & e1, expr const & e2, expr const & e3) { return mk_app({mk_or_left_comm_fn(), e1, e2, e3}); }
expr mk_imp_congrr_fn();
bool is_imp_congrr_fn(expr const & e);
inline expr mk_imp_congrr_th(expr const & e1, expr const & e2, expr const & e3, expr const & e4, expr const & e5, expr const & e6) { return mk_app({mk_imp_congrr_fn(), e1, e2, e3, e4, e5, e6}); }