]> 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 2d60662a4d8f19e3d3d5894110782878a9927092..939cecedec6486a27274a4ce6ac834888f7c58ce 100644 (file)
@@ -49,6 +49,43 @@ let rec insert (A:Set) (le: A → A → bool) x (l: list A) on 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
   [ nil ⇒ []
@@ -56,12 +93,10 @@ let rec insertionsort (A:Set) (le: A → A → bool) (l: list A) on l ≝
       let l'' ≝ insertionsort A le l' in
        insert A le he 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;
@@ -74,7 +109,8 @@ lemma ordered_injective:
       [ generalize in match Hcut; 
         apply andb_elim;
         elim (le s s1);
-        [ change with (ordered A le (s1::l2) = true \to ordered A le (s1::l2) = true).
+        [ simplify;
+          fold simplify (ordered ? le (s1::l2));
           intros; assumption;
         | simplify;
           intros (Habsurd);
@@ -94,56 +130,32 @@ lemma insert_sorted:
   (\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).
-  elim l;
-  [ 2: change with (ordered A le (match le x s with
-         [ true ⇒ x::(s::l1) | false ⇒ s::(insert A le x l1) ]) = true).
-       apply (bool_elim ? (le x s));
-       [ intros;
-         change with ((le x s \land ordered A le (s::l1)) = true);
-         apply andb_elim;
-         rewrite > H3;
-         assumption;
-       | generalize in match H2;
-         clear H1; clear H2;
-         generalize in match s;
-         clear s;
-         elim l1
-         [ simplify;
-           rewrite > (H x a H2);
-           reflexivity;
-         | change with ((ordered A le (a::match le x s with
-              [ true ⇒ x::s::l2 | false ⇒ s::(insert A le x l2) ])) = true).
-            apply (bool_elim ? (le x s));
-            [ intros;
-              change with ((le a x \land (le x s \land ordered A le (s::l2))) = true).
-              apply andb_elim;
-              rewrite > (H x a H3);
-              change with ((le x s \land ordered A le (s::l2)) = true).
-              apply andb_elim;
-              rewrite > H4;
-              apply (ordered_injective A le (a::s::l2));
-              assumption;
-            | intros;
-              change with ((le a s \land ordered A le (s::(insert A le x l2))) = true);
-              apply andb_elim;
-              change in H2 with ((le a s \land ordered A le (s::l2)) = true); 
-              generalize in match H2;
-              apply (andb_elim (le a s));
-              elim (le a s);
-              [ change with (ordered A le (s::l2) = true \to ordered A le (s::insert A le x l2) = true);
-                intros;
-                apply (H1 s);
-                assumption;
-              | simplify; intros; assumption
-              ]
-           ]
-         ]
+ intros 5 (A le H l x).
+ 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 > H5;
+        reflexivity
+      | elim (le x s); simplify;
+        [ rewrite > H5;
+          reflexivity
+        | simplify in H4;
+          rewrite > (andb_true_true ? ? H4);
+          reflexivity
+        ]
       ]
-   | simplify; reflexivity;
-   ]
+    | apply (ordered_injective ? ? ? H4)
+    ]
+  | reflexivity
+  | rewrite > H2;
+    rewrite > H4;
+    reflexivity
+  ].
 qed.
-
+  
 theorem insertionsort_sorted:
   ∀A:Set.
   ∀le:A → A → bool.∀eq:A → A → bool.
@@ -154,25 +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.
-
-(*
-theorem insertionsort_correct:
-  ∀A:Set.
-  ∀le:A → A → bool.∀eq:A → A → bool.
-  (∀a,b:A. le a b = false → le b a = true) \to
-  ∀l,l':list A.
-  l' = insertionsort A le l
-  \to ordered A le l' = true
-    ∧ ((\forall x:A. mem A eq x l = true \to mem A eq x l' = true)
-        ∧ (\forall x:A. mem A eq x l = true \to mem A eq x l' = true)).
-  intros;
-  split;
-  [ rewrite > H1; 
-    apply (insertionsort_sorted A le eq H);
-  | split;
-    [ TO BE CONTINUED ...
-*)
+qed.
\ No newline at end of file