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

feat(ring_theory/integral_domain): generalize card_fiber_eq_of_mem_range #17653

Open
wants to merge 23 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
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
6 changes: 6 additions & 0 deletions src/algebra/group/units.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/algebra/group/units.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2017 Kenny Lau. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kenny Lau, Mario Carneiro, Johannes Hölzl, Chris Hughes, Jens Wagemaker, Jon Eugster
Expand Down Expand Up @@ -490,6 +490,12 @@
@[to_additive] protected lemma mul_left_injective (h : is_unit b) : injective (* b) :=
λ _ _, h.mul_right_cancel

@[to_additive] protected lemma mul_right_eq_self (h : is_unit a) : a * b = a ↔ b = 1 :=
iff.trans (by rw [mul_one]) h.mul_right_inj

@[to_additive] protected lemma mul_left_eq_self (h : is_unit b) : a * b = b ↔ a = 1 :=
iff.trans (by rw [one_mul]) h.mul_left_inj

end monoid

variables [division_monoid M] {a : M}
Expand Down
8 changes: 6 additions & 2 deletions src/group_theory/coset.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/group_theory/coset.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2018 Mitchell Rowett. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mitchell Rowett, Scott Morrison
Expand Down Expand Up @@ -329,10 +329,14 @@
lemma mk_surjective : function.surjective $ @mk _ _ s := quotient.surjective_quotient_mk'

@[elab_as_eliminator, to_additive]
lemma induction_on {C : α ⧸ s → Prop} (x : α ⧸ s)
(H : ∀ z, C (quotient_group.mk z)) : C x :=
lemma induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (mk z)) : C x :=
quotient.induction_on' x H

@[elab_as_eliminator, to_additive]
lemma induction_on₂ {β} [group β] {t : subgroup β} {C : α ⧸ s → β ⧸ t → Prop}
(x : α ⧸ s) (y : β ⧸ t) (H : ∀ a b, C (mk a) (mk b)) : C x y :=
quotient.induction_on₂' x y H

@[to_additive]
instance : has_coe_t α (α ⧸ s) := ⟨mk⟩ -- note [use has_coe_t]

Expand Down
29 changes: 13 additions & 16 deletions src/group_theory/quotient_group.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/group_theory/quotient_group.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2018 Kevin Buzzard, Patrick Massot. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kevin Buzzard, Patrick Massot
Expand Down Expand Up @@ -43,11 +43,12 @@
-/

open function
universes u v
universes u v w

namespace quotient_group

variables {G : Type u} [group G] (N : subgroup G) [nN : N.normal] {H : Type v} [group H]
{M : Type w} [monoid M]
include nN

/-- The congruence relation generated by a normal subgroup. -/
Expand Down Expand Up @@ -99,12 +100,10 @@
rw [mul_one, subgroup.inv_mem_iff],
end

@[simp, to_additive]
lemma ker_mk : monoid_hom.ker (quotient_group.mk' N : G →* G ⧸ N) = N :=
@[simp, to_additive] lemma ker_mk : (quotient_group.mk' N : G →* G ⧸ N).ker = N :=
subgroup.ext eq_one_iff

@[to_additive]
lemma eq_iff_div_mem {N : subgroup G} [nN : N.normal] {x y : G} :
@[to_additive] lemma eq_iff_div_mem {N : subgroup G} [nN : N.normal] {x y : G} :
(x : G ⧸ N) = y ↔ x / y ∈ N :=
begin
refine eq_comm.trans (quotient_group.eq.trans _),
Expand Down Expand Up @@ -157,13 +156,7 @@
`N ⊆ f⁻¹(M)`."]
def map (M : subgroup H) [M.normal] (f : G →* H) (h : N ≤ M.comap f) :
G ⧸ N →* H ⧸ M :=
begin
refine quotient_group.lift N ((mk' M).comp f) _,
assume x hx,
refine quotient_group.eq.2 _,
rw [mul_one, subgroup.inv_mem_iff],
exact h hx,
end
quotient_group.lift N ((mk' M).comp f) $ by rwa [← (mk' M).comap_ker, ker_mk]

@[simp, to_additive] lemma map_coe (M : subgroup H) [M.normal] (f : G →* H) (h : N ≤ M.comap f)
(x : G) :
Expand Down Expand Up @@ -248,7 +241,9 @@

end congr

variables (φ : G →* H)
section ker_lift

variables (φ : G →* M) (ψ : G →* H)

open monoid_hom

Expand All @@ -269,8 +264,10 @@
assume a b (h : φ a = φ b), quotient.sound' $
by rw [left_rel_apply, mem_ker, φ.map_mul, ← h, φ.map_inv, inv_mul_self]

-- Note that `ker φ` isn't definitionally `ker (φ.range_restrict)`
-- so there is a bit of annoying code duplication here
@[simp, to_additive] lemma ker_lift_surjective : surjective (ker_lift φ) ↔ surjective φ :=
surjective_lift _ _

@[simp, to_additive] lemma range_ker_lift_eq : (ker_lift ψ).range = ψ.range := range_lift _ _

/-- The induced map from the quotient by the kernel to the range. -/
@[to_additive "The induced map from the quotient by the kernel to the range."]
Expand Down Expand Up @@ -314,7 +311,7 @@
/-- The canonical isomorphism `G/⊥ ≃* G`. -/
@[to_additive "The canonical isomorphism `G/⊥ ≃+ G`.", simps]
def quotient_bot : G ⧸ (⊥ : subgroup G) ≃* G :=
quotient_ker_equiv_of_right_inverse (monoid_hom.id G) id (λ x, rfl)
quotient_equiv_of_right_inverse (monoid_hom.id G) id ⊥ (monoid_hom.ker_id _) (λ x, rfl)

/-- The canonical isomorphism `G/(ker φ) ≃* H` induced by a surjection `φ : G →* H`.

Expand Down
30 changes: 30 additions & 0 deletions src/group_theory/subgroup/basic.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/group_theory/subgroup/basic.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2020 Kexing Ying. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kexing Ying
Expand Down Expand Up @@ -2011,6 +2011,36 @@
⟨λ x hx y, by rw [mem_ker, map_mul, map_mul, f.mem_ker.1 hx, mul_one,
map_mul_eq_one f (mul_inv_self y)]⟩

@[simp, to_additive] lemma preimage_mul_left_ker (f : G →* M) (x : G) :
((*) x) ⁻¹' f.ker = f ⁻¹' {f x⁻¹} :=
begin
ext y,
simp only [eq_iff, set.mem_preimage, set_like.mem_coe, set.mem_singleton_iff, inv_inv]
end

@[simp, to_additive] lemma preimage_mul_right_ker (f : G →* M) (x : G) :
(λ y, y * x) ⁻¹' f.ker = f ⁻¹' {f x⁻¹} :=
begin
ext y,
simpa only [eq_iff, set.mem_preimage, set_like.mem_coe, set.mem_singleton_iff, inv_inv]
using f.normal_ker.mem_comm_iff
end

@[simp, to_additive] lemma image_mul_left_ker (f : G →* M) (x : G) :
((*) x) '' f.ker = f ⁻¹' {f x} :=
by rw [← equiv.coe_mul_left, equiv.image_eq_preimage, equiv.mul_left_symm_apply,
preimage_mul_left_ker, inv_inv]

@[simp, to_additive] lemma image_mul_right_ker (f : G →* M) (x : G) :
(λ y, y * x) '' f.ker = f ⁻¹' {f x} :=
by rw [← equiv.coe_mul_right, equiv.image_eq_preimage, equiv.mul_right_symm_apply,
preimage_mul_right_ker, inv_inv]

@[simp, to_additive] lemma card_preimage_singleton (f : G →* M) (x : G)
[fintype f.ker] [fintype (f ⁻¹' {f x})] : fintype.card (f ⁻¹' {f x}) = fintype.card f.ker :=

Check failure on line 2040 in src/group_theory/subgroup/basic.lean

View workflow job for this annotation

GitHub Actions / Build mathlib

/home/lean/actions-runner/_work/mathlib/mathlib/src/group_theory/subgroup/basic.lean:2040:3: unknown identifier 'fintype'
fintype.card_congr $ (equiv.set_congr $ f.image_mul_left_ker x).symm.trans $
((equiv.mul_left x).image _).symm

end ker

section eq_locus
Expand Down
76 changes: 27 additions & 49 deletions src/ring_theory/integral_domain.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/ring_theory/integral_domain.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2020 Johan Commelin. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johan Commelin, Chris Hughes
Expand Down Expand Up @@ -144,20 +144,13 @@
instance [finite Rˣ] : is_cyclic Rˣ :=
is_cyclic_of_subgroup_is_domain (units.coe_hom R) $ units.ext

section

variables (S : subgroup Rˣ) [finite S]

/-- A finite subgroup of the units of an integral domain is cyclic. -/
instance subgroup_units_cyclic : is_cyclic S :=
begin
refine is_cyclic_of_subgroup_is_domain ⟨(coe : S → R), _, _⟩
(units.ext.comp subtype.val_injective),
{ simp },
{ intros, simp },
end
instance subgroup_units_cyclic (S : subgroup Rˣ) [finite S] : is_cyclic S :=
is_cyclic_of_subgroup_is_domain ((units.coe_hom R).comp S.subtype)
(units.ext.comp subtype.coe_injective)

end
instance is_domain.is_cyclic_quotient_ker [finite G] {f : G →* R} : is_cyclic (G ⧸ f.ker) :=
is_cyclic_of_subgroup_is_domain (quotient_group.ker_lift f)

section euclidean_division

Expand All @@ -184,55 +177,40 @@

variables [fintype G]

lemma card_fiber_eq_of_mem_range {H : Type*} [group H] [decidable_eq H]
(f : G →* H) {x y : H} (hx : x ∈ set.range f) (hy : y ∈ set.range f) :
(univ.filter $ λ g, f g = x).card = (univ.filter $ λ g, f g = y).card :=
begin
rcases hx with ⟨x, rfl⟩,
rcases hy with ⟨y, rfl⟩,
refine card_congr (λ g _, g * x⁻¹ * y) _ _ (λ g hg, ⟨g * y⁻¹ * x, _⟩),
{ simp only [mem_filter, one_mul, monoid_hom.map_mul, mem_univ, mul_right_inv,
eq_self_iff_true, monoid_hom.map_mul_inv, and_self, forall_true_iff] {contextual := tt} },
{ simp only [mul_left_inj, imp_self, forall_2_true_iff], },
{ simp only [true_and, mem_filter, mem_univ] at hg,
simp only [hg, mem_filter, one_mul, monoid_hom.map_mul, mem_univ, mul_right_inv,
eq_self_iff_true, exists_prop_of_true, monoid_hom.map_mul_inv, and_self,
mul_inv_cancel_right, inv_mul_cancel_right], }
end

/-- In an integral domain, a sum indexed by a nontrivial homomorphism from a finite group is zero.
-/
lemma sum_hom_units_eq_zero (f : G →* R) (hf : f ≠ 1) : ∑ g : G, f g = 0 :=
begin
classical,
obtain ⟨x, hx⟩ : ∃ x : monoid_hom.range f.to_hom_units,
∀ y : monoid_hom.range f.to_hom_units, y ∈ submonoid.powers x,
obtain ⟨f, rfl⟩ : ∃ f' : G →* Rˣ, (units.coe_hom R).comp f' = f,
from ⟨f.to_hom_units, fun_like.ext' rfl⟩,
obtain ⟨x, hx⟩ : ∃ x : f.range, ∀ y : f.range, y ∈ submonoid.powers x,
from is_cyclic.exists_monoid_generator,
have hx1 : x ≠ 1,
{ rintro rfl,
apply hf,
ext g,
rw [monoid_hom.one_apply],
cases hx ⟨f.to_hom_units g, g, rfl⟩ with n hn,
rwa [subtype.ext_iff, units.ext_iff, subtype.coe_mk, monoid_hom.coe_to_hom_units,
one_pow, eq_comm] at hn, },
replace hx1 : (x : R) - 1 ≠ 0,
from λ h, hx1 (subtype.eq (units.ext (sub_eq_zero.1 h))),
let c := (univ.filter (λ g, f.to_hom_units g = 1)).card,
calc ∑ g : G, f g
= ∑ g : G, f.to_hom_units g : rfl
... = ∑ u : Rˣ in univ.image f.to_hom_units,
(univ.filter (λ g, f.to_hom_units g = u)).card • u : sum_comp (coe : Rˣ → R) f.to_hom_units
... = ∑ u : Rˣ in univ.image f.to_hom_units, c • u :
sum_congr rfl (λ u hu, congr_arg2 _ _ rfl) -- remaining goal 1, proven below
... = ∑ b : monoid_hom.range f.to_hom_units, c • ↑b : finset.sum_subtype _
(by simp ) _
... = c • ∑ b : monoid_hom.range f.to_hom_units, (b : R) : smul_sum.symm
... = c • 0 : congr_arg2 _ rfl _ -- remaining goal 2, proven below
... = 0 : smul_zero _,
specialize hx ⟨f g, g, rfl⟩,
rw [submonoid.powers_one, submonoid.mem_bot, ← subtype.coe_inj, subtype.coe_mk] at hx,
rw [monoid_hom.comp_apply, monoid_hom.one_apply, hx, coe_one, map_one] },
replace hx1 : (x : R) - 1 ≠ 0, from sub_ne_zero.2 (λ h, hx1 $ subtype.ext $ units.ext h),
set c := fintype.card f.ker,
calc ∑ g : G, (f g : R) = ∑ u : Rˣ in univ.image f, c • u : eq.symm $ sum_image' _ $
λ g hg, _ -- remaining goal 1, proven below
... = _ : _,

-- calc ∑ g : G, f g = ∑ u : Rˣ in univ.image f,
-- (univ.filter (λ g, f.to_hom_units g = u)).card • u : sum_comp (coe : Rˣ → R) f.to_hom_units
-- ... = ∑ u : Rˣ in univ.image f, c • u :
-- sum_congr rfl (λ u hu, congr_arg2 _ _ rfl) -- remaining goal 1, proven below
-- ... = ∑ b : f.range, c • ↑b : finset.sum_subtype _
-- (by simp ) _
-- ... = c • ∑ b : monoid_hom.range f.to_hom_units, (b : R) : smul_sum.symm
-- ... = c • 0 : congr_arg2 _ rfl _ -- remaining goal 2, proven below
-- ... = 0 : smul_zero _,
{ -- remaining goal 1
show (univ.filter (λ (g : G), f.to_hom_units g = u)).card = c,
apply card_fiber_eq_of_mem_range f.to_hom_units,
apply f.to_hom_units.card_fiber_eq_of_mem_range,
{ simpa only [mem_image, mem_univ, exists_prop_of_true, set.mem_range] using hu, },
{ exact ⟨1, f.to_hom_units.map_one⟩ } },
-- remaining goal 2
Expand Down
Loading