X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Ffactorization.ma;h=92f64e56d9a9c7a556a21f1cacd6d60edaf3b81e;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=1ca684aed0692c05874db9dd9c1a5ca0174d3739;hpb=c445ba5534cccde19016c92660ab52777af221c0;p=helm.git diff --git a/helm/software/matita/library/nat/factorization.ma b/helm/software/matita/library/nat/factorization.ma index 1ca684aed..92f64e56d 100644 --- a/helm/software/matita/library/nat/factorization.ma +++ b/helm/software/matita/library/nat/factorization.ma @@ -61,7 +61,7 @@ cut (\exists i. nth_prime i = smallest_factor n); [ apply (trans_lt ? (S O)); [ unfold lt; apply le_n; | apply lt_SO_smallest_factor; assumption; ] - | letin x \def le.autobatch new. + | letin x \def le.autobatch. (* apply divides_smallest_factor_n; apply (trans_lt ? (S O)); @@ -74,6 +74,18 @@ cut (\exists i. nth_prime i = smallest_factor n); assumption; *) ] qed. +lemma divides_to_prime_divides : \forall n,m.1 < m \to m < n \to m \divides n \to + \exists p.p \leq m \land prime p \land p \divides n. +intros;apply (ex_intro ? ? (nth_prime (max_prime_factor m)));split + [split + [apply divides_to_le + [apply lt_to_le;assumption + |apply divides_max_prime_factor_n;assumption] + |apply prime_nth_prime;] + |apply (transitive_divides ? ? ? ? H2);apply divides_max_prime_factor_n; + assumption] +qed. + theorem divides_to_max_prime_factor : \forall n,m. (S O) < n \to O < m \to n \divides m \to max_prime_factor n \le max_prime_factor m. intros.unfold max_prime_factor. @@ -144,7 +156,7 @@ apply divides_max_prime_factor_n. assumption.unfold Not. intro. cut (r \mod (nth_prime (max_prime_factor n)) \neq O); - [unfold Not in Hcut1.autobatch new. + [unfold Not in Hcut1.autobatch. (* apply Hcut1.apply divides_to_mod_O; [ apply lt_O_nth_prime_n. @@ -514,7 +526,7 @@ split [apply lt_O_nth_prime_n |apply (lt_O_n_elim ? H). intro. - apply (witness ? ? (r*(nth_prime p \sup m))). + apply (witness ? ? ((r*(nth_prime p) \sup m))). rewrite < assoc_times. rewrite < sym_times in \vdash (? ? ? (? % ?)). rewrite > exp_n_SO in \vdash (? ? ? (? (? ? %) ?)). @@ -636,21 +648,19 @@ theorem eq_defactorize_aux_to_eq: \forall f,g:nat_fact.\forall i:nat. defactorize_aux f i = defactorize_aux g i \to f = g. intro. elim f. -generalize in match H. -elim g. +elim g in H ⊢ %. apply eq_f. apply inj_S. apply (inj_exp_r (nth_prime i)). apply lt_SO_nth_prime_n. assumption. apply False_ind. -apply (not_eq_nf_last_nf_cons n2 n n1 i H2). -generalize in match H1. -elim g. +apply (not_eq_nf_last_nf_cons n2 n n1 i H1). +elim g in H1 ⊢ %. apply False_ind. apply (not_eq_nf_last_nf_cons n1 n2 n i). apply sym_eq. assumption. -simplify in H3. -generalize in match H3. +simplify in H2. +generalize in match H2. apply (nat_elim2 (\lambda n,n2. ((nth_prime i) \sup n)*(defactorize_aux n1 (S i)) = ((nth_prime i) \sup n2)*(defactorize_aux n3 (S i)) \to @@ -682,7 +692,7 @@ change with [ (nf_last m) \Rightarrow m | (nf_cons m g) \Rightarrow m ] = m). rewrite > Hcut.simplify.reflexivity. -apply H4.simplify in H5. +apply H3.simplify in H4. apply (inj_times_r1 (nth_prime i)). apply lt_O_nth_prime_n. rewrite < assoc_times.rewrite < assoc_times.assumption. @@ -700,22 +710,21 @@ injective nat_fact_all nat defactorize. unfold injective. change with (\forall f,g.defactorize f = defactorize g \to f=g). intro.elim f. -generalize in match H.elim g. +elim g in H ⊢ %. (* zero - zero *) reflexivity. (* zero - one *) simplify in H1. apply False_ind. -apply (not_eq_O_S O H1). +apply (not_eq_O_S O H). (* zero - proper *) simplify in H1. apply False_ind. apply (not_le_Sn_n O). -rewrite > H1 in \vdash (? ? %). +rewrite > H in \vdash (? ? %). change with (O < defactorize_aux n O). apply lt_O_defactorize_aux. -generalize in match H. -elim g. +elim g in H ⊢ %. (* one - zero *) simplify in H1. apply False_ind. @@ -726,28 +735,28 @@ reflexivity. simplify in H1. apply False_ind. apply (not_le_Sn_n (S O)). -rewrite > H1 in \vdash (? ? %). +rewrite > H in \vdash (? ? %). change with ((S O) < defactorize_aux n O). apply lt_SO_defactorize_aux. -generalize in match H.elim g. +elim g in H ⊢ %. (* proper - zero *) simplify in H1. apply False_ind. apply (not_le_Sn_n O). -rewrite < H1 in \vdash (? ? %). +rewrite < H in \vdash (? ? %). change with (O < defactorize_aux n O). apply lt_O_defactorize_aux. (* proper - one *) simplify in H1. apply False_ind. apply (not_le_Sn_n (S O)). -rewrite < H1 in \vdash (? ? %). +rewrite < H in \vdash (? ? %). change with ((S O) < defactorize_aux n O). apply lt_SO_defactorize_aux. (* proper - proper *) apply eq_f. apply (injective_defactorize_aux O). -exact H1. +exact H. qed. theorem factorize_defactorize: