1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basics/bool.ma".
16 include "ground_2/lib/relations.ma".
17 include "ground_2/notation/constructors/no_0.ma".
18 include "ground_2/notation/constructors/yes_0.ma".
20 (* BOOLEAN PROPERTIES *******************************************************)
22 interpretation "boolean false" 'no = false.
24 interpretation "boolean true" 'yes = true.
26 (* Basic properties *********************************************************)
28 lemma commutative_orb: commutative … orb.
31 lemma orb_true_dx: ∀b. (b ∨ Ⓣ) = Ⓣ.
34 lemma orb_true_sn: ∀b. (Ⓣ ∨ b) = Ⓣ.
37 lemma commutative_andb: commutative … andb.
40 lemma andb_false_dx: ∀b. (b ∧ Ⓕ) = Ⓕ.
43 lemma andb_false_sn: ∀b. (Ⓕ ∧ b) = Ⓕ.
46 lemma eq_bool_dec: ∀b1,b2:bool. Decidable (b1 = b2).
47 * * /2 width=1 by or_introl/
48 @or_intror #H destruct
51 (* Basic inversion lemmas ***************************************************)
53 lemma orb_inv_false_dx: ∀b1,b2:bool. (b1 ∨ b2) = Ⓕ → b1 = Ⓕ ∧ b2 = Ⓕ.
54 * normalize /2 width=1 by conj/ #b2 #H destruct
57 lemma andb_inv_true_dx: ∀b1,b2:bool. (b1 ∧ b2) = Ⓣ → b1 = Ⓣ ∧ b2 = Ⓣ.
58 * normalize /2 width=1 by conj/ #b2 #H destruct