2014-12-22 15:33:29 -05:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Author: Leonardo de Moura
|
|
|
|
|
-/
|
2014-11-30 20:34:12 -08:00
|
|
|
|
import logic.eq
|
2014-08-19 19:32:44 -07:00
|
|
|
|
|
2014-07-22 09:49:54 -07:00
|
|
|
|
namespace bool
|
2015-01-26 11:31:12 -08:00
|
|
|
|
local attribute bor [reducible]
|
|
|
|
|
local attribute band [reducible]
|
2014-07-04 22:22:26 -07:00
|
|
|
|
|
2014-09-04 22:31:52 -07:00
|
|
|
|
theorem dichotomy (b : bool) : b = ff ∨ b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem cond_ff [simp] {A : Type} (t e : A) : cond ff t e = e :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-04 22:22:26 -07:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem cond_tt [simp] {A : Type} (t e : A) : cond tt t e = t :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem eq_tt_of_ne_ff : ∀ {a : bool}, a ≠ ff → a = tt :=
|
|
|
|
|
by rec_simp
|
2015-03-04 17:56:50 -08:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem eq_ff_of_ne_tt : ∀ {a : bool}, a ≠ tt → a = ff :=
|
|
|
|
|
by rec_simp
|
2015-03-04 17:56:50 -08:00
|
|
|
|
|
|
|
|
|
theorem absurd_of_eq_ff_of_eq_tt {B : Prop} {a : bool} (H₁ : a = ff) (H₂ : a = tt) : B :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2015-03-04 17:56:50 -08:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem tt_bor [simp] (a : bool) : bor tt a = tt :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2014-10-21 14:08:07 -07:00
|
|
|
|
notation a || b := bor a b
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem bor_tt [simp] (a : bool) : a || tt = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem ff_bor [simp] (a : bool) : ff || a = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem bor_ff [simp] (a : bool) : a || ff = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem bor_self [simp] (a : bool) : a || a = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
theorem bor_comm [simp] (a b : bool) : a || b = b || a :=
|
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem bor_assoc [simp] (a b c : bool) : (a || b) || c = a || (b || c) :=
|
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem bor_left_comm [simp] (a b c : bool) : a || (b || c) = b || (a || c) :=
|
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-03-06 19:20:48 -08:00
|
|
|
|
theorem or_of_bor_eq {a b : bool} : a || b = tt → a = tt ∨ b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-27 12:50:57 -07:00
|
|
|
|
|
2015-03-30 04:55:28 -07:00
|
|
|
|
theorem bor_inl {a b : bool} (H : a = tt) : a || b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2015-03-30 04:55:28 -07:00
|
|
|
|
|
|
|
|
|
theorem bor_inr {a b : bool} (H : b = tt) : a || b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2015-03-30 04:55:28 -07:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem ff_band [simp] (a : bool) : ff && a = ff :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem tt_band [simp] (a : bool) : tt && a = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem band_ff [simp] (a : bool) : a && ff = ff :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem band_tt [simp] (a : bool) : a && tt = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem band_self [simp] (a : bool) : a && a = a :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem band_comm [simp] (a b : bool) : a && b = b && a :=
|
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem band_assoc [simp] (a b c : bool) : (a && b) && c = a && (b && c) :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
theorem band_left_comm [simp] (a b c : bool) : a && (b && c) = b && (a && c) :=
|
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-03-06 19:20:48 -08:00
|
|
|
|
theorem band_elim_left {a b : bool} (H : a && b = tt) : a = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-27 12:50:57 -07:00
|
|
|
|
|
2015-03-30 04:55:28 -07:00
|
|
|
|
theorem band_intro {a b : bool} (H₁ : a = tt) (H₂ : b = tt) : a && b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2015-03-30 04:55:28 -07:00
|
|
|
|
|
2015-03-06 19:20:48 -08:00
|
|
|
|
theorem band_elim_right {a b : bool} (H : a && b = tt) : b = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem bnot_false [simp] : bnot ff = tt :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2015-12-30 11:22:58 -08:00
|
|
|
|
theorem bnot_true [simp] : bnot tt = ff :=
|
2014-09-04 22:31:52 -07:00
|
|
|
|
rfl
|
2014-07-19 20:08:52 +01:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
theorem bnot_bnot [simp] (a : bool) : bnot (bnot a) = a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
2015-07-03 17:38:23 -07:00
|
|
|
|
theorem eq_tt_of_bnot_eq_ff {a : bool} : bnot a = ff → a = tt :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
2015-07-03 17:38:23 -07:00
|
|
|
|
|
|
|
|
|
theorem eq_ff_of_bnot_eq_tt {a : bool} : bnot a = tt → a = ff :=
|
2016-01-01 13:52:42 -08:00
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
definition bxor : bool → bool → bool
|
|
|
|
|
| ff ff := ff
|
|
|
|
|
| ff tt := tt
|
|
|
|
|
| tt ff := tt
|
|
|
|
|
| tt tt := ff
|
|
|
|
|
|
|
|
|
|
lemma ff_bxor_ff [simp] : bxor ff ff = ff := rfl
|
|
|
|
|
lemma ff_bxor_tt [simp] : bxor ff tt = tt := rfl
|
|
|
|
|
lemma tt_bxor_ff [simp] : bxor tt ff = tt := rfl
|
|
|
|
|
lemma tt_bxor_tt [simp] : bxor tt tt = ff := rfl
|
|
|
|
|
|
|
|
|
|
lemma bxor_self [simp] (a : bool) : bxor a a = ff :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma bxor_ff [simp] (a : bool) : bxor a ff = a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma bxor_tt [simp] (a : bool) : bxor a tt = bnot a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma ff_bxor [simp] (a : bool) : bxor ff a = a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma tt_bxor [simp] (a : bool) : bxor tt a = bnot a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma bxor_comm [simp] (a b : bool) : bxor a b = bxor b a :=
|
|
|
|
|
by rec_simp
|
|
|
|
|
|
|
|
|
|
lemma bxor_assoc [simp] (a b c : bool) : bxor (bxor a b) c = bxor a (bxor b c) :=
|
|
|
|
|
by rec_simp
|
2015-11-29 23:06:06 -08:00
|
|
|
|
|
2016-01-01 13:52:42 -08:00
|
|
|
|
lemma bxor_left_comm [simp] (a b c : bool) : bxor a (bxor b c) = bxor b (bxor a c) :=
|
|
|
|
|
by rec_simp
|
2015-02-11 12:49:27 -08:00
|
|
|
|
end bool
|