]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/bool.ma
Splitted DeqSets in their own file. Notation for memb to hide the type.
[helm.git] / matita / matita / lib / basics / bool.ma
index 7e874f41264897f1abf497b754ad5c8ae18fd8c5..43a3369a393d0c9e4725610beac64ceb868b6e6a 100644 (file)
@@ -82,9 +82,9 @@ definition xorb : bool → bool → bool ≝
   [ 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
+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 19 e \nbsp break 'then' \nbsp term 19 t \nbsp break 'else' \nbsp term 48 f \nbsp)" non associative with precedence 19
+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 with [ true ⇒ $t | false ⇒ $f]  }.
 
 theorem bool_to_decidable_eq: 
@@ -95,23 +95,3 @@ theorem true_or_false:
 ∀b:bool. b = true ∨ b = 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)
-}.
-
-*)