X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Fbool.ma;h=0d790efc54c590f0e5e8b4f47257d5b9ad99ad3f;hb=636c25914e83819c2f529edc891a7eb899499a97;hp=7b26064856545b510fabb55073aa2393bdd00d5a;hpb=844794906f5a9f97406d72b5305ce13fb75acf94;p=helm.git diff --git a/matita/matita/lib/basics/bool.ma b/matita/matita/lib/basics/bool.ma index 7b2606485..0d790efc5 100644 --- a/matita/matita/lib/basics/bool.ma +++ b/matita/matita/lib/basics/bool.ma @@ -37,6 +37,14 @@ theorem notb_notb: ∀b:bool. notb (notb b) = b. theorem injective_notb: injective bool bool notb. #b1 #b2 #H // qed. +theorem noteq_to_eqnot: ∀b1,b2. b1 ≠ b2 → b1 = notb b2. +* * // #H @False_ind /2/ +qed. + +theorem eqnot_to_noteq: ∀b1,b2. b1 = notb b2 → b1 ≠ b2. +* * normalize // #_ @(not_to_not … not_eq_true_false) // +qed. + definition andb : bool → bool → bool ≝ λb1,b2:bool. match b1 with [ true ⇒ b2 | false ⇒ false ]. @@ -46,6 +54,10 @@ theorem andb_elim: ∀ b1,b2:bool. ∀ P:bool → Prop. match b1 with [ true ⇒ P b2 | false ⇒ P false] → P (b1 ∧ b2). #b1 #b2 #P (elim b1) normalize // qed. +theorem true_to_andb_true: ∀b1,b2. b1 = true → b2 = true → (b1 ∧ b2) = true. +#b1 cases b1 normalize // +qed. + theorem andb_true_l: ∀ b1,b2. (b1 ∧ b2) = true → b1 = true. #b1 (cases b1) normalize // qed. @@ -76,11 +88,16 @@ lemma orb_true_l: ∀b1,b2:bool. (b1 ∨ b2) = true → (b1 = true) ∨ (b2 = true). * normalize /2/ qed. -definition if_then_else: ∀A:Type[0]. bool → A → A → A ≝ -λA.λb.λ P,Q:A. match b with [ true ⇒ P | false ⇒ Q]. +definition xorb : bool → bool → bool ≝ +λb1,b2:bool. + match b1 with + [ true ⇒ match b2 with [ true ⇒ false | false ⇒ true ] + | false ⇒ match b2 with [ true ⇒ true | false ⇒ false ]]. -notation "'if' term 19 e 'then' term 19 t 'else' term 19 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }. -interpretation "if_then_else" 'if_then_else e t f = (if_then_else ? e t f). +notation > "'if' term 46 e 'then' term 46 t 'else' term 46 f" non associative with precedence 46 + for @{ match $e in bool with [ true ⇒ $t | false ⇒ $f] }. +notation < "hvbox('if' \nbsp term 46 e \nbsp break 'then' \nbsp term 46 t \nbsp break 'else' \nbsp term 49 f \nbsp)" non associative with precedence 46 + for @{ match $e return $T with [ true ⇒ $t | false ⇒ $f] }. theorem bool_to_decidable_eq: ∀b1,b2:bool. decidable (b1=b2). @@ -91,22 +108,3 @@ theorem true_or_false: #b (cases b) /2/ qed. -(****** DeqSet: a set with a decidbale equality ******) - -record DeqSet : Type[1] ≝ { carr :> Type[0]; - eqb: carr → carr → bool; - eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y) -}. - -notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }. -interpretation "eqb" 'eqb a b = (eqb ? a b). - - -(****** EnumSet: a DeqSet with an enumeration function ****** - -record EnumSet : Type[1] ≝ { carr :> DeqSet; - enum: carr - eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y) -}. - -*) \ No newline at end of file