]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/minimization.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / minimization.ma
index 4dfda0ba69f014bdb7030dfd7d8782d46ab30817..0abed5ad354319674d89e049ca833f804cdaaeec 100644 (file)
@@ -26,7 +26,7 @@ let rec max i f \def
 
 theorem max_O_f : \forall f: nat \to bool. max O f = O.
 intro. simplify.
-elim f O.
+elim (f O).
 simplify.reflexivity.
 simplify.reflexivity.
 qed. 
@@ -50,9 +50,9 @@ 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).
+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.
@@ -62,37 +62,38 @@ 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.
+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 (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)).
 
 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.
+apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
 simplify.assumption.
 simplify.
-apply bool_ind (\lambda b:bool.
-(f (S n1) = b) \to (f ([\lambda b:bool.nat] match b in bool with
+apply (bool_ind (\lambda b:bool.
+(f (S n1) = b) \to (f (match b in bool with
 [ true \Rightarrow (S n1)
-| false  \Rightarrow (max n1 f)])) = true).
+| false  \Rightarrow (max n1 f)])) = true)).
 simplify.intro.assumption.
 simplify.intro.apply H.
 elim H1.elim H3.generalize in match H5.
-apply le_n_Sm_elim a n1 H4.
+apply (le_n_Sm_elim a n1 H4).
 intros.
-apply ex_intro nat ? a.
+apply (ex_intro nat ? a).
 split.apply le_S_S_to_le.assumption.assumption.
 intros.apply False_ind.apply not_eq_true_false.
 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
@@ -102,17 +103,16 @@ 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.
+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 (max_S_max f n1).
 elim H3.
-absurd m \le S n1.assumption.
+absurd (m \le S n1).assumption.
 apply lt_to_not_le.rewrite < H6.assumption.
 elim H3.
-apply le_n_Sm_elim m n1 H2.
+apply (le_n_Sm_elim m n1 H2).
 intro.
 apply H.rewrite < H6.assumption.
 apply le_S_S_to_le.assumption.
@@ -133,14 +133,13 @@ definition min : nat \to (nat \to bool) \to nat \def
 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.
-elim f i.
-simplify.reflexivity.
+elim (f i).reflexivity.
 simplify.reflexivity.
 qed.
 
 theorem min_O_f : \forall f:nat \to bool.
 min O f = O.
-intro.apply min_aux_O_f f O.
+intro.apply (min_aux_O_f f O).
 qed.
 
 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
@@ -152,27 +151,27 @@ 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.
-cut a = m.
-rewrite > min_aux_O_f f.rewrite < Hcut.assumption.
-apply antisym_le a m .assumption.rewrite > minus_n_O m.assumption.
+cut (a = m).
+rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
+apply (antisym_le a m).assumption.rewrite > (minus_n_O m).assumption.
 simplify.
-apply bool_ind (\lambda b:bool.
-(f (m-(S n)) = b) \to (f ([\lambda b:bool.nat] match b in bool with
+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).
+| false  \Rightarrow (min_aux n m f)])) = true)).
 simplify.intro.assumption.
 simplify.intro.apply H.
 elim H1.elim H3.elim H4.
 elim (le_to_or_lt_eq (m-(S n)) a H6). 
-apply ex_intro nat ? a.
+apply (ex_intro nat ? a).
 split.split.
 apply lt_minus_S_n_to_le_minus_n.assumption.
 assumption.assumption.
-absurd f a = false.rewrite < H8.assumption.
+absurd (f a = false).rewrite < H8.assumption.
 rewrite > H5.
 apply not_eq_true_false.
 reflexivity.
@@ -181,15 +180,15 @@ 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.
 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.
+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 (min_aux_S f n1 n).
 elim H3.
-absurd n - S n1 \le m.assumption.
+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.
+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.
@@ -200,13 +199,24 @@ theorem le_min_aux : \forall f:nat \to bool.
 \forall n,off:nat. (n-off) \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.
+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.
 qed.
 
+theorem le_min_aux_r : \forall f:nat \to bool. 
+\forall n,off:nat. (min_aux off n f) \le n.
+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.