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/notation/constructors/no_0.ma".
17 include "ground_2/notation/constructors/yes_0.ma".
19 (* BOOLEAN PROPERTIES *******************************************************)
21 interpretation "boolean false" 'no = false.
23 interpretation "boolean true" 'yes = true.
25 lemma orb_false_r: ∀b1,b2:bool. (b1 ∨ b2) = false → b1 = false ∧ b2 = false.
26 * normalize /2 width=1 by conj/ #b2 #H destruct
29 lemma commutative_orb: commutative … orb.