]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/minimization.ma
generalize no more required before elim
[helm.git] / helm / software / matita / library / nat / minimization.ma
index bc2cc70591a3de6e10d5d1f4604e5958046a32d6..5b1552dc6777de3c1a9ef298b846b62796af70ff 100644 (file)
@@ -236,17 +236,16 @@ 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
@@ -385,7 +384,7 @@ 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);
@@ -422,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.
@@ -437,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).