Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(Mathlib.NumberTheory.Cyclotomic.Three): add various results #13798

Closed
wants to merge 33 commits into from
Closed
Show file tree
Hide file tree
Changes from 15 commits
Commits
Show all changes
33 commits
Select commit Hold shift + click to select a range
0b135a1
add fintypeQuotienttoIntegerSubOne
riccardobrasca Jun 13, 2024
fdddcf7
add card_quotient_toInteger_sub_one
riccardobrasca Jun 13, 2024
4c64691
add ringEquivOfPrime
riccardobrasca Jun 13, 2024
403721c
Update Mathlib/NumberTheory/Cyclotomic/Rat.lean
riccardobrasca Jun 13, 2024
80a7a52
Update Mathlib/NumberTheory/Cyclotomic/Rat.lean
riccardobrasca Jun 13, 2024
5062527
add Finset.univ_of_card_eq_three
riccardobrasca Jun 13, 2024
a0f4a47
add coe_toInteger
riccardobrasca Jun 13, 2024
96001ad
add charP_of_card_eq_prime
riccardobrasca Jun 13, 2024
558df8c
add comment
riccardobrasca Jun 13, 2024
f175627
doc
riccardobrasca Jun 13, 2024
e34ee41
add lambda_pow_four_dvd_cube_sub_one_or_add_one_of_lambda_not_dvd
riccardobrasca Jun 13, 2024
3615b27
use better notation
riccardobrasca Jun 13, 2024
b4f6fe9
typo
riccardobrasca Jun 13, 2024
a681633
use Fact
riccardobrasca Jun 13, 2024
baf2ce9
fix build
riccardobrasca Jun 13, 2024
aeafa8e
use Finite and Nat.card
riccardobrasca Jun 17, 2024
80e9313
fix name
riccardobrasca Jun 17, 2024
55152d7
fix name
riccardobrasca Jun 17, 2024
1aea1c6
golf
riccardobrasca Jun 17, 2024
fec9005
use mul_dvd_mul_left
riccardobrasca Jun 17, 2024
fe84a5d
golf
riccardobrasca Jun 17, 2024
7b42ecb
add ringEquivOfPrime_eq_ringEquiv
riccardobrasca Jun 17, 2024
649761b
Update Mathlib/Data/ZMod/Basic.lean
riccardobrasca Jun 17, 2024
515b3ff
Merge remote-tracking branch 'origin/master' into RB/flt_lemmas
riccardobrasca Jun 17, 2024
8c342d3
Update Mathlib/NumberTheory/Cyclotomic/Three.lean
riccardobrasca Jun 18, 2024
a295aff
Update Mathlib/NumberTheory/Cyclotomic/Three.lean
riccardobrasca Jun 18, 2024
3fab841
Update Mathlib/NumberTheory/Cyclotomic/Three.lean
riccardobrasca Jun 18, 2024
4294249
Update Mathlib/NumberTheory/Cyclotomic/Three.lean
riccardobrasca Jun 18, 2024
aafcade
Update Mathlib/NumberTheory/Cyclotomic/Three.lean
riccardobrasca Jun 18, 2024
4e3fdd2
useless
riccardobrasca Jun 18, 2024
918ba7c
use le instead of eq
riccardobrasca Jun 25, 2024
10f5c5f
fix build
riccardobrasca Jun 25, 2024
f68c15c
fix name
riccardobrasca Jun 25, 2024
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
5 changes: 5 additions & 0 deletions Mathlib/Algebra/CharP/CharAndCard.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,3 +82,8 @@ theorem not_isUnit_prime_of_dvd_card {R : Type*} [CommRing R] [Fintype R] (p :
mt (isUnit_iff_not_dvd_char R p).mp
(Classical.not_not.mpr ((prime_dvd_char_iff_dvd_card p).mpr hp))
#align not_is_unit_prime_of_dvd_card not_isUnit_prime_of_dvd_card

lemma charP_of_card_eq_prime {R : Type*} [NonAssocRing R] [Fintype R] (p : ℕ) [hp : Fact p.Prime]
(hR : Fintype.card R = p) : CharP R p :=
have := Fintype.one_lt_card_iff_nontrivial.1 (hR ▸ hp.1.one_lt)
(CharP.charP_iff_prime_eq_zero hp.1).2 (hR ▸ Nat.cast_card_eq_zero R)
8 changes: 8 additions & 0 deletions Mathlib/Data/ZMod/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -491,6 +491,14 @@ noncomputable def ringEquiv [Fintype R] (h : Fintype.card R = n) : ZMod n ≃+*
RingEquiv.ofBijective _ (ZMod.castHom_bijective R h)
#align zmod.ring_equiv ZMod.ringEquiv

/-- The unique ring isomorphism between `ZMod p` and a ring `R` of cardinality a prime `p`. -/
noncomputable def ringEquivOfPrime [Fintype R] {p : ℕ} (hp : p.Prime) (hR : Fintype.card R = p) :
ZMod p ≃+* R :=
have : Nontrivial R := Fintype.one_lt_card_iff_nontrivial.1 (hR ▸ hp.one_lt)
-- The following line exists as `charP_of_card_eq_prime` in `Mathlib.Algebra.CharP.CharAndCard`.
have : CharP R p := (CharP.charP_iff_prime_eq_zero hp).2 (hR ▸ Nat.cast_card_eq_zero R)
ZMod.ringEquiv R hR

/-- The identity between `ZMod m` and `ZMod n` when `m = n`, as a ring isomorphism. -/
def ringEquivCongr {m n : ℕ} (h : m = n) : ZMod m ≃+* ZMod n := by
cases' m with m <;> cases' n with n
Expand Down
20 changes: 20 additions & 0 deletions Mathlib/NumberTheory/Cyclotomic/Rat.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,6 +182,26 @@ noncomputable def integralPowerBasis [IsCyclotomicExtension {p ^ k} ℚ K]
/-- Abbreviation to see a primitive root of unity as a member of the ring of integers. -/
abbrev toInteger {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) : 𝓞 K := ⟨ζ, hζ.isIntegral k.pos⟩

lemma coe_toInteger {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) : hζ.toInteger.1 = ζ := rfl

/-- `𝓞 K ⧸ Ideal.span {ζ - 1}` is finite. -/
noncomputable
def fintypeQuotienttoIntegerSubOne [NumberField K] {k : ℕ+} (hk : 1 < k)
(hζ : IsPrimitiveRoot ζ k) : Fintype (𝓞 K ⧸ Ideal.span {hζ.toInteger - 1}) := by
refine Ideal.fintypeQuotientOfFreeOfNeBot _ (fun h ↦ ?_)
simp only [Ideal.span_singleton_eq_bot, sub_eq_zero, ← Subtype.coe_inj] at h
exact hζ.ne_one hk (RingOfIntegers.ext_iff.1 h)

/-- We have that `𝓞 K ⧸ Ideal.span {ζ - 1}` has cardinality equal to the norm of `ζ - 1`.

See the results below to compute this norm in various cases. -/
lemma card_quotient_toInteger_sub_one [NumberField K] {k : ℕ+} (hk : 1 < k)
(hζ : IsPrimitiveRoot ζ k) :
letI _ := hζ.fintypeQuotienttoIntegerSubOne hk
Fintype.card (𝓞 K ⧸ Ideal.span {hζ.toInteger - 1}) =
(Algebra.norm ℤ (hζ.toInteger - 1)).natAbs := by
rw [← Submodule.cardQuot_apply, ← Ideal.absNorm_apply, Ideal.absNorm_span_singleton]

lemma toInteger_isPrimitiveRoot {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) :
IsPrimitiveRoot hζ.toInteger k :=
IsPrimitiveRoot.of_map_of_injective (by exact hζ) RingOfIntegers.coe_injective
Expand Down
84 changes: 82 additions & 2 deletions Mathlib/NumberTheory/Cyclotomic/Three.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ Authors: Riccardo Brasca, Pietro Monticone
import Mathlib.NumberTheory.Cyclotomic.Embeddings
import Mathlib.NumberTheory.Cyclotomic.Rat
import Mathlib.NumberTheory.NumberField.Units.DirichletTheorem
import Mathlib.RingTheory.Fintype

/-!
# Third Cyclotomic Field
Expand All @@ -33,7 +34,9 @@ namespace IsCyclotomicExtension.Rat.Three
variable {K : Type*} [Field K] [NumberField K] [IsCyclotomicExtension {3} ℚ K]
variable {ζ : K} (hζ : IsPrimitiveRoot ζ ↑(3 : ℕ+)) (u : (𝓞 K)ˣ)
local notation3 "η" => (IsPrimitiveRoot.isUnit (hζ.toInteger_isPrimitiveRoot) (by decide)).unit
local notation3 "λ" => (η : 𝓞 K) - 1
local notation3 "λ" => hζ.toInteger - 1

lemma coe_eta : (η : 𝓞 K) = hζ.toInteger := rfl

/-- Let `u` be a unit in `(𝓞 K)ˣ`, then `u ∈ [1, -1, η, -η, η^2, -η^2]`. -/
-- Here `List` is more convenient than `Finset`, even if further from the informal statement.
Expand Down Expand Up @@ -70,7 +73,8 @@ theorem Units.mem : u ∈ [1, -1, η, -η, η ^ 2, -η ^ 2] := by
/-- We have that `λ ^ 2 = -3 * η`. -/
private lemma lambda_sq : λ ^ 2 = -3 * η := by
ext
calc (λ ^ 2 : K) = η ^ 2 + η + 1 - 3 * η := by ring
calc (λ ^ 2 : K) = η ^ 2 + η + 1 - 3 * η := by
simp only [RingOfIntegers.map_mk, IsUnit.unit_spec]; ring
_ = 0 - 3 * η := by simpa using hζ.isRoot_cyclotomic (by decide)
_ = -3 * η := by ring

Expand Down Expand Up @@ -109,3 +113,79 @@ theorem eq_one_or_neg_one_of_unit_of_congruent (hcong : ∃ n : ℤ, λ ^ 2 ∣
have : (hζ.pow_of_coprime 2 (by decide)).toInteger = hζ.toInteger ^ 2 := by ext; simp
simp only [this, PNat.val_ofNat, Nat.cast_ofNat, mul_neg, Int.cast_neg, ← neg_add, ←
sub_eq_iff_eq_add.1 hx, Units.val_neg, val_pow_eq_pow_val, IsUnit.unit_spec, neg_neg]

variable (x : 𝓞 K)

/-- Let `(x : 𝓞 K)`. Then we have that `λ` divides one amongst `x`, `x - 1` and `x + 1`. -/
lemma dvd_or_dvd_sub_one_or_dvd_add_one : λ ∣ x ∨ λ ∣ x - 1 ∨ λ ∣ x + 1 := by
classical
let _ := hζ.fintypeQuotienttoIntegerSubOne (by decide)
have := Finset.mem_univ (Ideal.Quotient.mk (Ideal.span {λ}) x)
rw [Finset.univ_of_card_eq_three] at this
· simp only [Finset.mem_insert, Finset.mem_singleton] at this
rcases this with (h | h | h)
· left
exact Ideal.mem_span_singleton.1 <| Ideal.Quotient.eq_zero_iff_mem.1 h
· right; left
refine Ideal.mem_span_singleton.1 <| Ideal.Quotient.eq_zero_iff_mem.1 ?_
rw [RingHom.map_sub, h, RingHom.map_one, sub_self]
· right; right
refine Ideal.mem_span_singleton.1 <| Ideal.Quotient.eq_zero_iff_mem.1 ?_
rw [RingHom.map_add, h, RingHom.map_one, add_left_neg]
· rw [hζ.card_quotient_toInteger_sub_one, hζ.norm_toInteger_sub_one_of_prime_ne_two' (by decide)]
simp

/-- We have that `η ^ 2 + η + 1 = 0`. -/
lemma eta_eq_add_eta_add_one : (η : 𝓞 K) ^ 2 + η + 1 = 0 := by
ext; simpa using hζ.isRoot_cyclotomic (by decide)

/-- We have that `x ^ 3 - 1 = (x - 1) * (x - η) * (x - η ^ 2)`. -/
lemma cube_sub_one_eq_mul : x ^ 3 - 1 = (x - 1) * (x - η) * (x - η ^ 2) := by
symm
calc _ = x ^ 3 - x ^ 2 * (η ^ 2 + η + 1) + x * (η ^ 2 + η + η ^ 3) - η ^ 3 := by ring
_ = x ^ 3 - x ^ 2 * (η ^ 2 + η + 1) + x * (η ^ 2 + η + 1) - 1 := by
simp [show hζ.toInteger ^ 3 = 1 from hζ.toInteger_isPrimitiveRoot.pow_eq_one]
_ = x ^ 3 - 1 := by rw [eta_eq_add_eta_add_one hζ]; ring

/-- We have that `λ` divides `x * (x - 1) * (x - (η + 1))`. -/
lemma lambda_dvd_mul_sub_one_mul_sub_eta_add_one : λ ∣ x * (x - 1) * (x - (η + 1)) := by
rcases dvd_or_dvd_sub_one_or_dvd_add_one hζ x with (h | h | h)
· exact dvd_mul_of_dvd_left (dvd_mul_of_dvd_left h _) _
· exact dvd_mul_of_dvd_left (dvd_mul_of_dvd_right h _) _
· refine dvd_mul_of_dvd_right ?_ _
rw [show x - (η + 1) = x + 1 - (η - 1 + 3) by ring]
exact dvd_sub h <| dvd_add dvd_rfl hζ.toInteger_sub_one_dvd_prime'

/-- If `λ` divides `x - 1`, then `λ ^ 4` divides `x ^ 3 - 1`. -/
lemma lambda_pow_four_dvd_cube_sub_one_of_dvd_sub_one {x : 𝓞 K} (h : λ ∣ x - 1) :
λ ^ 4 ∣ x ^ 3 - 1 := by
obtain ⟨y, hy⟩ := h
have : x ^ 3 - 1 = λ ^ 3 * (y * (y - 1) * (y - (η + 1))) := by
calc _ = (x - 1) * (x - 1 - λ) * (x - 1 - λ * (η + 1)) := by
simp only [coe_eta, cube_sub_one_eq_mul hζ x]; ring
_ = _ := by rw [hy]; ring
rw [this, show λ ^ 4 = λ ^ 3 * λ by ring]
exact mul_dvd_mul dvd_rfl (lambda_dvd_mul_sub_one_mul_sub_eta_add_one hζ y)

/-- If `λ` divides `x + 1`, then `λ ^ 4` divides `x ^ 3 + 1`. -/
lemma lambda_pow_four_dvd_cube_add_one_of_dvd_add_one {x : 𝓞 K} (h : λ ∣ x + 1) :
λ ^ 4 ∣ x ^ 3 + 1 := by
replace h : λ ∣ -x - 1 := by
obtain ⟨y, hy⟩ := h
refine ⟨-y, ?_⟩
rw [mul_neg, ← hy]
ring
obtain ⟨y, hy⟩ := lambda_pow_four_dvd_cube_sub_one_of_dvd_sub_one hζ h
refine ⟨-y, ?_⟩
rw [mul_neg, ← hy]
ring

/-- If `λ` does not divide `x`, then `λ ^ 4` divides `x ^ 3 - 1` or `x ^ 3 + 1`. -/
lemma lambda_pow_four_dvd_cube_sub_one_or_add_one_of_lambda_not_dvd {x : 𝓞 K} (h : ¬ λ ∣ x) :
λ ^ 4 ∣ x ^ 3 - 1 ∨ λ ^ 4 ∣ x ^ 3 + 1 := by
rcases dvd_or_dvd_sub_one_or_dvd_add_one hζ x with (H | H | H)
· contradiction
· left
exact lambda_pow_four_dvd_cube_sub_one_of_dvd_sub_one hζ H
· right
exact lambda_pow_four_dvd_cube_add_one_of_dvd_add_one hζ H
4 changes: 2 additions & 2 deletions Mathlib/NumberTheory/FLT/Three.lean
Original file line number Diff line number Diff line change
Expand Up @@ -120,8 +120,8 @@ open NumberField
variable {K : Type*} [Field K] [NumberField K] [IsCyclotomicExtension {3} ℚ K]
variable {ζ : K} (hζ : IsPrimitiveRoot ζ (3 : ℕ+))

local notation3 "η" => hζ.toInteger
local notation3 "λ" => η - 1
local notation3 "η" => (IsPrimitiveRoot.isUnit (hζ.toInteger_isPrimitiveRoot) (by decide)).unit
local notation3 "λ" => hζ.toInteger - 1

/-- `FermatLastTheoremForThreeGen` is the statement that `a ^ 3 + b ^ 3 = u * c ^ 3` has no
nontrivial solutions in `𝓞 K` for all `u : (𝓞 K)ˣ` such that `¬ λ ∣ a`, `¬ λ ∣ b` and `λ ∣ c`.
Expand Down
38 changes: 38 additions & 0 deletions Mathlib/RingTheory/Fintype.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Scott Morrison
-/
import Mathlib.Data.Fintype.Units
import Mathlib.Data.ZMod.Basic

#align_import ring_theory.fintype from "leanprover-community/mathlib"@"1126441d6bccf98c81214a0780c73d499f6721fe"

Expand All @@ -12,6 +13,43 @@ import Mathlib.Data.Fintype.Units
-/


open Finset ZMod

variable {R : Type*} [Ring R] [Fintype R] [DecidableEq R]

lemma Finset.univ_of_card_eq_two (h : Fintype.card R = 2) :
(univ : Finset R) = {0, 1} := by
have : Nontrivial R := by
refine Fintype.one_lt_card_iff_nontrivial.1 ?_
rw [h]
norm_num
refine (eq_of_subset_of_card_le (fun _ _ ↦ mem_univ _) ?_).symm
rw [card_univ, h, card_insert_of_not_mem, card_singleton]
rw [mem_singleton]
exact zero_ne_one

lemma Finset.univ_of_card_eq_three (h : Fintype.card R = 3) :
(univ : Finset R) = {0, 1, -1} := by
have : Nontrivial R := by
refine Fintype.one_lt_card_iff_nontrivial.1 ?_
rw [h]
norm_num
refine (eq_of_subset_of_card_le (fun _ _ ↦ mem_univ _) ?_).symm
rw [card_univ, h, card_insert_of_not_mem, card_insert_of_not_mem, card_singleton]
· rw [mem_singleton]
intro H
rw [← add_eq_zero_iff_eq_neg, one_add_one_eq_two] at H
apply_fun (ringEquivOfPrime R Nat.prime_three h).symm at H
simp only [map_ofNat, map_zero] at H
replace H : ((2 : ℕ) : ZMod 3) = 0 := H
rw [natCast_zmod_eq_zero_iff_dvd] at H
norm_num at H
· intro h
simp only [mem_insert, mem_singleton, zero_eq_neg] at h
rcases h with (h | h)
· exact zero_ne_one h
· exact zero_ne_one h.symm

open scoped Classical

theorem card_units_lt (M₀ : Type*) [MonoidWithZero M₀] [Nontrivial M₀] [Fintype M₀] :
Expand Down
Loading