]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 28 Feb 2012 09:29:28 +0000 (09:29 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 28 Feb 2012 09:29:28 +0000 (09:29 +0000)
weblib/tutorial/chapter4.ma

index 1be877197f9cff9659a5e5573d508e9af1e2ef2d..d2dc7c84338e11e6490b4e9a6b25bb87e351f7c4 100644 (file)
@@ -1,7 +1,8 @@
 (* In this Chapter we shall develop a naif theory of sets represented as characteristic
 predicates over some universe \ 5code\ 6A\ 5/code\ 6, that is as objects of type A→Prop. *)
 
-include "basics/logic.ma".
+include "basics/types.ma".
+include "basics/bool.ma".
 
 (**** For instance the empty set is defined by the always function predicate *)
 
@@ -86,9 +87,9 @@ lemma eqP_substract_l: ∀U.∀A,B,C:U →Prop.
   B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61 C  → A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 B \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61 A \ 5a title="substraction" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6 C.
 #U #A #B #C #eqBC #a @\ 5a href="cic:/matita/basics/logic/iff_and_l.def(2)"\ 6iff_and_l\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/iff_not.def(4)"\ 6iff_not\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
-(* We can now prove several properties of the previous set-theoretic 
-operations. In particular, union is commutative and associative, and 
-the empty set is an identity element: *)
+(* We can now prove several properties of the previous set-theoretic operations. 
+In particular, union is commutative and associative, and the empty set is an 
+identity element: *)
 
 lemma union_empty_r: ∀U.∀A:U→Prop. 
   A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="empty set" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61 A.
@@ -104,7 +105,7 @@ lemma union_assoc: ∀U.∀A,B,C:U → Prop.
 #S #A #B #C #w % [* [* /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] | * [/\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ | * /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/Or.con(0,1,2)"\ 6or_introl\ 5/a\ 6\ 5a href="cic:/matita/basics/logic/Or.con(0,2,2)"\ 6or_intror\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/]
 qed.   
 
-(* In the same way we prove commutativity and associativity for set
+(* In the same way we prove commutativity and associativity for set 
 interesection *)
 
 lemma cap_comm : ∀U.∀A,B:U →Prop. 
@@ -126,9 +127,8 @@ lemma cap_idemp: ∀U.∀A:U →Prop.
   A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 A \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61 A.
 #U #A #a % [* // | /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] qed. 
 
-(* We conclude our examples with a couple of distributivity theorems,
-and a characterization of substraction in terms of interesection and
-complementation. *)
+(* We conclude our examples with a couple of distributivity theorems, and a 
+characterization of substraction in terms of interesection and complementation. *)
 
 lemma distribute_intersect : ∀U.∀A,B,C:U→Prop. 
   (A \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 B) \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61 (A \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C) \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (B \ 5a title="intersection" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 C).
@@ -144,60 +144,108 @@ lemma substract_def:∀U.∀A,B:U→Prop. A\ 5a title="substraction" href="cic:/fa
 #U #A #B #w normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
 qed.
 
-(****** DeqSet: a set with a decidbale equality ******)
+(* In several situation it is important to assume to have a decidable equality 
+between elements of a set U, namely a boolean function eqb: U→U→bool such that
+for any pair of elements a and b in U, (eqb x y) is true if and only if x=y. 
+A set equipped with such an equality is called a DeqSet: *)
 
 record DeqSet : Type[1] ≝ { carr :> Type[0];
-   eqb: carr → carr → bool;
-   eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y)
+   eqb: carr → carr → \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6;
+   eqb_true: ∀x,y. (eqb x y \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (x \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 y)
 }.
 
+(* We use the notation == to denote the decidable equality, to distinguish it
+from the propositional equality. In particular, a term of the form a==b is a 
+boolean, while a=b is a proposition. *)
+
 notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }.
 interpretation "eqb" 'eqb a b = (eqb ? a b).
 
+(* It is convenient to have a simple way to reflect a proof of the fact 
+that (eqb a b) is true into a proof of the proposition (a = b); to this aim, 
+we introduce two operators "\P" and "\b". *)
+
 notation "\P H" non associative with precedence 90 
   for @{(proj1 … (eqb_true ???) $H)}. 
 
 notation "\b H" non associative with precedence 90 
   for @{(proj2 … (eqb_true ???) $H)}. 
   
-lemma eqb_false: ∀S:DeqSet.∀a,b:S. 
-  (eqb ? a b) = false ↔ a ≠ b.
+(* If H:eqb a b = true, then \P H: a = b, and conversely if h:a = b, then
+\b h: eqb a b = true. Let us see an example of their use: the following 
+statement asserts that we can reflect a proof that eqb a b is false into
+a proof of the proposition a ≠ b. *)
+
+lemma eqb_false: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀a,b:S. 
+  (\ 5a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 \ 5a title="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 a \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b.
+
+(* We start the proof introducing the hypothesis, and then split the "if" and
+"only if" cases *)
 #S #a #b % #H 
-  [@(not_to_not … not_eq_true_false) #H1 <H @sym_eq @(\b H1)
-  |cases (true_or_false (eqb ? a b)) // #H1 @False_ind @(absurd … (\P H1) H)
+
+(* The latter is easily reduced to prove the goal true=false under the assumption
+H1: a = b *)
+  [@(\ 5a href="cic:/matita/basics/logic/not_to_not.def(3)"\ 6not_to_not\ 5/a\ 6 … \ 5a href="cic:/matita/basics/bool/not_eq_true_false.def(3)"\ 6not_eq_true_false\ 5/a\ 6) #H1 
+  
+(* since by assumption H false is equal to (a==b), by rewriting we obtain the goal 
+true=(a==b) that is just the boolean version of H1 *) 
+
+  <H @\ 5a href="cic:/matita/basics/logic/sym_eq.def(2)"\ 6sym_eq\ 5/a\ 6 @(\b H1)
+
+(* In the "if" case, we proceed by cases over the boolean equality (a==b); if 
+(a==b) is false, the goal is trivial; the other case is absurd, since if (a==b) is
+true, then by reflection a=b, while by hypothesis a≠b *)
+  
+ |cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b)) // #H1 @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 @(\ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6 … (\P H1) H)
   ]
 qed.
  
+(* We also introduce two operators "\Pf" and "\bf" to reflect a proof
+of (a==b)=false into a proof of a≠b, and vice-versa *) 
+
 notation "\Pf H" non associative with precedence 90 
   for @{(proj1 … (eqb_false ???) $H)}. 
 
 notation "\bf H" non associative with precedence 90 
   for @{(proj2 … (eqb_false ???) $H)}. 
-  
-lemma dec_eq: ∀S:DeqSet.∀a,b:S. a = b ∨ a ≠ b.
-#S #a #b cases (true_or_false (eqb ? a b)) #H
+
+(* The following statement proves that propositional equality in a 
+DeqSet is decidable in the traditional sense, namely either a=b or a≠b *)
+
+ lemma dec_eq: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀a,b:S. a \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 b \ 5a title="logical or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 a \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b.
+#S #a #b cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"\ 6eqb\ 5/a\ 6 ? a b)) #H
   [%1 @(\P H) | %2 @(\Pf H)]
 qed.
 
+(* A simple example of a set with a decidable equality is bool. We first define 
+the boolean equality beqb, that is just the xand function, then prove that 
+beqb b1 b2 is true if and only if b1=b2, and finally build the type DeqBool by 
+instantiating the DeqSet record with the previous information *)
+
 definition beqb ≝ λb1,b2.
-  match b1 with [ true ⇒ b2 | false ⇒ notb b2].
+  match b1 with [ true ⇒ b2 | false ⇒ \ 5a href="cic:/matita/basics/bool/notb.def(1)"\ 6notb\ 5/a\ 6 b2].
 
 notation < "a == b" non associative with precedence 45 for @{beqb $a $b }.
-lemma beqb_true: ∀b1,b2. iff (beqb b1 b2 = true) (b1 = b2).
-#b1 #b2 cases b1 cases b2 normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace conj\ 5/span\ 6\ 5/span\ 6/
+
+lemma beqb_true: ∀b1,b2. \ 5a href="cic:/matita/basics/logic/iff.def(1)"\ 6iff\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 b1 b2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6) (b1 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 b2).
+#b1 #b2 cases b1 cases b2 normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
 qed. 
 
-definition DeqBool ≝ mk_DeqSet bool beqb beqb_true.
+definition DeqBool ≝ \ 5a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter4/beqb_true.def(4)"\ 6beqb_true\ 5/a\ 6.
 
-unification hint  0 ≔ ; 
-    X ≟ mk_DeqSet bool beqb beqb_true
+(* At this point, we would expect to be able to prove things like the
+following: for any boolean b, if b==false is true then b=false. *)
+
+(* unification hint  0 \ 5a href="cic:/fakeuri.def(1)" title="hint_decl_Type1"\ 6\ 5/a\ 6 ; 
+    X ≟ \ 5a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter4/beqb_true.def(4)"\ 6beqb_true\ 5/a\ 6
 (* ---------------------------------------- *) ⊢ 
-    bool ≡ carr X.
+    \ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 ≡ \ 5a href="cic:/matita/tutorial/chapter4/carr.fix(0,0,2)"\ 6carr\ 5/a\ 6 X.
     
 unification hint  0 ≔ b1,b2:bool; 
     X ≟ mk_DeqSet bool beqb beqb_true
 (* ---------------------------------------- *) ⊢ 
-    beqb b1 b2 ≡ eqb X b1 b2.
+    beqb b1 b2 ≡ eqb X b1 b2. *)
     
 example exhint: ∀b:bool. (b == false) = true → b = false. 
 #b #H @(\P H).