Skip to content

Commit

Permalink
refactor(GroupTheory/SpecificGroups/Cyclic): Switch from Fintype to…
Browse files Browse the repository at this point in the history
… `Finite` (#19109)

This PR switches most of `GroupTheory/SpecificGroups/Cyclic.lean` from `Fintype` to `Finite`.
  • Loading branch information
tb65536 committed Nov 18, 2024
1 parent 6d297a4 commit 66df24a
Show file tree
Hide file tree
Showing 9 changed files with 92 additions and 109 deletions.
10 changes: 6 additions & 4 deletions Mathlib/FieldTheory/Finite/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -172,6 +172,7 @@ theorem sum_subgroup_units [Ring K] [NoZeroDivisors K]
theorem sum_subgroup_pow_eq_zero [CommRing K] [NoZeroDivisors K]
{G : Subgroup Kˣ} [Fintype G] {k : ℕ} (k_pos : k ≠ 0) (k_lt_card_G : k < Fintype.card G) :
∑ x : G, ((x : Kˣ) : K) ^ k = 0 := by
rw [← Nat.card_eq_fintype_card] at k_lt_card_G
nontriviality K
have := NoZeroDivisors.to_isDomain K
rcases (exists_pow_ne_one_of_isCyclic k_pos k_lt_card_G) with ⟨a, ha⟩
Expand Down Expand Up @@ -255,7 +256,7 @@ theorem cast_card_eq_zero : (q : K) = 0 := by
theorem forall_pow_eq_one_iff (i : ℕ) : (∀ x : Kˣ, x ^ i = 1) ↔ q - 1 ∣ i := by
classical
obtain ⟨x, hx⟩ := IsCyclic.exists_generator (α := Kˣ)
rw [← Fintype.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hx,
rw [← Nat.card_eq_fintype_card, ← Nat.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hx,
orderOf_dvd_iff_pow_eq_one]
constructor
· intro h; apply h
Expand Down Expand Up @@ -586,11 +587,12 @@ theorem unit_isSquare_iff (hF : ringChar F ≠ 2) (a : Fˣ) :
push_cast
exact FiniteField.pow_card_sub_one_eq_one (y : F) (Units.ne_zero y)
· subst a; intro h
have key : 2 * (Fintype.card F / 2) ∣ n * (Fintype.card F / 2) := by
rw [← Nat.card_eq_fintype_card] at hodd h
have key : 2 * (Nat.card F / 2) ∣ n * (Nat.card F / 2) := by
rw [← pow_mul] at h
rw [hodd, ← Fintype.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hg]
rw [hodd, ← Nat.card_units, ← orderOf_eq_card_of_forall_mem_zpowers hg]
apply orderOf_dvd_of_pow_eq_one h
have : 0 < Fintype.card F / 2 := Nat.div_pos Fintype.one_lt_card (by norm_num)
have : 0 < Nat.card F / 2 := Nat.div_pos Finite.one_lt_card (by norm_num)
obtain ⟨m, rfl⟩ := Nat.dvd_of_mul_dvd_mul_right this key
refine ⟨g ^ m, ?_⟩
dsimp
Expand Down
2 changes: 2 additions & 0 deletions Mathlib/GroupTheory/Exponent.lean
Original file line number Diff line number Diff line change
Expand Up @@ -501,9 +501,11 @@ variable [Group G]
lemma Group.one_lt_exponent [Finite G] [Nontrivial G] : 1 < Monoid.exponent G :=
Monoid.one_lt_exponent

@[to_additive]
theorem Group.exponent_dvd_card [Fintype G] : Monoid.exponent G ∣ Fintype.card G :=
Monoid.exponent_dvd.mpr <| fun _ => orderOf_dvd_card

@[to_additive]
theorem Group.exponent_dvd_nat_card : Monoid.exponent G ∣ Nat.card G :=
Monoid.exponent_dvd.mpr orderOf_dvd_natCard

Expand Down
166 changes: 73 additions & 93 deletions Mathlib/GroupTheory/SpecificGroups/Cyclic.lean

Large diffs are not rendered by default.

2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/SpecificGroups/Quaternion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -210,7 +210,7 @@ theorem orderOf_xa [NeZero n] (i : ZMod (2 * n)) : orderOf (xa i) = 4 := by
/-- In the special case `n = 1`, `Quaternion 1` is a cyclic group (of order `4`). -/
theorem quaternionGroup_one_isCyclic : IsCyclic (QuaternionGroup 1) := by
apply isCyclic_of_orderOf_eq_card
· rw [card, mul_one]
· rw [Nat.card_eq_fintype_card, card, mul_one]
exact orderOf_xa 0

/-- If `0 < n`, then `a 1` has order `2 * n`.
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -119,7 +119,7 @@ theorem toFun_spec'' (g : L ≃+* L) {n : ℕ} [NeZero n] {t : L} (ht : IsPrimit
/-- If g(t)=t^c for all roots of unity, then c=χ(g). -/
theorem toFun_unique (g : L ≃+* L) (c : ZMod (Fintype.card (rootsOfUnity n L)))
(hc : ∀ t : rootsOfUnity n L, g (t : Lˣ) = (t ^ c.val : Lˣ)) : c = χ₀ n g := by
apply IsCyclic.ext rfl (fun ζ ↦ ?_)
apply IsCyclic.ext Nat.card_eq_fintype_card (fun ζ ↦ ?_)
specialize hc ζ
suffices ((ζ ^ c.val : Lˣ) : L) = (ζ ^ (χ₀ n g).val : Lˣ) by exact_mod_cast this
rw [← toFun_spec g ζ, hc]
Expand Down
3 changes: 2 additions & 1 deletion Mathlib/NumberTheory/LucasPrimality.lean
Original file line number Diff line number Diff line change
Expand Up @@ -54,7 +54,8 @@ theorem reverse_lucas_primality (p : ℕ) (hP : p.Prime) :
have : Fact p.Prime := ⟨hP⟩
obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := (ZMod p)ˣ)
have h1 : orderOf g = p - 1 := by
rwa [orderOf_eq_card_of_forall_mem_zpowers hg, ← Nat.prime_iff_card_units]
rwa [orderOf_eq_card_of_forall_mem_zpowers hg, Nat.card_eq_fintype_card,
← Nat.prime_iff_card_units]
have h2 := tsub_pos_iff_lt.2 hP.one_lt
rw [← orderOf_injective (Units.coeHom _) Units.ext _, orderOf_eq_iff h2] at h1
refine ⟨g, h1.1, fun q hq hqd ↦ ?_⟩
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/MulChar/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ noncomputable def ofRootOfUnity {ζ : Rˣ} (hζ : ζ ∈ rootsOfUnity (Fintype.c
have : orderOf ζ ∣ Fintype.card Mˣ :=
orderOf_dvd_iff_pow_eq_one.mpr <| (mem_rootsOfUnity _ ζ).mp hζ
refine ofUnitHom <| monoidHomOfForallMemZpowers hg <| this.trans <| dvd_of_eq ?_
rw [orderOf_generator_eq_natCard hg, Nat.card_eq_fintype_card]
rw [orderOf_eq_card_of_forall_mem_zpowers hg, Nat.card_eq_fintype_card]

lemma ofRootOfUnity_spec {ζ : Rˣ} (hζ : ζ ∈ rootsOfUnity (Fintype.card Mˣ) R)
{g : Mˣ} (hg : ∀ x, x ∈ Subgroup.zpowers g) :
Expand Down
11 changes: 5 additions & 6 deletions Mathlib/RingTheory/RootsOfUnity/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -251,12 +251,12 @@ namespace IsCyclic
`n` into another group `G'` to the group of `n`th roots of unity in `G'` determined by a generator
`g` of `G`. It sends `φ : G →* G'` to `φ g`. -/
noncomputable
def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] [Fintype G] {g : G}
def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] {g : G}
(hg : ∀ (x : G), x ∈ Subgroup.zpowers g) (G' : Type*) [CommGroup G'] :
(G →* G') ≃* rootsOfUnity (Fintype.card G) G' where
(G →* G') ≃* rootsOfUnity (Nat.card G) G' where
toFun φ := ⟨(IsUnit.map φ <| Group.isUnit g).unit, by
simp only [mem_rootsOfUnity, Units.ext_iff, Units.val_pow_eq_pow_val, IsUnit.unit_spec,
← map_pow, pow_card_eq_one, map_one, Units.val_one]⟩
← map_pow, pow_card_eq_one', map_one, Units.val_one]⟩
invFun ζ := monoidHomOfForallMemZpowers hg (g' := (ζ.val : G')) <| by
simpa only [orderOf_eq_card_of_forall_mem_zpowers hg, orderOf_dvd_iff_pow_eq_one,
← Units.val_pow_eq_pow_val, Units.val_eq_one] using ζ.prop
Expand All @@ -270,12 +270,11 @@ def monoidHomMulEquivRootsOfUnityOfGenerator {G : Type*} [CommGroup G] [Fintype

/-- The group of group homomorphisms from a finite cyclic group `G` of order `n` into another
group `G'` is (noncanonically) isomorphic to the group of `n`th roots of unity in `G'`. -/
lemma monoidHom_mulEquiv_rootsOfUnity (G : Type*) [CommGroup G] [Finite G] [IsCyclic G]
lemma monoidHom_mulEquiv_rootsOfUnity (G : Type*) [CommGroup G] [IsCyclic G]
(G' : Type*) [CommGroup G'] :
Nonempty <| (G →* G') ≃* rootsOfUnity (Nat.card G) G' := by
obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := G)
have : Fintype G := Fintype.ofFinite _
exact ⟨Nat.card_eq_fintype_card (α := G) ▸ monoidHomMulEquivRootsOfUnityOfGenerator hg G'⟩
exact ⟨monoidHomMulEquivRootsOfUnityOfGenerator hg G'⟩

end IsCyclic

Expand Down
3 changes: 1 addition & 2 deletions Mathlib/RingTheory/RootsOfUnity/EnoughRootsOfUnity.lean
Original file line number Diff line number Diff line change
Expand Up @@ -64,8 +64,7 @@ lemma natCard_rootsOfUnity (M : Type*) [CommMonoid M] (n : ℕ) [NeZero n]
[HasEnoughRootsOfUnity M n] :
Nat.card (rootsOfUnity n M) = n := by
obtain ⟨ζ, h⟩ := exists_primitiveRoot M n
have : Fintype <| rootsOfUnity n M := Fintype.ofFinite _
rw [Nat.card_eq_fintype_card, ← IsCyclic.exponent_eq_card]
rw [← IsCyclic.exponent_eq_card]
refine dvd_antisymm ?_ ?_
· exact Monoid.exponent_dvd_of_forall_pow_eq_one fun g ↦ OneMemClass.coe_eq_one.mp g.prop
· nth_rewrite 1 [h.eq_orderOf]
Expand Down

0 comments on commit 66df24a

Please sign in to comment.