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

[Merged by Bors] - feat(number_theory/prime_counting): The prime counting function #9080

Closed
wants to merge 267 commits into from
Closed
Show file tree
Hide file tree
Changes from 250 commits
Commits
Show all changes
267 commits
Select commit Hold shift + click to select a range
c0c0d6e
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
6c27bf2
fix import
kim-em Aug 21, 2021
ce20ad0
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
c9e159a
not include the upper bound in count
YaelDillies Aug 21, 2021
3c849aa
Merge branch 'count' of https://github.com/leanprover-community/mathl…
YaelDillies Aug 21, 2021
3122b4d
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
7b84209
nth_mem_of_infinite_aux
kmill Aug 21, 2021
31fac8c
minor
kim-em Aug 21, 2021
63f6e6b
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
7b3e299
Merge branch 'count' of https://github.com/leanprover-community/mathl…
kmill Aug 21, 2021
019e30a
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
28271e6
fixed accidental breakage
kmill Aug 21, 2021
cbd5356
count_succ
YaelDillies Aug 21, 2021
ff7a9f1
Merge branch 'count' of https://github.com/leanprover-community/mathl…
YaelDillies Aug 21, 2021
4866256
Merge branch 'count' of https://github.com/leanprover-community/mathl…
kmill Aug 21, 2021
4b7ebef
fixed errors
kmill Aug 21, 2021
3bb8df2
moved count_succ_eq_succ_count_iff to where it could be proved
kmill Aug 21, 2021
b7c3c18
same thing for count_succ_eq_count_iff
kmill Aug 21, 2021
bbe03d3
implemented exists_gt_of_infinite just in case, even though it's not …
kmill Aug 21, 2021
2550f9c
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
9e012cd
a little of nth_mem_of_le_card
kmill Aug 21, 2021
f83a442
Merge branch 'count' of https://github.com/leanprover-community/mathl…
kmill Aug 21, 2021
0c0484d
remove a -1
SymmetryUnbroken Aug 21, 2021
45d1fb3
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 21, 2021
05547af
remove a -1 from order_iso_nat
SymmetryUnbroken Aug 21, 2021
bac4ea2
prove the count_succ stuff
YaelDillies Aug 21, 2021
e308233
some remaining +- 1
YaelDillies Aug 21, 2021
ce49224
improve count_succ + related
ericrbg Aug 21, 2021
87a6ac9
Merge remote-tracking branch 'refs/remotes/origin/count' into count
ericrbg Aug 21, 2021
01ae0c8
more tidying
ericrbg Aug 21, 2021
b673c6d
count_zero
YaelDillies Aug 21, 2021
8503cda
exists_gt_of_infinite
ericrbg Aug 21, 2021
2fbc7fd
random stuff
ericrbg Aug 21, 2021
8d37f6d
fixed defeq abuse
kmill Aug 21, 2021
05e8d64
recovered slightly proof of exists_gt_of_infinite (sorry eric)
kmill Aug 21, 2021
aac6fb7
simp considerations
kmill Aug 21, 2021
97f348a
curly braces to see structure
kmill Aug 21, 2021
6232bc4
remove useless line
ericrbg Aug 21, 2021
48b124e
seems hard
kim-em Aug 22, 2021
5a12ba5
progress
kim-em Aug 22, 2021
177b3bb
more?
kim-em Aug 22, 2021
66bf3b2
commenting out unneeded lemmas
kim-em Aug 22, 2021
15354b2
rearrange
kim-em Aug 22, 2021
db6b8f9
Merge branch 'master' into count
ericrbg Aug 22, 2021
8d060b0
quick crappy docs
ericrbg Aug 22, 2021
02824a0
some more random stuff
ericrbg Aug 22, 2021
250e9dc
`count_le_card` + misc
ericrbg Aug 22, 2021
2060ccd
filled a sorry
kmill Aug 22, 2021
5014a10
tidy i guess
ericrbg Aug 22, 2021
de971c1
some more finagling, nth_count is nearly there
ericrbg Aug 22, 2021
baf6d17
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Aug 24, 2021
9d5a29d
work on count_nth_gc
SymmetryUnbroken Aug 24, 2021
9baa0a7
moved the nth_count_le lemma
SymmetryUnbroken Aug 24, 2021
667c75f
close one sorry with one line
SymmetryUnbroken Aug 24, 2021
a1f3977
some golfing of Vlad's stuff + random
ericrbg Aug 24, 2021
e315f9a
nth_count progress
ericrbg Aug 24, 2021
d682ec3
tiny
ericrbg Aug 24, 2021
7a95101
cleanup, and add a lemma we should prove
kim-em Aug 25, 2021
f01c6cd
use `reflect_lt`
ericrbg Aug 25, 2021
b0639d8
finish the proof of galois connection
SymmetryUnbroken Aug 25, 2021
0b5235d
reorder lemmas
kim-em Aug 25, 2021
d93e322
minor
kim-em Aug 25, 2021
1647619
work on nth_succ_of_zero
SymmetryUnbroken Aug 27, 2021
0f1e52c
change Inf_plus; close a simple sorry
SymmetryUnbroken Aug 27, 2021
169a483
work on nth_succ_of_zero
SymmetryUnbroken Aug 27, 2021
9af788d
progress on Inf_plus, golf
ericrbg Aug 28, 2021
da4fe2b
proved Inf_plus
SymmetryUnbroken Aug 28, 2021
4d77121
minor
SymmetryUnbroken Aug 28, 2021
e49f6f6
leave only one Inf_plus
SymmetryUnbroken Aug 28, 2021
1240107
prove count_nth_of_infinite
SymmetryUnbroken Aug 28, 2021
63d4609
filled in proof of count_eq_card_finset
kmill Aug 28, 2021
1560860
cleaned up gross proof i'd written for count_eq_card
kmill Aug 28, 2021
7c514f9
add a condition to nth_count_le
SymmetryUnbroken Aug 29, 2021
c7e7f62
prove nth_mem_of_le_card_aux
SymmetryUnbroken Aug 29, 2021
1cf67db
nth_set_card
SymmetryUnbroken Aug 29, 2021
1b20068
prove nth_set_nonempty_of_lt_card
SymmetryUnbroken Aug 29, 2021
e8d2ac2
prove nth_mem_of_lt_card_finite
SymmetryUnbroken Aug 29, 2021
a0e1fe3
prove count_nth_of_lt_card_finite
SymmetryUnbroken Aug 30, 2021
59b866f
prove lemmas for count_nth_of_lt_card_finite
SymmetryUnbroken Aug 30, 2021
8d4051a
prove nth_count_of_infinte
SymmetryUnbroken Aug 30, 2021
0ea0bf1
prove nth_count
SymmetryUnbroken Aug 31, 2021
8b998f3
proof refinement, deleted old nth_count, deleted count_injective_of_i…
kmill Aug 31, 2021
8c25d89
more set_of normalization
kmill Aug 31, 2021
547a5cb
line length linter
kmill Sep 1, 2021
07b88d3
prove nth_count_le
SymmetryUnbroken Sep 1, 2021
2ce6a84
prove nth_count_le
SymmetryUnbroken Sep 6, 2021
731e611
knock out some easy sorries
ericrbg Sep 6, 2021
e0352ff
clean up
SymmetryUnbroken Sep 7, 2021
573e4f7
Merge branch 'count' of https://github.com/leanprover-community/mathl…
SymmetryUnbroken Sep 7, 2021
507c325
Add file for prime counting function
BoltonBailey Sep 8, 2021
174f706
progress
BoltonBailey Sep 8, 2021
4d89c01
work on nth_lt_of_lt_count
SymmetryUnbroken Sep 8, 2021
7e60cfd
remove nth_succ_of_zero
SymmetryUnbroken Sep 8, 2021
cfd17c2
prove nth_lt_of_lt_count, remove prove nth_le_of_le_count
SymmetryUnbroken Sep 14, 2021
f470ff3
prove le_nth_of_count_le
SymmetryUnbroken Sep 14, 2021
1646575
prove nth_nonzero_of_ge_nonzero
SymmetryUnbroken Sep 15, 2021
b6994ff
work on nth_eq_order_iso_of_nat
SymmetryUnbroken Sep 15, 2021
0db62d1
prove nth_eq_order_iso_of_nat
SymmetryUnbroken Sep 15, 2021
4adc580
Merge remote-tracking branch 'origin/master' into count
YaelDillies Sep 16, 2021
9dea5da
golf
YaelDillies Sep 16, 2021
759b337
Merge branch 'master' into count
ericrbg Sep 23, 2021
7c01b94
add copyright header
ericrbg Sep 23, 2021
37e4574
golf + restructure
ericrbg Sep 24, 2021
3854cb0
move lemmas (and realise some already exist!)
ericrbg Sep 24, 2021
7a975de
move last thing
ericrbg Sep 24, 2021
9247761
Merge branch 'master' into count
ericrbg Sep 24, 2021
32b0def
change direction to more general linear bound
BoltonBailey Sep 24, 2021
db4312d
Merge branch 'master' into count
ericrbg Sep 24, 2021
f6eb09d
move last lemma
ericrbg Sep 24, 2021
749a1b3
consistent rw style, 100char
ericrbg Sep 24, 2021
46f295d
prove count_add
YaelDillies Sep 25, 2021
275ee5b
general golf + dropping some cardinal lemmas
YaelDillies Sep 25, 2021
1b4bf3f
revert revert set_of
YaelDillies Oct 1, 2021
a3635cc
break long line
YaelDillies Oct 1, 2021
ee74b67
Merge remote-tracking branch 'origin/master' into count
YaelDillies Oct 12, 2021
f9957c0
Merge branch 'master' into BoltonBailey/prime-counting
Smaug123 Oct 15, 2021
d055a47
Expand a couple of sorries
Smaug123 Oct 15, 2021
e4ca167
more progress on bound
BoltonBailey Oct 16, 2021
e133e47
anothe potentially useful todo
BoltonBailey Oct 16, 2021
cb531f4
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Oct 25, 2021
aa918af
fix lines too long
BoltonBailey Oct 25, 2021
1ab44fb
new lemmas and todos
BoltonBailey Oct 25, 2021
6ad4c89
replace work with sorry
BoltonBailey Oct 25, 2021
5bf3301
Merge branch 'master' into count
YaelDillies Oct 31, 2021
f879efb
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Nov 8, 2021
db6b755
proof work
BoltonBailey Nov 8, 2021
cee0fdb
changes
BoltonBailey Nov 8, 2021
63af1dc
found lemma in library
BoltonBailey Nov 8, 2021
ac68f03
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Nov 19, 2021
0d3b2c4
add filter_erase
BoltonBailey Nov 19, 2021
b2e7f92
Merge branch 'BoltonBailey/filter_erase' of github.com:leanprover-com…
BoltonBailey Nov 19, 2021
874b53b
use filter erase
BoltonBailey Nov 19, 2021
12a5719
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Nov 28, 2021
e28c284
remove todo
BoltonBailey Nov 28, 2021
409b67a
refactor
BoltonBailey Nov 29, 2021
1f73234
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Nov 29, 2021
0858715
remove todo
BoltonBailey Nov 29, 2021
e976266
finished lemma
BoltonBailey Nov 29, 2021
1a00464
complete proof
BoltonBailey Nov 29, 2021
93ebfee
finish eq_or_coprime_of_lt_prime
BoltonBailey Nov 29, 2021
60670a0
another lemma
BoltonBailey Nov 29, 2021
f8e2166
golf
BoltonBailey Nov 29, 2021
8fd3083
golf
BoltonBailey Nov 29, 2021
5c28394
golf Ico_eq_insert_Ico_succ
BoltonBailey Nov 29, 2021
9ac7b38
golf coprime_add_iff_coprime
BoltonBailey Nov 29, 2021
c0ea254
note todos with coprimes
BoltonBailey Nov 29, 2021
10da518
remove assumption
BoltonBailey Nov 29, 2021
23a08a5
todo note
BoltonBailey Nov 29, 2021
1dcf05e
rename
BoltonBailey Nov 29, 2021
b42f059
modularize split_range
BoltonBailey Nov 30, 2021
f8f738a
Generalize split_ico to typeclass
BoltonBailey Nov 30, 2021
062b725
arrow unicode
BoltonBailey Nov 30, 2021
62a53fc
fix lines too long
BoltonBailey Nov 30, 2021
c4eaff1
remove unnecessary typeclass
BoltonBailey Nov 30, 2021
bfa40e6
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Nov 30, 2021
e15234e
use library function
BoltonBailey Nov 30, 2021
fb2cb95
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Dec 2, 2021
39ed768
delete lemmas spun out
BoltonBailey Dec 2, 2021
c53bdaf
delete lemma from library
BoltonBailey Dec 3, 2021
77d6bf0
golfing lemmas
BoltonBailey Dec 3, 2021
043af94
more golfing
BoltonBailey Dec 3, 2021
352ee25
golf
BoltonBailey Dec 3, 2021
1f1ad63
golf
BoltonBailey Dec 3, 2021
8fe5287
fix line too long
BoltonBailey Dec 3, 2021
f6fabdd
golf
BoltonBailey Dec 3, 2021
afefa39
remove unnecessary lemmas
BoltonBailey Dec 3, 2021
58cd49e
delete old comments
BoltonBailey Dec 3, 2021
6e2bd31
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Dec 3, 2021
59c267e
simplification
BoltonBailey Dec 3, 2021
ca37fd6
Update prime_counting.lean
BoltonBailey Dec 4, 2021
43df3f2
cleanup
BoltonBailey Dec 4, 2021
9c507bb
Merge branch 'BoltonBailey/prime-counting' of github.com:leanprover-c…
BoltonBailey Dec 4, 2021
79579de
generalize
BoltonBailey Dec 4, 2021
2902652
fix line
BoltonBailey Dec 4, 2021
c91948a
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Dec 4, 2021
6ff8cd9
docs
BoltonBailey Dec 4, 2021
38469de
remove extraneous/unneeded imports
BoltonBailey Dec 4, 2021
c3c2864
remove completed todo
BoltonBailey Dec 4, 2021
e7c64ca
move intermediate lemmas to other files
BoltonBailey Dec 5, 2021
016f4e8
added application
BoltonBailey Dec 5, 2021
b3f0c7c
Update src/data/nat/count.lean
BoltonBailey Dec 5, 2021
9064f16
Update src/data/nat/lattice.lean
BoltonBailey Dec 5, 2021
395c717
Update src/data/nat/count.lean
BoltonBailey Dec 5, 2021
7b9e932
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Dec 5, 2021
9fa1097
revert low level change
BoltonBailey Dec 5, 2021
b3b49ee
fix
BoltonBailey Dec 5, 2021
b21cac9
fix proofs
BoltonBailey Dec 5, 2021
c0ee4db
comment out lemmas
BoltonBailey Dec 5, 2021
e9d5abf
update main theorem
BoltonBailey Dec 5, 2021
6b8aa74
fix name in main results
BoltonBailey Dec 5, 2021
d984827
fix other theorem
BoltonBailey Dec 6, 2021
4f2621e
golfing
BoltonBailey Dec 6, 2021
7765343
fix lattice.lean
BoltonBailey Dec 6, 2021
97eafda
docs for nth vs order_iso_of_nat
BoltonBailey Dec 6, 2021
7ffe8b3
Merge branch 'count' of github.com:leanprover-community/mathlib into …
BoltonBailey Dec 6, 2021
19584ba
complete explicit bound
BoltonBailey Dec 6, 2021
8ca3640
golf
BoltonBailey Dec 6, 2021
ec11436
golf
BoltonBailey Dec 6, 2021
296e25e
golf
BoltonBailey Dec 7, 2021
8236831
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Dec 19, 2021
b7fa947
revert lattice
BoltonBailey Dec 19, 2021
2e77afa
fix simp
BoltonBailey Dec 19, 2021
adfb0d4
revert
BoltonBailey Dec 20, 2021
64dd895
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Jan 13, 2022
dd9344c
fix totient proof
BoltonBailey Jan 13, 2022
3fc72a5
fix namespaces
BoltonBailey Jan 13, 2022
013d398
remove todos
BoltonBailey Jan 13, 2022
3a88613
rewrite proof and dedent
BoltonBailey Jan 13, 2022
89023d3
more finset
BoltonBailey Jan 13, 2022
5ffead5
golf
BoltonBailey Jan 13, 2022
246d619
new version
BoltonBailey Jan 14, 2022
41ff5b8
comments and golf
BoltonBailey Jan 14, 2022
c9f3ac2
Merge branch 'master' into BoltonBailey/prime-counting
BoltonBailey Jan 15, 2022
40626dd
golf
BoltonBailey Jan 15, 2022
ac10003
arrows
BoltonBailey Jan 15, 2022
5c2089d
Merge branch 'BoltonBailey/prime-counting' of github.com:leanprover-c…
BoltonBailey Jan 15, 2022
0dbd8d9
dedent
BoltonBailey Jan 15, 2022
d68f144
rename
BoltonBailey Jan 15, 2022
8b7f2a4
rcases tactic
BoltonBailey Jan 15, 2022
a806495
simplify
BoltonBailey Jan 15, 2022
136707c
elaborate doc
BoltonBailey Jan 15, 2022
9e48bb4
protect def of nat.card
BoltonBailey Jan 15, 2022
c7743d8
rewrite coprimes
BoltonBailey Jan 15, 2022
a32a57e
remove explicit bound
BoltonBailey Jan 18, 2022
9fa600d
dedent calc
BoltonBailey Jan 27, 2022
55bb1c3
remove todo
BoltonBailey Jan 27, 2022
c04918b
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Jan 27, 2022
494d171
remove comma
BoltonBailey Jan 27, 2022
ed80701
j renamed to i
BoltonBailey Jan 27, 2022
92c5f0e
simp only
BoltonBailey Jan 27, 2022
818294e
golf
BoltonBailey Jan 27, 2022
ffffa79
golf
BoltonBailey Jan 27, 2022
9c6fc3f
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 28, 2022
97ac1cd
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 28, 2022
79c7ed5
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 28, 2022
50ddecb
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 28, 2022
0e64417
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 28, 2022
a3a2ce3
unprivate and move
BoltonBailey Jan 28, 2022
d186256
shorten lemma
BoltonBailey Jan 28, 2022
d18d320
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Jan 28, 2022
8bb49f9
golf
BoltonBailey Jan 28, 2022
1437421
to equals
BoltonBailey Jan 28, 2022
b5e4683
to le
BoltonBailey Jan 28, 2022
254131a
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Jan 29, 2022
9d9da9f
Update src/data/nat/totient.lean
BoltonBailey Jan 29, 2022
bb15b1d
Update src/data/nat/totient.lean
BoltonBailey Jan 29, 2022
da3aa4c
implicit variables
BoltonBailey Jan 29, 2022
f4ede29
Merge branch 'BoltonBailey/prime-counting' of github.com:leanprover-c…
BoltonBailey Jan 29, 2022
b267bdf
fix call
BoltonBailey Jan 29, 2022
f678f6c
Update src/number_theory/prime_counting.lean
BoltonBailey Jan 29, 2022
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
43 changes: 41 additions & 2 deletions src/data/nat/totient.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ import algebra.big_operators.basic
import data.nat.prime
import data.zmod.basic
import ring_theory.multiplicity
import data.nat.periodic
import algebra.char_p.two

/-!
Expand All @@ -26,7 +27,7 @@ namespace nat

/-- Euler's totient function. This counts the number of naturals strictly less than `n` which are
coprime with `n`. -/
def totient (n : ℕ) : ℕ := ((range n).filter (nat.coprime n)).card
def totient (n : ℕ) : ℕ := ((range n).filter n.coprime).card

localized "notation `φ` := nat.totient" in nat

Expand All @@ -35,7 +36,7 @@ localized "notation `φ` := nat.totient" in nat
@[simp] theorem totient_one : φ 1 = 1 :=
by simp [totient]

lemma totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter (nat.coprime n)).card := rfl
lemma totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter (coprime n)).card := rfl
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved

lemma totient_le (n : ℕ) : φ n ≤ n :=
calc totient n ≤ (range n).card : card_filter_le _ _
Expand All @@ -57,6 +58,44 @@ lemma totient_pos : ∀ {n : ℕ}, 0 < n → 0 < φ n
| 1 := by simp [totient]
| (n+2) := λ h, card_pos.2 ⟨1, mem_filter.2 ⟨mem_range.2 dec_trivial, coprime_one_right _⟩⟩

lemma filter_coprime_Ico_eq_totient (a n : ℕ) :
((Ico n (n+a)).filter (coprime a)).card = totient a :=
begin
rw [totient, filter_Ico_card_eq_of_periodic, count_eq_card_filter_range],
exact periodic_coprime a,
end

lemma Ico_filter_coprime_le (a k n : ℕ) (a_pos : 0 < a) :
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
((Ico k (k + n)).filter (coprime a)).card ≤ totient a * (n / a + 1) :=
begin
conv_lhs { rw ←nat.mod_add_div n a },
induction n / a with i ih,
{ rw ←filter_coprime_Ico_eq_totient a k,
simp only [add_zero, mul_one, mul_zero, le_of_lt (mod_lt n a_pos)],
mono,
refine monotone_filter_left a.coprime _,
simp only [finset.le_eq_subset],
exact Ico_subset_Ico rfl.le (add_le_add_left (le_of_lt (mod_lt n a_pos)) k), },
simp only [mul_succ],
simp_rw ←add_assoc at ih ⊢,
calc (filter a.coprime (Ico k (k + n % a + a * i + a))).card
= (filter a.coprime (Ico k (k + n % a + a * i)
∪ Ico (k + n % a + a * i) (k + n % a + a * i + a))).card :
begin
congr,
rw Ico_union_Ico_eq_Ico,
rw add_assoc,
exact le_self_add,
exact le_self_add,
end
... ≤ (filter a.coprime (Ico k (k + n % a + a * i))).card + a.totient :
begin
rw [filter_union, ←filter_coprime_Ico_eq_totient a (k + n % a + a * i)],
apply card_union_le,
end
... ≤ a.totient * i + a.totient + a.totient : add_le_add_right ih (totient a),
end

open zmod

/-- Note this takes an explicit `fintype ((zmod n)ˣ)` argument to avoid trouble with instance
Expand Down
86 changes: 86 additions & 0 deletions src/number_theory/prime_counting.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,86 @@
/-
Copyright (c) 2021 Bolton Bailey. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bolton Bailey
-/

import data.nat.prime
import data.nat.totient
import algebra.periodic
import data.finset.locally_finite
import data.nat.count


/-!
# The Prime Counting Function

In this file we define the prime counting function: the function on natural numbers that returns
the number of primes less than or equal to its input.

## Main Results

The main definitions for this file are

- `prime_counting`: The prime counting function π
- `prime_counting'`: π(n - 1)

We then prove that these are monotone. The last main theorem is an upper bound on `π'` which arises
by observing that all numbers greater than `k` and not coprime to `k` are not prime, and so only at
most `φ(k)/k` fraction of the numbers from `k` to `n` are prime.
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved

## Notation

We use the standard notation `π` to represent the prime counting function (and `π'` to represent
the reindexed version).

-/

namespace nat
open finset

/--
A variant of the traditional prime counting function which gives the number of primes
*strictly* less than the input. More convenient for avoiding off-by-one errors.
-/
def prime_counting' : ℕ → ℕ := nat.count prime

/-- The prime counting function: Returns the number of primes less than or equal to the input. -/
def prime_counting (n : ℕ) : ℕ := prime_counting' (n + 1)

localized "notation `π` := nat.prime_counting" in nat
localized "notation `π'` := nat.prime_counting'" in nat

lemma monotone_prime_counting' : monotone prime_counting' := count_monotone prime

lemma monotone_prime_counting : monotone prime_counting :=
λ a b a_le_b, monotone_prime_counting' (add_le_add_right a_le_b 1)

/-- A linear upper bound on the size of the `prime_counting'` function -/
lemma prime_counting'_add_le (n k a : ℕ) (h0 : 0 < a) (h1 : a < k) :
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
π' (k + n) ≤ π' k + nat.totient a * (n / a + 1) :=
calc π' (k + n)
≤ ((range k).filter (prime)).card + ((Ico k (k + n)).filter (prime)).card :
begin
rw [prime_counting', count_eq_card_filter_range, range_eq_Ico,
←Ico_union_Ico_eq_Ico (zero_le k) (le_self_add), filter_union],
apply card_union_le,
end
... ≤ π' k + ((Ico k (k + n)).filter (prime)).card :
by rw [prime_counting', count_eq_card_filter_range]
... ≤ π' k + ((Ico k (k + n)).filter (coprime a)).card :
begin
refine add_le_add_left (card_le_of_subset _) k.prime_counting',
simp only [subset_iff, and_imp, mem_filter, mem_Ico],
intros p succ_k_le_p p_lt_n p_prime,
split,
{ exact ⟨succ_k_le_p, p_lt_n⟩, },
{ rw coprime_comm,
exact coprime_of_lt_prime h0 (gt_of_ge_of_gt succ_k_le_p h1) p_prime, },
end
... ≤ π' k + totient a * (n / a + 1) :
begin
rw [add_le_add_iff_left],
exact Ico_filter_coprime_le a k n h0,
end

end nat
6 changes: 3 additions & 3 deletions src/set_theory/fincard.lean
Original file line number Diff line number Diff line change
Expand Up @@ -26,13 +26,13 @@ namespace nat

/-- `nat.card α` is the cardinality of `α` as a natural number.
If `α` is infinite, `nat.card α = 0`. -/
def card (α : Type*) : ℕ := (mk α).to_nat
protected def card (α : Type*) : ℕ := (mk α).to_nat

@[simp]
lemma card_eq_fintype_card [fintype α] : card α = fintype.card α := mk_to_nat_eq_card
lemma card_eq_fintype_card [fintype α] : nat.card α = fintype.card α := mk_to_nat_eq_card

@[simp]
lemma card_eq_zero_of_infinite [infinite α] : card α = 0 := mk_to_nat_of_infinite
lemma card_eq_zero_of_infinite [infinite α] : nat.card α = 0 := mk_to_nat_of_infinite

end nat

Expand Down