2014-08-07 18:36:44 +00:00
|
|
|
|
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
2014-07-02 15:36:05 +00:00
|
|
|
|
-- Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
-- Author: Leonardo de Moura
|
2014-12-01 04:34:12 +00:00
|
|
|
|
import logic.eq
|
2014-10-02 00:51:17 +00:00
|
|
|
|
open eq eq.ops decidable
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
namespace bool
|
2014-11-07 16:41:14 +00:00
|
|
|
|
reducible bor band rec_on
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-09-05 05:31:52 +00:00
|
|
|
|
theorem dichotomy (b : bool) : b = ff ∨ b = tt :=
|
|
|
|
|
cases_on b (or.inl rfl) (or.inr rfl)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem cond.ff {A : Type} (t e : A) : cond ff t e = e :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem cond.tt {A : Type} (t e : A) : cond tt t e = t :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-09-05 05:31:52 +00:00
|
|
|
|
theorem ff_ne_tt : ¬ ff = tt :=
|
|
|
|
|
assume H : ff = tt, absurd
|
2014-10-05 16:50:55 +00:00
|
|
|
|
(calc true = cond tt true false : !cond.tt⁻¹
|
2014-09-05 05:31:52 +00:00
|
|
|
|
... = cond ff true false : {H⁻¹}
|
2014-11-07 16:41:14 +00:00
|
|
|
|
... = false : cond.ff)
|
2014-09-05 05:31:52 +00:00
|
|
|
|
true_ne_false
|
2014-07-05 07:43:10 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.tt_left (a : bool) : bor tt a = tt :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-21 21:08:07 +00:00
|
|
|
|
notation a || b := bor a b
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.tt_right (a : bool) : a || tt = tt :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.ff_left (a : bool) : ff || a = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.ff_right (a : bool) : a || ff = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.id (a : bool) : a || a = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.comm (a b : bool) : a || b = b || a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a
|
|
|
|
|
(cases_on b rfl rfl)
|
|
|
|
|
(cases_on b rfl rfl)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.assoc (a b c : bool) : (a || b) || c = a || (b || c) :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a
|
2014-10-05 16:50:55 +00:00
|
|
|
|
(calc (ff || b) || c = b || c : {!bor.ff_left}
|
|
|
|
|
... = ff || (b || c) : !bor.ff_left⁻¹)
|
|
|
|
|
(calc (tt || b) || c = tt || c : {!bor.tt_left}
|
|
|
|
|
... = tt : !bor.tt_left
|
|
|
|
|
... = tt || (b || c) : !bor.tt_left⁻¹)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bor.to_or {a b : bool} : a || b = tt → a = tt ∨ b = tt :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rec_on a
|
|
|
|
|
(assume H : ff || b = tt,
|
2014-10-05 16:50:55 +00:00
|
|
|
|
have Hb : b = tt, from !bor.ff_left ▸ H,
|
2014-09-05 05:31:52 +00:00
|
|
|
|
or.inr Hb)
|
|
|
|
|
(assume H, or.inl rfl)
|
2014-07-27 19:50:57 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.ff_left (a : bool) : ff && a = ff :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.tt_left (a : bool) : tt && a = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.ff_right (a : bool) : a && ff = ff :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.tt_right (a : bool) : a && tt = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.id (a : bool) : a && a = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.comm (a b : bool) : a && b = b && a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a
|
|
|
|
|
(cases_on b rfl rfl)
|
|
|
|
|
(cases_on b rfl rfl)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.assoc (a b c : bool) : (a && b) && c = a && (b && c) :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a
|
2014-10-05 16:50:55 +00:00
|
|
|
|
(calc (ff && b) && c = ff && c : {!band.ff_left}
|
|
|
|
|
... = ff : !band.ff_left
|
|
|
|
|
... = ff && (b && c) : !band.ff_left⁻¹)
|
|
|
|
|
(calc (tt && b) && c = b && c : {!band.tt_left}
|
|
|
|
|
... = tt && (b && c) : !band.tt_left⁻¹)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.eq_tt_elim_left {a b : bool} (H : a && b = tt) : a = tt :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
or.elim (dichotomy a)
|
|
|
|
|
(assume H0 : a = ff,
|
|
|
|
|
absurd
|
2014-10-05 16:50:55 +00:00
|
|
|
|
(calc ff = ff && b : !band.ff_left⁻¹
|
2014-09-05 05:31:52 +00:00
|
|
|
|
... = a && b : {H0⁻¹}
|
|
|
|
|
... = tt : H)
|
|
|
|
|
ff_ne_tt)
|
|
|
|
|
(assume H1 : a = tt, H1)
|
2014-07-27 19:50:57 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem band.eq_tt_elim_right {a b : bool} (H : a && b = tt) : b = tt :=
|
|
|
|
|
band.eq_tt_elim_left (!band.comm ⬝ H)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bnot.bnot (a : bool) : bnot (bnot a) = a :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
cases_on a rfl rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bnot.false : bnot ff = tt :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-05 16:50:55 +00:00
|
|
|
|
theorem bnot.true : bnot tt = ff :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-10-02 16:00:34 +00:00
|
|
|
|
protected definition is_inhabited [instance] : inhabited bool :=
|
2014-09-05 05:31:52 +00:00
|
|
|
|
inhabited.mk ff
|
2014-08-01 00:48:51 +00:00
|
|
|
|
|
2014-09-30 16:02:37 +00:00
|
|
|
|
protected definition has_decidable_eq [instance] : decidable_eq bool :=
|
2014-09-09 23:07:07 +00:00
|
|
|
|
take a b : bool,
|
2014-09-08 05:22:04 +00:00
|
|
|
|
rec_on a
|
|
|
|
|
(rec_on b (inl rfl) (inr ff_ne_tt))
|
2014-09-09 23:07:07 +00:00
|
|
|
|
(rec_on b (inr (ne.symm ff_ne_tt)) (inl rfl))
|
2014-11-07 16:41:14 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
end bool
|