]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/arithmetics/primes.ma
* Almost ready for release 0.99.1.
[helm.git] / matita / matita / lib / arithmetics / primes.ma
index 2e2ef98fbc18a2c113ddac96aa996cc96a2debfb..7c5df29f35151cebe5d3c6664368c8ee9f2ece73 100644 (file)
@@ -255,18 +255,17 @@ smallest_factor n = (min n 2 (λm.(eqb (n \mod m) O))).
 #n #lt1n normalize >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.
@@ -339,16 +338,16 @@ definition primeb ≝ λn:nat.
   (leb 2 n) ∧ (eqb (smallest_factor n) n).
 
 (* it works! *)
-theorem example4 : primeb 3 = true.
+example example4 : primeb 3 = true.
 normalize // qed.
 
-theorem example5 : primeb 6 = false.
+example example5 : primeb 6 = false.
 normalize // qed.
 
-theorem example6 : primeb 11 = true.
+example example6 : primeb 11 = true.
 normalize // qed.
 
-theorem example7 : primeb 17 = true.
+example example7 : primeb 17 = true.
 normalize // qed.
 
 theorem primeb_true_to_prime : ∀n:nat.
@@ -446,14 +445,14 @@ lemma nth_primeS: ∀n.nth_prime (S n) =
 // 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 
@@ -481,7 +480,7 @@ change with
 (let previous_prime ≝ (nth_prime n) in
  let upper_bound ≝ S previous_prime! in
  S previous_prime ≤ min upper_bound (S previous_prime) primeb)
-apply le_min_l
+@le_min_l
 qed.
 
 theorem lt_SO_nth_prime_n : ∀n:nat. 1 < nth_prime n.