]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/minimization.ma
Adapted to new applyS.
[helm.git] / helm / software / matita / library / nat / minimization.ma
index 8b4b83c2fbc7d910fc9c3d91dbd4dc130e713c30..5b1552dc6777de3c1a9ef298b846b62796af70ff 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/minimization".
-
 include "nat/minus.ma".
 
 let rec max i f \def
@@ -72,6 +70,76 @@ 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.
+
+
+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. 
 intros 2.
@@ -95,23 +163,101 @@ 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.
 elim n.absurd (le m O).assumption.
 cut (O < m).apply (lt_O_n_elim m Hcut).exact not_le_Sn_O.
 rewrite < (max_O_f f).assumption.
-generalize in match H1.
-elim (max_S_max f n1).
-elim H3.
+elim (max_S_max f n1) in H1 ⊢ %.
+elim H1.
 absurd (m \le S n1).assumption.
-apply lt_to_not_le.rewrite < H6.assumption.
-elim H3.
+apply lt_to_not_le.rewrite < H5.assumption.
+elim H1.
 apply (le_n_Sm_elim m n1 H2).
 intro.
-apply H.rewrite < H6.assumption.
+apply H.rewrite < H5.assumption.
 apply le_S_S_to_le.assumption.
-intro.rewrite > H7.assumption.
+intro.rewrite > H6.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.
@@ -219,10 +365,26 @@ apply not_eq_true_false.
 reflexivity.
 qed.
 
+theorem f_min_true: \forall f:nat \to bool. \forall m:nat.
+(\exists i. le i m \land f i = true) \to
+f (min m f) = true.
+intros.unfold min.
+apply f_min_aux_true.
+elim H.clear H.elim H1.clear H1.
+apply (ex_intro ? ? a).
+split
+  [split
+    [apply le_O_n
+    |rewrite < plus_n_O.assumption
+    ]
+  |assumption
+  ]
+qed.
+
 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
 \forall n,off,m:nat. n \leq m \to m < (min_aux off n f) \to f m = false.
 intros 3.
-generalize in match n; clear n.
+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);
@@ -259,8 +421,7 @@ qed.
 lemma le_min_aux : \forall f:nat \to bool. 
 \forall n,off:nat. n \leq (min_aux off n f).
 intros 3.
-generalize in match n. clear n.
-elim off.
+elim off in n ⊢ %.
 rewrite > (min_aux_O_f f n1).apply le_n.
 elim (min_aux_S f n n1).
 elim H1.rewrite > H3.apply le_n.
@@ -274,8 +435,7 @@ qed.
 theorem le_min_aux_r : \forall f:nat \to bool. 
 \forall n,off:nat. (min_aux off n f) \le n+off.
 intros.
-generalize in match n. clear n.
-elim off.simplify.
+elim off in n ⊢ %.simplify.
 elim (f n1).simplify.rewrite < plus_n_O.apply le_n.
 simplify.rewrite < plus_n_O.apply le_n.
 simplify.elim (f n1).