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

Commit c10d1b1

Browse files
feat(ring_theory/polynomial/cyclotomic): add order_of_root_cyclotomic (#5151)
Two lemmas about roots of cyclotomic polynomials modulo `p`. `order_of_root_cyclotomic` is the main algebraic tool to prove the existence of infinitely many primes congruent to `1` modulo `n`. Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent c939c9e commit c10d1b1

File tree

2 files changed

+110
-1
lines changed

2 files changed

+110
-1
lines changed

src/number_theory/divisors.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -126,6 +126,19 @@ begin
126126
exact nat.le_of_dvd (nat.succ_pos m),
127127
end
128128

129+
lemma divisors_subset_of_dvd {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) : divisors m ⊆ divisors n :=
130+
finset.subset_iff.2 $ λ x hx, nat.mem_divisors.mpr (⟨dvd.trans (nat.mem_divisors.mp hx).1 h, hzero⟩)
131+
132+
lemma divisors_subset_proper_divisors {m : ℕ} (hzero : n ≠ 0) (h : m ∣ n) (hdiff : m ≠ n) :
133+
divisors m ⊆ proper_divisors n :=
134+
begin
135+
apply finset.subset_iff.2,
136+
intros x hx,
137+
exact nat.mem_proper_divisors.2 (⟨dvd.trans (nat.mem_divisors.1 hx).1 h,
138+
lt_of_le_of_lt (divisor_le hx) (lt_of_le_of_ne (divisor_le (nat.mem_divisors.2
139+
⟨h, hzero⟩)) hdiff)⟩)
140+
end
141+
129142
@[simp]
130143
lemma divisors_zero : divisors 0 = ∅ := by { ext, simp }
131144

src/ring_theory/polynomial/cyclotomic.lean

Lines changed: 97 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -10,6 +10,7 @@ import algebra.polynomial.big_operators
1010
import number_theory.arithmetic_function
1111
import data.polynomial.lifts
1212
import analysis.complex.roots_of_unity
13+
import field_theory.separable
1314

1415
/-!
1516
# Cyclotomic polynomials.
@@ -455,6 +456,20 @@ begin
455456
exact monic.ne_zero prod_monic (degree_eq_bot.1 h)
456457
end
457458

459+
/-- If `m` is a proper divisor of `n`, then `X ^ m - 1` divides
460+
`∏ i in nat.proper_divisors n, cyclotomic i R`. -/
461+
lemma X_pow_sub_one_dvd_prod_cyclotomic (R : Type*) [comm_ring R] {n m : ℕ} (hpos : 0 < n)
462+
(hm : m ∣ n) (hdiff : m ≠ n) : X ^ m - 1 ∣ ∏ i in nat.proper_divisors n, cyclotomic i R :=
463+
begin
464+
replace hm := nat.mem_proper_divisors.2 ⟨hm, lt_of_le_of_ne (nat.divisor_le (nat.mem_divisors.2
465+
⟨hm, (ne_of_lt hpos).symm⟩)) hdiff⟩,
466+
rw [← finset.sdiff_union_of_subset (nat.divisors_subset_proper_divisors (ne_of_lt hpos).symm
467+
(nat.mem_proper_divisors.1 hm).1 (ne_of_lt (nat.mem_proper_divisors.1 hm).2)),
468+
finset.prod_union finset.sdiff_disjoint, prod_cyclotomic_eq_X_pow_sub_one
469+
(nat.pos_of_mem_proper_divisors hm)],
470+
exact ⟨(∏ (x : ℕ) in n.proper_divisors \ m.divisors, cyclotomic x R), by rw mul_comm⟩
471+
end
472+
458473
/-- If there is a primitive `n`-th root of unity in `K`, then
459474
`cyclotomic n K = ∏ μ in primitive_roots n R, (X - C μ)`. In particular,
460475
`cyclotomic n K = cyclotomic' n K` -/
@@ -510,7 +525,7 @@ begin
510525
end
511526

512527
/-- The constant term of `cyclotomic n R` is `1` if `2 ≤ n`. -/
513-
lemma cyclotomic_coeff_zero {R : Type*} [comm_ring R] (n : ℕ) (hn : 2 ≤ n) :
528+
lemma cyclotomic_coeff_zero (R : Type*) [comm_ring R] {n : ℕ} (hn : 2 ≤ n) :
514529
(cyclotomic n R).coeff 0 = 1 :=
515530
begin
516531
induction n using nat.strong_induction_on with n hi,
@@ -543,6 +558,87 @@ begin
543558
exact neg_inj.mp (eq.symm heq)
544559
end
545560

561+
/-- If `(a : ℕ)` is a root of `cyclotomic n (zmod p)`, where `p` is a prime, then `a` and `p` are
562+
coprime. -/
563+
lemma coprime_of_root_cyclotomic {n : ℕ} (hpos : 0 < n) {p : ℕ} [hprime : fact p.prime] {a : ℕ}
564+
(hroot : is_root (cyclotomic n (zmod p)) (nat.cast_ring_hom (zmod p) a)) :
565+
a.coprime p :=
566+
begin
567+
apply nat.coprime.symm,
568+
rw [nat.prime.coprime_iff_not_dvd hprime],
569+
by_contra h,
570+
replace h := (zmod.nat_coe_zmod_eq_zero_iff_dvd a p).2 h,
571+
rw [is_root.def, ring_hom.eq_nat_cast, h, ← coeff_zero_eq_eval_zero] at hroot,
572+
by_cases hone : n = 1,
573+
{ simp only [hone, cyclotomic_one, zero_sub, coeff_one_zero, coeff_X_zero, neg_eq_zero,
574+
one_ne_zero, coeff_sub] at hroot,
575+
exact hroot },
576+
rw [cyclotomic_coeff_zero (zmod p) (nat.succ_le_of_lt (lt_of_le_of_ne
577+
(nat.succ_le_of_lt hpos) (ne.symm hone)))] at hroot,
578+
exact one_ne_zero hroot
579+
end
580+
546581
end cyclotomic
547582

583+
section order
584+
585+
/-- If `(a : ℕ)` is a root of `cyclotomic n (zmod p)`, then the multiplicative order of `a` modulo
586+
`p` divides `n`. -/
587+
lemma order_of_root_cyclotomic_dvd {n : ℕ} (hpos : 0 < n) {p : ℕ} [hprime : fact p.prime]
588+
{a : ℕ} (hroot : is_root (cyclotomic n (zmod p)) (nat.cast_ring_hom (zmod p) a)) :
589+
order_of (zmod.unit_of_coprime a (coprime_of_root_cyclotomic hpos hroot)) ∣ n :=
590+
begin
591+
apply order_of_dvd_of_pow_eq_one,
592+
suffices hpow : eval (nat.cast_ring_hom (zmod p) a) (X ^ n - 1 : polynomial (zmod p)) = 0,
593+
{ simp only [eval_X, eval_one, eval_pow, eval_sub, ring_hom.eq_nat_cast] at hpow,
594+
apply units.coe_eq_one.1,
595+
simp only [sub_eq_zero.mp hpow, zmod.cast_unit_of_coprime, units.coe_pow] },
596+
rw [is_root.def] at hroot,
597+
rw [← prod_cyclotomic_eq_X_pow_sub_one hpos (zmod p),
598+
nat.divisors_eq_proper_divisors_insert_self_of_pos hpos,
599+
finset.prod_insert nat.proper_divisors.not_self_mem, eval_mul, hroot, zero_mul]
600+
end
601+
602+
/-- If `(a : ℕ)` is a root of `cyclotomic n (zmod p)`, where `p` is a prime that does not divide
603+
`n`, then the multiplicative order of `a` modulo `p` is exactly `n`. -/
604+
lemma order_of_root_cyclotomic {n : ℕ} (hpos : 0 < n) {p : ℕ} [hprime : fact p.prime] {a : ℕ}
605+
(hn : ¬ p ∣ n) (hroot : is_root (cyclotomic n (zmod p)) (nat.cast_ring_hom (zmod p) a)) :
606+
order_of (zmod.unit_of_coprime a (coprime_of_root_cyclotomic hpos hroot)) = n :=
607+
begin
608+
set m := order_of (zmod.unit_of_coprime a (coprime_of_root_cyclotomic hpos hroot)),
609+
have ha := coprime_of_root_cyclotomic hpos hroot,
610+
have hdivcycl : map (int.cast_ring_hom (zmod p)) (X - a) ∣ (cyclotomic n (zmod p)),
611+
{ replace hrootdiv := dvd_iff_is_root.2 hroot,
612+
simp only [C_eq_nat_cast, ring_hom.eq_nat_cast] at hrootdiv,
613+
simp only [hrootdiv, map_nat_cast, map_X, map_sub] },
614+
by_contra hdiff,
615+
have hdiv : map (int.cast_ring_hom (zmod p)) (X - a) ∣
616+
∏ i in nat.proper_divisors n, cyclotomic i (zmod p),
617+
{ suffices hdivm : map (int.cast_ring_hom (zmod p)) (X - a) ∣ X ^ m - 1,
618+
{ exact dvd_trans hdivm (X_pow_sub_one_dvd_prod_cyclotomic (zmod p) hpos
619+
(order_of_root_cyclotomic_dvd hpos hroot) hdiff) },
620+
rw [map_sub, map_X, map_nat_cast, ← C_eq_nat_cast, dvd_iff_is_root, is_root.def, eval_sub,
621+
eval_pow, eval_one, eval_X, sub_eq_zero, ← zmod.cast_unit_of_coprime a ha, ← units.coe_pow,
622+
units.coe_eq_one],
623+
exact pow_order_of_eq_one (zmod.unit_of_coprime a ha) },
624+
have habs : (map (int.cast_ring_hom (zmod p)) (X - a)) ^ 2 ∣ X ^ n - 1,
625+
{ obtain ⟨P, hP⟩ := hdivcycl,
626+
obtain ⟨Q, hQ⟩ := hdiv,
627+
rw [← prod_cyclotomic_eq_X_pow_sub_one hpos, nat.divisors_eq_proper_divisors_insert_self_of_pos
628+
hpos, finset.prod_insert nat.proper_divisors.not_self_mem, hP, hQ],
629+
exact ⟨P * Q, by ring⟩ },
630+
have hnzero : ↑n ≠ (0 : (zmod p)),
631+
{ intro ha,
632+
exact hn (int.coe_nat_dvd.1 ((zmod.int_coe_zmod_eq_zero_iff_dvd n p).1 ha)) },
633+
rw [pow_two] at habs,
634+
replace habs := squarefree_X_pow_sub_C (1 : (zmod p)) hnzero one_ne_zero
635+
(map (int.cast_ring_hom (zmod p)) (X - a)) habs,
636+
simp only [map_nat_cast, map_X, map_sub] at habs,
637+
replace habs := degree_eq_zero_of_is_unit habs,
638+
rw [← C_eq_nat_cast, degree_X_sub_C] at habs,
639+
norm_cast at habs
640+
end
641+
642+
end order
643+
548644
end polynomial

0 commit comments

Comments
 (0)