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-08-01 00:48:51 +00:00
|
|
|
|
|
2014-08-28 01:39:55 +00:00
|
|
|
|
import logic.core.connectives logic.classes.decidable logic.classes.inhabited
|
2014-08-20 22:49:44 +00:00
|
|
|
|
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open eq_ops decidable
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
|
|
|
|
inductive bool : Type :=
|
2014-08-22 22:46:10 +00:00
|
|
|
|
ff : bool,
|
|
|
|
|
tt : bool
|
2014-07-19 19:09:47 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
namespace bool
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-22 00:54:50 +00:00
|
|
|
|
theorem cases_on {p : bool → Prop} (b : bool) (H0 : p ff) (H1 : p tt) : p b :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec H0 H1 b
|
2014-07-04 21:25:44 +00:00
|
|
|
|
|
2014-08-20 22:49:44 +00:00
|
|
|
|
theorem bool_inhabited [instance] : inhabited bool :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
inhabited.mk ff
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
definition cond {A : Type} (b : bool) (t e : A) :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec e t b
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem dichotomy (b : bool) : b = ff ∨ b = tt :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
cases_on b (or.inl (eq.refl ff)) (or.inr (eq.refl tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem cond_ff {A : Type} (t e : A) : cond ff t e = e :=
|
2014-08-26 05:54:44 +00:00
|
|
|
|
rfl
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem cond_tt {A : Type} (t e : A) : cond tt t e = t :=
|
2014-08-26 05:54:44 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem ff_ne_tt : ¬ ff = tt :=
|
|
|
|
|
assume H : ff = tt, absurd
|
|
|
|
|
(calc true = cond tt true false : (cond_tt _ _)⁻¹
|
|
|
|
|
... = cond ff true false : {H⁻¹}
|
|
|
|
|
... = false : cond_ff _ _)
|
|
|
|
|
true_ne_false
|
2014-07-05 07:43:10 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem decidable_eq [instance] (a b : bool) : decidable (a = b) :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec
|
2014-09-04 23:36:06 +00:00
|
|
|
|
(rec (inl (eq.refl ff)) (inr ff_ne_tt) b)
|
2014-09-05 01:41:06 +00:00
|
|
|
|
(rec (inr (ne.symm ff_ne_tt)) (inl (eq.refl tt)) b)
|
2014-07-29 02:58:57 +00:00
|
|
|
|
a
|
2014-07-27 20:17:55 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
definition bor (a b : bool) :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec (rec ff tt b) tt a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_tt_left (a : bool) : bor tt a = tt :=
|
2014-08-26 05:54:44 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-22 23:36:47 +00:00
|
|
|
|
infixl `||` := bor
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_tt_right (a : bool) : a || tt = tt :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff || tt)) (eq.refl (tt || tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_ff_left (a : bool) : ff || a = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff || ff)) (eq.refl (ff || tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_ff_right (a : bool) : a || ff = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff || ff)) (eq.refl (tt || ff))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_id (a : bool) : a || a = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff || ff)) (eq.refl (tt || tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_comm (a b : bool) : a || b = b || a :=
|
2014-08-22 00:54:50 +00:00
|
|
|
|
cases_on a
|
2014-09-04 23:36:06 +00:00
|
|
|
|
(cases_on b (eq.refl (ff || ff)) (eq.refl (ff || tt)))
|
|
|
|
|
(cases_on b (eq.refl (tt || ff)) (eq.refl (tt || tt)))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_assoc (a b c : bool) : (a || b) || c = a || (b || c) :=
|
2014-08-22 00:54:50 +00:00
|
|
|
|
cases_on a
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(calc (ff || b) || c = b || c : {bor_ff_left b}
|
|
|
|
|
... = ff || (b || c) : bor_ff_left (b || c)⁻¹)
|
|
|
|
|
(calc (tt || b) || c = tt || c : {bor_tt_left b}
|
|
|
|
|
... = tt : bor_tt_left c
|
|
|
|
|
... = tt || (b || c) : bor_tt_left (b || c)⁻¹)
|
2014-07-27 19:50:57 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bor_to_or {a b : bool} : a || b = tt → a = tt ∨ b = tt :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(assume H : ff || b = tt,
|
|
|
|
|
have Hb : b = tt, from (bor_ff_left b) ▸ H,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.inr Hb)
|
|
|
|
|
(assume H, or.inl (eq.refl tt))
|
2014-07-29 02:58:57 +00:00
|
|
|
|
a
|
2014-07-27 20:17:55 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
definition band (a b : bool) :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec ff (rec ff tt b) a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-22 23:36:47 +00:00
|
|
|
|
infixl `&&` := band
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_ff_left (a : bool) : ff && a = ff :=
|
2014-08-26 05:54:44 +00:00
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_tt_left (a : bool) : tt && a = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (tt && ff)) (eq.refl (tt && tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_ff_right (a : bool) : a && ff = ff :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff && ff)) (eq.refl (tt && ff))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_tt_right (a : bool) : a && tt = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff && tt)) (eq.refl (tt && tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_id (a : bool) : a && a = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (ff && ff)) (eq.refl (tt && tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_comm (a b : bool) : a && b = b && a :=
|
2014-08-22 00:54:50 +00:00
|
|
|
|
cases_on a
|
2014-09-04 23:36:06 +00:00
|
|
|
|
(cases_on b (eq.refl (ff && ff)) (eq.refl (ff && tt)))
|
|
|
|
|
(cases_on b (eq.refl (tt && ff)) (eq.refl (tt && tt)))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_assoc (a b c : bool) : (a && b) && c = a && (b && c) :=
|
2014-08-22 00:54:50 +00:00
|
|
|
|
cases_on a
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(calc (ff && b) && c = ff && c : {band_ff_left b}
|
|
|
|
|
... = ff : band_ff_left c
|
|
|
|
|
... = ff && (b && c) : band_ff_left (b && c)⁻¹)
|
|
|
|
|
(calc (tt && b) && c = b && c : {band_tt_left b}
|
|
|
|
|
... = tt && (b && c) : band_tt_left (b && c)⁻¹)
|
2014-07-27 19:50:57 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_eq_tt_elim_left {a b : bool} (H : a && b = tt) : a = tt :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim (dichotomy a)
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(assume H0 : a = ff,
|
2014-08-28 01:34:09 +00:00
|
|
|
|
absurd
|
2014-07-29 02:58:57 +00:00
|
|
|
|
(calc ff = ff && b : (band_ff_left _)⁻¹
|
|
|
|
|
... = a && b : {H0⁻¹}
|
|
|
|
|
... = tt : H)
|
|
|
|
|
ff_ne_tt)
|
|
|
|
|
(assume H1 : a = tt, H1)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem band_eq_tt_elim_right {a b : bool} (H : a && b = tt) : b = tt :=
|
2014-09-05 01:41:06 +00:00
|
|
|
|
band_eq_tt_elim_left (eq.trans (band_comm b a) H)
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition bnot (a : bool) := rec tt ff a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-22 23:36:47 +00:00
|
|
|
|
notation `!` x:max := bnot x
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-29 02:58:57 +00:00
|
|
|
|
theorem bnot_bnot (a : bool) : !!a = a :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
cases_on a (eq.refl (!!ff)) (eq.refl (!!tt))
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-26 05:54:44 +00:00
|
|
|
|
theorem bnot_false : !ff = tt :=
|
|
|
|
|
rfl
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-08-26 05:54:44 +00:00
|
|
|
|
theorem bnot_true : !tt = ff :=
|
|
|
|
|
rfl
|
2014-08-01 00:48:51 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
end bool
|