]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/minimization.ma
Towards chebyshev.
[helm.git] / matita / library / nat / minimization.ma
index 0abed5ad354319674d89e049ca833f804cdaaeec..22d5030ddb37a12d032f66b489cd87aa7bc01780 100644 (file)
@@ -72,10 +72,75 @@ intro.simplify.rewrite < H3.
 rewrite > H2.simplify.apply le_n.
 qed.
 
+theorem max_f_g: \forall f,g,n. (\forall i. i \le n \to f i = g i) \to
+max n f = max n g.
+intros 3.
+elim n
+  [simplify.
+   rewrite > (H O)
+    [reflexivity
+    |apply le_n
+    ]
+  |simplify.
+   rewrite > H
+    [rewrite > H1
+      [reflexivity
+      |apply le_n
+      ]
+    |intros.
+     apply H1.
+     apply le_S.
+     assumption
+    ]
+  ]
+qed.
+
+theorem le_max_f_max_g: \forall f,g,n. (\forall i. i \le n \to f i = true \to g i =true) \to
+max n f \le max n g.
+intros 3.
+elim n
+  [simplify.
+   elim (f O);apply le_O_n
+  |simplify.
+   apply (bool_elim ? (f (S n1)));intro
+    [rewrite > (H1 (S n1) ? H2)
+      [apply le_n
+      |apply le_n
+      ]
+    |cases (g(S n1))
+      [simplify.
+       apply le_S.
+       apply le_max_n
+      |simplify.
+       apply H.
+       intros.
+       apply H1
+        [apply le_S.assumption
+        |assumption
+        ]
+      ]
+    ]
+  ]
+qed.
 
-definition max_spec \def \lambda f:nat \to bool.\lambda n: nat.
-\exists i. (le i n) \land (f i = true) \to
-(f n) = true \land (\forall i. i < n \to (f i = false)).
+
+theorem max_O : \forall f:nat \to bool. \forall n:nat.
+(\forall i:nat. le i n \to f i = false) \to max n f = O.
+intros 2.elim n
+  [simplify.rewrite > H
+    [reflexivity
+    |apply le_O_n
+    ]
+  |simplify.rewrite > H1
+    [simplify.apply H.
+     intros.
+     apply H1.
+     apply le_S.
+     assumption
+    |apply le_n
+    ]
+  ]
+qed.
 
 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
 (\exists i:nat. le i n \land f i = true) \to f (max n f) = true. 
@@ -100,6 +165,73 @@ rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
 reflexivity.
 qed.
 
+theorem exists_forall_le:\forall f,n. 
+(\exists i. i \le n \land f i = true) \lor
+(\forall i. i \le n \to f i = false).
+intros.
+elim n
+  [apply (bool_elim ? (f O));intro
+    [left.apply (ex_intro ? ? O).
+     split[apply le_n|assumption]
+    |right.intros.
+     apply (le_n_O_elim ? H1).
+     assumption
+    ]
+  |elim H
+    [elim H1.elim H2.
+     left.apply (ex_intro ? ? a).
+     split[apply le_S.assumption|assumption]
+    |apply (bool_elim ? (f (S n1)));intro
+      [left.apply (ex_intro ? ? (S n1)).
+       split[apply le_n|assumption]
+      |right.intros.
+       elim (le_to_or_lt_eq ? ? H3)
+        [apply H1.
+         apply le_S_S_to_le.
+         apply H4
+        |rewrite > H4.
+         assumption
+        ]
+      ]
+    ]
+  ]
+qed.
+     
+theorem exists_max_forall_false:\forall f,n. 
+((\exists i. i \le n \land f i = true) \land (f (max n f) = true))\lor
+((\forall i. i \le n \to f i = false) \land (max n f) = O).
+intros.
+elim (exists_forall_le f n)
+  [left.split
+    [assumption
+    |apply f_max_true.assumption
+    ]
+  |right.split
+    [assumption
+    |apply max_O.assumption
+    ]
+  ]
+qed.
+
+theorem false_to_lt_max: \forall f,n,m.O < n \to
+f n = false \to max m f \le n \to max m f < n.
+intros.
+elim (le_to_or_lt_eq ? ? H2)
+  [assumption
+  |elim (exists_max_forall_false f m)
+    [elim H4.
+     apply False_ind.
+     apply not_eq_true_false.
+     rewrite < H6.
+     rewrite > H3.
+     assumption
+    |elim H4.
+     rewrite > H6.
+     assumption
+    ]
+  ]
+qed.
+
 theorem lt_max_to_false : \forall f:nat \to bool. 
 \forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
 intros 2.
@@ -119,20 +251,78 @@ apply le_S_S_to_le.assumption.
 intro.rewrite > H7.assumption.
 qed.
 
+theorem f_false_to_le_max: \forall f,n,p. (∃i:nat.i≤n∧f i=true) \to
+(\forall m. p < m \to f m = false)
+\to max n f \le p.
+intros.
+apply not_lt_to_le.intro.
+apply not_eq_true_false.
+rewrite < (H1 ? H2).
+apply sym_eq.
+apply f_max_true.
+assumption.
+qed.
+
+definition max_spec \def \lambda f:nat \to bool.\lambda n,m: nat.
+m \le n \land (f m)=true \land (\forall i. m < i \to i \le n \to (f i = false)).
+
+theorem max_spec_to_max: \forall f:nat \to bool.\forall n,m:nat.
+max_spec f n m \to max n f = m.
+intros 2.
+elim n
+  [elim H.elim H1.apply (le_n_O_elim ? H3).
+   apply max_O_f
+  |elim H1.
+   elim (max_S_max f n1)
+    [elim H4.
+     rewrite > H6.
+     apply le_to_le_to_eq
+      [apply not_lt_to_le.
+       unfold Not.intro.
+       apply not_eq_true_false.
+       rewrite < H5.
+       apply H3
+        [assumption|apply le_n]
+      |elim H2.assumption
+      ]
+    |elim H4.
+     rewrite > H6.
+     apply H.
+     elim H2.
+     split
+      [split
+        [elim (le_to_or_lt_eq ? ? H7)
+          [apply le_S_S_to_le.assumption
+          |apply False_ind.
+           apply not_eq_true_false.
+           rewrite < H8.
+           rewrite > H9.
+           assumption
+          ]
+        |assumption
+        ]
+      |intros.
+       apply H3
+        [assumption|apply le_S.assumption]
+      ]
+    ]
+  ]
+qed.
 let rec min_aux off n f \def
-  match f (n-off) with 
-  [ true \Rightarrow (n-off)
+  match f n with 
+  [ true \Rightarrow n
   | false \Rightarrow 
       match off with
       [ O \Rightarrow n
-      | (S p) \Rightarrow min_aux p n f]].
+      | (S p) \Rightarrow min_aux p (S n) f]].
 
 definition min : nat \to (nat \to bool) \to nat \def
-\lambda n.\lambda f. min_aux n n f.
+\lambda n.\lambda f. min_aux n O f.
 
 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
 min_aux O i f = i.
-intros.simplify.rewrite < minus_n_O.
+intros.simplify.
 elim (f i).reflexivity.
 simplify.reflexivity.
 qed.
@@ -143,34 +333,35 @@ intro.apply (min_aux_O_f f O).
 qed.
 
 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
-(f (n -(S i)) = true \land min_aux (S i) n f = (n - (S i))) \lor 
-(f (n -(S i)) = false \land min_aux (S i) n f = min_aux i n f).
-intros.simplify.elim (f (n - (S i))).
+((f n) = true \land min_aux (S i) n f = n) \lor 
+((f n) = false \land min_aux (S i) n f = min_aux i (S n) f).
+intros.simplify.elim (f n).
 simplify.left.split.reflexivity.reflexivity.
 simplify.right.split.reflexivity.reflexivity.
 qed.
 
 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
-(\exists i. le (m-off) i \land le i m \land f i = true) \to
+(\exists i. le m i \land le i (off + m) \land f i = true) \to
 f (min_aux off m f) = true. 
 intros 2.
 elim off.elim H.elim H1.elim H2.
 cut (a = m).
 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
-apply (antisym_le a m).assumption.rewrite > (minus_n_O m).assumption.
+apply (antisym_le a m).assumption.assumption.
 simplify.
 apply (bool_ind (\lambda b:bool.
-(f (m-(S n)) = b) \to (f (match b in bool with
-[ true \Rightarrow m-(S n)
-| false  \Rightarrow (min_aux n m f)])) = true)).
-simplify.intro.assumption.
-simplify.intro.apply H.
+(f m = b) \to (f (match b in bool with
+[ true \Rightarrow m
+| false  \Rightarrow (min_aux n (S m) f)])) = true)).
+intro; assumption.
+intro. apply H.
 elim H1.elim H3.elim H4.
-elim (le_to_or_lt_eq (m-(S n)) a H6). 
+elim (le_to_or_lt_eq ? a H6). 
 apply (ex_intro nat ? a).
 split.split.
-apply lt_minus_S_n_to_le_minus_n.assumption.
-assumption.assumption.
+assumption.
+rewrite < plus_n_Sm; assumption.
+assumption.
 absurd (f a = false).rewrite < H8.assumption.
 rewrite > H5.
 apply not_eq_true_false.
@@ -178,45 +369,67 @@ reflexivity.
 qed.
 
 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
-\forall n,off,m:nat. (n-off) \leq m \to m < (min_aux off n f) \to f m = false.
+\forall n,off,m:nat. n \leq m \to m < (min_aux off n f) \to f m = false.
 intros 3.
-elim off.absurd (le n m).rewrite > minus_n_O.assumption.
-apply lt_to_not_le.rewrite < (min_aux_O_f f n).assumption.
-generalize in match H1.
-elim (min_aux_S f n1 n).
-elim H3.
-absurd (n - S n1 \le m).assumption.
-apply lt_to_not_le.rewrite < H6.assumption.
-elim H3.
-elim (le_to_or_lt_eq (n -(S n1)) m).
-apply H.apply lt_minus_S_n_to_le_minus_n.assumption.
-rewrite < H6.assumption.
-rewrite < H7.assumption.
-assumption.
+generalize in match n; clear n.
+elim off.absurd (le n1 m).assumption.
+apply lt_to_not_le.rewrite < (min_aux_O_f f n1).assumption.
+elim (le_to_or_lt_eq ? ? H1);
+ [ unfold lt in H3;
+   apply (H (S n1));
+   [ assumption
+   | elim (min_aux_S f n n1);
+     [ elim H4;
+       elim (not_le_Sn_n n1);
+       rewrite > H6 in H2;
+       apply (lt_to_le (S n1) n1 ?).
+       apply (le_to_lt_to_lt (S n1) m n1 ? ?);
+        [apply (H3).
+        |apply (H2).
+        ]
+     | elim H4;
+       rewrite < H6;
+       assumption
+     ]
+   ]
+ | rewrite < H3 in H2 ⊢ %.
+   elim (min_aux_S f n n1);
+    [ elim H4;
+      rewrite > H6 in H2;
+      unfold lt in H2;
+      elim (not_le_Sn_n ? H2)
+    | elim H4;
+      assumption
+    ]
+ ]
 qed.
 
-theorem le_min_aux : \forall f:nat \to bool. 
-\forall n,off:nat. (n-off) \leq (min_aux off n f).
+
+lemma le_min_aux : \forall f:nat \to bool. 
+\forall n,off:nat. n \leq (min_aux off n f).
 intros 3.
-elim off.rewrite < minus_n_O.
-rewrite > (min_aux_O_f f n).apply le_n.
-elim (min_aux_S f n1 n).
+generalize in match n. clear n.
+elim off.
+rewrite > (min_aux_O_f f n1).apply le_n.
+elim (min_aux_S f n n1).
 elim H1.rewrite > H3.apply le_n.
 elim H1.rewrite > H3.
-apply (trans_le (n-(S n1)) (n-n1)).
-apply monotonic_le_minus_r.
-apply le_n_Sn.
-assumption.
+apply (transitive_le ? (S n1));
+ [ apply le_n_Sn
+ | apply (H (S n1))
+ ]
 qed.
 
 theorem le_min_aux_r : \forall f:nat \to bool. 
-\forall n,off:nat. (min_aux off n f) \le n.
+\forall n,off:nat. (min_aux off n f) \le n+off.
 intros.
-elim off.simplify.rewrite < minus_n_O.
-elim (f n).simplify.apply le_n.
-simplify.apply le_n.
-simplify.elim (f (n -(S n1))).
-simplify.apply le_plus_to_minus.
-rewrite < sym_plus.apply le_plus_n.
-simplify.assumption.
-qed.
+generalize in match n. clear n.
+elim off.simplify.
+elim (f n1).simplify.rewrite < plus_n_O.apply le_n.
+simplify.rewrite < plus_n_O.apply le_n.
+simplify.elim (f n1).
+simplify.
+apply (le_plus_n_r (S n) n1).
+simplify.rewrite < plus_n_Sm.
+apply (H (S n1)).
+qed.
\ No newline at end of file