Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 3061d18

Browse files
committed
feat(data/nat/{nth,prime}): add facts about primes (#12560)
Gives `{p | prime p}.infinite` as well as `infinite_of_not_bdd_above` lemma. Also gives simp lemmas for `prime_counting'`.
1 parent de4d14c commit 3061d18

File tree

4 files changed

+47
-2
lines changed

4 files changed

+47
-2
lines changed

src/data/nat/nth.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -146,6 +146,16 @@ lemma nth_mem_of_infinite (hp : (set_of p).infinite) (n : ℕ) : p (nth p n) :=
146146
lemma nth_strict_mono (hp : (set_of p).infinite) : strict_mono (nth p) :=
147147
λ a b, (nth_mem_of_infinite_aux p hp b).2 _
148148

149+
lemma nth_injective_of_infinite (hp : (set_of p).infinite) : function.injective (nth p) :=
150+
begin
151+
intros m n h,
152+
wlog h' : m ≤ n,
153+
rw le_iff_lt_or_eq at h',
154+
obtain (h' | rfl) := h',
155+
{ simpa [h] using nth_strict_mono p hp h' },
156+
{ refl },
157+
end
158+
149159
lemma nth_monotone (hp : (set_of p).infinite) : monotone (nth p) :=
150160
(nth_strict_mono p hp).monotone
151161

src/data/nat/prime.lean

Lines changed: 16 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@ import data.list.prime
77
import data.list.sort
88
import data.nat.gcd
99
import data.nat.sqrt
10+
import data.set.finite
1011
import tactic.norm_num
1112
import tactic.wlog
1213

@@ -20,7 +21,8 @@ This file deals with prime numbers: natural numbers `p ≥ 2` whose only divisor
2021
- `nat.prime`: the predicate that expresses that a natural number `p` is prime
2122
- `nat.primes`: the subtype of natural numbers that are prime
2223
- `nat.min_fac n`: the minimal prime factor of a natural number `n ≠ 1`
23-
- `nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers
24+
- `nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers.
25+
This also appears as `nat.not_bdd_above_set_of_prime` and `nat.infinite_set_of_prime`.
2426
- `nat.factors n`: the prime factorization of `n`
2527
- `nat.factors_unique`: uniqueness of the prime factorisation
2628
* `nat.prime_iff`: `nat.prime` coincides with the general definition of `prime`
@@ -392,6 +394,19 @@ have np : n ≤ p, from le_of_not_ge $ λ h,
392394
pp.not_dvd_one h₂,
393395
⟨p, np, pp⟩
394396

397+
/-- A version of `nat.exists_infinite_primes` using the `bdd_above` predicate. -/
398+
lemma not_bdd_above_set_of_prime : ¬ bdd_above {p | prime p} :=
399+
begin
400+
rw not_bdd_above_iff,
401+
intro n,
402+
obtain ⟨p, hi, hp⟩ := exists_infinite_primes n.succ,
403+
exact ⟨p, hp, hi⟩,
404+
end
405+
406+
/-- A version of `nat.exists_infinite_primes` using the `set.infinite` predicate. -/
407+
lemma infinite_set_of_prime : {p | prime p}.infinite :=
408+
set.infinite_of_not_bdd_above not_bdd_above_set_of_prime
409+
395410
lemma prime.eq_two_or_odd {p : ℕ} (hp : prime p) : p = 2 ∨ p % 2 = 1 :=
396411
p.mod_two_eq_zero_or_one.imp_left
397412
(λ h, ((hp.eq_one_or_self_of_dvd 2 (dvd_of_mod_eq_zero h)).resolve_left dec_trivial).symm)

src/data/set/finite.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -828,6 +828,13 @@ finite.induction_on H
828828
(by simp only [bUnion_empty, bdd_above_empty, ball_empty_iff])
829829
(λ a s ha _ hs, by simp only [bUnion_insert, ball_insert_iff, bdd_above_union, hs])
830830

831+
lemma infinite_of_not_bdd_above : ¬ bdd_above s → s.infinite :=
832+
begin
833+
contrapose!,
834+
rw not_infinite,
835+
apply finite.bdd_above,
836+
end
837+
831838
end
832839

833840
section
@@ -843,6 +850,13 @@ lemma finite.bdd_below_bUnion {I : set β} {S : β → set α} (H : finite I) :
843850
(bdd_below (⋃i∈I, S i)) ↔ (∀i ∈ I, bdd_below (S i)) :=
844851
@finite.bdd_above_bUnion (order_dual α) _ _ _ _ _ H
845852

853+
lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite :=
854+
begin
855+
contrapose!,
856+
rw not_infinite,
857+
apply finite.bdd_below,
858+
end
859+
846860
end
847861

848862
end set

src/number_theory/prime_counting.lean

Lines changed: 7 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ import data.nat.totient
99
import algebra.periodic
1010
import data.finset.locally_finite
1111
import data.nat.count
12-
12+
import data.nat.nth
1313

1414
/-!
1515
# The Prime Counting Function
@@ -56,6 +56,12 @@ lemma monotone_prime_counting' : monotone prime_counting' := count_monotone prim
5656
lemma monotone_prime_counting : monotone prime_counting :=
5757
λ a b a_le_b, monotone_prime_counting' (add_le_add_right a_le_b 1)
5858

59+
@[simp] lemma prime_counting'_nth_eq (n : ℕ) : π' (nth prime n) = n :=
60+
count_nth_of_infinite _ infinite_set_of_prime _
61+
62+
@[simp] lemma prime_nth_prime (n : ℕ) : prime (nth prime n) :=
63+
nth_mem_of_infinite _ infinite_set_of_prime _
64+
5965
/-- A linear upper bound on the size of the `prime_counting'` function -/
6066
lemma prime_counting'_add_le {a k : ℕ} (h0 : 0 < a) (h1 : a < k) (n : ℕ) :
6167
π' (k + n) ≤ π' k + nat.totient a * (n / a + 1) :=

0 commit comments

Comments
 (0)