]> matita.cs.unibo.it Git - helm.git/commitdiff
New proof based on an hand-made functional induction.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 21 Nov 2005 18:07:47 +0000 (18:07 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 21 Nov 2005 18:07:47 +0000 (18:07 +0000)
helm/matita/library/list/sort.ma

index ffd1e21872cd85c53c38fd3c9c518bc3a381a211..2716235d0914f0cd0a4da052a5d45b58cbc9583b 100644 (file)
@@ -110,126 +110,72 @@ 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 (A le H l x K).
- letin P ≝ (\lambda ll. ordered A le ll = true).
- fold simplify (P (insert A le x l)).
+ intros 5 (A le H l x).
  apply (
   let rec insert_ind (l: list A) \def
-  match l in list return λli. l = li → P (insert A le x li) with
-   [ nil ⇒ (? : l = [] → P [x])
+  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' → P (match b with 
-             [ true ⇒ x::he::l'
-             | false ⇒ he::(insert A le x l') ])
+        λ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' → P (x::he::l'))
+        [ true ⇒
+            (? :
+              le x he = true →
+               l = he::l' →
+                ordered ? le (he::l') = true →
+                 ordered ? le (x::he::l') = true)
         | false ⇒
-            (? : \forall lrec. P (insert A le x lrec) \to
-                  le x he = false → l = he::l' → P (he::(insert A le x l')))
-            l' (insert_ind l')
+           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 > H1;
-    apply andb_elim; simplify;
-    generalize in match K; clear K;
-    rewrite > H2; intro;
-    apply H3
-  | 
+  [ rewrite > insert_ind;
+    [ generalize in match (H ? ? H1); clear H1; intro;
+      generalize in match H3; clear H3;
+      elim l'; simplify;
+      [ rewrite > H4;
+        reflexivity
+      | elim (le x s); simplify;
+        [ rewrite > H4;
+          reflexivity
+        | simplify in H3;
+          rewrite > (andb_true_true ? ? H3);
+          reflexivity
+        ]
+      ]
+      
+    | apply (ordered_injective ? ? ? H3)
+    ]
+  | rewrite > H1;
+    rewrite > H3;
+    reflexivity
   | reflexivity
-  ]. 
-    
-    
-    
-
-  
-  
-  [ rewrite > H1; simplify;
-    generalize in match (ordered_injective A le l K);
-    rewrite > H2; simplify; intro; change with (ordered A le l' = true).
-    elim l'; simplify;
-     [ reflexivity 
-     | 
-          
-    
-  rewrite > H1; simplify.
-    elim l'; [ reflexivity | simplify; 
-  | simplify.
-  | reflexivity.
   ].
-*)
-
-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).
-  elim l;
-  [ 2: simplify;
-       apply (bool_elim ? (le x s));
-       [ intros;
-         simplify;
-         fold simplify (ordered ? le (s::l1));
-         apply andb_elim;
-         rewrite > H3;
-         assumption;
-       | change with (le x s = false → ordered ? le (s::insert A le x l1) = true);
-         generalize in match H2;
-         clear H1; clear H2;
-         generalize in match s;
-         clear s;
-         elim l1
-         [ simplify;
-           rewrite > (H x a H2);
-           reflexivity;
-         | simplify in \vdash (? ? (? ? ? %) ?);
-            change with (ordered A le (a::(insert A le x (s::l2))) = true);
-            simplify;
-            apply (bool_elim ? (le x s));
-            [ intros;
-              simplify;
-              fold simplify (ordered A le (s::l2));
-              apply andb_elim;
-              rewrite > (H x a H3);
-              simplify;
-              fold simplify (ordered A le (s::l2));
-              apply andb_elim;
-              rewrite > H4;
-              apply (ordered_injective A le (a::s::l2));
-              assumption;
-            | intros;
-              simplify;
-              fold simplify (ordered A le (s::(insert A le x l2)));
-              apply andb_elim;
-              simplify in H2;
-              fold simplify (ordered A le (s::l2)) in H2;
-              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
-              ]
-           ]
-         ]
-      ]
-   | simplify; reflexivity;
-   ]
 qed.
-
+  
 theorem insertionsort_sorted:
   ∀A:Set.
   ∀le:A → A → bool.∀eq:A → A → bool.