]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/minimization.ma
...
[helm.git] / helm / matita / library / nat / minimization.ma
index 8aa1deed30cb12b96ba981c4b0f326b868578706..72812cb651ecf447983be299aa9e94d00794af64 100644 (file)
@@ -15,7 +15,6 @@
 set "baseuri" "cic:/matita/nat/minimization".
 
 include "nat/minus.ma".
-include "datatypes/bool.ma".
 
 let rec max i f \def
   match (f i) with 
@@ -40,13 +39,46 @@ simplify.left.split.reflexivity.reflexivity.
 simplify.right.split.reflexivity.reflexivity.
 qed.
 
+theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
+max n f \le n.
+intros.elim n.rewrite > max_O_f.apply le_n.
+simplify.elim (f (S n1)).simplify.apply le_n.
+simplify.apply le_S.assumption.
+qed.
+
+theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
+n\le m  \to max n f \le max m f.
+intros.elim H.
+apply le_n.
+apply trans_le ? (max n1 f).apply H2.
+cut (f (S n1) = true \land max (S n1) f = (S n1)) \lor 
+(f (S n1) = false \land max (S n1) f = max n1 f).
+elim Hcut.elim H3.
+rewrite > H5.
+apply le_S.apply le_max_n.
+elim H3.rewrite > H5.apply le_n.
+apply max_S_max.
+qed.
+
+theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
+m\le n \to f m = true \to m \le max n f.
+intros 3.elim n.apply le_n_O_elim m H.
+apply le_O_n.
+apply le_n_Sm_elim m n1 H1.
+intro.apply trans_le ? (max n1 f).
+apply H.apply le_S_S_to_le.assumption.assumption.
+apply le_to_le_max.apply le_n_Sn.
+intro.simplify.rewrite < H3. 
+rewrite > H2.simplify.apply le_n.
+qed.
+
+
 definition max_spec \def \lambda f:nat \to bool.\lambda n: nat.
-ex nat (\lambda i:nat. (le i n) \land (f i = true)) \to
+\exists i. (le i n) \land (f i = true) \to
 (f n) = true \land (\forall i. i < n \to (f i = false)).
 
-(* perche' si blocca per mezzo minuto qui ??? *)
 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
-ex nat (\lambda i:nat. (le i n) \land (f i = true)) \to f (max n f) = true. 
+(\exists i:nat. le i n \land f i = true) \to f (max n f) = true. 
 intros 2.
 elim n.elim H.elim H1.generalize in match H3.
 apply le_n_O_elim a H2.intro.simplify.rewrite > H4.
@@ -55,8 +87,7 @@ simplify.
 apply bool_ind (\lambda b:bool.
 (f (S n1) = b) \to (f ([\lambda b:bool.nat] match b in bool with
 [ true \Rightarrow (S n1)
-| false  \Rightarrow (max n1 f)])) = true) ? ? ?.
-reflexivity.
+| false  \Rightarrow (max n1 f)])) = true).
 simplify.intro.assumption.
 simplify.intro.apply H.
 elim H1.elim H3.generalize in match H5.
@@ -64,8 +95,9 @@ apply le_n_Sm_elim a n1 H4.
 intros.
 apply ex_intro nat ? a.
 split.apply le_S_S_to_le.assumption.assumption.
-intros.apply False_ind.apply not_eq_true_false ?.
+intros.apply False_ind.apply not_eq_true_false.
 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
+reflexivity.
 qed.
 
 theorem lt_max_to_false : \forall f:nat \to bool. 
@@ -75,7 +107,6 @@ 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.
-(* ?? non posso generalizzare su un goal implicativo ?? *)
 elim max_S_max f n1.
 elim H3.
 absurd m \le S n1.assumption.
@@ -121,7 +152,7 @@ simplify.right.split.reflexivity.reflexivity.
 qed.
 
 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
-ex nat (\lambda i:nat. (le (m-off) i) \land (le i m) \land (f i = true)) \to
+(\exists i. le (m-off) i \land le i m \land f i = true) \to
 f (min_aux off m f) = true. 
 intros 2.
 elim off.elim H.elim H1.elim H2.
@@ -132,8 +163,7 @@ simplify.
 apply bool_ind (\lambda b:bool.
 (f (m-(S n)) = b) \to (f ([\lambda b:bool.nat] match b in bool with
 [ true \Rightarrow m-(S n)
-| false  \Rightarrow (min_aux n m f)])) = true) ? ? ?.
-reflexivity.
+| false  \Rightarrow (min_aux n m f)])) = true).
 simplify.intro.assumption.
 simplify.intro.apply H.
 elim H1.elim H3.elim H4.
@@ -145,6 +175,7 @@ assumption.assumption.
 absurd f a = false.rewrite < H8.assumption.
 rewrite > H5.
 apply not_eq_true_false.
+reflexivity.
 qed.
 
 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
@@ -160,8 +191,9 @@ 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.assumption.
+rewrite < H6.assumption.
 rewrite < H7.assumption.
+assumption.
 qed.
 
 theorem le_min_aux : \forall f:nat \to bool. 
@@ -172,7 +204,7 @@ rewrite > min_aux_O_f f n.apply le_n.
 elim min_aux_S f n1 n.
 elim H1.rewrite > H3.apply le_n.
 elim H1.rewrite > H3.
-apply trans_le (n-(S n1)) (n-n1) ?.
+apply trans_le (n-(S n1)) (n-n1).
 apply monotonic_le_minus_r.
 apply le_n_Sn.
 assumption.