X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fprimes.ma;h=68e511e982617f293eedfc320aaa6be1ec888447;hb=1efc4c2c7be1e4aff0ccccabf905d45795b3865f;hp=2c4897d06d00152d33bd03a98c1a810985f447f8;hpb=53452958508001e7af3090695b619fe92135fb9e;p=helm.git diff --git a/matita/matita/lib/arithmetics/primes.ma b/matita/matita/lib/arithmetics/primes.ma index 2c4897d06..68e511e98 100644 --- a/matita/matita/lib/arithmetics/primes.ma +++ b/matita/matita/lib/arithmetics/primes.ma @@ -81,7 +81,7 @@ theorem eq_mod_to_divides: ∀n,m,q. O < q → |@(quotient ?? ((div n q)-(div m q))) >distributive_times_minus >commutative_times >(commutative_times q) cut((n/q)*q = n - (n \mod q)) [//] #H - >H >minus_minus >eqmod >commutative_plus + >H >minus_plus >eqmod >commutative_plus lt_to_leb_false // qed. -(* it works ! -theorem example1 : smallest_factor 3 = 3. +example example1 : smallest_factor 3 = 3. normalize // qed. -theorem example2: smallest_factor 4 = 2. +example example2: smallest_factor 4 = 2. normalize // qed. -theorem example3 : smallest_factor 7 = 7. +example example3 : smallest_factor 7 = 7. normalize // -qed. *) +qed. theorem le_SO_smallest_factor: ∀n:nat. n ≤ 1 → smallest_factor n = n. @@ -291,7 +289,7 @@ theorem divides_smallest_factor_n : ∀n:nat. 0 < n → #n #posn (cases (le_to_or_lt_eq … posn)) [#lt1n @mod_O_to_divides [@lt_O_smallest_factor //] >smallest_factor_to_min // @eqb_true_to_eq @f_min_true - @(ex_intro … n) % /2/ @eq_to_eqb_true /2/ + @(ex_intro … n) % /2/ |#H H //] +#n #H cut ((leb 2 n ∧ (eqb (smallest_factor n) n)) = false) [@H] @leb_elim [#_ #H1 @(not_to_not … (prime_to_smallest_factor … )) @eqb_false_to_not_eq @H1 @@ -440,20 +438,20 @@ let rec nth_prime n ≝ match n with min upper_bound (S previous_prime) primeb]. lemma nth_primeS: ∀n.nth_prime (S n) = - let previous_prime ≝ nth_prime n in + (let previous_prime ≝ nth_prime n in let upper_bound ≝ S previous_prime! in - min upper_bound (S previous_prime) primeb. + min upper_bound (S previous_prime) primeb). // qed. (* it works *) -theorem example11 : nth_prime 2 = 5. +example example11 : nth_prime 2 = 5. normalize // qed. -theorem example12: nth_prime 3 = 7. +example example12: nth_prime 3 = 7. normalize // qed. -theorem example13 : nth_prime 4 = 11. +example example13 : nth_prime 4 = 11. normalize // qed. (* done in 13.1369330883s -- qualcosa non va: // lentissimo @@ -465,7 +463,7 @@ qed. theorem prime_nth_prime : ∀n:nat.prime (nth_prime n). #n (cases n) [@primeb_true_to_prime // - |#m >nth_primeS @primeb_true_to_prime @f_min_true + |#m >nth_primeS @primeb_true_to_prime @(f_min_true primeb) @(ex_intro nat ? (smallest_factor (S (nth_prime m)!))) % [% // @le_S_S @(transitive_le … (le_smallest_factor_n …))