X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fgcd.ma;h=9bbce8144ff8885a7c7346bdb9b7937983bd5f7e;hb=947ee89dec9e60561dfac3ce7e1842f35f178cb8;hp=ded9d4843ac6cf3cb4da2b4f564568db2df2e02b;hpb=db9c252cc8adb9243892203805b203bafe486bfc;p=helm.git diff --git a/helm/software/matita/library/nat/gcd.ma b/helm/software/matita/library/nat/gcd.ma index ded9d4843..9bbce8144 100644 --- a/helm/software/matita/library/nat/gcd.ma +++ b/helm/software/matita/library/nat/gcd.ma @@ -749,6 +749,7 @@ apply gcd_O_to_eq_O.apply sym_eq.assumption. apply le_to_or_lt_eq.apply le_O_n. qed. +(* primes and divides *) theorem divides_times_to_divides: \forall n,p,q:nat.prime n \to n \divides p*q \to n \divides p \lor n \divides q. intros. @@ -802,6 +803,30 @@ cut (n \divides p \lor n \ndivides p) ] qed. +theorem divides_exp_to_divides: +\forall p,n,m:nat. prime p \to +p \divides n \sup m \to p \divides n. +intros 3.elim m.simplify in H1. +apply (transitive_divides p (S O)).assumption. +apply divides_SO_n. +cut (p \divides n \lor p \divides n \sup n1). +elim Hcut.assumption. +apply H.assumption.assumption. +apply divides_times_to_divides.assumption. +exact H2. +qed. + +theorem divides_exp_to_eq: +\forall p,q,m:nat. prime p \to prime q \to +p \divides q \sup m \to p = q. +intros. +unfold prime in H1. +elim H1.apply H4. +apply (divides_exp_to_divides p q m). +assumption.assumption. +unfold prime in H.elim H.assumption. +qed. + theorem eq_gcd_times_SO: \forall m,n,p:nat. O < n \to O < p \to gcd m n = (S O) \to gcd m p = (S O) \to gcd m (n*p) = (S O). intros. @@ -877,4 +902,5 @@ cut (n \divides p \lor n \ndivides p) |apply (decidable_divides n p). assumption. ] -qed. \ No newline at end of file +qed. +