X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Flib%2Fbool.ma;h=7b827baf913f99f6b091556ce883e9213ee0aa41;hb=6604a232815858a6c75dd25ac45abd68438077ff;hp=69bc66c87f7f8e1a579948f56812fa860ce72905;hpb=68b4f2490c12139c03760b39895619e63b0f38c9;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/lib/bool.ma b/matita/matita/contribs/lambdadelta/ground/lib/bool.ma index 69bc66c87..7b827baf9 100644 --- a/matita/matita/contribs/lambdadelta/ground/lib/bool.ma +++ b/matita/matita/contribs/lambdadelta/ground/lib/bool.ma @@ -17,43 +17,20 @@ include "ground/lib/relations.ma". include "ground/notation/functions/no_0.ma". include "ground/notation/functions/yes_0.ma". -(* BOOLEAN PROPERTIES *******************************************************) +(* BOOLEANS *****************************************************************) -interpretation "boolean false" 'no = false. +interpretation + "false (booleans)" + 'no = false. -interpretation "boolean true" 'yes = true. +interpretation + "true (booleans)" + 'yes = true. -(* Basic properties *********************************************************) +(* Advanced constructions ***************************************************) -lemma commutative_orb: commutative … orb. -* * // qed. - -lemma orb_true_dx: ∀b. (b ∨ Ⓣ) = Ⓣ. -* // qed. - -lemma orb_true_sn: ∀b. (Ⓣ ∨ b) = Ⓣ. -// qed. - -lemma commutative_andb: commutative … andb. -* * // qed. - -lemma andb_false_dx: ∀b. (b ∧ Ⓕ) = Ⓕ. -* // qed. - -lemma andb_false_sn: ∀b. (Ⓕ ∧ b) = Ⓕ. -// qed. - -lemma eq_bool_dec: ∀b1,b2:bool. Decidable (b1 = b2). +lemma eq_bool_dec (b1:bool) (b2:bool): + Decidable (b1 = b2). * * /2 width=1 by or_introl/ @or_intror #H destruct qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma orb_inv_false_dx: ∀b1,b2:bool. (b1 ∨ b2) = Ⓕ → b1 = Ⓕ ∧ b2 = Ⓕ. -* normalize /2 width=1 by conj/ #b2 #H destruct -qed-. - -lemma andb_inv_true_dx: ∀b1,b2:bool. (b1 ∧ b2) = Ⓣ → b1 = Ⓣ ∧ b2 = Ⓣ. -* normalize /2 width=1 by conj/ #b2 #H destruct -qed-.