]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/list/sort.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / list / sort.ma
index 2716235d0914f0cd0a4da052a5d45b58cbc9583b..939cecedec6486a27274a4ce6ac834888f7c58ce 100644 (file)
@@ -48,26 +48,43 @@ let rec insert (A:Set) (le: A → A → bool) x (l: list A) on l ≝
        | false ⇒ he::(insert A le x l')
        ]
   ].
-(*
-theorem insert_ind:
- ∀A:Set. ∀le: A → A → bool. ∀x:A.
-  ∀P: list A → Prop.
-   ∀l:list A. P (insert A le x l).
- intros (A le x P H l).
- apply (
- let rec insert_ind (l: list A) ≝
-  match l in list return λl.P (insert A le x l) with
-   [ nil ⇒ (? : P [x])
-   | (cons he l') ⇒
-       match le x he return λb.P (match b with [ true ⇒ x::he::l' | false ⇒ he::(insert A le x l') ]) with
-        [ true ⇒ (H2 : P (x::he::l'))
-        | false ⇒ (? : P (he::(insert A le x l')))
-        ]
-   ]
- in
-  insert_ind l).
+
+lemma insert_ind :
+ ∀A:Set. ∀le: A → A → bool. ∀x.
+  ∀P:(list A → list A → Prop).
+   ∀H:(∀l: list A. l=[] → P [] [x]).
+    ∀H2:
+    (∀l: list A. ∀he. ∀l'. P l' (insert ? le x l') →
+      le x he = false → l=he::l' → P (he::l') (he::(insert ? le x l'))).
+     ∀H3:
+     (∀l: list A. ∀he. ∀l'. P l' (insert ? le x l') →
+       le x he = true → l=he::l' → P (he::l') (x::he::l')).
+    ∀l:list A. P l (insert ? le x l).
+  intros.
+  apply (
+    let rec insert_ind (l: list A) \def
+    match l in list
+    return
+      λli.
+       l = li → P li (insert ? le x li)
+    with
+     [ nil ⇒ H l
+     | (cons he l') ⇒
+         match le x he
+         return
+          λb. le x he = b → l = he::l' →
+           P (he::l')
+            (match b with 
+              [ true ⇒ x::he::l'
+              | false ⇒ he::(insert ? le x l') ])
+         with
+          [ true ⇒ H2 l he l' (insert_ind l')
+          | false ⇒ H1 l he l' (insert_ind l')
+          ]
+         (refl_eq ? (le x he))
+     ] (refl_eq ? l) in insert_ind l).
 qed.
-*)
+
 
 let rec insertionsort (A:Set) (le: A → A → bool) (l: list A) on l ≝
  match l with
@@ -78,10 +95,8 @@ let rec insertionsort (A:Set) (le: A → A → bool) (l: list A) on l ≝
   ].
 
 lemma ordered_injective:
-  ∀ A:Set. ∀ le:A → A → bool.
-  ∀ l:list A.
-  ordered A le l = true
-  \to ordered A le (tail A l) = true.
+  ∀A:Set. ∀le:A → A → bool.
+   ∀l:list A. ordered A le l = true → ordered A le (tail A l) = true.
   intros 3 (A le l).
   elim l
   [ simplify; reflexivity;
@@ -110,69 +125,34 @@ lemma ordered_injective:
   ].
 qed.
 
-
 lemma insert_sorted:
   \forall A:Set. \forall le:A\to A\to bool.
   (\forall a,b:A. le a b = false \to le b a = true) \to
   \forall l:list A. \forall x:A.
     ordered A le l = true \to ordered A le (insert A le x l) = true.
  intros 5 (A le H l x).
- apply (
-  let rec insert_ind (l: list A) \def
-  match l in list
-  return
-    λli.
-     l = li → ordered ? le li = true →
-      ordered ? le (insert A le x li) = true
-  with
-   [ nil ⇒ (? : l = [] → ordered ? le [] = true → ordered ? le [x] = true)
-   | (cons he l') ⇒
-       match le x he
-       return
-        λb. le x he = b → l = he::l' →
-         ordered ? le (he::l') = true → ordered ? le
-          (match b with 
-            [ true ⇒ x::he::l'
-            | false ⇒ he::(insert A le x l') ]) = true
-       with
-        [ true ⇒
-            (? :
-              le x he = true →
-               l = he::l' →
-                ordered ? le (he::l') = true →
-                 ordered ? le (x::he::l') = true)
-        | false ⇒
-           let res ≝ insert_ind l' in
-            (? :                  
-                  le x he = false → l = he::l' →
-                   ordered ? le (he::l') = true →
-                    ordered ? le (he::(insert ? le x l')) = true)           
-        ]
-       (refl_eq ? (le x he))
-   ] (refl_eq ? l) in insert_ind l);
-  
-  intros; simplify;
-  [ rewrite > insert_ind;
-    [ generalize in match (H ? ? H1); clear H1; intro;
-      generalize in match H3; clear H3;
+ apply (insert_ind ? ? ? (λl,il. ordered ? le l = true → ordered ? le il = true));
+ clear l; intros; simplify; intros;
+  [2: rewrite > H1;
+    [ generalize in match (H ? ? H2); clear H2; intro;
+      generalize in match H4; clear H4;
       elim l'; simplify;
-      [ rewrite > H4;
+      [ rewrite > H5;
         reflexivity
       | elim (le x s); simplify;
-        [ rewrite > H4;
+        [ rewrite > H5;
           reflexivity
-        | simplify in H3;
-          rewrite > (andb_true_true ? ? H3);
+        | simplify in H4;
+          rewrite > (andb_true_true ? ? H4);
           reflexivity
         ]
       ]
-      
-    | apply (ordered_injective ? ? ? H3)
+    | apply (ordered_injective ? ? ? H4)
     ]
-  | rewrite > H1;
-    rewrite > H3;
-    reflexivity
   | reflexivity
+  | rewrite > H2;
+    rewrite > H4;
+    reflexivity
   ].
 qed.
   
@@ -186,7 +166,7 @@ theorem insertionsort_sorted:
   elim l;
   [ simplify;
     reflexivity;
-  | apply (insert_sorted A le le_tot (insertionsort A le l1) s);
+  | apply (insert_sorted ? ? le_tot (insertionsort ? le l1) s);
     assumption;
   ]
 qed.
\ No newline at end of file