2014-07-02 15:36:05 +00: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-07-19 19:09:47 +00:00
|
|
|
|
import logic decidable
|
2014-07-25 05:49:12 +00:00
|
|
|
|
using eq_proofs
|
2014-07-19 19:09:47 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
namespace bool
|
|
|
|
|
inductive bool : Type :=
|
|
|
|
|
| b0 : bool
|
|
|
|
|
| b1 : bool
|
2014-07-19 19:08:52 +00:00
|
|
|
|
notation `'0`:max := b0
|
|
|
|
|
notation `'1`:max := b1
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem induction_on {p : bool → Prop} (b : bool) (H0 : p '0) (H1 : p '1) : p b
|
|
|
|
|
:= bool_rec H0 H1 b
|
2014-07-04 21:25:44 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem inhabited_bool [instance] : inhabited bool
|
2014-07-04 21:25:44 +00:00
|
|
|
|
:= inhabited_intro b0
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
definition cond {A : Type} (b : bool) (t e : A)
|
|
|
|
|
:= bool_rec e t b
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem dichotomy (b : bool) : b = '0 ∨ b = '1
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on b (or_intro_left _ (refl '0)) (or_intro_right _ (refl '1))
|
|
|
|
|
|
|
|
|
|
theorem cond_b0 {A : Type} (t e : A) : cond '0 t e = e
|
|
|
|
|
:= refl (cond '0 t e)
|
2014-07-05 05:22:26 +00:00
|
|
|
|
|
2014-07-19 19:08:52 +00:00
|
|
|
|
theorem cond_b1 {A : Type} (t e : A) : cond '1 t e = t
|
|
|
|
|
:= refl (cond '1 t e)
|
|
|
|
|
|
|
|
|
|
theorem b0_ne_b1 : ¬ '0 = '1
|
2014-07-25 05:14:15 +00:00
|
|
|
|
:= assume H : '0 = '1, absurd
|
2014-07-25 05:49:12 +00:00
|
|
|
|
(calc true = cond '1 true false : (cond_b1 _ _)⁻¹
|
|
|
|
|
... = cond '0 true false : {H⁻¹}
|
2014-07-19 19:08:52 +00:00
|
|
|
|
... = false : cond_b0 _ _)
|
2014-07-25 05:14:15 +00:00
|
|
|
|
true_ne_false
|
2014-07-05 07:43:10 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
definition bor (a b : bool)
|
|
|
|
|
:= bool_rec (bool_rec '0 '1 b) '1 a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_b1_left (a : bool) : bor '1 a = '1
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= refl (bor '1 a)
|
|
|
|
|
|
|
|
|
|
infixl `||`:65 := bor
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_b1_right (a : bool) : a || '1 = '1
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 || '1)) (refl ('1 || '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_b0_left (a : bool) : '0 || a = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 || '0)) (refl ('0 || '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_b0_right (a : bool) : a || '0 = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 || '0)) (refl ('1 || '0))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_id (a : bool) : a || a = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 || '0)) (refl ('1 || '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bor_swap (a b : bool) : a || b = b || a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a
|
|
|
|
|
(induction_on b (refl ('0 || '0)) (refl ('0 || '1)))
|
|
|
|
|
(induction_on b (refl ('1 || '0)) (refl ('1 || '1)))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
definition band (a b : bool)
|
|
|
|
|
:= bool_rec '0 (bool_rec '0 '1 b) a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
|
|
|
|
infixl `&&`:75 := band
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_b0_left (a : bool) : '0 && a = '0
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= refl ('0 && a)
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_b1_left (a : bool) : '1 && a = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('1 && '0)) (refl ('1 && '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_b0_right (a : bool) : a && '0 = '0
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 && '0)) (refl ('1 && '0))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_b1_right (a : bool) : a && '1 = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 && '1)) (refl ('1 && '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_id (a : bool) : a && a = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl ('0 && '0)) (refl ('1 && '1))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_swap (a b : bool) : a && b = b && a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a
|
|
|
|
|
(induction_on b (refl ('0 && '0)) (refl ('0 && '1)))
|
|
|
|
|
(induction_on b (refl ('1 && '0)) (refl ('1 && '1)))
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_eq_b1_elim_left {a b : bool} (H : a && b = '1) : a = '1
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= or_elim (dichotomy a)
|
|
|
|
|
(assume H0 : a = '0,
|
|
|
|
|
absurd_elim (a = '1)
|
2014-07-25 05:49:12 +00:00
|
|
|
|
(calc '0 = '0 && b : (band_b0_left _)⁻¹
|
|
|
|
|
... = a && b : {H0⁻¹}
|
2014-07-19 19:08:52 +00:00
|
|
|
|
... = '1 : H)
|
|
|
|
|
b0_ne_b1)
|
|
|
|
|
(assume H1 : a = '1, H1)
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem band_eq_b1_elim_right {a b : bool} (H : a && b = '1) : b = '1
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= band_eq_b1_elim_left (trans (band_swap b a) H)
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
definition bnot (a : bool) := bool_rec '1 '0 a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
|
|
|
|
|
prefix `!`:85 := bnot
|
|
|
|
|
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem bnot_bnot (a : bool) : !!a = a
|
2014-07-19 19:08:52 +00:00
|
|
|
|
:= induction_on a (refl (!!'0)) (refl (!!'1))
|
|
|
|
|
|
|
|
|
|
theorem bnot_false : !'0 = '1
|
|
|
|
|
:= refl _
|
|
|
|
|
|
|
|
|
|
theorem bnot_true : !'1 = '0
|
|
|
|
|
:= refl _
|
|
|
|
|
|
2014-07-19 19:09:47 +00:00
|
|
|
|
using decidable
|
2014-07-22 16:49:54 +00:00
|
|
|
|
theorem decidable_eq [instance] (a b : bool) : decidable (a = b)
|
|
|
|
|
:= bool_rec
|
|
|
|
|
(bool_rec (inl (refl '0)) (inr b0_ne_b1) b)
|
|
|
|
|
(bool_rec (inr (not_eq_symm b0_ne_b1)) (inl (refl '1)) b)
|
2014-07-19 19:09:47 +00:00
|
|
|
|
a
|
2014-07-02 15:36:05 +00:00
|
|
|
|
end
|