X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fprimes.ma;h=0186d4324e3890666660282de63d5dc619a83520;hb=10f29fdd78ee089a9a94446207b543d33d6c851c;hp=a95b2e88fddda2344f0fafc448593de57404d470;hpb=21d7aa4df8d5d4bfe1073720ea4f9410e9cec879;p=helm.git diff --git a/helm/software/matita/library/nat/primes.ma b/helm/software/matita/library/nat/primes.ma index a95b2e88f..0186d4324 100644 --- a/helm/software/matita/library/nat/primes.ma +++ b/helm/software/matita/library/nat/primes.ma @@ -190,6 +190,125 @@ rewrite > H2.rewrite < H3. simplify.exact (not_le_Sn_n O). qed. +(*a variant of or_div_mod *) +theorem or_div_mod1: \forall n,q. O < q \to +(divides q (S n)) \land S n = (S (div n q)) * q \lor +(\lnot (divides q (S n)) \land S n= (div n q) * q + S (n\mod q)). +intros.elim (or_div_mod n q H);elim H1 + [left.split + [apply (witness ? ? (S (n/q))). + rewrite > sym_times.assumption + |assumption + ] + |right.split + [intro. + apply (not_eq_O_S (n \mod q)). + (* come faccio a fare unfold nelleipotesi ? *) + cut ((S n) \mod q = O) + [rewrite < Hcut. + apply (div_mod_spec_to_eq2 (S n) q (div (S n) q) (mod (S n) q) (div n q) (S (mod n q))) + [apply div_mod_spec_div_mod. + assumption + |apply div_mod_spec_intro;assumption + ] + |apply divides_to_mod_O;assumption + ] + |assumption + ] + ] +qed. + +theorem divides_to_div: \forall n,m.divides n m \to m/n*n = m. +intro. +elim (le_to_or_lt_eq O n (le_O_n n)) + [rewrite > plus_n_O. + rewrite < (divides_to_mod_O ? ? H H1). + apply sym_eq. + apply div_mod. + assumption + |elim H1. + generalize in match H2. + rewrite < H. + simplify. + intro. + rewrite > H3. + reflexivity + ] +qed. + +theorem divides_div: \forall d,n. divides d n \to divides (n/d) n. +intros. +apply (witness ? ? d). +apply sym_eq. +apply divides_to_div. +assumption. +qed. + +theorem div_div: \forall n,d:nat. O < n \to divides d n \to +n/(n/d) = d. +intros. +apply (inj_times_l1 (n/d)) + [apply (lt_times_n_to_lt d) + [apply (divides_to_lt_O ? ? H H1). + |rewrite > divides_to_div;assumption + ] + |rewrite > divides_to_div + [rewrite > sym_times. + rewrite > divides_to_div + [reflexivity + |assumption + ] + |apply (witness ? ? d). + apply sym_eq. + apply divides_to_div. + assumption + ] + ] +qed. + +theorem divides_to_eq_times_div_div_times: \forall a,b,c:nat. +O \lt b \to c \divides b \to a * (b /c) = (a*b)/c. +intros. +elim H1. +rewrite > H2. +rewrite > (sym_times c n2). +cut(O \lt c) +[ rewrite > (lt_O_to_div_times n2 c) + [ rewrite < assoc_times. + rewrite > (lt_O_to_div_times (a *n2) c) + [ reflexivity + | assumption + ] + | assumption + ] +| apply (divides_to_lt_O c b); + assumption. +] +qed. + +theorem eq_div_plus: \forall n,m,d. O < d \to +divides d n \to divides d m \to +(n + m ) / d = n/d + m/d. +intros. +elim H1. +elim H2. +rewrite > H3.rewrite > H4. +rewrite < distr_times_plus. +rewrite > sym_times. +rewrite > sym_times in ⊢ (? ? ? (? (? % ?) ?)). +rewrite > sym_times in ⊢ (? ? ? (? ? (? % ?))). +rewrite > lt_O_to_div_times + [rewrite > lt_O_to_div_times + [rewrite > lt_O_to_div_times + [reflexivity + |assumption + ] + |assumption + ] + |assumption + ] +qed. + (* boolean divides *) definition divides_b : nat \to nat \to bool \def \lambda n,m :nat. (eqb (m \mod n) O). @@ -223,7 +342,8 @@ intros 2.apply (nat_case n) [apply (nat_case m) [intro.apply divides_n_n |simplify.intros.apply False_ind. - apply not_eq_true_false.apply sym_eq.assumption + apply not_eq_true_false.apply sym_eq. + assumption ] |intros. apply divides_b_true_to_divides1 @@ -281,6 +401,22 @@ elim (divides_b n m).reflexivity. absurd (n \divides m).assumption.assumption. qed. +theorem divides_to_divides_b_true1 : \forall n,m:nat. +O < m \to n \divides m \to divides_b n m = true. +intro. +elim (le_to_or_lt_eq O n (le_O_n n)) + [apply divides_to_divides_b_true + [assumption|assumption] + |apply False_ind. + rewrite < H in H2. + elim H2. + simplify in H3. + apply (not_le_Sn_O O). + rewrite > H3 in H1. + assumption + ] +qed. + theorem not_divides_to_divides_b_false: \forall n,m:nat. O < n \to \lnot(n \divides m) \to (divides_b n m) = false. intros. @@ -293,6 +429,18 @@ absurd (n \divides m).assumption.assumption. reflexivity. qed. +theorem divides_b_div_true: +\forall d,n. O < n \to + divides_b d n = true \to divides_b (n/d) n = true. +intros. +apply divides_to_divides_b_true1 + [assumption + |apply divides_div. + apply divides_b_true_to_divides. + assumption + ] +qed. + theorem divides_b_true_to_lt_O: \forall n,m. O < n \to divides_b m n = true \to O < m. intros. elim (le_to_or_lt_eq ? ? (le_O_n m)) @@ -389,6 +537,11 @@ theorem prime_to_lt_O: \forall p. prime p \to O < p. intros.elim H.apply lt_to_le.assumption. qed. +theorem prime_to_lt_SO: \forall p. prime p \to S O < p. +intros.elim H. +assumption. +qed. + (* smallest factor *) definition smallest_factor : nat \to nat \def \lambda n:nat.