]> matita.cs.unibo.it Git - helm.git/blobdiff - weblib/tutorial/chapter5.ma
commit by user andrea
[helm.git] / weblib / tutorial / chapter5.ma
index bc01f3c7021d350d8a490f3e5fb99f1aabbbfe32..8d4fc9f6c7f489c30fa468516fac2fba50c0e930 100644 (file)
@@ -4,7 +4,7 @@ effectively searching an element of that set inside a data structure. In this
 section we shall define several boolean functions acting on lists of elements in 
 a DeqSet, and prove some of their properties.*)
 
-include "basics/list.ma".
+include "basics/list.ma". 
 include "tutorial/chapter4.ma".
 
 (* The first function we define is an effective version of the membership relation,
@@ -14,8 +14,8 @@ l.*)
 let rec memb (S:DeqSet) (x:S) (l: list\ 5span class="error" title="Parse error: RPAREN expected after [term] (in [arg])"\ 6\ 5/span\ 6 S) on l  ≝
   match l with
   [ nil ⇒ false
-  | cons a tl ⇒ (x =\ 5span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"\ 6\ 5/span\ 6= a) ∨ memb S x tl
-  ].
+  | cons a tl ⇒ (x == a) ∨ memb S x tl
+  ]\ 5span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"\ 6\ 5/span\ 6\ 5span class="error" title="No choices for ID nil"\ 6\ 5/span\ 6.
 
 notation < "\memb x l" non associative with precedence 90 for @{'memb $x $l}.
 interpretation "boolean membership" 'memb a l = (memb ? a l).
@@ -40,7 +40,7 @@ lemma memb_hd: ∀S,a,l. memb S a (a::l) = true.
 qed.
 
 lemma memb_cons: ∀S,a,b,l. 
-  memb S a l = true → memb\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6\ 5/a\ 6 S a (b::l) = true.
+  memb S a l = true → memb\ 5span class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"\ 6\ 5/span\ 6 S a (b::l) = true.
 #S #a #b #l normalize cases (a==b) normalize // 
 qed.
 
@@ -50,13 +50,13 @@ lemma memb_single: ∀S,a,x. memb S a (x::[]) = true → a = x.
 qed.
 
 lemma memb_append: ∀S,a,l1,l2. 
-memb S a (l1@\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6l2) = true → memb S a l1= true ∨ memb S a l2 = true.
-#S #a #l1 elim l1 normalize [#l2 #H %2 //] 
+memb S a (l1@\ 5span class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"\ 6\ 5/span\ 6l2) = true → memb S a l1= true ∨ memb S a l2 = true.
+#S #a #l1\ 5span class="error" title="Parse error: illegal begin of statement"\ 6\ 5/span\ 6\ 5span class="error" title="Parse error: illegal begin of statement"\ 6\ 5/span\ 6 elim l1 normalize [#l2 #H %2 //] 
 #b #tl #Hind #l2 cases (a==b) normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace orb_true_l\ 5/span\ 6\ 5/span\ 6
 qed. 
 
 lemma memb_append_l1: ∀S,a,l1,l2. 
- memb S a l1= true → memb S a (l1@\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6l2) = true.
+ memb S a l1= true → memb S a (l1@l2) = true.
 #S #a #l1 elim l1 normalize
   [normalize #le #abs @False_ind /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace absurd\ 5/span\ 6\ 5/span\ 6/
   |#b #tl #Hind #l2 cases (a==b) normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6
@@ -69,7 +69,7 @@ lemma memb_append_l2: ∀S,a,l1,l2.
 #b #tl #Hind #l2 cases (a==b) normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6
 qed. 
 
-lemma memb_exists: ∀S,a,l.memb S a l = true\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6\ 5/a\ 6 → ∃l1,l2.l=l1@(a::l2).
+lemma memb_exists: ∀S,a,l.memb S a l = true\ 5span class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"\ 6\ 5/span\ 6 → ∃l1,l2.l=l1@(a::l2).
 #S #a #l elim l [normalize #abs @False_ind /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace absurd\ 5/span\ 6\ 5/span\ 6/]
 #b #tl #Hind #H cases (orb_true_l … H)
   [#eqba @(ex_intro … (nil S)) @(ex_intro … tl) >(\P eqba) //
@@ -79,13 +79,13 @@ lemma memb_exists: ∀S,a,l.memb S a l = true\ 5a href="cic:/matita/basics/bool/bo
 qed.
 
 lemma not_memb_to_not_eq: ∀S,a,b,l. 
- memb S a l = false\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6\ 5/a\ 6 → memb S b l = true → a==b = false.
+ memb S a l = false → memb S b l = true → a==b = false.
 #S #a #b #l cases (true_or_false (a==b)) // 
 #eqab >(\P eqab) #H >H #abs @False_ind /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace absurd\ 5/span\ 6\ 5/span\ 6/
 qed. 
  
 lemma memb_map: ∀S1,S2,f,a,l. memb S1 a l= true → 
-  memb S2 (f a) (map … f l) =\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 true.
+  memb S2 (f a) (map … f l) = true.
 #S1 #S2 #f #a #l elim l normalize [//]
 #x #tl #memba cases (true_or_false (a==x))
   [#eqx >eqx >(\P eqx) >(\b (refl … (f x))) normalize //
@@ -97,7 +97,7 @@ lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2.
   memb S1 a1 l1 = true → memb S2 a2 l2 = true →
   memb S3 (op a1 a2) (compose S1 S2 S3 op l1 l2) = true.
 #S1 #S2 #S3 #op #a1 #a2 #l1 elim l1 [normalize //]
-#x #tl #Hind #l2 #memba1 #memba2 cases (orb_true_l\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6\ 5/a\ 6 … memba1)
+#x #tl #Hind #l2 #memba1 #memba2 cases (orb_true_l … memba1)
   [#eqa1 >(\P eqa1) @memb_append_l1 @memb_map // 
   |#membtl @memb_append_l2 @Hind //
   ]
@@ -125,13 +125,13 @@ let rec unique_append (S:DeqSet) (l1,l2: list S) on l1 ≝
   ].
 
 axiom unique_append_elim: ∀S:DeqSet.∀P: S → Prop.∀l1,l2. 
-(∀x. memb S x l1 =\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 true → P x) → (∀x. memb S x l2 = true → P x) →
+(∀x. memb S x l1 =\ 5span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"\ 6\ 5/span\ 6 true → P x) → (∀x. memb S x l2 = true → P x) →
 ∀x. memb S x (unique_append S l1 l2) = true → P x. 
 
 lemma unique_append_unique: ∀S,l1,l2. uniqueb S l2 = true →
   uniqueb S (unique_append S l1 l2) = true.
 #S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2
-cases (true_or_false\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6\ 5/a\ 6 … (memb S a (unique_append S tl l2))) 
+cases (true_or_false … (memb S a (unique_append S tl l2))) 
 #H >H normalize [@Hind //] >H normalize @Hind //
 qed.
 
@@ -143,7 +143,7 @@ lemma sublist_length: ∀S,l1,l2.
  uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|.
 #S #l1 elim l1 // 
 #a #tl #Hind #l2 #unique #sub
-cut (∃\ 5a title="exists" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //]
+cut (∃l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //]
 * #l3 * #l4 #eql2 >eql2 >length_append normalize 
 applyS le_S_S <length_append @Hind [@(andb_true_r … unique)]
 >eql2 in sub; #sub #x #membx 
@@ -151,7 +151,7 @@ cases (memb_append … (sub x (orb_true_r2 … membx)))
   [#membxl3 @memb_append_l1 //
   |#membxal4 cases (orb_true_l … membxal4)
     [#eqxa @False_ind lapply (andb_true_l … unique)
-     <(\P eqxa) >membx normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace absurd\ 5/span\ 6\ 5/span\ 6/ |#membxl4 @memb_append_l2\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6\ 5/a\ 6 //
+     <(\P eqxa) >membx normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace absurd\ 5/span\ 6\ 5/span\ 6/ |#membxl4 @memb_append_l2 //
     ]
   ]
 qed.
@@ -171,7 +171,7 @@ qed.
 lemma sublist_unique_append_l2: 
   ∀S,l1,l2. sublist S l2 (unique_append S l1 l2).
 #S #l1 elim l1 [normalize //] #x #tl #Hind normalize 
-#l2 #a cases (memb S x (unique_append\ 5a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"\ 6\ 5/a\ 6 S tl l2)) normalize
+#l2 #a cases (memb S x (unique_append S tl l2)) normalize
 [@Hind | cases (a==x) normalize // @Hind]
 qed.
 
@@ -185,7 +185,7 @@ lemma decidable_sublist:∀S,l1,l2.
         [#eqax >(\P eqax) // |@subtl]
       |%2 @(not_to_not … (eqnot_to_noteq … true memba)) #H1 @H1 @memb_hd
       ]
-    |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons\ 5a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"\ 6\ 5/a\ 6 //
+    |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons //
     ] 
   ]
 qed.
@@ -202,7 +202,7 @@ cases (true_or_false (a==b)) #eqab
 qed. 
   
 lemma memb_filter_memb: ∀S,f,a,l. 
-  memb S a (filter S f l) = true → memb\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6\ 5/a\ 6 S a l = true.
+  memb S a (filter S f l) = true → memb S a l = true.
 #S #f #a #l elim l [normalize //]
 #b #tl #Hind normalize (cases (f b)) normalize 
 cases (a==b) normalize // @Hind
@@ -210,7 +210,7 @@ qed.
   
 lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true → 
 memb S x l = true ∧ (f x = true).
-/\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace conj, filter_true, memb_filter_memb\ 5a href="cic:/matita/tutorial/chapter5/memb_filter_memb.def(5)"\ 6\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
+/\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace conj, memb_filter_memb, filter_true\ 5span class="error" title="disambiguation error"\ 6\ 5/span\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
 lemma memb_filter_l: ∀S,f,x,l. (f x = true) → memb S x l = true →
 memb S x (filter ? f l) = true.
@@ -223,7 +223,7 @@ qed.
 
 (********************* exists *****************)
 
-let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool\ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6\ 5/a\ 6 ≝
+let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool ≝
 match l with
 [ nil ⇒ false
 | cons h t ⇒ orb (p h) (exists A p t)