X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fprimes.ma;h=67d8254641379f309669e8dc0fa3ee6e1001caea;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=f7d6970063c4cf9b06a1e771c73b50c876bc1894;hpb=db9c252cc8adb9243892203805b203bafe486bfc;p=helm.git diff --git a/helm/software/matita/library/nat/primes.ma b/helm/software/matita/library/nat/primes.ma index f7d697006..67d825464 100644 --- a/helm/software/matita/library/nat/primes.ma +++ b/helm/software/matita/library/nat/primes.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/primes". - include "nat/div_and_mod.ma". include "nat/minimization.ma". include "nat/sigma_and_pi.ma". @@ -22,9 +20,8 @@ include "nat/factorial.ma". inductive divides (n,m:nat) : Prop \def witness : \forall p:nat.m = times n p \to divides n m. -interpretation "divides" 'divides n m = (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m). -interpretation "not divides" 'ndivides n m = - (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/primes/divides.ind#xpointer(1/1) n m)). +interpretation "divides" 'divides n m = (divides n m). +interpretation "not divides" 'ndivides n m = (Not (divides n m)). theorem reflexive_divides : reflexive nat divides. unfold reflexive. @@ -83,36 +80,36 @@ qed. theorem divides_plus: \forall n,p,q:nat. n \divides p \to n \divides q \to n \divides p+q. intros. -elim H.elim H1. apply (witness n (p+q) (n2+n1)). +elim H.elim H1. apply (witness n (p+q) (n1+n2)). rewrite > H2.rewrite > H3.apply sym_eq.apply distr_times_plus. qed. theorem divides_minus: \forall n,p,q:nat. divides n p \to divides n q \to divides n (p-q). intros. -elim H.elim H1. apply (witness n (p-q) (n2-n1)). +elim H.elim H1. apply (witness n (p-q) (n1-n2)). rewrite > H2.rewrite > H3.apply sym_eq.apply distr_times_minus. qed. theorem divides_times: \forall n,m,p,q:nat. n \divides p \to m \divides q \to n*m \divides p*q. intros. -elim H.elim H1. apply (witness (n*m) (p*q) (n2*n1)). +elim H.elim H1. apply (witness (n*m) (p*q) (n1*n2)). rewrite > H2.rewrite > H3. -apply (trans_eq nat ? (n*(m*(n2*n1)))). -apply (trans_eq nat ? (n*(n2*(m*n1)))). +apply (trans_eq nat ? (n*(m*(n1*n2)))). +apply (trans_eq nat ? (n*(n1*(m*n2)))). apply assoc_times. apply eq_f. -apply (trans_eq nat ? ((n2*m)*n1)). +apply (trans_eq nat ? ((n1*m)*n2)). apply sym_eq. apply assoc_times. -rewrite > (sym_times n2 m).apply assoc_times. +rewrite > (sym_times n1 m).apply assoc_times. apply sym_eq. apply assoc_times. qed. theorem transitive_divides: transitive ? divides. unfold. intros. -elim H.elim H1. apply (witness x z (n2*n)). +elim H.elim H1. apply (witness x z (n1*n)). rewrite > H3.rewrite > H2. apply assoc_times. qed. @@ -152,7 +149,7 @@ qed. theorem antisymmetric_divides: antisymmetric nat divides. unfold antisymmetric.intros.elim H. elim H1. -apply (nat_case1 n2).intro. +apply (nat_case1 n1).intro. rewrite > H3.rewrite > H2.rewrite > H4. rewrite < times_n_O.reflexivity. intros. @@ -169,11 +166,11 @@ qed. (* divides le *) theorem divides_to_le : \forall n,m. O < m \to n \divides m \to n \le m. -intros. elim H1.rewrite > H2.cut (O < n2). -apply (lt_O_n_elim n2 Hcut).intro.rewrite < sym_times. +intros. elim H1.rewrite > H2.cut (O < n1). +apply (lt_O_n_elim n1 Hcut).intro.rewrite < sym_times. simplify.rewrite < sym_plus. apply le_plus_n. -elim (le_to_or_lt_eq O n2). +elim (le_to_or_lt_eq O n1). assumption. absurd (O H2.rewrite < H3.rewrite < times_n_O. @@ -190,8 +187,33 @@ rewrite > H2.rewrite < H3. simplify.exact (not_le_Sn_n O). qed. - -(*divides and div*) +(*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. @@ -211,6 +233,14 @@ elim (le_to_or_lt_eq O n (le_O_n n)) ] 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. @@ -238,11 +268,11 @@ O \lt b \to c \divides b \to a * (b /c) = (a*b)/c. intros. elim H1. rewrite > H2. -rewrite > (sym_times c n2). +rewrite > (sym_times c n1). cut(O \lt c) -[ rewrite > (lt_O_to_div_times n2 c) +[ rewrite > (lt_O_to_div_times n1 c) [ rewrite < assoc_times. - rewrite > (lt_O_to_div_times (a *n2) c) + rewrite > (lt_O_to_div_times (a *n1) c) [ reflexivity | assumption ] @@ -253,6 +283,28 @@ cut(O \lt c) ] 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 @@ -287,7 +339,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 @@ -345,6 +398,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. @@ -357,6 +426,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)) @@ -453,6 +534,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.