X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fiteration2.ma;h=3a4128e70d4727d81aad445bf59036a715b6511d;hb=aa5f71baeba0299c0d29be01798f7a1ad13656f9;hp=6dbbfd572a2ac824de1d69bff781d06eb18009cc;hpb=9edc9828ff3c3d074ccc5547a53eadf092c186d8;p=helm.git diff --git a/helm/software/matita/library/nat/iteration2.ma b/helm/software/matita/library/nat/iteration2.ma index 6dbbfd572..3a4128e70 100644 --- a/helm/software/matita/library/nat/iteration2.ma +++ b/helm/software/matita/library/nat/iteration2.ma @@ -12,14 +12,11 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/iteration2". - include "nat/primes.ma". include "nat/ord.ma". include "nat/generic_iter_p.ma". include "nat/count.ma".(*necessary just to use bool_to_nat and bool_to_nat_andb*) - (* sigma_p on nautral numbers is a specialization of sigma_p_gen *) definition sigma_p: nat \to (nat \to bool) \to (nat \to nat) \to nat \def \lambda n, p, g. (iter_p_gen n p nat g O plus). @@ -289,17 +286,16 @@ theorem le_sigma_p: (\forall i. i < n \to p i = true \to g1 i \le g2 i ) \to sigma_p n p g1 \le sigma_p n p g2. intros. -generalize in match H. -elim n +elim n in H ⊢ % [apply le_n. |apply (bool_elim ? (p n1));intros [rewrite > true_to_sigma_p_Sn [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %) [apply le_plus - [apply H2[apply le_n|assumption] - |apply H1. + [apply H1[apply le_n|assumption] + |apply H. intros. - apply H2[apply le_S.assumption|assumption] + apply H1[apply le_S.assumption|assumption] ] |assumption ] @@ -307,9 +303,9 @@ elim n ] |rewrite > false_to_sigma_p_Sn [rewrite > false_to_sigma_p_Sn in ⊢ (? ? %) - [apply H1. + [apply H. intros. - apply H2[apply le_S.assumption|assumption] + apply H1[apply le_S.assumption|assumption] |assumption ] |assumption @@ -325,25 +321,24 @@ theorem le_sigma_p1: bool_to_nat (p1 i)*(g1 i) \le bool_to_nat (p2 i)*g2 i) \to sigma_p n p1 g1 \le sigma_p n p2 g2. intros. -generalize in match H. -elim n +elim n in H ⊢ % [apply le_n. |apply (bool_elim ? (p1 n1));intros [apply (bool_elim ? (p2 n1));intros [rewrite > true_to_sigma_p_Sn [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %) [apply le_plus - [lapply (H2 n1) as H5 - [rewrite > H3 in H5. - rewrite > H4 in H5. + [lapply (H1 n1) as H5 + [rewrite > H2 in H5. + rewrite > H3 in H5. simplify in H5. rewrite < plus_n_O in H5. rewrite < plus_n_O in H5. assumption |apply le_S_S.apply le_n ] - |apply H1.intros. - apply H2.apply le_S.assumption + |apply H.intros. + apply H1.apply le_S.assumption ] |assumption ] @@ -353,16 +348,16 @@ elim n [rewrite > false_to_sigma_p_Sn in ⊢ (? ? %) [change in ⊢ (? ? %) with (O + sigma_p n1 p2 g2). apply le_plus - [lapply (H2 n1) as H5 - [rewrite > H3 in H5. - rewrite > H4 in H5. + [lapply (H1 n1) as H5 + [rewrite > H2 in H5. + rewrite > H3 in H5. simplify in H5. rewrite < plus_n_O in H5. assumption |apply le_S_S.apply le_n ] - |apply H1.intros. - apply H2.apply le_S.assumption + |apply H.intros. + apply H1.apply le_S.assumption ] |assumption ] @@ -374,16 +369,16 @@ elim n [rewrite > true_to_sigma_p_Sn in ⊢ (? ? %) [change in ⊢ (? % ?) with (O + sigma_p n1 p1 g1). apply le_plus - [lapply (H2 n1) as H5 - [rewrite > H3 in H5. - rewrite > H4 in H5. + [lapply (H1 n1) as H5 + [rewrite > H2 in H5. + rewrite > H3 in H5. simplify in H5. rewrite < plus_n_O in H5. assumption |apply le_S_S.apply le_n ] - |apply H1.intros. - apply H2.apply le_S.assumption + |apply H.intros. + apply H1.apply le_S.assumption ] |assumption ] @@ -391,8 +386,8 @@ elim n ] |rewrite > false_to_sigma_p_Sn [rewrite > false_to_sigma_p_Sn in ⊢ (? ? %) - [apply H1.intros. - apply H2.apply le_S.assumption + [apply H.intros. + apply H1.apply le_S.assumption |assumption ] |assumption @@ -400,7 +395,7 @@ elim n ] ] ] -qed. +qed. theorem lt_sigma_p: \forall n:nat. \forall p:nat \to bool. \forall g1,g2:nat \to nat. @@ -629,32 +624,33 @@ elim n qed. theorem sigma_p_plus_1: \forall n:nat. \forall f,g:nat \to nat. -sigma_p n (\lambda b:nat. true) (\lambda a:nat.(f a) + (g a)) = -sigma_p n (\lambda b:nat. true) f + sigma_p n (\lambda b:nat. true) g. +\forall p. +sigma_p n p (\lambda a:nat.(f a) + (g a)) = +sigma_p n p f + sigma_p n p g. intros. elim n [ simplify. reflexivity -| rewrite > true_to_sigma_p_Sn - [ rewrite > (true_to_sigma_p_Sn n1 (\lambda c:nat.true) f) - [ rewrite > (true_to_sigma_p_Sn n1 (\lambda c:nat.true) g) - [ rewrite > assoc_plus in \vdash (? ? ? %). - rewrite < assoc_plus in \vdash (? ? ? (? ? %)). - rewrite < sym_plus in \vdash (? ? ? (? ? (? % ?))). - rewrite > assoc_plus in \vdash (? ? ? (? ? %)). - rewrite < assoc_plus in \vdash (? ? ? %). - apply eq_f. - assumption - | reflexivity - ] - | reflexivity - ] - | reflexivity - ] -] +| apply (bool_elim ? (p n1)); intro; + [ rewrite > true_to_sigma_p_Sn + [ rewrite > (true_to_sigma_p_Sn n1 p f) + [ rewrite > (true_to_sigma_p_Sn n1 p g) + [ rewrite > assoc_plus in \vdash (? ? ? %). + rewrite < assoc_plus in \vdash (? ? ? (? ? %)). + rewrite < sym_plus in \vdash (? ? ? (? ? (? % ?))). + rewrite > assoc_plus in \vdash (? ? ? (? ? %)). + rewrite < assoc_plus in \vdash (? ? ? %). + apply eq_f. + assumption]]] + assumption + | rewrite > false_to_sigma_p_Sn + [ rewrite > (false_to_sigma_p_Sn n1 p f) + [ rewrite > (false_to_sigma_p_Sn n1 p g) + [assumption]]] + assumption +]] qed. - theorem eq_sigma_p_sigma_p_times1 : \forall n,m:nat.\forall f:nat \to nat. sigma_p (n*m) (\lambda x:nat.true) f = sigma_p m (\lambda x:nat.true)