From 8895dd8bc3a7129ffbc5693c30df6dc5c4c0f6ac Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 02:10:50 +0100 Subject: [PATCH 001/116] first stab at hom out of pi group --- src/group_theory/pi_hom.lean | 233 +++++++++++++++++++++++++++++++++++ 1 file changed, 233 insertions(+) create mode 100644 src/group_theory/pi_hom.lean diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean new file mode 100644 index 0000000000000..7e94518ab920d --- /dev/null +++ b/src/group_theory/pi_hom.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2021 Jocchim Breitner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joachim Breitner +-/ +import group_theory.subgroup.basic +import tactic.group +import group_theory.general_commutator +import group_theory.order_of_element +import data.finset.noncomm_prod + +variables {G : Type*} [group G] + +lemma mul_eq_one_of_disjoint + {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) + (heq : x * y = 1) : x = 1 ∧ y = 1 := +begin + have : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), + subst this, + have hy := H₂.inv_mem_iff.mp hy, + have : x ∈ H₁ ⊓ H₂, by { simp, cc }, + rw [hdis.eq_bot, subgroup.mem_bot] at this, + subst this, + simp +end + +namespace pi_hom + +variables {I : Type*} [fintype I] {H : I → Type*} [∀ i, group (H i)] (ϕ : Π (i : I), H i →* G) + +variables (f g : Π (i : I), H i) + +-- A bit like `finset.noncomm_prodi`, but without the dependent Prop arg that makes things tricky +noncomputable +def fun_on (S : finset I) : G := (list.map (λ i, ϕ i (f i)) S.to_list).prod + +noncomputable +def to_fun : G := fun_on ϕ f finset.univ + +@[simp] +lemma fun_on_empty : fun_on ϕ f ∅ = 1 := by simp [fun_on] + +variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)) ] +include hcomm + +lemma fun_on_perm {S : finset I} {l : list I} (h : S.to_list ~ l) : + fun_on ϕ f S = (list.map (λ i, ϕ i (f i)) l).prod := +begin + unfold fun_on, + have hnd : l.nodup := h.nodup_iff.mp (finset.nodup_to_list S), + induction h with x l₁ l₂ hl₁l₂ ih x y l₁ l₁ l₂ l₃ hl₁l₂ hl₂l₃ ih12 ih23, + { reflexivity }, + { specialize ih (list.nodup_cons.mp hnd).2, simpa }, + { simp only [list.map, list.prod_cons, ← mul_assoc], + congr' 1, + apply fact.elim hcomm, + simp only [list.nodup_cons, list.mem_cons_iff] at hnd, + cc, }, + { specialize ih12 (hl₂l₃.nodup_iff.mpr hnd), + specialize ih23 hnd, + rw [ih12, ih23], } +end + +@[simp] +lemma fun_on_cons (S : finset I) (i : I) (h : i ∉ S) : + fun_on ϕ f (finset.cons i S h) = ϕ i (f i) * fun_on ϕ f S := +by { rw fun_on_perm ϕ f (finset.to_list_cons h), simp [fun_on] } + +@[simp] +lemma fun_on_one (S : finset I) : fun_on ϕ 1 S = 1 := +begin + induction S using finset.cons_induction_on with i S hnmem ih, + { simp }, + { simp [ih] } +end + +@[simp] +lemma to_fun_one : to_fun ϕ 1 = 1 := fun_on_one _ _ + +lemma fun_on_commutes (S : finset I) (i : I) (hnmem : i ∉ S) : + ϕ i (g i) * fun_on ϕ f S = fun_on ϕ f S * ϕ i (g i) := +begin + induction S using finset.cons_induction_on with j S hnmem ih, + { simp, }, + { repeat { rw fun_on_cons }, + have hij : i ≠ j, by {simp at hnmem, cc}, + have hiS : i ∉ S, by {simp at hnmem, cc}, + calc ϕ i (g i) * (ϕ j (f j) * fun_on ϕ f S) + = (ϕ i (g i) * ϕ j (f j)) * fun_on ϕ f S : by rw ← mul_assoc + ... = (ϕ j (f j) * ϕ i (g i)) * fun_on ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} + ... = ϕ j (f j) * (ϕ i (g i) * fun_on ϕ f S) : by rw mul_assoc + ... = ϕ j (f j) * (fun_on ϕ f S * ϕ i (g i)) : by { congr' 1, apply (ih hiS) } + ... = (ϕ j (f j) * fun_on ϕ f S) * ϕ i (g i) : by rw ← mul_assoc + } +end + +@[simp] +lemma fun_on_mul (S : finset I) : fun_on ϕ (f * g) S = fun_on ϕ f S * fun_on ϕ g S := +begin + induction S using finset.cons_induction_on with i S hnmem ih, + { simp, }, + { repeat { rw [fun_on_cons] }, + rw ih, clear ih, + simp only [pi.mul_apply, map_mul], + repeat { rw mul_assoc }, congr' 1, + repeat { rw ← mul_assoc }, congr' 1, + exact (fun_on_commutes _ _ _ S i hnmem), } +end + +lemma fun_on_in_sup_range (S : finset I) : + fun_on ϕ f S ∈ ⨆ (i : I) (H_1 : i ∈ S), (ϕ i).range := +begin + induction S using finset.cons_induction_on with i S hnmem ih, + { simp, }, + { repeat { rw fun_on_cons }, + refine (subgroup.mul_mem _ _ _), + { apply (subgroup.mem_Sup_of_mem), { use i }, { simp, }, }, + { have : (∀ j : I, j ∈ S → j ∈ finset.cons i S hnmem), + by { intro, simp, intro h, right, exact h}, + have : (⨆ (i_1 : I) (H_1 : i_1 ∈ S), (ϕ i_1).range) ≤ + (⨆ (i_1 : I) (H_1 : i_1 ∈ finset.cons i S hnmem), (ϕ i_1).range) := + @bsupr_le_bsupr' _ _ _ (λ j, j ∈ S) (λ j, j ∈ finset.cons i S hnmem) this _, + exact (this ih), + } + } +end + +lemma to_fun_in_sup_range : to_fun ϕ f ∈ ⨆ (i : I), (ϕ i).range := +bsupr_le_supr _ (λ i, (ϕ i).range) (fun_on_in_sup_range ϕ f finset.univ) + +@[simp] +lemma to_fun_mul : to_fun ϕ (f * g) = to_fun ϕ f * to_fun ϕ g := fun_on_mul _ _ _ _ + +noncomputable +def hom : (Π (i : I), H i) →* G := +{ to_fun := to_fun ϕ, + map_one' := to_fun_one _, + map_mul' := λ f g, to_fun_mul _ f g, } + +omit hcomm +def just_one [decidable_eq I] (i : I) (y : H i) : Π (i : I), H i := + λ j, if h : j = i then by { subst h; exact y} else 1 + +@[simp] +lemma just_one_eq [decidable_eq I] (i : I) (y : H i) : just_one i y i = y := +by { unfold just_one, simp } + +@[simp] +lemma just_one_ne [decidable_eq I] (i : I) (y : H i) (j : I) (h : i ≠ j) : + just_one i y j = (1 : H j) := +by { unfold just_one, have : ¬ (j = i), by cc, simp [this], } + +include hcomm + +lemma fun_on_just_one [decidable_eq I] (i : I) (y : H i) (S : finset I) : + fun_on ϕ (just_one i y) S = if i ∈ S then ϕ i y else 1 := +begin + induction S using finset.cons_induction_on with j S hnmem ih, + { simp, }, + { repeat { rw fun_on_cons }, + by_cases (i = j), + { + subst h, + rw ih, + simp only [just_one_eq, mul_ite, mul_one, finset.cons_eq_insert, finset.mem_insert, + eq_self_iff_true, true_or, if_true, ite_eq_right_iff, mul_left_eq_self], + intro i, contradiction, + }, + { change i ≠ j at h, + simp [h], + exact ih, } } +end + +lemma to_fun_just_one [decidable_eq I] (i : I) (y : H i) : + to_fun ϕ (just_one i y) = ϕ i y := +begin + unfold to_fun, + rw fun_on_just_one ϕ i y _, + simp, +end + +lemma range_eq [decidable_eq I] : (hom ϕ).range = (⨆ (i : I), (ϕ i).range) := +begin + apply le_antisymm, + { rintro x ⟨f, rfl⟩, + exact (to_fun_in_sup_range ϕ f), }, + { refine (supr_le _), + rintro i x ⟨y, rfl⟩, + exact ⟨_, to_fun_just_one _ i y⟩, } +end + + +variables (hind : complete_lattice.independent (λ i, (ϕ i).range)) +variables (hinj : ∀ i, function.injective (ϕ i)) + +include hind +include hinj + +lemma injective_of_independent : function.injective (hom ϕ) := +begin + apply (monoid_hom.ker_eq_bot_iff _).mp, + apply eq_bot_iff.mpr, + intro f, + show fun_on ϕ f finset.univ = 1 → f = 1, + suffices : fun_on ϕ f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), + by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), + induction (finset.univ : finset I) using finset.cons_induction_on with i S hnmem ih, + { simp }, + { intro heq1, + rw fun_on_cons at heq1, + have hnmem' : i ∉ (S : set I), by simpa, + have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, + { apply mul_eq_one_of_disjoint (complete_lattice.independent.disjoint_bsupr hind hnmem') _ _ heq1, + { simp, }, + { apply fun_on_in_sup_range, }, }, + rcases heq1' with ⟨ heq1i, heq1S ⟩, + specialize ih heq1S, + intros i h, + simp only [finset.mem_cons] at h, + rcases h with ⟨rfl | _⟩, + { apply hinj i, simpa, }, + { exact (ih _ h), } } +end + +noncomputable +def mul_equiv [decidable_eq I ]: + (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := +begin + rw ← range_eq, + exact (monoid_hom.of_injective (injective_of_independent _ hind hinj)), +end + +end pi_hom From 47cf2cd8f011148529d2b84ff04b1f966bb723c3 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 02:21:42 +0100 Subject: [PATCH 002/116] Please lint a bit --- src/group_theory/pi_hom.lean | 23 +++++++++++++---------- 1 file changed, 13 insertions(+), 10 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 7e94518ab920d..33e233770cdd3 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -9,6 +9,11 @@ import group_theory.general_commutator import group_theory.order_of_element import data.finset.noncomm_prod +/-! +# TODO + +-/ + variables {G : Type*} [group G] lemma mul_eq_one_of_disjoint @@ -31,9 +36,12 @@ variables {I : Type*} [fintype I] {H : I → Type*} [∀ i, group (H i)] (ϕ : variables (f g : Π (i : I), H i) -- A bit like `finset.noncomm_prodi`, but without the dependent Prop arg that makes things tricky + +/-- TODO -/ noncomputable def fun_on (S : finset I) : G := (list.map (λ i, ϕ i (f i)) S.to_list).prod +/-- TODO -/ noncomputable def to_fun : G := fun_on ϕ f finset.univ @@ -90,8 +98,7 @@ begin ... = (ϕ j (f j) * ϕ i (g i)) * fun_on ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} ... = ϕ j (f j) * (ϕ i (g i) * fun_on ϕ f S) : by rw mul_assoc ... = ϕ j (f j) * (fun_on ϕ f S * ϕ i (g i)) : by { congr' 1, apply (ih hiS) } - ... = (ϕ j (f j) * fun_on ϕ f S) * ϕ i (g i) : by rw ← mul_assoc - } + ... = (ϕ j (f j) * fun_on ϕ f S) * ϕ i (g i) : by rw ← mul_assoc } end @[simp] @@ -120,9 +127,7 @@ begin have : (⨆ (i_1 : I) (H_1 : i_1 ∈ S), (ϕ i_1).range) ≤ (⨆ (i_1 : I) (H_1 : i_1 ∈ finset.cons i S hnmem), (ϕ i_1).range) := @bsupr_le_bsupr' _ _ _ (λ j, j ∈ S) (λ j, j ∈ finset.cons i S hnmem) this _, - exact (this ih), - } - } + exact (this ih), } } end lemma to_fun_in_sup_range : to_fun ϕ f ∈ ⨆ (i : I), (ϕ i).range := @@ -159,13 +164,11 @@ begin { simp, }, { repeat { rw fun_on_cons }, by_cases (i = j), - { - subst h, + { subst h, rw ih, simp only [just_one_eq, mul_ite, mul_one, finset.cons_eq_insert, finset.mem_insert, eq_self_iff_true, true_or, if_true, ite_eq_right_iff, mul_left_eq_self], - intro i, contradiction, - }, + intro i, contradiction, }, { change i ≠ j at h, simp [h], exact ih, } } @@ -210,7 +213,7 @@ begin rw fun_on_cons at heq1, have hnmem' : i ∉ (S : set I), by simpa, have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, - { apply mul_eq_one_of_disjoint (complete_lattice.independent.disjoint_bsupr hind hnmem') _ _ heq1, + { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, { apply fun_on_in_sup_range, }, }, rcases heq1' with ⟨ heq1i, heq1S ⟩, From 637058fcabd397505e417f0615a8233361806e5d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 02:22:03 +0100 Subject: [PATCH 003/116] time flies --- src/group_theory/pi_hom.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 33e233770cdd3..1e23e9500495b 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -1,5 +1,5 @@ /- -Copyright (c) 2021 Jocchim Breitner. All rights reserved. +Copyright (c) 2022 Jocchim Breitner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joachim Breitner -/ From 363f6ffe03b250b97a090ef389d17315b6818c90 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 02:22:45 +0100 Subject: [PATCH 004/116] shuffling whitespace --- src/group_theory/pi_hom.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 1e23e9500495b..844fb9c83afa7 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -226,7 +226,7 @@ begin end noncomputable -def mul_equiv [decidable_eq I ]: +def mul_equiv [decidable_eq I] : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := begin rw ← range_eq, From 204dea2c2f15382a712ac1ddf46dffd6aaf71601 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 02:30:42 +0100 Subject: [PATCH 005/116] Get rid of ugly have --- src/group_theory/pi_hom.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 844fb9c83afa7..58845397a09d9 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -122,12 +122,8 @@ begin { repeat { rw fun_on_cons }, refine (subgroup.mul_mem _ _ _), { apply (subgroup.mem_Sup_of_mem), { use i }, { simp, }, }, - { have : (∀ j : I, j ∈ S → j ∈ finset.cons i S hnmem), - by { intro, simp, intro h, right, exact h}, - have : (⨆ (i_1 : I) (H_1 : i_1 ∈ S), (ϕ i_1).range) ≤ - (⨆ (i_1 : I) (H_1 : i_1 ∈ finset.cons i S hnmem), (ϕ i_1).range) := - @bsupr_le_bsupr' _ _ _ (λ j, j ∈ S) (λ j, j ∈ finset.cons i S hnmem) this _, - exact (this ih), } } + { refine (@bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih), + by { intro, simp, intro h, right, exact h}, } } end lemma to_fun_in_sup_range : to_fun ϕ f ∈ ⨆ (i : I), (ϕ i).range := From e42f4aa6f6d8227094fb5624e26ac2c8bb32866d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 11:22:22 +0100 Subject: [PATCH 006/116] Base development on noncomm_prod --- src/group_theory/pi_hom.lean | 82 +++++++++++++++++------------------- 1 file changed, 38 insertions(+), 44 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 58845397a09d9..52074ff51cc60 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -31,55 +31,49 @@ end namespace pi_hom -variables {I : Type*} [fintype I] {H : I → Type*} [∀ i, group (H i)] (ϕ : Π (i : I), H i →* G) +-- We have an family of group +variables {I : Type*} [fintype I] [decidable_eq I] {H : I → Type*} [∀ i, group (H i)] -variables (f g : Π (i : I), H i) +-- And morphism ϕ into G +variables (ϕ : Π (i : I), H i →* G) --- A bit like `finset.noncomm_prodi`, but without the dependent Prop arg that makes things tricky +-- We assume that the elements of different morphism commute +-- Since we need this all over the place we wrap it up in `fact` +variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)) ] +include hcomm -/-- TODO -/ -noncomputable -def fun_on (S : finset I) : G := (list.map (λ i, ϕ i (f i)) S.to_list).prod +-- Elements of `Π (i : I), H i` are called `f` and `g` here +variables (f g : Π (i : I), H i) -/-- TODO -/ -noncomputable +/-- A wrapper around `finset.noncomm_prod` that discharges the commutativiy requirement using +`hcomm` -/ +def fun_on (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ + by { rintros i - j -, by_cases (i = j), { subst h }, { exact hcomm.elim _ _ h _ _} } + +/-- The product of `ϕ i (f i)` for all `i : I` -/ def to_fun : G := fun_on ϕ f finset.univ @[simp] lemma fun_on_empty : fun_on ϕ f ∅ = 1 := by simp [fun_on] -variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)) ] -include hcomm - -lemma fun_on_perm {S : finset I} {l : list I} (h : S.to_list ~ l) : - fun_on ϕ f S = (list.map (λ i, ϕ i (f i)) l).prod := -begin - unfold fun_on, - have hnd : l.nodup := h.nodup_iff.mp (finset.nodup_to_list S), - induction h with x l₁ l₂ hl₁l₂ ih x y l₁ l₁ l₂ l₃ hl₁l₂ hl₂l₃ ih12 ih23, - { reflexivity }, - { specialize ih (list.nodup_cons.mp hnd).2, simpa }, - { simp only [list.map, list.prod_cons, ← mul_assoc], - congr' 1, - apply fact.elim hcomm, - simp only [list.nodup_cons, list.mem_cons_iff] at hnd, - cc, }, - { specialize ih12 (hl₂l₃.nodup_iff.mpr hnd), - specialize ih23 hnd, - rw [ih12, ih23], } -end +@[simp] +lemma fun_on_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : + fun_on ϕ f (insert i S) = ϕ i (f i) * fun_on ϕ f S := +finset.noncomm_prod_insert_of_not_mem _ _ _ _ h +/- @[simp] lemma fun_on_cons (S : finset I) (i : I) (h : i ∉ S) : fun_on ϕ f (finset.cons i S h) = ϕ i (f i) * fun_on ϕ f S := -by { rw fun_on_perm ϕ f (finset.to_list_cons h), simp [fun_on] } +by { rw finset.cons_eq_insert i S h, exact (fun_on_insert_of_not_mem _ _ _ _ h) } +-/ @[simp] lemma fun_on_one (S : finset I) : fun_on ϕ 1 S = 1 := begin induction S using finset.cons_induction_on with i S hnmem ih, { simp }, - { simp [ih] } + { simp [ih, hnmem], } end @[simp] @@ -88,11 +82,12 @@ lemma to_fun_one : to_fun ϕ 1 = 1 := fun_on_one _ _ lemma fun_on_commutes (S : finset I) (i : I) (hnmem : i ∉ S) : ϕ i (g i) * fun_on ϕ f S = fun_on ϕ f S * ϕ i (g i) := begin - induction S using finset.cons_induction_on with j S hnmem ih, + induction S using finset.induction_on with j S hnmem' ih, { simp, }, - { repeat { rw fun_on_cons }, - have hij : i ≠ j, by {simp at hnmem, cc}, - have hiS : i ∉ S, by {simp at hnmem, cc}, + { simp only [fun_on_insert_of_not_mem _ _ _ _ hnmem'], + + have hij : i ≠ j, by {simp at hnmem, tauto}, + have hiS : i ∉ S, by {simp at hnmem, tauto}, calc ϕ i (g i) * (ϕ j (f j) * fun_on ϕ f S) = (ϕ i (g i) * ϕ j (f j)) * fun_on ϕ f S : by rw ← mul_assoc ... = (ϕ j (f j) * ϕ i (g i)) * fun_on ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} @@ -104,9 +99,9 @@ end @[simp] lemma fun_on_mul (S : finset I) : fun_on ϕ (f * g) S = fun_on ϕ f S * fun_on ϕ g S := begin - induction S using finset.cons_induction_on with i S hnmem ih, + induction S using finset.induction_on with i S hnmem ih, { simp, }, - { repeat { rw [fun_on_cons] }, + { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], rw ih, clear ih, simp only [pi.mul_apply, map_mul], repeat { rw mul_assoc }, congr' 1, @@ -117,9 +112,9 @@ end lemma fun_on_in_sup_range (S : finset I) : fun_on ϕ f S ∈ ⨆ (i : I) (H_1 : i ∈ S), (ϕ i).range := begin - induction S using finset.cons_induction_on with i S hnmem ih, + induction S using finset.induction_on with i S hnmem ih, { simp, }, - { repeat { rw fun_on_cons }, + { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], refine (subgroup.mul_mem _ _ _), { apply (subgroup.mem_Sup_of_mem), { use i }, { simp, }, }, { refine (@bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih), @@ -132,7 +127,6 @@ bsupr_le_supr _ (λ i, (ϕ i).range) (fun_on_in_sup_range ϕ f finset.univ) @[simp] lemma to_fun_mul : to_fun ϕ (f * g) = to_fun ϕ f * to_fun ϕ g := fun_on_mul _ _ _ _ -noncomputable def hom : (Π (i : I), H i) →* G := { to_fun := to_fun ϕ, map_one' := to_fun_one _, @@ -156,9 +150,9 @@ include hcomm lemma fun_on_just_one [decidable_eq I] (i : I) (y : H i) (S : finset I) : fun_on ϕ (just_one i y) S = if i ∈ S then ϕ i y else 1 := begin - induction S using finset.cons_induction_on with j S hnmem ih, + induction S using finset.induction_on with j S hnmem ih, { simp, }, - { repeat { rw fun_on_cons }, + { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], by_cases (i = j), { subst h, rw ih, @@ -203,10 +197,10 @@ begin show fun_on ϕ f finset.univ = 1 → f = 1, suffices : fun_on ϕ f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), - induction (finset.univ : finset I) using finset.cons_induction_on with i S hnmem ih, + induction (finset.univ : finset I) using finset.induction_on with i S hnmem ih, { simp }, { intro heq1, - rw fun_on_cons at heq1, + simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem] at heq1, have hnmem' : i ∉ (S : set I), by simpa, have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, @@ -215,7 +209,7 @@ begin rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, - simp only [finset.mem_cons] at h, + simp only [finset.mem_insert] at h, rcases h with ⟨rfl | _⟩, { apply hinj i, simpa, }, { exact (ih _ h), } } From 2aef604e9afab11406e3f03e9864971bff2a9cab Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 11:26:37 +0100 Subject: [PATCH 007/116] Redundant [decidable_eq I] --- src/group_theory/pi_hom.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 52074ff51cc60..23167de254762 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -133,21 +133,21 @@ def hom : (Π (i : I), H i) →* G := map_mul' := λ f g, to_fun_mul _ f g, } omit hcomm -def just_one [decidable_eq I] (i : I) (y : H i) : Π (i : I), H i := +def just_one (i : I) (y : H i) : Π (i : I), H i := λ j, if h : j = i then by { subst h; exact y} else 1 @[simp] -lemma just_one_eq [decidable_eq I] (i : I) (y : H i) : just_one i y i = y := +lemma just_one_eq (i : I) (y : H i) : just_one i y i = y := by { unfold just_one, simp } @[simp] -lemma just_one_ne [decidable_eq I] (i : I) (y : H i) (j : I) (h : i ≠ j) : +lemma just_one_ne (i : I) (y : H i) (j : I) (h : i ≠ j) : just_one i y j = (1 : H j) := by { unfold just_one, have : ¬ (j = i), by cc, simp [this], } include hcomm -lemma fun_on_just_one [decidable_eq I] (i : I) (y : H i) (S : finset I) : +lemma fun_on_just_one (i : I) (y : H i) (S : finset I) : fun_on ϕ (just_one i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, @@ -164,7 +164,7 @@ begin exact ih, } } end -lemma to_fun_just_one [decidable_eq I] (i : I) (y : H i) : +lemma to_fun_just_one (i : I) (y : H i) : to_fun ϕ (just_one i y) = ϕ i y := begin unfold to_fun, @@ -172,7 +172,7 @@ begin simp, end -lemma range_eq [decidable_eq I] : (hom ϕ).range = (⨆ (i : I), (ϕ i).range) := +lemma range_eq : (hom ϕ).range = (⨆ (i : I), (ϕ i).range) := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, From 7e18e9b2a3eb642d25b534e5fadd308e9563b769 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 12:12:27 +0100 Subject: [PATCH 008/116] Specialize the theorem to normal subgroups and subgroup inclusion which is what I need for https://github.com/leanprover-community/mathlib/issues/11723 --- src/group_theory/pi_hom.lean | 38 ++++++++++++++++++++++++++++++++++-- 1 file changed, 36 insertions(+), 2 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 23167de254762..06caff5bfe0e1 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -216,11 +216,45 @@ begin end noncomputable -def mul_equiv [decidable_eq I] : - (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := +def mul_equiv : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := begin rw ← range_eq, exact (monoid_hom.of_injective (injective_of_independent _ hind hinj)), end end pi_hom + +lemma pairwise_elements_commute_of_normal_of_independent + {I : Type*} {H : I → subgroup G} + (hnorm : ∀ i, (H i).normal) (hind : complete_lattice.independent H) : + (∀ i j : I, i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) := +begin + intros i j hne x y hx hy, + have : H i ⊓ H j ≤ ⊥ := complete_lattice.independent.disjoint hind hne, + have : ⁅H i, H j⁆ ≤ ⊥ := le_trans (general_commutator_le_inf _ _) this, + have : x * y * x ⁻¹ * y ⁻¹ = 1, + by { rw [← subgroup.mem_bot], exact this (general_commutator_containment _ _ hx hy), }, + have : (x * y * x ⁻¹ * y ⁻¹) * (y * x) = y * x, by { simp [this] }, + show x * y = y * x, by simpa [mul_assoc] using this, +end + +noncomputable +def internal_product + {I : Type*} [fintype I] [decidable_eq I] {H : I → subgroup G} + (hnorm : ∀ i, (H i).normal) (hind : complete_lattice.independent H) : + (Π (i : I), H i) ≃* (⨆ (i : I), H i : subgroup G) := +begin + haveI : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), + commute ((H i).subtype x) ((H j).subtype y)) := fact.mk + begin + intros i j hne x y, + induction x with x hx, + induction y with y hy, + exact pairwise_elements_commute_of_normal_of_independent hnorm hind i j hne x y hx hy, + end, + have : (⨆ (i : I), H i: subgroup G) = (⨆ (i : I), (H i).subtype.range : subgroup G), by simp, + rw this, clear this, + refine (pi_hom.mul_equiv _ _ _), + { simpa using hind, }, + { exact λ _, subtype.val_injective, } +end From cda4df81f0fc7b40c8bf567ae0b25dba63be0717 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 16:57:35 +0100 Subject: [PATCH 009/116] feat(group_theory/subgroup/basic): add commute_of_normal_of_disjoint --- src/group_theory/subgroup/basic.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 4966a57852438..7bf6b86037809 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2467,6 +2467,25 @@ begin rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this, end +/-- Elements of disjoint, normal subgroups commute -/ +@[to_additive] lemma commute_of_normal_of_disjoint + (H₁ H₂ : subgroup G) (hH₁ : H₁.normal) (hH₂ : H₂.normal) (hdis : disjoint H₁ H₂) + (x y : G) (hx : x ∈ H₁) (hy : y ∈ H₂) : + commute x y := +begin + suffices : x * y * x⁻¹ * y⁻¹ = 1, + { show x * y = y * x, by { rw [mul_assoc, mul_eq_one_iff_eq_inv] at this, simpa } }, + apply hdis, split, + { show x * y * x⁻¹ * y⁻¹ ∈ H₁, + repeat { rw mul_assoc }, + apply H₁.mul_mem hx, + repeat { rw ← mul_assoc }, + apply (hH₁.conj_mem _ (H₁.inv_mem hx)), }, + { show x * y * x⁻¹ * y⁻¹ ∈ H₂, + apply H₂.mul_mem _ (H₂.inv_mem hy), + apply (hH₂.conj_mem _ hy), } +end + end subgroup_normal end subgroup From 945d984f408058b39c9c1da9470c9ecf7ec70a0f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 19:54:12 +0100 Subject: [PATCH 010/116] Stash coprimality argument --- src/group_theory/pi_hom.lean | 170 ++++++++++++++++++++++++++++++++--- 1 file changed, 157 insertions(+), 13 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 06caff5bfe0e1..2fa62bfcb6693 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -8,14 +8,41 @@ import tactic.group import group_theory.general_commutator import group_theory.order_of_element import data.finset.noncomm_prod - +import ring_theory.coprime.lemmas /-! # TODO -/ +open_locale big_operators + +lemma coprime_prod_left + {I : Type*} + {x : ℕ} {s : I → ℕ} {t : finset I} : + (∀ (i : I), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : I) in t, s i) x := +begin + intro h, + rw ← nat.is_coprime_iff_coprime, + have := @nat.cast_prod _ ℤ _ s t, + simp [ -nat.cast_prod ] at this, + rw this, + apply is_coprime.prod_left, + intros i hi, + rw nat.is_coprime_iff_coprime, + apply h i hi, +end + + variables {G : Type*} [group G] +@[simp] +lemma order_of_inv (x : G) : order_of x⁻¹ = order_of x := +begin + apply nat.dvd_antisymm; rewrite order_of_dvd_iff_pow_eq_one, + { rw [inv_pow, pow_order_of_eq_one, one_inv], }, + { nth_rewrite 0 ← (inv_inv x), rw [inv_pow, pow_order_of_eq_one, one_inv], } +end + lemma mul_eq_one_of_disjoint {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) (heq : x * y = 1) : x = 1 ∧ y = 1 := @@ -80,7 +107,7 @@ end lemma to_fun_one : to_fun ϕ 1 = 1 := fun_on_one _ _ lemma fun_on_commutes (S : finset I) (i : I) (hnmem : i ∉ S) : - ϕ i (g i) * fun_on ϕ f S = fun_on ϕ f S * ϕ i (g i) := + commute (ϕ i (g i)) (fun_on ϕ f S) := begin induction S using finset.induction_on with j S hnmem' ih, { simp, }, @@ -182,14 +209,42 @@ begin exact ⟨_, to_fun_just_one _ i y⟩, } end +lemma fun_on_pow (k : ℕ) (S : finset I) : (fun_on ϕ f S) ^ k = fun_on ϕ (f ^ k) S := +begin + induction S using finset.induction_on with i S hnmem ih, + { simp }, + { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + rw [(fun_on_commutes ϕ f f S i hnmem).mul_pow, ih, pi.pow_apply, map_pow] }, +end -variables (hind : complete_lattice.independent (λ i, (ϕ i).range)) -variables (hinj : ∀ i, function.injective (ϕ i)) +lemma fun_on_eq_one_of_eq_one (S : finset I) (h : ∀ x ∈ S, f x = 1) : fun_on ϕ f S = 1 := +begin + induction S using finset.induction_on with i S hnmem ih, + { simp }, + { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + rw [h _ (finset.mem_insert_self _ _), ih (λ i h', h i (finset.mem_insert_of_mem h'))], + simp, }, +end -include hind -include hinj +open_locale big_operators +lemma order_of_fun_on_div_prod_card (S : finset I) [hfin : ∀ i, fintype (H i)]: + order_of (fun_on ϕ f S) ∣ (∏ i in S, fintype.card (H i)) := +begin + rw order_of_dvd_iff_pow_eq_one, + rw fun_on_pow, + apply fun_on_eq_one_of_eq_one, + intros i hmem, + simp only [pi.pow_apply, pi.one_apply], + rw ← order_of_dvd_iff_pow_eq_one, + calc order_of (f i) ∣ fintype.card (H i) : order_of_dvd_card_univ + ... ∣ (∏ (i : I) in S, fintype.card (H i)) : finset.dvd_prod_of_mem _ hmem, +end -lemma injective_of_independent : function.injective (hom ϕ) := +/-- The skelleton of a injectivitiy proof -/ +lemma injective_aux + (hpred : ∀ (f : Π i : I, H i) i S, + i ∉ S → ϕ i (f i) * fun_on ϕ f S = 1 → f i = 1 ∧ fun_on ϕ f S = 1 ) : + function.injective (hom ϕ) := begin apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, @@ -201,20 +256,84 @@ begin { simp }, { intro heq1, simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem] at heq1, - have hnmem' : i ∉ (S : set I), by simpa, - have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, - { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, - { simp, }, - { apply fun_on_in_sup_range, }, }, + have heq1' : f i = 1 ∧ fun_on ϕ f S = 1 := hpred f i S hnmem heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, simp only [finset.mem_insert] at h, rcases h with ⟨rfl | _⟩, - { apply hinj i, simpa, }, + { exact heq1i }, { exact (ih _ h), } } end +section injective_of_coprime_order + + +variables [∀ i, fintype (H i)] +variables (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) +variables (hinj : ∀ i, function.injective (ϕ i)) + +include hcoprime +include hinj + +open_locale big_operators +lemma injective_of_coprime_order : function.injective (hom ϕ) := injective_aux ϕ +begin + intros f i S hnmem heq1, + let x := ϕ i (f i), + let y := fun_on ϕ f S, + let p := ∏ (i : I) in S, fintype.card (H i), + let p' := ∏ (i : I) in S, (↑(fintype.card (H i)) : ℤ), + have h1 := calc order_of x = order_of (f i) : order_of_injective _ (hinj i) _ + ... ∣ fintype.card (H i) : order_of_dvd_card_univ, + have h2 := calc order_of x = order_of (y⁻¹) : congr_arg _ (eq_inv_of_mul_eq_one heq1) + ... = order_of y : order_of_inv _ + ... ∣ p : order_of_fun_on_div_prod_card ϕ f S , + have hcop : p.coprime (fintype.card (H i)), + { apply coprime_prod_left, + intros j hmem, + apply hcoprime, + rintro rfl, contradiction, }, + have hx : ϕ i (f i) = 1, + { have := nat.dvd_gcd h2 h1, + unfold nat.coprime at hcop, + simpa [hcop] using this, }, + have hf : f i = 1, by { apply hinj i, simpa using hx }, + have hy : fun_on ϕ f S = 1, by simpa only [hx, one_mul] using heq1, + exact ⟨ hf, hy ⟩ +end + +noncomputable +def mul_equiv_of_coprime_order : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := +begin + rw ← range_eq, + exact (monoid_hom.of_injective (injective_of_coprime_order _ hcoprime hinj)), +end + +end injective_of_coprime_order + +section injective_of_independent + +variables (hind : complete_lattice.independent (λ i, (ϕ i).range)) +variables (hinj : ∀ i, function.injective (ϕ i)) + +include hind +include hinj + +lemma injective_of_independent : function.injective (hom ϕ) := injective_aux ϕ +begin + intros f i S hnmem heq1, + have hnmem' : i ∉ (S : set I), by simpa, + have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, + { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, + { simp, }, + { apply fun_on_in_sup_range, }, }, + rcases heq1' with ⟨ heq1i, heq1S ⟩, + split, + { apply hinj i, simpa, }, + { exact heq1S } , +end + noncomputable def mul_equiv : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := begin @@ -222,6 +341,8 @@ begin exact (monoid_hom.of_injective (injective_of_independent _ hind hinj)), end +end injective_of_independent + end pi_hom lemma pairwise_elements_commute_of_normal_of_independent @@ -258,3 +379,26 @@ begin { simpa using hind, }, { exact λ _, subtype.val_injective, } end + + +open_locale classical -- so that we know that subgroups are fintype + +noncomputable +def internal_product_of_coprime + [fintype G] + {I : Type*} [fintype I] {H : I → subgroup G} + (hnorm : ∀ i, (H i).normal) + (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : + (Π (i : I), H i) ≃* (⨆ (i : I), H i : subgroup G) := +begin + haveI : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), + commute ((H i).subtype x) ((H j).subtype y)) := fact.mk + begin + sorry, + end, + + have : (⨆ (i : I), H i: subgroup G) = (⨆ (i : I), (H i).subtype.range : subgroup G), by simp, + rw this, clear this, + refine (pi_hom.mul_equiv_of_coprime_order _ hcoprime _), + { exact λ _, subtype.val_injective, } +end From 723e30a3b8edf262652d81edd97c631a6d35ba30 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 31 Jan 2022 19:59:06 +0100 Subject: [PATCH 011/116] Injectivity via coprimality of groups. More suited for sylow group. --- src/group_theory/pi_hom.lean | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 2fa62bfcb6693..5cb4049611fe9 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -359,6 +359,7 @@ begin show x * y = y * x, by simpa [mul_assoc] using this, end +/-- Subgroups of a finite group form an internal product if they are normal, and independent -/ noncomputable def internal_product {I : Type*} [fintype I] [decidable_eq I] {H : I → subgroup G} @@ -381,20 +382,26 @@ begin end -open_locale classical -- so that we know that subgroups are fintype +open_locale classical -- so that we know that subgroups of fintypes are fintype +/-- Subgroups of a finite group form an internal product if they are normal, pairwise disjoint, +and have pairwise coprime orders -/ noncomputable def internal_product_of_coprime [fintype G] {I : Type*} [fintype I] {H : I → subgroup G} (hnorm : ∀ i, (H i).normal) - (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : + (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) + (hdisjoint : ∀ i j, i ≠ j → disjoint (H i) (H j)) : (Π (i : I), H i) ≃* (⨆ (i : I), H i : subgroup G) := begin haveI : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := fact.mk begin - sorry, + simp, + rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, + refine subgroup.commute_of_normal_of_disjoint _ _ (hnorm _) (hnorm _) _ x y hx hy, + exact (hdisjoint i j hne) end, have : (⨆ (i : I), H i: subgroup G) = (⨆ (i : I), (H i).subtype.range : subgroup G), by simp, From 88398c279be9c96bdb505d3bc5ff2c873f1737df Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 15:16:28 +0100 Subject: [PATCH 012/116] Big clean-up and documentation --- src/group_theory/pi_hom.lean | 431 ++++++++++++++++++----------------- 1 file changed, 218 insertions(+), 213 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 5cb4049611fe9..a909111dee234 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -10,12 +10,43 @@ import group_theory.order_of_element import data.finset.noncomm_prod import ring_theory.coprime.lemmas /-! -# TODO + +# Canonical homomorphism from a pi group + +This file defines the construction of the canoncial homomorphism from a product group. + +Given a family of morphisms `ϕ i : H i →* G` for each `i : I` where elements in the +image of different morphism commute, we obtain a canoncial morphism +`pi_hom.hom : (Π (i : I), H i) →* G` that coincides with `ϕ`` + +## Main definitions + +* `pi_hom.hom : (Π (i : I), H i) →* G` is the main homomorphism +* `pi_hom_restr.hom (S : fintype S): (Π (i : I), H i) →* G` is the homomorphism restricted to the set + `S`, and mostly of internal interest in this file, to help with induction +* `subgroup_pi_hom.hom : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` and + the subgroup embedding. + +## Main theorems + +* `pi_hom.hom` conicides with each `ϕ` +* `pi_hom.range`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).range` +* The range of `subgroup_pi_hom.hom` is `⨆ (i : I), H i` +* `pi_hom.pow`: `pi_hom.hom` commutes with potentation. +* `pi_hom.injective_of_independent`: `pi_hom.hom` is injective if the `ϕ` are injective and the + ranges of the `ϕ` are independent. +* `independent_range_of_coprime_order`: If the `ϕ` are injective and the `H i` have coprime orders, + then the ranges are independent. +* `independent_of_coprime_order`: If commuting, normal `H i` have coprime orders, they are + independent. +* `pairwise_elements_commute_of_normal_of_independent`: Normal, independent subgroups have + commuting elements. -/ open_locale big_operators +-- A lot of faff just to transport `is_coprime.prod_left` from ℤ to ℕ lemma coprime_prod_left {I : Type*} {x : ℕ} {s : I → ℕ} {t : finset I} : @@ -33,8 +64,11 @@ begin end -variables {G : Type*} [group G] +section with_group + +parameters {G : Type*} [group G] +-- TODO: Move to suitable file @[simp] lemma order_of_inv (x : G) : order_of x⁻¹ = order_of x := begin @@ -43,6 +77,7 @@ begin { nth_rewrite 0 ← (inv_inv x), rw [inv_pow, pow_order_of_eq_one, one_inv], } end +-- TODO: Move to suitable file? Or too specialized? lemma mul_eq_one_of_disjoint {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) (heq : x * y = 1) : x = 1 ∧ y = 1 := @@ -56,183 +91,151 @@ begin simp end -namespace pi_hom +section family_of_groups --- We have an family of group -variables {I : Type*} [fintype I] [decidable_eq I] {H : I → Type*} [∀ i, group (H i)] +-- We have an family of groups +parameters {I : Type*} [fintype I] [decidable_eq I] {H : I → Type*} [∀ i, group (H i)] -- And morphism ϕ into G -variables (ϕ : Π (i : I), H i →* G) +parameters (ϕ : Π (i : I), H i →* G) -- We assume that the elements of different morphism commute -- Since we need this all over the place we wrap it up in `fact` -variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)) ] +-- TODO: Worth making this a real `class` over `ϕ`? +parameters [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y))] include hcomm -- Elements of `Π (i : I), H i` are called `f` and `g` here variables (f g : Π (i : I), H i) -/-- A wrapper around `finset.noncomm_prod` that discharges the commutativiy requirement using -`hcomm` -/ -def fun_on (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases (i = j), { subst h }, { exact hcomm.elim _ _ h _ _} } -/-- The product of `ϕ i (f i)` for all `i : I` -/ -def to_fun : G := fun_on ϕ f finset.univ +namespace pi_hom_restr -@[simp] -lemma fun_on_empty : fun_on ϕ f ∅ = 1 := by simp [fun_on] +-- In this section, we restrict the hom to a set S +variables (S : finset I) + +/-- The underlying function of `pi_hom_restr.hom` -/ +def to_fun (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ + by { rintros i - j -, by_cases (i = j), { subst h }, { exact hcomm.elim _ _ h _ _} } @[simp] -lemma fun_on_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : - fun_on ϕ f (insert i S) = ϕ i (f i) * fun_on ϕ f S := -finset.noncomm_prod_insert_of_not_mem _ _ _ _ h +lemma to_fun_empty : to_fun f ∅ = 1 := by simp [to_fun] -/- @[simp] -lemma fun_on_cons (S : finset I) (i : I) (h : i ∉ S) : - fun_on ϕ f (finset.cons i S h) = ϕ i (f i) * fun_on ϕ f S := -by { rw finset.cons_eq_insert i S h, exact (fun_on_insert_of_not_mem _ _ _ _ h) } --/ +lemma to_fun_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : + to_fun f (insert i S) = ϕ i (f i) * to_fun f S := +finset.noncomm_prod_insert_of_not_mem _ _ _ _ h @[simp] -lemma fun_on_one (S : finset I) : fun_on ϕ 1 S = 1 := +lemma to_fun_one : to_fun 1 S = 1 := begin induction S using finset.cons_induction_on with i S hnmem ih, - { simp }, - { simp [ih, hnmem], } + { simp }, { simp [ih, hnmem], } end -@[simp] -lemma to_fun_one : to_fun ϕ 1 = 1 := fun_on_one _ _ - -lemma fun_on_commutes (S : finset I) (i : I) (hnmem : i ∉ S) : - commute (ϕ i (g i)) (fun_on ϕ f S) := +lemma to_fun_commutes (i : I) (hnmem : i ∉ S) : + commute (ϕ i (g i)) (to_fun f S) := begin induction S using finset.induction_on with j S hnmem' ih, { simp, }, - { simp only [fun_on_insert_of_not_mem _ _ _ _ hnmem'], + { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], have hij : i ≠ j, by {simp at hnmem, tauto}, have hiS : i ∉ S, by {simp at hnmem, tauto}, - calc ϕ i (g i) * (ϕ j (f j) * fun_on ϕ f S) - = (ϕ i (g i) * ϕ j (f j)) * fun_on ϕ f S : by rw ← mul_assoc - ... = (ϕ j (f j) * ϕ i (g i)) * fun_on ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} - ... = ϕ j (f j) * (ϕ i (g i) * fun_on ϕ f S) : by rw mul_assoc - ... = ϕ j (f j) * (fun_on ϕ f S * ϕ i (g i)) : by { congr' 1, apply (ih hiS) } - ... = (ϕ j (f j) * fun_on ϕ f S) * ϕ i (g i) : by rw ← mul_assoc } + calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) -- TODO: Why do I have to mention `ϕ`? + = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ f S : by rw ← mul_assoc + ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} + ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ f S) : by rw mul_assoc + ... = ϕ j (f j) * (to_fun ϕ f S * ϕ i (g i)) : by { congr' 1, apply (ih hiS) } + ... = (ϕ j (f j) * to_fun ϕ f S) * ϕ i (g i) : by rw ← mul_assoc } end @[simp] -lemma fun_on_mul (S : finset I) : fun_on ϕ (f * g) S = fun_on ϕ f S * fun_on ϕ g S := +lemma to_fun_mul (S : finset I) : to_fun (f * g) S = to_fun f S * to_fun g S := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, - { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], rw ih, clear ih, simp only [pi.mul_apply, map_mul], repeat { rw mul_assoc }, congr' 1, repeat { rw ← mul_assoc }, congr' 1, - exact (fun_on_commutes _ _ _ S i hnmem), } + exact (to_fun_commutes _ _ _ S i hnmem), } end -lemma fun_on_in_sup_range (S : finset I) : - fun_on ϕ f S ∈ ⨆ (i : I) (H_1 : i ∈ S), (ϕ i).range := + +lemma to_fun_in_sup_range (S : finset I) : + to_fun f S ∈ ⨆ i ∈ S, (ϕ i).range := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, - { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], refine (subgroup.mul_mem _ _ _), { apply (subgroup.mem_Sup_of_mem), { use i }, { simp, }, }, { refine (@bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih), by { intro, simp, intro h, right, exact h}, } } end -lemma to_fun_in_sup_range : to_fun ϕ f ∈ ⨆ (i : I), (ϕ i).range := -bsupr_le_supr _ (λ i, (ϕ i).range) (fun_on_in_sup_range ϕ f finset.univ) - -@[simp] -lemma to_fun_mul : to_fun ϕ (f * g) = to_fun ϕ f * to_fun ϕ g := fun_on_mul _ _ _ _ - +/-- The canonical homomorphism from a pi group, restricted to a subset -/ def hom : (Π (i : I), H i) →* G := -{ to_fun := to_fun ϕ, +{ to_fun := λ f, to_fun f S, map_one' := to_fun_one _, - map_mul' := λ f g, to_fun_mul _ f g, } - -omit hcomm -def just_one (i : I) (y : H i) : Π (i : I), H i := - λ j, if h : j = i then by { subst h; exact y} else 1 - -@[simp] -lemma just_one_eq (i : I) (y : H i) : just_one i y i = y := -by { unfold just_one, simp } - -@[simp] -lemma just_one_ne (i : I) (y : H i) (j : I) (h : i ≠ j) : - just_one i y j = (1 : H j) := -by { unfold just_one, have : ¬ (j = i), by cc, simp [this], } - -include hcomm + map_mul' := λ f g, to_fun_mul f g S, } -lemma fun_on_just_one (i : I) (y : H i) (S : finset I) : - fun_on ϕ (just_one i y) S = if i ∈ S then ϕ i y else 1 := +lemma to_fun_update_one (i : I) (y : H i) (S : finset I) : + to_fun (function.update 1 i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, { simp, }, - { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], by_cases (i = j), - { subst h, - rw ih, - simp only [just_one_eq, mul_ite, mul_one, finset.cons_eq_insert, finset.mem_insert, - eq_self_iff_true, true_or, if_true, ite_eq_right_iff, mul_left_eq_self], - intro i, contradiction, }, + { subst h, rw ih, simp [hnmem], }, { change i ≠ j at h, - simp [h], + simp [h, hnmem, function.update_noteq (ne_comm.mp h)], exact ih, } } end -lemma to_fun_just_one (i : I) (y : H i) : - to_fun ϕ (just_one i y) = ϕ i y := -begin - unfold to_fun, - rw fun_on_just_one ϕ i y _, - simp, -end +@[simp] +lemma hom_update_one (i : I) (y : H i): + hom S (function.update 1 i y) = if i ∈ S then ϕ i y else 1 := to_fun_update_one _ _ _ -lemma range_eq : (hom ϕ).range = (⨆ (i : I), (ϕ i).range) := +lemma range : (hom S).range = ⨆ i ∈ S, (ϕ i).range := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, - exact (to_fun_in_sup_range ϕ f), }, - { refine (supr_le _), - rintro i x ⟨y, rfl⟩, - exact ⟨_, to_fun_just_one _ i y⟩, } + exact (to_fun_in_sup_range ϕ f S), }, + { refine (bsupr_le _), + rintro i hmem x ⟨y, rfl⟩, + use (function.update 1 i y), + simp [hmem], + } end -lemma fun_on_pow (k : ℕ) (S : finset I) : (fun_on ϕ f S) ^ k = fun_on ϕ (f ^ k) S := +lemma pow (k : ℕ) : (hom S f) ^ k = hom S (f ^ k) := begin + change (to_fun ϕ f S) ^ k = to_fun ϕ (f ^ k) S, induction S using finset.induction_on with i S hnmem ih, { simp }, - { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], - rw [(fun_on_commutes ϕ f f S i hnmem).mul_pow, ih, pi.pow_apply, map_pow] }, + { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], + rw [(to_fun_commutes ϕ f f S i hnmem).mul_pow, ih, pi.pow_apply, map_pow] }, end -lemma fun_on_eq_one_of_eq_one (S : finset I) (h : ∀ x ∈ S, f x = 1) : fun_on ϕ f S = 1 := +lemma hom_eq_one_of_eq_one (h : ∀ x ∈ S, f x = 1) : hom S f = 1 := begin + change to_fun ϕ f S = 1, induction S using finset.induction_on with i S hnmem ih, { simp }, - { simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem], + { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], rw [h _ (finset.mem_insert_self _ _), ih (λ i h', h i (finset.mem_insert_of_mem h'))], simp, }, end -open_locale big_operators -lemma order_of_fun_on_div_prod_card (S : finset I) [hfin : ∀ i, fintype (H i)]: - order_of (fun_on ϕ f S) ∣ (∏ i in S, fintype.card (H i)) := +lemma order_of_hom_dvd_prod_card [hfin : ∀ i, fintype (H i)]: + order_of (hom S f) ∣ (∏ i in S, fintype.card (H i)) := begin rw order_of_dvd_iff_pow_eq_one, - rw fun_on_pow, - apply fun_on_eq_one_of_eq_one, + rw pow, + apply hom_eq_one_of_eq_one, intros i hmem, simp only [pi.pow_apply, pi.one_apply], rw ← order_of_dvd_iff_pow_eq_one, @@ -240,115 +243,162 @@ begin ... ∣ (∏ (i : I) in S, fintype.card (H i)) : finset.dvd_prod_of_mem _ hmem, end -/-- The skelleton of a injectivitiy proof -/ -lemma injective_aux - (hpred : ∀ (f : Π i : I, H i) i S, - i ∉ S → ϕ i (f i) * fun_on ϕ f S = 1 → f i = 1 ∧ fun_on ϕ f S = 1 ) : - function.injective (hom ϕ) := +end pi_hom_restr + +namespace pi_hom + +/-- The product of `ϕ i (f i)` for all `i : I` -/ +def to_fun : G := pi_hom_restr.to_fun ϕ f finset.univ + +/-- The canonical homomorphism from a pi group -/ +def hom : (Π (i : I), H i) →* G := pi_hom_restr.hom ϕ finset.univ + +@[simp] +lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = ϕ i y := +by { change pi_hom_restr.hom ϕ finset.univ (function.update 1 i y) = ϕ i y, simp } + +lemma range : hom.range = ⨆ i : I, (ϕ i).range := +by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } + +lemma pow (k : ℕ) : (hom f) ^ k = hom (f ^ k) := pi_hom_restr.pow _ _ _ _ + +lemma injective_of_independent + (hind : complete_lattice.independent (λ i, (ϕ i).range)) + (hinj : ∀ i, function.injective (ϕ i)) : + function.injective hom := begin apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, intro f, - show fun_on ϕ f finset.univ = 1 → f = 1, - suffices : fun_on ϕ f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), + show pi_hom_restr.to_fun ϕ f finset.univ = 1 → f = 1, + suffices : pi_hom_restr.to_fun ϕ f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), induction (finset.univ : finset I) using finset.induction_on with i S hnmem ih, { simp }, { intro heq1, - simp only [ fun_on_insert_of_not_mem _ _ _ _ hnmem] at heq1, - have heq1' : f i = 1 ∧ fun_on ϕ f S = 1 := hpred f i S hnmem heq1, + simp only [ pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, + have hnmem' : i ∉ (S : set I), by simpa, + have heq1' : ϕ i (f i) = 1 ∧ pi_hom_restr.to_fun ϕ f S = 1, + { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, + { simp, }, + { apply pi_hom_restr.to_fun_in_sup_range, }, }, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, simp only [finset.mem_insert] at h, rcases h with ⟨rfl | _⟩, - { exact heq1i }, + { apply hinj i, simpa, }, { exact (ih _ h), } } end -section injective_of_coprime_order - - -variables [∀ i, fintype (H i)] -variables (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) -variables (hinj : ∀ i, function.injective (ϕ i)) - -include hcoprime -include hinj +end pi_hom -open_locale big_operators -lemma injective_of_coprime_order : function.injective (hom ϕ) := injective_aux ϕ +lemma independent_range_of_coprime_order [∀ i, fintype (H i)] + (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) + (hinj : ∀ i, function.injective (ϕ i)) : + complete_lattice.independent (λ i, (ϕ i).range) := begin - intros f i S hnmem heq1, - let x := ϕ i (f i), - let y := fun_on ϕ f S, + rintros i f ⟨hxi, hxp⟩, + simp only [ne.def, subgroup.coe_to_submonoid, set_like.mem_coe, + monoid_hom.coe_range, set.mem_range] at hxi hxp, + rcases hxi with ⟨ y, rfl ⟩, + let S := finset.erase finset.univ i, + have hnmem : i ∉ S := finset.not_mem_erase i finset.univ, + have : (⨆ (j : I) (x : ¬j = i), (ϕ j).range) = (⨆ j ∈ S, (ϕ j).range) := + begin + congr' 1, + ext i, + congr' 2, + apply supr_congr_Prop, + { simp [S] }, + { exact λ _, rfl } + end, + rw this at hxp, clear this, + rw ← @pi_hom_restr.range _ _ _ _ _ _ _ _ hcomm at hxp, + cases hxp with f heq1, + + let x := ϕ i y, + let z := pi_hom_restr.hom ϕ S f, + change z = x at heq1, let p := ∏ (i : I) in S, fintype.card (H i), - let p' := ∏ (i : I) in S, (↑(fintype.card (H i)) : ℤ), - have h1 := calc order_of x = order_of (f i) : order_of_injective _ (hinj i) _ + have h1 := calc order_of x = order_of y : order_of_injective _ (hinj i) _ ... ∣ fintype.card (H i) : order_of_dvd_card_univ, - have h2 := calc order_of x = order_of (y⁻¹) : congr_arg _ (eq_inv_of_mul_eq_one heq1) - ... = order_of y : order_of_inv _ - ... ∣ p : order_of_fun_on_div_prod_card ϕ f S , + have h2 := calc order_of x = order_of z : by rw [heq1] + ... ∣ p : pi_hom_restr.order_of_hom_dvd_prod_card ϕ f S , + have hcop : p.coprime (fintype.card (H i)), - { apply coprime_prod_left, - intros j hmem, - apply hcoprime, - rintro rfl, contradiction, }, - have hx : ϕ i (f i) = 1, + { apply coprime_prod_left, intros j hmem, apply hcoprime, rintro rfl, contradiction, }, + have hx : ϕ i y = 1, { have := nat.dvd_gcd h2 h1, unfold nat.coprime at hcop, simpa [hcop] using this, }, - have hf : f i = 1, by { apply hinj i, simpa using hx }, - have hy : fun_on ϕ f S = 1, by simpa only [hx, one_mul] using heq1, - exact ⟨ hf, hy ⟩ + simpa using hx, end -noncomputable -def mul_equiv_of_coprime_order : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := -begin - rw ← range_eq, - exact (monoid_hom.of_injective (injective_of_coprime_order _ hcoprime hinj)), +end family_of_groups + +namespace subgroup_pi_hom + +section subgroup_pi_hom + +-- We have an family of subgroups +parameters {I : Type*} [fintype I] [decidable_eq I] {H : I → subgroup G} + +-- Elements of `Π (i : I), H i` are called `f` and `g` here +variables (f g : Π (i : I), H i) + +section commuting_subgroups + +-- We assume that the elements of different subgroups commute +parameters (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) + +include hcomm + +lemma hcomm_subtype : + fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := +fact.mk begin + rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, + simp only [subgroup.coe_subtype, subgroup.coe_mk], + exact hcomm i j hne x y hx hy, end -end injective_of_coprime_order +/-- The canonical homomorphism from a pi group of subgroups -/ +def hom : (Π (i : I), H i) →* G := @pi_hom.hom _ _ _ _ _ _ _ (λ i, (H i).subtype) hcomm_subtype -section injective_of_independent +@[simp] +lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = y := +by apply pi_hom.hom_update_one -variables (hind : complete_lattice.independent (λ i, (ϕ i).range)) -variables (hinj : ∀ i, function.injective (ϕ i)) +lemma range : hom.range = ⨆ i : I, H i := +by { unfold hom, simp [pi_hom.range] } -include hind -include hinj +lemma pow (k : ℕ) : (hom f) ^ k = hom (f ^ k) := +by apply pi_hom_restr.pow -lemma injective_of_independent : function.injective (hom ϕ) := injective_aux ϕ +lemma injective_of_independent (hind : complete_lattice.independent H) : + function.injective hom := begin - intros f i S hnmem heq1, - have hnmem' : i ∉ (S : set I), by simpa, - have heq1' : ϕ i (f i) = 1 ∧ fun_on ϕ f S = 1, - { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, - { simp, }, - { apply fun_on_in_sup_range, }, }, - rcases heq1' with ⟨ heq1i, heq1S ⟩, - split, - { apply hinj i, simpa, }, - { exact heq1S } , + apply pi_hom.injective_of_independent, + { simpa using hind }, + { intro i, exact subtype.coe_injective } end -noncomputable -def mul_equiv : (Π (i : I), H i) ≃* (⨆ (i : I), (ϕ i).range : subgroup G) := +lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] + (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : + complete_lattice.independent H := begin - rw ← range_eq, - exact (monoid_hom.of_injective (injective_of_independent _ hind hinj)), + haveI := hcomm_subtype hcomm, + have := independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime + (λ i , subtype.coe_injective), + simpa using this, end -end injective_of_independent - -end pi_hom +end commuting_subgroups lemma pairwise_elements_commute_of_normal_of_independent {I : Type*} {H : I → subgroup G} (hnorm : ∀ i, (H i).normal) (hind : complete_lattice.independent H) : - (∀ i j : I, i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) := + (∀ i j : I, i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) := begin intros i j hne x y hx hy, have : H i ⊓ H j ≤ ⊥ := complete_lattice.independent.disjoint hind hne, @@ -359,53 +409,8 @@ begin show x * y = y * x, by simpa [mul_assoc] using this, end -/-- Subgroups of a finite group form an internal product if they are normal, and independent -/ -noncomputable -def internal_product - {I : Type*} [fintype I] [decidable_eq I] {H : I → subgroup G} - (hnorm : ∀ i, (H i).normal) (hind : complete_lattice.independent H) : - (Π (i : I), H i) ≃* (⨆ (i : I), H i : subgroup G) := -begin - haveI : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), - commute ((H i).subtype x) ((H j).subtype y)) := fact.mk - begin - intros i j hne x y, - induction x with x hx, - induction y with y hy, - exact pairwise_elements_commute_of_normal_of_independent hnorm hind i j hne x y hx hy, - end, - have : (⨆ (i : I), H i: subgroup G) = (⨆ (i : I), (H i).subtype.range : subgroup G), by simp, - rw this, clear this, - refine (pi_hom.mul_equiv _ _ _), - { simpa using hind, }, - { exact λ _, subtype.val_injective, } -end - - -open_locale classical -- so that we know that subgroups of fintypes are fintype +end subgroup_pi_hom -/-- Subgroups of a finite group form an internal product if they are normal, pairwise disjoint, -and have pairwise coprime orders -/ -noncomputable -def internal_product_of_coprime - [fintype G] - {I : Type*} [fintype I] {H : I → subgroup G} - (hnorm : ∀ i, (H i).normal) - (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) - (hdisjoint : ∀ i j, i ≠ j → disjoint (H i) (H j)) : - (Π (i : I), H i) ≃* (⨆ (i : I), H i : subgroup G) := -begin - haveI : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), - commute ((H i).subtype x) ((H j).subtype y)) := fact.mk - begin - simp, - rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, - refine subgroup.commute_of_normal_of_disjoint _ _ (hnorm _) (hnorm _) _ x y hx hy, - exact (hdisjoint i j hne) - end, +end subgroup_pi_hom - have : (⨆ (i : I), H i: subgroup G) = (⨆ (i : I), (H i).subtype.range : subgroup G), by simp, - rw this, clear this, - refine (pi_hom.mul_equiv_of_coprime_order _ hcoprime _), - { exact λ _, subtype.val_injective, } -end +end with_group From ba94eebc3497f7618585cf72343a9d4235834f17 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 15:19:02 +0100 Subject: [PATCH 013/116] Prune imports --- src/group_theory/pi_hom.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index a909111dee234..2109e3e975691 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -3,12 +3,11 @@ Copyright (c) 2022 Jocchim Breitner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joachim Breitner -/ -import group_theory.subgroup.basic -import tactic.group import group_theory.general_commutator import group_theory.order_of_element import data.finset.noncomm_prod import ring_theory.coprime.lemmas + /-! # Canonical homomorphism from a pi group From 9d87cf4f0bd055a9d7cf19932d04b31d2fd411ca Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 15:22:37 +0100 Subject: [PATCH 014/116] Remove kinda unused lemma --- src/group_theory/pi_hom.lean | 25 ++++--------------------- src/group_theory/subgroup/basic.lean | 27 ++++++++------------------- 2 files changed, 12 insertions(+), 40 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 2109e3e975691..71b8fdf513f10 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -21,8 +21,8 @@ image of different morphism commute, we obtain a canoncial morphism ## Main definitions * `pi_hom.hom : (Π (i : I), H i) →* G` is the main homomorphism -* `pi_hom_restr.hom (S : fintype S): (Π (i : I), H i) →* G` is the homomorphism restricted to the set - `S`, and mostly of internal interest in this file, to help with induction +* `pi_hom_restr.hom (S : fintype S): (Π (i : I), H i) →* G` is the homomorphism restricted to the + set `S`, and mostly of internal interest in this file, to allow inductive proofs. * `subgroup_pi_hom.hom : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` and the subgroup embedding. @@ -38,8 +38,6 @@ image of different morphism commute, we obtain a canoncial morphism then the ranges are independent. * `independent_of_coprime_order`: If commuting, normal `H i` have coprime orders, they are independent. -* `pairwise_elements_commute_of_normal_of_independent`: Normal, independent subgroups have - commuting elements. -/ @@ -206,8 +204,7 @@ begin { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, use (function.update 1 i y), - simp [hmem], - } + simp [hmem], } end lemma pow (k : ℕ) : (hom S f) ^ k = hom S (f ^ k) := @@ -392,21 +389,7 @@ begin simpa using this, end -end commuting_subgroups - -lemma pairwise_elements_commute_of_normal_of_independent - {I : Type*} {H : I → subgroup G} - (hnorm : ∀ i, (H i).normal) (hind : complete_lattice.independent H) : - (∀ i j : I, i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) := -begin - intros i j hne x y hx hy, - have : H i ⊓ H j ≤ ⊥ := complete_lattice.independent.disjoint hind hne, - have : ⁅H i, H j⁆ ≤ ⊥ := le_trans (general_commutator_le_inf _ _) this, - have : x * y * x ⁻¹ * y ⁻¹ = 1, - by { rw [← subgroup.mem_bot], exact this (general_commutator_containment _ _ hx hy), }, - have : (x * y * x ⁻¹ * y ⁻¹) * (y * x) = y * x, by { simp [this] }, - show x * y = y * x, by simpa [mul_assoc] using this, -end +end commuting_subgroups end subgroup_pi_hom diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 7bf6b86037809..9858947968d56 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1228,6 +1228,14 @@ begin exact hg (ϕ h), end +lemma _root_.comm_group.center_eq_top {G : Type*} [comm_group G] : center G = ⊤ := +by { rw [eq_top_iff'], intros x y, exact mul_comm y x } + +/-- A group is commutative if the center is the whole group -/ +def _root_.group.comm_group_of_center_eq_top (h : center G = ⊤) : comm_group G := +{ mul_comm := by { rw eq_top_iff' at h, intros x y, exact h y x }, + .. (_ : group G) } + variables {G} (H) /-- The `normalizer` of `H` is the largest subgroup of `G` inside which `H` is normal. -/ @[to_additive "The `normalizer` of `H` is the largest subgroup of `G` inside which `H` is normal."] @@ -2467,25 +2475,6 @@ begin rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this, end -/-- Elements of disjoint, normal subgroups commute -/ -@[to_additive] lemma commute_of_normal_of_disjoint - (H₁ H₂ : subgroup G) (hH₁ : H₁.normal) (hH₂ : H₂.normal) (hdis : disjoint H₁ H₂) - (x y : G) (hx : x ∈ H₁) (hy : y ∈ H₂) : - commute x y := -begin - suffices : x * y * x⁻¹ * y⁻¹ = 1, - { show x * y = y * x, by { rw [mul_assoc, mul_eq_one_iff_eq_inv] at this, simpa } }, - apply hdis, split, - { show x * y * x⁻¹ * y⁻¹ ∈ H₁, - repeat { rw mul_assoc }, - apply H₁.mul_mem hx, - repeat { rw ← mul_assoc }, - apply (hH₁.conj_mem _ (H₁.inv_mem hx)), }, - { show x * y * x⁻¹ * y⁻¹ ∈ H₂, - apply H₂.mul_mem _ (H₂.inv_mem hy), - apply (hH₂.conj_mem _ hy), } -end - end subgroup_normal end subgroup From 29f828a611d3614c38101f67d461598ac64e31eb Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 15:28:09 +0100 Subject: [PATCH 015/116] whitespace --- src/group_theory/pi_hom.lean | 12 +++--------- 1 file changed, 3 insertions(+), 9 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 71b8fdf513f10..6dc0ef7a6bf7d 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -9,7 +9,6 @@ import data.finset.noncomm_prod import ring_theory.coprime.lemmas /-! - # Canonical homomorphism from a pi group This file defines the construction of the canoncial homomorphism from a product group. @@ -60,7 +59,6 @@ begin apply h i hi, end - section with_group parameters {G : Type*} [group G] @@ -105,7 +103,6 @@ include hcomm -- Elements of `Π (i : I), H i` are called `f` and `g` here variables (f g : Π (i : I), H i) - namespace pi_hom_restr -- In this section, we restrict the hom to a set S @@ -160,7 +157,6 @@ begin exact (to_fun_commutes _ _ _ S i hnmem), } end - lemma to_fun_in_sup_range (S : finset I) : to_fun f S ∈ ⨆ i ∈ S, (ϕ i).range := begin @@ -310,7 +306,7 @@ begin { exact λ _, rfl } end, rw this at hxp, clear this, - rw ← @pi_hom_restr.range _ _ _ _ _ _ _ _ hcomm at hxp, + rw ← pi_hom_restr.range at hxp, cases hxp with f heq1, let x := ϕ i y, @@ -319,9 +315,8 @@ begin let p := ∏ (i : I) in S, fintype.card (H i), have h1 := calc order_of x = order_of y : order_of_injective _ (hinj i) _ ... ∣ fintype.card (H i) : order_of_dvd_card_univ, - have h2 := calc order_of x = order_of z : by rw [heq1] + have h2 := calc order_of x = order_of z : by rw heq1 ... ∣ p : pi_hom_restr.order_of_hom_dvd_prod_card ϕ f S , - have hcop : p.coprime (fintype.card (H i)), { apply coprime_prod_left, intros j hmem, apply hcoprime, rintro rfl, contradiction, }, have hx : ϕ i y = 1, @@ -347,10 +342,9 @@ section commuting_subgroups -- We assume that the elements of different subgroups commute parameters (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) - include hcomm -lemma hcomm_subtype : +instance hcomm_subtype : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := fact.mk begin rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, From 05e4029c893d8bcb8f95355a5703dfadd24a0662 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 15:50:54 +0100 Subject: [PATCH 016/116] Ups --- src/group_theory/pi_hom.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 6dc0ef7a6bf7d..5c4a0d5170944 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -344,7 +344,7 @@ section commuting_subgroups parameters (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) include hcomm -instance hcomm_subtype : +lemma hcomm_subtype : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := fact.mk begin rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, From e73b58cf95f7fc56346dfb752ba5d483023f7e55 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 22:02:14 +0100 Subject: [PATCH 017/116] Make it lint-clean --- src/group_theory/pi_hom.lean | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 5c4a0d5170944..f0d6a7554fc69 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -20,7 +20,7 @@ image of different morphism commute, we obtain a canoncial morphism ## Main definitions * `pi_hom.hom : (Π (i : I), H i) →* G` is the main homomorphism -* `pi_hom_restr.hom (S : fintype S): (Π (i : I), H i) →* G` is the homomorphism restricted to the +* `pi_hom_restr.hom (S : finset S) : (Π (i : I), H i) →* G` is the homomorphism restricted to the set `S`, and mostly of internal interest in this file, to allow inductive proofs. * `subgroup_pi_hom.hom : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` and the subgroup embedding. @@ -41,6 +41,7 @@ image of different morphism commute, we obtain a canoncial morphism -/ open_locale big_operators +open_locale classical -- A lot of faff just to transport `is_coprime.prod_left` from ℤ to ℕ lemma coprime_prod_left @@ -89,7 +90,8 @@ end section family_of_groups -- We have an family of groups -parameters {I : Type*} [fintype I] [decidable_eq I] {H : I → Type*} [∀ i, group (H i)] +parameters {I : Type*} [hfin : fintype I] +parameters {H : I → Type*} [∀ i, group (H i)] -- And morphism ϕ into G parameters (ϕ : Π (i : I), H i →* G) @@ -239,6 +241,8 @@ end pi_hom_restr namespace pi_hom +include hfin + /-- The product of `ϕ i (f i)` for all `i : I` -/ def to_fun : G := pi_hom_restr.to_fun ϕ f finset.univ @@ -285,6 +289,8 @@ end end pi_hom +include hfin + lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) (hinj : ∀ i, function.injective (ϕ i)) : @@ -333,7 +339,7 @@ namespace subgroup_pi_hom section subgroup_pi_hom -- We have an family of subgroups -parameters {I : Type*} [fintype I] [decidable_eq I] {H : I → subgroup G} +parameters {I : Type*} [hfin : fintype I] [hdec : decidable_eq I] {H : I → subgroup G} -- Elements of `Π (i : I), H i` are called `f` and `g` here variables (f g : Π (i : I), H i) @@ -352,8 +358,10 @@ fact.mk begin exact hcomm i j hne x y hx hy, end +include hfin + /-- The canonical homomorphism from a pi group of subgroups -/ -def hom : (Π (i : I), H i) →* G := @pi_hom.hom _ _ _ _ _ _ _ (λ i, (H i).subtype) hcomm_subtype +def hom : (Π (i : I), H i) →* G := @pi_hom.hom _ _ _ _ _ _ (λ i, (H i).subtype) hcomm_subtype @[simp] lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = y := From fdda033194d5c2dc80f330adafb9d46f654d7040 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 1 Feb 2022 22:06:16 +0100 Subject: [PATCH 018/116] Less underscores --- src/group_theory/pi_hom.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index f0d6a7554fc69..d0e36d4f42a04 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -361,7 +361,8 @@ end include hfin /-- The canonical homomorphism from a pi group of subgroups -/ -def hom : (Π (i : I), H i) →* G := @pi_hom.hom _ _ _ _ _ _ (λ i, (H i).subtype) hcomm_subtype +def hom : (Π (i : I), H i) →* G := + let _ := hcomm_subtype in by exactI pi_hom.hom (λ i, (H i).subtype) @[simp] lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = y := From 4b7f97f0b0fe8f1c0d4ab379f61c93fa7c4ded23 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 9 Feb 2022 09:44:23 +0100 Subject: [PATCH 019/116] Update src/group_theory/pi_hom.lean --- src/group_theory/pi_hom.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index d0e36d4f42a04..5f1c002205ea2 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -78,8 +78,7 @@ lemma mul_eq_one_of_disjoint {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) (heq : x * y = 1) : x = 1 ∧ y = 1 := begin - have : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), - subst this, + obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), have hy := H₂.inv_mem_iff.mp hy, have : x ∈ H₁ ⊓ H₂, by { simp, cc }, rw [hdis.eq_bot, subgroup.mem_bot] at this, From d576d7ca922e4553d758e40411035db6c6d06230 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 10 Feb 2022 11:33:36 +0100 Subject: [PATCH 020/116] Proof cosmetics --- src/group_theory/pi_hom.lean | 51 ++++++++++++++++-------------------- 1 file changed, 22 insertions(+), 29 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 5f1c002205ea2..7316adf9c0e58 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -111,10 +111,10 @@ variables (S : finset I) /-- The underlying function of `pi_hom_restr.hom` -/ def to_fun (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases (i = j), { subst h }, { exact hcomm.elim _ _ h _ _} } + by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm.elim _ _ h _ _} } @[simp] -lemma to_fun_empty : to_fun f ∅ = 1 := by simp [to_fun] +lemma to_fun_empty : to_fun f ∅ = 1 := finset.noncomm_prod_empty _ _ @[simp] lemma to_fun_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : @@ -139,9 +139,9 @@ begin have hiS : i ∉ S, by {simp at hnmem, tauto}, calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) -- TODO: Why do I have to mention `ϕ`? = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ f S : by rw ← mul_assoc - ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by {congr' 1, apply (fact.elim hcomm _ _ hij)} + ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by { congr' 1, apply fact.elim hcomm _ _ hij } ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ f S) : by rw mul_assoc - ... = ϕ j (f j) * (to_fun ϕ f S * ϕ i (g i)) : by { congr' 1, apply (ih hiS) } + ... = ϕ j (f j) * (to_fun ϕ f S * ϕ i (g i)) : by { congr' 1, apply ih hiS } ... = (ϕ j (f j) * to_fun ϕ f S) * ϕ i (g i) : by rw ← mul_assoc } end @@ -150,7 +150,7 @@ lemma to_fun_mul (S : finset I) : to_fun (f * g) S = to_fun f S * to_fun g S := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, - { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], + { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], rw ih, clear ih, simp only [pi.mul_apply, map_mul], repeat { rw mul_assoc }, congr' 1, @@ -163,11 +163,11 @@ lemma to_fun_in_sup_range (S : finset I) : begin induction S using finset.induction_on with i S hnmem ih, { simp, }, - { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], + { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], refine (subgroup.mul_mem _ _ _), - { apply (subgroup.mem_Sup_of_mem), { use i }, { simp, }, }, - { refine (@bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih), - by { intro, simp, intro h, right, exact h}, } } + { apply subgroup.mem_Sup_of_mem, { use i }, { simp, }, }, + { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih, + exact λ i, finset.mem_insert_of_mem } } end /-- The canonical homomorphism from a pi group, restricted to a subset -/ @@ -185,8 +185,7 @@ begin by_cases (i = j), { subst h, rw ih, simp [hnmem], }, { change i ≠ j at h, - simp [h, hnmem, function.update_noteq (ne_comm.mp h)], - exact ih, } } + simpa [h, hnmem, function.update_noteq (ne_comm.mp h)] using ih, } } end @[simp] @@ -232,8 +231,8 @@ begin intros i hmem, simp only [pi.pow_apply, pi.one_apply], rw ← order_of_dvd_iff_pow_eq_one, - calc order_of (f i) ∣ fintype.card (H i) : order_of_dvd_card_univ - ... ∣ (∏ (i : I) in S, fintype.card (H i)) : finset.dvd_prod_of_mem _ hmem, + calc order_of (f i) ∣ fintype.card (H i) : order_of_dvd_card_univ + ... ∣ (∏ (i : I) in S, fintype.card (H i)) : finset.dvd_prod_of_mem _ hmem, end end pi_hom_restr @@ -250,7 +249,7 @@ def hom : (Π (i : I), H i) →* G := pi_hom_restr.hom ϕ finset.univ @[simp] lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = ϕ i y := -by { change pi_hom_restr.hom ϕ finset.univ (function.update 1 i y) = ϕ i y, simp } +by { show pi_hom_restr.hom ϕ finset.univ (function.update 1 i y) = ϕ i y, simp } lemma range : hom.range = ⨆ i : I, (ϕ i).range := by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } @@ -301,7 +300,7 @@ begin rcases hxi with ⟨ y, rfl ⟩, let S := finset.erase finset.univ i, have hnmem : i ∉ S := finset.not_mem_erase i finset.univ, - have : (⨆ (j : I) (x : ¬j = i), (ϕ j).range) = (⨆ j ∈ S, (ϕ j).range) := + have : (⨆ (j : I) (_ : ¬j = i), (ϕ j).range) = (⨆ j ∈ S, (ϕ j).range) := begin congr' 1, ext i, @@ -318,16 +317,15 @@ begin let z := pi_hom_restr.hom ϕ S f, change z = x at heq1, let p := ∏ (i : I) in S, fintype.card (H i), - have h1 := calc order_of x = order_of y : order_of_injective _ (hinj i) _ - ... ∣ fintype.card (H i) : order_of_dvd_card_univ, + have h1 := calc order_of x = order_of y : order_of_injective _ (hinj i) _ + ... ∣ fintype.card (H i) : order_of_dvd_card_univ, have h2 := calc order_of x = order_of z : by rw heq1 - ... ∣ p : pi_hom_restr.order_of_hom_dvd_prod_card ϕ f S , + ... ∣ p : pi_hom_restr.order_of_hom_dvd_prod_card ϕ f S , have hcop : p.coprime (fintype.card (H i)), { apply coprime_prod_left, intros j hmem, apply hcoprime, rintro rfl, contradiction, }, have hx : ϕ i y = 1, - { have := nat.dvd_gcd h2 h1, - unfold nat.coprime at hcop, - simpa [hcop] using this, }, + { unfold nat.coprime at hcop, + simpa [hcop] using nat.dvd_gcd h2 h1, }, simpa using hx, end @@ -351,11 +349,7 @@ include hcomm lemma hcomm_subtype : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := -fact.mk begin - rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, - simp only [subgroup.coe_subtype, subgroup.coe_mk], - exact hcomm i j hne x y hx hy, -end +fact.mk (by { rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy }) include hfin @@ -368,7 +362,7 @@ lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = y := by apply pi_hom.hom_update_one lemma range : hom.range = ⨆ i : I, H i := -by { unfold hom, simp [pi_hom.range] } +by simp [hom, pi_hom.range] lemma pow (k : ℕ) : (hom f) ^ k = hom (f ^ k) := by apply pi_hom_restr.pow @@ -386,9 +380,8 @@ lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent H := begin haveI := hcomm_subtype hcomm, - have := independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime + simpa using independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime (λ i , subtype.coe_injective), - simpa using this, end end commuting_subgroups From 3cdf4b9086d60ebaa5826c67130c2cfdb141b8ac Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 10 Feb 2022 11:36:47 +0100 Subject: [PATCH 021/116] Move order of inv --- src/group_theory/order_of_element.lean | 8 ++++++++ src/group_theory/pi_hom.lean | 9 --------- 2 files changed, 8 insertions(+), 9 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index df41cfd566366..c7a03578625b7 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -317,6 +317,14 @@ begin order_of_dvd_iff_pow_eq_one] } end +@[simp, to_additive] +lemma order_of_inv (x : G) : order_of x⁻¹ = order_of x := +begin + apply nat.dvd_antisymm; rewrite order_of_dvd_iff_pow_eq_one, + { rw [inv_pow, pow_order_of_eq_one, one_inv], }, + { nth_rewrite 0 ← (inv_inv x), rw [inv_pow, pow_order_of_eq_one, one_inv], } +end + @[simp, norm_cast, to_additive] lemma order_of_subgroup {H : subgroup G} (y: H) : order_of (y : G) = order_of y := order_of_injective H.subtype subtype.coe_injective y diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 7316adf9c0e58..054b167147dfa 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -64,15 +64,6 @@ section with_group parameters {G : Type*} [group G] --- TODO: Move to suitable file -@[simp] -lemma order_of_inv (x : G) : order_of x⁻¹ = order_of x := -begin - apply nat.dvd_antisymm; rewrite order_of_dvd_iff_pow_eq_one, - { rw [inv_pow, pow_order_of_eq_one, one_inv], }, - { nth_rewrite 0 ← (inv_inv x), rw [inv_pow, pow_order_of_eq_one, one_inv], } -end - -- TODO: Move to suitable file? Or too specialized? lemma mul_eq_one_of_disjoint {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) From d33876f0142b181a0fd909470ae1273dfe73874b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 11 Feb 2022 17:42:25 +0100 Subject: [PATCH 022/116] Apply suggestions from code review Co-authored-by: Junyan Xu --- src/group_theory/order_of_element.lean | 6 +----- src/group_theory/pi_hom.lean | 18 ++++-------------- 2 files changed, 5 insertions(+), 19 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index c7a03578625b7..200801434e4a5 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -319,11 +319,7 @@ end @[simp, to_additive] lemma order_of_inv (x : G) : order_of x⁻¹ = order_of x := -begin - apply nat.dvd_antisymm; rewrite order_of_dvd_iff_pow_eq_one, - { rw [inv_pow, pow_order_of_eq_one, one_inv], }, - { nth_rewrite 0 ← (inv_inv x), rw [inv_pow, pow_order_of_eq_one, one_inv], } -end +by simp [order_of_eq_order_of_iff] @[simp, norm_cast, to_additive] lemma order_of_subgroup {H : subgroup G} (y: H) : order_of (y : G) = order_of y := diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 054b167147dfa..72c7b2eac38a6 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -15,7 +15,7 @@ This file defines the construction of the canoncial homomorphism from a product Given a family of morphisms `ϕ i : H i →* G` for each `i : I` where elements in the image of different morphism commute, we obtain a canoncial morphism -`pi_hom.hom : (Π (i : I), H i) →* G` that coincides with `ϕ`` +`pi_hom.hom : (Π (i : I), H i) →* G` that coincides with `ϕ` ## Main definitions @@ -27,7 +27,7 @@ image of different morphism commute, we obtain a canoncial morphism ## Main theorems -* `pi_hom.hom` conicides with each `ϕ` +* `pi_hom.hom` conicides with `ϕ i` when restricted to `H i` * `pi_hom.range`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).range` * The range of `subgroup_pi_hom.hom` is `⨆ (i : I), H i` * `pi_hom.pow`: `pi_hom.hom` commutes with potentation. @@ -48,17 +48,7 @@ lemma coprime_prod_left {I : Type*} {x : ℕ} {s : I → ℕ} {t : finset I} : (∀ (i : I), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : I) in t, s i) x := -begin - intro h, - rw ← nat.is_coprime_iff_coprime, - have := @nat.cast_prod _ ℤ _ s t, - simp [ -nat.cast_prod ] at this, - rw this, - apply is_coprime.prod_left, - intros i hi, - rw nat.is_coprime_iff_coprime, - apply h i hi, -end +finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) section with_group @@ -128,7 +118,7 @@ begin have hij : i ≠ j, by {simp at hnmem, tauto}, have hiS : i ∉ S, by {simp at hnmem, tauto}, - calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) -- TODO: Why do I have to mention `ϕ`? + calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ f S : by rw ← mul_assoc ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by { congr' 1, apply fact.elim hcomm _ _ hij } ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ f S) : by rw mul_assoc From f56d66d97d05f7f6d271459308e5f18ba8f81615 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 11 Feb 2022 17:53:21 +0100 Subject: [PATCH 023/116] Use monoid_hom.single, not function.update --- src/group_theory/pi_hom.lean | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 72c7b2eac38a6..e4eb5cfe867af 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -157,8 +157,8 @@ def hom : (Π (i : I), H i) →* G := map_one' := to_fun_one _, map_mul' := λ f g, to_fun_mul f g S, } -lemma to_fun_update_one (i : I) (y : H i) (S : finset I) : - to_fun (function.update 1 i y) S = if i ∈ S then ϕ i y else 1 := +lemma to_fun_single (i : I) (y : H i) (S : finset I) : + to_fun (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, { simp, }, @@ -166,12 +166,12 @@ begin by_cases (i = j), { subst h, rw ih, simp [hnmem], }, { change i ≠ j at h, - simpa [h, hnmem, function.update_noteq (ne_comm.mp h)] using ih, } } + simpa [h] using ih, } } end @[simp] -lemma hom_update_one (i : I) (y : H i): - hom S (function.update 1 i y) = if i ∈ S then ϕ i y else 1 := to_fun_update_one _ _ _ +lemma hom_single (i : I) (y : H i): + hom S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ lemma range : (hom S).range = ⨆ i ∈ S, (ϕ i).range := begin @@ -180,7 +180,7 @@ begin exact (to_fun_in_sup_range ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, - use (function.update 1 i y), + use (monoid_hom.single _ i y), simp [hmem], } end @@ -229,8 +229,8 @@ def to_fun : G := pi_hom_restr.to_fun ϕ f finset.univ def hom : (Π (i : I), H i) →* G := pi_hom_restr.hom ϕ finset.univ @[simp] -lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = ϕ i y := -by { show pi_hom_restr.hom ϕ finset.univ (function.update 1 i y) = ϕ i y, simp } +lemma hom_single (i : I) (y : H i): hom (monoid_hom.single _ i y) = ϕ i y := +by { show pi_hom_restr.hom ϕ finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } lemma range : hom.range = ⨆ i : I, (ϕ i).range := by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } @@ -339,8 +339,8 @@ def hom : (Π (i : I), H i) →* G := let _ := hcomm_subtype in by exactI pi_hom.hom (λ i, (H i).subtype) @[simp] -lemma hom_update_one (i : I) (y : H i): hom (function.update 1 i y) = y := -by apply pi_hom.hom_update_one +lemma hom_single (i : I) (y : H i): hom (monoid_hom.single _ i y) = y := +by apply pi_hom.hom_single lemma range : hom.range = ⨆ i : I, H i := by simp [hom, pi_hom.range] From e810447f7e0abfdb3a364f4c833f13e13664a377 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 12 Feb 2022 04:42:47 -0500 Subject: [PATCH 024/116] simplify independence proof and remove lemmas --- src/group_theory/order_of_element.lean | 5 ++ src/group_theory/pi_hom.lean | 97 +++++++------------------- 2 files changed, 29 insertions(+), 73 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 200801434e4a5..7cc1fe49360bc 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -136,6 +136,11 @@ is_periodic_pt.minimal_period_dvd ((is_periodic_pt_mul_iff_pow_eq_one _).mpr h) lemma order_of_dvd_iff_pow_eq_one {n : ℕ} : order_of x ∣ n ↔ x ^ n = 1 := ⟨λ h, by rw [pow_eq_mod_order_of, nat.mod_eq_zero_of_dvd h, pow_zero], order_of_dvd_of_pow_eq_one⟩ +@[to_additive map_add_order] +lemma map_order {H : Type*} [monoid H] (ψ : G →* H) (x : G) : + order_of (ψ x) ∣ order_of x := +by { apply order_of_dvd_of_pow_eq_one, rw [←map_pow, pow_order_of_eq_one], apply map_one } + @[to_additive exists_nsmul_eq_self_of_coprime] lemma exists_pow_eq_self_of_coprime (h : n.coprime (order_of x)) : ∃ m : ℕ, (x ^ n) ^ m = x := diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index e4eb5cfe867af..817792e05d255 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -3,11 +3,9 @@ Copyright (c) 2022 Jocchim Breitner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joachim Breitner -/ -import group_theory.general_commutator import group_theory.order_of_element import data.finset.noncomm_prod -import ring_theory.coprime.lemmas - +import data.fintype.card /-! # Canonical homomorphism from a pi group @@ -43,7 +41,6 @@ image of different morphism commute, we obtain a canoncial morphism open_locale big_operators open_locale classical --- A lot of faff just to transport `is_coprime.prod_left` from ℤ to ℕ lemma coprime_prod_left {I : Type*} {x : ℕ} {s : I → ℕ} {t : finset I} : @@ -66,6 +63,10 @@ begin subst this, simp end +-- I think it's worth keeping it and moving to appropriate file +lemma mul_eq_one_iff_disjoint {H₁ H₂ : subgroup G} : + disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := +by sorry section family_of_groups @@ -92,7 +93,7 @@ variables (S : finset I) /-- The underlying function of `pi_hom_restr.hom` -/ def to_fun (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm.elim _ _ h _ _} } + by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm.elim _ _ h _ _ } } @[simp] lemma to_fun_empty : to_fun f ∅ = 1 := finset.noncomm_prod_empty _ _ @@ -105,8 +106,8 @@ finset.noncomm_prod_insert_of_not_mem _ _ _ _ h @[simp] lemma to_fun_one : to_fun 1 S = 1 := begin - induction S using finset.cons_induction_on with i S hnmem ih, - { simp }, { simp [ih, hnmem], } + induction S using finset.cons_induction_on with i S hnmem ih, + { simp }, { simp [ih, hnmem], } end lemma to_fun_commutes (i : I) (hnmem : i ∉ S) : @@ -115,7 +116,6 @@ begin induction S using finset.induction_on with j S hnmem' ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], - have hij : i ≠ j, by {simp at hnmem, tauto}, have hiS : i ∉ S, by {simp at hnmem, tauto}, calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) @@ -184,37 +184,8 @@ begin simp [hmem], } end -lemma pow (k : ℕ) : (hom S f) ^ k = hom S (f ^ k) := -begin - change (to_fun ϕ f S) ^ k = to_fun ϕ (f ^ k) S, - induction S using finset.induction_on with i S hnmem ih, - { simp }, - { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], - rw [(to_fun_commutes ϕ f f S i hnmem).mul_pow, ih, pi.pow_apply, map_pow] }, -end - -lemma hom_eq_one_of_eq_one (h : ∀ x ∈ S, f x = 1) : hom S f = 1 := -begin - change to_fun ϕ f S = 1, - induction S using finset.induction_on with i S hnmem ih, - { simp }, - { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], - rw [h _ (finset.mem_insert_self _ _), ih (λ i h', h i (finset.mem_insert_of_mem h'))], - simp, }, -end - -lemma order_of_hom_dvd_prod_card [hfin : ∀ i, fintype (H i)]: - order_of (hom S f) ∣ (∏ i in S, fintype.card (H i)) := -begin - rw order_of_dvd_iff_pow_eq_one, - rw pow, - apply hom_eq_one_of_eq_one, - intros i hmem, - simp only [pi.pow_apply, pi.one_apply], - rw ← order_of_dvd_iff_pow_eq_one, - calc order_of (f i) ∣ fintype.card (H i) : order_of_dvd_card_univ - ... ∣ (∏ (i : I) in S, fintype.card (H i)) : finset.dvd_prod_of_mem _ hmem, -end +lemma pow (k : ℕ) : (hom S f) ^ k = hom S (f ^ k) := by simp +/- uses map_pow, and now unnecessary -/ end pi_hom_restr @@ -271,43 +242,24 @@ end pi_hom include hfin lemma independent_range_of_coprime_order [∀ i, fintype (H i)] - (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) - (hinj : ∀ i, function.injective (ϕ i)) : + (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := begin rintros i f ⟨hxi, hxp⟩, simp only [ne.def, subgroup.coe_to_submonoid, set_like.mem_coe, monoid_hom.coe_range, set.mem_range] at hxi hxp, - rcases hxi with ⟨ y, rfl ⟩, - let S := finset.erase finset.univ i, - have hnmem : i ∉ S := finset.not_mem_erase i finset.univ, - have : (⨆ (j : I) (_ : ¬j = i), (ϕ j).range) = (⨆ j ∈ S, (ϕ j).range) := - begin - congr' 1, - ext i, - congr' 2, - apply supr_congr_Prop, - { simp [S] }, - { exact λ _, rfl } - end, - rw this at hxp, clear this, - rw ← pi_hom_restr.range at hxp, - cases hxp with f heq1, - - let x := ϕ i y, - let z := pi_hom_restr.hom ϕ S f, - change z = x at heq1, - let p := ∏ (i : I) in S, fintype.card (H i), - have h1 := calc order_of x = order_of y : order_of_injective _ (hinj i) _ - ... ∣ fintype.card (H i) : order_of_dvd_card_univ, - have h2 := calc order_of x = order_of z : by rw heq1 - ... ∣ p : pi_hom_restr.order_of_hom_dvd_prod_card ϕ f S , - have hcop : p.coprime (fintype.card (H i)), - { apply coprime_prod_left, intros j hmem, apply hcoprime, rintro rfl, contradiction, }, - have hx : ϕ i y = 1, - { unfold nat.coprime at hcop, - simpa [hcop] using nat.dvd_gcd h2 h1, }, - simpa using hx, + rw [supr_subtype', ← @pi_hom.range _ _ _ _ _ _ _ _] at hxp, + rotate, apply_instance, + { constructor, intros j₁ j₂ hj, apply hcomm.1, intro h, apply hj, ext, exact h }, + cases hxp with g hgf, cases hxi with g' hg'f, + have hxi : order_of f ∣ fintype.card (H i), + { rw ← hg'f, exact (map_order _ _).trans order_of_dvd_card_univ }, + have hxp : order_of f ∣ ∏ j : {j // j ≠ i}, fintype.card (H j), + { subst hgf, apply (map_order _ _).trans, rw ← fintype.card_pi, + convert order_of_dvd_card_univ, apply_instance }, + change f = 1, rw [← pow_one f, ← order_of_dvd_iff_pow_eq_one], + convert ← nat.dvd_gcd hxp hxi, rw ← nat.coprime_iff_gcd_eq_one, + apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, end end family_of_groups @@ -361,8 +313,7 @@ lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent H := begin haveI := hcomm_subtype hcomm, - simpa using independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime - (λ i , subtype.coe_injective), + simpa using independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime, end end commuting_subgroups From 7cc8b50c2f7c2fc93caf54886b8056dc9e33edb0 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Feb 2022 17:46:02 +0100 Subject: [PATCH 025/116] Define hom for monoids there is some unsatisfying redundancy wrt. `range` and `mrange`. Also proved `mul_eq_one_iff_disjoint` to make this free of sorry. --- src/group_theory/pi_hom.lean | 185 +++++++++++++++++++++-------------- 1 file changed, 112 insertions(+), 73 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 817792e05d255..a3cb7b37d75e5 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -6,34 +6,35 @@ Authors: Joachim Breitner import group_theory.order_of_element import data.finset.noncomm_prod import data.fintype.card + /-! -# Canonical homomorphism from a pi group +# Canonical homomorphism from a pi monoid -This file defines the construction of the canoncial homomorphism from a product group. +This file defines the construction of the canoncial homomorphism from a family of monoids. -Given a family of morphisms `ϕ i : H i →* G` for each `i : I` where elements in the +Given a family of morphisms `ϕ i : N i →* M` for each `i : I` where elements in the image of different morphism commute, we obtain a canoncial morphism -`pi_hom.hom : (Π (i : I), H i) →* G` that coincides with `ϕ` +`pi_hom.hom : (Π (i : I), N i) →* M` that coincides with `ϕ` ## Main definitions -* `pi_hom.hom : (Π (i : I), H i) →* G` is the main homomorphism -* `pi_hom_restr.hom (S : finset S) : (Π (i : I), H i) →* G` is the homomorphism restricted to the +* `pi_hom.hom : (Π (i : I), N i) →* M` is the main homomorphism +* `pi_hom_restr.hom (S : finset S) : (Π (i : I), N i) →* M` is the homomorphism restricted to the set `S`, and mostly of internal interest in this file, to allow inductive proofs. * `subgroup_pi_hom.hom : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` and the subgroup embedding. ## Main theorems -* `pi_hom.hom` conicides with `ϕ i` when restricted to `H i` +* `pi_hom.hom` conicides with `ϕ i` when restricted to `N i` +* `pi_hom.mrange`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).mrange` * `pi_hom.range`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).range` * The range of `subgroup_pi_hom.hom` is `⨆ (i : I), H i` -* `pi_hom.pow`: `pi_hom.hom` commutes with potentation. -* `pi_hom.injective_of_independent`: `pi_hom.hom` is injective if the `ϕ` are injective and the - ranges of the `ϕ` are independent. -* `independent_range_of_coprime_order`: If the `ϕ` are injective and the `H i` have coprime orders, - then the ranges are independent. -* `independent_of_coprime_order`: If commuting, normal `H i` have coprime orders, they are +* `pi_hom.injective_of_independent`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` + are injective and the ranges of the `ϕ` are independent. +* `pi_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges are + independent. +* `independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, they are independent. -/ @@ -47,44 +48,43 @@ lemma coprime_prod_left (∀ (i : I), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : I) in t, s i) x := finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) -section with_group - -parameters {G : Type*} [group G] - --- TODO: Move to suitable file? Or too specialized? -lemma mul_eq_one_of_disjoint - {H₁ H₂ : subgroup G} (hdis : disjoint H₁ H₂) {x y : G} (hx : x ∈ H₁) (hy : y ∈ H₂) - (heq : x * y = 1) : x = 1 ∧ y = 1 := -begin - obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), - have hy := H₂.inv_mem_iff.mp hy, - have : x ∈ H₁ ⊓ H₂, by { simp, cc }, - rw [hdis.eq_bot, subgroup.mem_bot] at this, - subst this, - simp -end -- I think it's worth keeping it and moving to appropriate file -lemma mul_eq_one_iff_disjoint {H₁ H₂ : subgroup G} : +lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := -by sorry +begin + split, + { intros hdis x y hx hy heq, + obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), + have hy := H₂.inv_mem_iff.mp hy, + have : x ∈ H₁ ⊓ H₂, by { simp, cc }, + rw [hdis.eq_bot, subgroup.mem_bot] at this, + subst this, + simp }, + { rintros h x ⟨hx1, hx2⟩, + obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, + exact rfl, + }, +end -section family_of_groups +section family_of_monoids --- We have an family of groups +parameters {M : Type*} [monoid M] + +-- We have a family of monoids +-- The fintype assumption is not always used, but declared here, to keep things in order parameters {I : Type*} [hfin : fintype I] -parameters {H : I → Type*} [∀ i, group (H i)] +parameters {N : I → Type*} [∀ i, monoid (N i)] --- And morphism ϕ into G -parameters (ϕ : Π (i : I), H i →* G) +-- And morphisms ϕ into G +parameters (ϕ : Π (i : I), N i →* M) -- We assume that the elements of different morphism commute -- Since we need this all over the place we wrap it up in `fact` --- TODO: Worth making this a real `class` over `ϕ`? -parameters [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y))] +parameters [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y))] include hcomm --- Elements of `Π (i : I), H i` are called `f` and `g` here -variables (f g : Π (i : I), H i) +-- We use `f` and `g` to denote elements of `Π (i : I), N i` +variables (f g : Π (i : I), N i) namespace pi_hom_restr @@ -92,7 +92,7 @@ namespace pi_hom_restr variables (S : finset I) /-- The underlying function of `pi_hom_restr.hom` -/ -def to_fun (S : finset I) : G := finset.noncomm_prod S (λ i, ϕ i (f i)) $ +def to_fun (S : finset I) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm.elim _ _ h _ _ } } @[simp] @@ -118,7 +118,7 @@ begin { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], have hij : i ≠ j, by {simp at hnmem, tauto}, have hiS : i ∉ S, by {simp at hnmem, tauto}, - calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : G)) + calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : M)) = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ f S : by rw ← mul_assoc ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by { congr' 1, apply fact.elim hcomm _ _ hij } ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ f S) : by rw mul_assoc @@ -139,25 +139,25 @@ begin exact (to_fun_commutes _ _ _ S i hnmem), } end -lemma to_fun_in_sup_range (S : finset I) : - to_fun f S ∈ ⨆ i ∈ S, (ϕ i).range := +lemma to_fun_in_sup_mrange (S : finset I) : + to_fun f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], - refine (subgroup.mul_mem _ _ _), - { apply subgroup.mem_Sup_of_mem, { use i }, { simp, }, }, - { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih, + refine (submonoid.mul_mem _ _ _), + { apply submonoid.mem_Sup_of_mem, { use i }, { simp, }, }, + { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).mrange) _ ih, exact λ i, finset.mem_insert_of_mem } } end /-- The canonical homomorphism from a pi group, restricted to a subset -/ -def hom : (Π (i : I), H i) →* G := +def hom : (Π (i : I), N i) →* M := { to_fun := λ f, to_fun f S, map_one' := to_fun_one _, map_mul' := λ f g, to_fun_mul f g S, } -lemma to_fun_single (i : I) (y : H i) (S : finset I) : +lemma to_fun_single (i : I) (y : N i) (S : finset I) : to_fun (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, @@ -170,23 +170,20 @@ begin end @[simp] -lemma hom_single (i : I) (y : H i): +lemma hom_single (i : I) (y : N i): hom S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ -lemma range : (hom S).range = ⨆ i ∈ S, (ϕ i).range := +lemma mrange : (hom S).mrange = ⨆ i ∈ S, (ϕ i).mrange := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, - exact (to_fun_in_sup_range ϕ f S), }, + exact (to_fun_in_sup_mrange ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, use (monoid_hom.single _ i y), simp [hmem], } end -lemma pow (k : ℕ) : (hom S f) ^ k = hom S (f ^ k) := by simp -/- uses map_pow, and now unnecessary -/ - end pi_hom_restr namespace pi_hom @@ -194,24 +191,72 @@ namespace pi_hom include hfin /-- The product of `ϕ i (f i)` for all `i : I` -/ -def to_fun : G := pi_hom_restr.to_fun ϕ f finset.univ +def to_fun : M := pi_hom_restr.to_fun ϕ f finset.univ /-- The canonical homomorphism from a pi group -/ -def hom : (Π (i : I), H i) →* G := pi_hom_restr.hom ϕ finset.univ +def hom : (Π (i : I), N i) →* M := pi_hom_restr.hom ϕ finset.univ @[simp] -lemma hom_single (i : I) (y : H i): hom (monoid_hom.single _ i y) = ϕ i y := +lemma hom_single (i : I) (y : N i): hom (monoid_hom.single _ i y) = ϕ i y := by { show pi_hom_restr.hom ϕ finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } -lemma range : hom.range = ⨆ i : I, (ϕ i).range := -by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } +lemma mrange : hom.mrange = ⨆ i : I, (ϕ i).mrange := +by { show (pi_hom_restr.hom ϕ finset.univ).mrange = _, simp [pi_hom_restr.mrange] } + +end pi_hom + +end family_of_monoids -lemma pow (k : ℕ) : (hom f) ^ k = hom (f ^ k) := pi_hom_restr.pow _ _ _ _ +section family_of_groups + +variables {G : Type*} [group G] +variables {I : Type*} [hfin : fintype I] +variables {H : I → Type*} [∀ i, group (H i)] +variables (ϕ : Π (i : I), H i →* G) +variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y))] +include hcomm + +-- We use `f` and `g` to denote elements of `Π (i : I), H i` +variables (f g : Π (i : I), H i) + +-- The subgroup version of to_fun_in_sup_mrange +lemma pi_hom_restr.to_fun_in_sup_range (S : finset I) : + pi_hom_restr.to_fun ϕ f S ∈ ⨆ i ∈ S, (ϕ i).range := +begin + induction S using finset.induction_on with i S hnmem ih, + { simp, }, + { simp only [pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem], + refine (subgroup.mul_mem _ _ _), + { apply subgroup.mem_Sup_of_mem, { use i }, { simp, }, }, + { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih, + exact λ i, finset.mem_insert_of_mem } } +end + +-- The subgroup version of pi_hom_restr.mrange +lemma pi_hom_restr.range (S : finset I) : + (pi_hom_restr.hom ϕ S).range = ⨆ i ∈ S, (ϕ i).range := +begin + apply le_antisymm, + { rintro x ⟨f, rfl⟩, + exact (pi_hom_restr.to_fun_in_sup_range ϕ f S), }, + { refine (bsupr_le _), + rintro i hmem x ⟨y, rfl⟩, + use (monoid_hom.single _ i y), + simp [hmem], } +end + +include hfin + +namespace pi_hom + +-- The subgroup version of pi_hom.mrange +lemma range : (pi_hom.hom ϕ).range = ⨆ i : I, (ϕ i).range := +by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } lemma injective_of_independent (hind : complete_lattice.independent (λ i, (ϕ i).range)) (hinj : ∀ i, function.injective (ϕ i)) : - function.injective hom := + function.injective (pi_hom.hom ϕ):= begin apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, @@ -225,7 +270,7 @@ begin simp only [ pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, have hnmem' : i ∉ (S : set I), by simpa, have heq1' : ϕ i (f i) = 1 ∧ pi_hom_restr.to_fun ϕ f S = 1, - { apply mul_eq_one_of_disjoint (hind.disjoint_bsupr hnmem') _ _ heq1, + { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, { apply pi_hom_restr.to_fun_in_sup_range, }, }, rcases heq1' with ⟨ heq1i, heq1S ⟩, @@ -237,10 +282,6 @@ begin { exact (ih _ h), } } end -end pi_hom - -include hfin - lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := @@ -262,6 +303,8 @@ begin apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, end +end pi_hom + end family_of_groups namespace subgroup_pi_hom @@ -269,6 +312,7 @@ namespace subgroup_pi_hom section subgroup_pi_hom -- We have an family of subgroups +parameters {G : Type*} [group G] parameters {I : Type*} [hfin : fintype I] [hdec : decidable_eq I] {H : I → subgroup G} -- Elements of `Π (i : I), H i` are called `f` and `g` here @@ -297,9 +341,6 @@ by apply pi_hom.hom_single lemma range : hom.range = ⨆ i : I, H i := by simp [hom, pi_hom.range] -lemma pow (k : ℕ) : (hom f) ^ k = hom (f ^ k) := -by apply pi_hom_restr.pow - lemma injective_of_independent (hind : complete_lattice.independent H) : function.injective hom := begin @@ -313,7 +354,7 @@ lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent H := begin haveI := hcomm_subtype hcomm, - simpa using independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime, + simpa using pi_hom.independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime, end end commuting_subgroups @@ -321,5 +362,3 @@ end commuting_subgroups end subgroup_pi_hom end subgroup_pi_hom - -end with_group From e1257d0b9ddf01bc2ede2f2db69f74db0b95b484 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Feb 2022 18:10:50 +0100 Subject: [PATCH 026/116] Apply suggestions from code review --- src/group_theory/pi_hom.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index a3cb7b37d75e5..8f072ecb65fa8 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -62,8 +62,7 @@ begin simp }, { rintros h x ⟨hx1, hx2⟩, obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, - exact rfl, - }, + exact rfl, }, end section family_of_monoids From 194253f078da76417cb96b2fa91d1f8e7dd59679 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 12 Feb 2022 13:49:08 -0500 Subject: [PATCH 027/116] polish proof --- src/group_theory/pi_hom.lean | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 8f072ecb65fa8..8e2a0c3f00e31 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -285,18 +285,15 @@ lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := begin - rintros i f ⟨hxi, hxp⟩, - simp only [ne.def, subgroup.coe_to_submonoid, set_like.mem_coe, - monoid_hom.coe_range, set.mem_range] at hxi hxp, + rintros i f ⟨hxi, hxp⟩, dsimp at hxi hxp, rw [supr_subtype', ← @pi_hom.range _ _ _ _ _ _ _ _] at hxp, rotate, apply_instance, - { constructor, intros j₁ j₂ hj, apply hcomm.1, intro h, apply hj, ext, exact h }, + { split, intros _ _ hj, apply hcomm.1, exact hj ∘ subtype.ext }, cases hxp with g hgf, cases hxi with g' hg'f, have hxi : order_of f ∣ fintype.card (H i), { rw ← hg'f, exact (map_order _ _).trans order_of_dvd_card_univ }, have hxp : order_of f ∣ ∏ j : {j // j ≠ i}, fintype.card (H j), - { subst hgf, apply (map_order _ _).trans, rw ← fintype.card_pi, - convert order_of_dvd_card_univ, apply_instance }, + { rw [← hgf, ← fintype.card_pi], exact (map_order _ _).trans order_of_dvd_card_univ }, change f = 1, rw [← pow_one f, ← order_of_dvd_iff_pow_eq_one], convert ← nat.dvd_gcd hxp hxi, rw ← nat.coprime_iff_gcd_eq_one, apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, From 50712bb584328cbdaf1ed63010d7558eb55f14c6 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 14 Feb 2022 11:07:09 +0100 Subject: [PATCH 028/116] =?UTF-8?q?Don=E2=80=99t=20use=20fact=20nor=20para?= =?UTF-8?q?meters?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/group_theory/pi_hom.lean | 120 +++++++++++++++++++---------------- 1 file changed, 65 insertions(+), 55 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 8e2a0c3f00e31..03e5d799452e7 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -10,7 +10,7 @@ import data.fintype.card /-! # Canonical homomorphism from a pi monoid -This file defines the construction of the canoncial homomorphism from a family of monoids. +This file defines txinput set-prop 11 "libinput Scroll Method Enabled" 0, 0, 1 #he construction of the canoncial homomorphism from a family of monoids. Given a family of morphisms `ϕ i : N i →* M` for each `i : I` where elements in the image of different morphism commute, we obtain a canoncial morphism @@ -67,19 +67,18 @@ end section family_of_monoids -parameters {M : Type*} [monoid M] +variables {M : Type*} [monoid M] -- We have a family of monoids -- The fintype assumption is not always used, but declared here, to keep things in order -parameters {I : Type*} [hfin : fintype I] -parameters {N : I → Type*} [∀ i, monoid (N i)] +variables {I : Type*} [hfin : fintype I] +variables {N : I → Type*} [∀ i, monoid (N i)] -- And morphisms ϕ into G -parameters (ϕ : Π (i : I), N i →* M) +variables (ϕ : Π (i : I), N i →* M) -- We assume that the elements of different morphism commute --- Since we need this all over the place we wrap it up in `fact` -parameters [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y))] +variables (hcomm : ∀ (i j : I), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y)) include hcomm -- We use `f` and `g` to denote elements of `Π (i : I), N i` @@ -92,41 +91,44 @@ variables (S : finset I) /-- The underlying function of `pi_hom_restr.hom` -/ def to_fun (S : finset I) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm.elim _ _ h _ _ } } + by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } } + +variable {hcomm} @[simp] -lemma to_fun_empty : to_fun f ∅ = 1 := finset.noncomm_prod_empty _ _ +lemma to_fun_empty : to_fun ϕ hcomm f ∅ = 1 := finset.noncomm_prod_empty _ _ @[simp] lemma to_fun_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : - to_fun f (insert i S) = ϕ i (f i) * to_fun f S := + to_fun ϕ hcomm f (insert i S) = ϕ i (f i) * to_fun ϕ hcomm f S := finset.noncomm_prod_insert_of_not_mem _ _ _ _ h @[simp] -lemma to_fun_one : to_fun 1 S = 1 := +lemma to_fun_one : to_fun ϕ hcomm 1 S = 1 := begin induction S using finset.cons_induction_on with i S hnmem ih, { simp }, { simp [ih, hnmem], } end lemma to_fun_commutes (i : I) (hnmem : i ∉ S) : - commute (ϕ i (g i)) (to_fun f S) := + commute (ϕ i (g i)) (to_fun ϕ hcomm f S) := begin induction S using finset.induction_on with j S hnmem' ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], have hij : i ≠ j, by {simp at hnmem, tauto}, have hiS : i ∉ S, by {simp at hnmem, tauto}, - calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ f S : M)) - = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ f S : by rw ← mul_assoc - ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ f S : by { congr' 1, apply fact.elim hcomm _ _ hij } - ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ f S) : by rw mul_assoc - ... = ϕ j (f j) * (to_fun ϕ f S * ϕ i (g i)) : by { congr' 1, apply ih hiS } - ... = (ϕ j (f j) * to_fun ϕ f S) * ϕ i (g i) : by rw ← mul_assoc } + calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ hcomm f S : M)) + = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ hcomm f S : by rw ← mul_assoc + ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ hcomm f S : by { congr' 1, apply hcomm _ _ hij } + ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ hcomm f S) : by rw mul_assoc + ... = ϕ j (f j) * (to_fun ϕ hcomm f S * ϕ i (g i)) : by { congr' 1, apply ih hiS } + ... = (ϕ j (f j) * to_fun ϕ hcomm f S) * ϕ i (g i) : by rw ← mul_assoc } end @[simp] -lemma to_fun_mul (S : finset I) : to_fun (f * g) S = to_fun f S * to_fun g S := +lemma to_fun_mul (S : finset I) : + to_fun ϕ hcomm (f * g) S = to_fun ϕ hcomm f S * to_fun ϕ hcomm g S := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, @@ -139,7 +141,7 @@ begin end lemma to_fun_in_sup_mrange (S : finset I) : - to_fun f S ∈ ⨆ i ∈ S, (ϕ i).mrange := + to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, @@ -150,14 +152,18 @@ begin exact λ i, finset.mem_insert_of_mem } } end +variable (hcomm) + /-- The canonical homomorphism from a pi group, restricted to a subset -/ def hom : (Π (i : I), N i) →* M := -{ to_fun := λ f, to_fun f S, - map_one' := to_fun_one _, - map_mul' := λ f g, to_fun_mul f g S, } +{ to_fun := λ f, to_fun ϕ hcomm f S, + map_one' := to_fun_one ϕ _, + map_mul' := λ f g, to_fun_mul ϕ f g S, } + +variable {hcomm} lemma to_fun_single (i : I) (y : N i) (S : finset I) : - to_fun (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := + to_fun ϕ hcomm (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, { simp, }, @@ -170,9 +176,9 @@ end @[simp] lemma hom_single (i : I) (y : N i): - hom S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ + hom ϕ hcomm S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ _ -lemma mrange : (hom S).mrange = ⨆ i ∈ S, (ϕ i).mrange := +lemma mrange : (hom ϕ hcomm S).mrange = ⨆ i ∈ S, (ϕ i).mrange := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, @@ -189,18 +195,19 @@ namespace pi_hom include hfin -/-- The product of `ϕ i (f i)` for all `i : I` -/ -def to_fun : M := pi_hom_restr.to_fun ϕ f finset.univ +variable (hcomm) /-- The canonical homomorphism from a pi group -/ -def hom : (Π (i : I), N i) →* M := pi_hom_restr.hom ϕ finset.univ +def hom : (Π (i : I), N i) →* M := pi_hom_restr.hom ϕ hcomm finset.univ + +variable {hcomm} @[simp] -lemma hom_single (i : I) (y : N i): hom (monoid_hom.single _ i y) = ϕ i y := -by { show pi_hom_restr.hom ϕ finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +lemma hom_single (i : I) (y : N i): hom ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := +by { show pi_hom_restr.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } -lemma mrange : hom.mrange = ⨆ i : I, (ϕ i).mrange := -by { show (pi_hom_restr.hom ϕ finset.univ).mrange = _, simp [pi_hom_restr.mrange] } +lemma mrange : (hom ϕ hcomm).mrange = ⨆ i : I, (ϕ i).mrange := +by { show (pi_hom_restr.hom ϕ hcomm finset.univ).mrange = _, simp [pi_hom_restr.mrange] } end pi_hom @@ -212,7 +219,7 @@ variables {G : Type*} [group G] variables {I : Type*} [hfin : fintype I] variables {H : I → Type*} [∀ i, group (H i)] variables (ϕ : Π (i : I), H i →* G) -variables [hcomm : fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y))] +variables {hcomm : ∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)} include hcomm -- We use `f` and `g` to denote elements of `Π (i : I), H i` @@ -220,7 +227,7 @@ variables (f g : Π (i : I), H i) -- The subgroup version of to_fun_in_sup_mrange lemma pi_hom_restr.to_fun_in_sup_range (S : finset I) : - pi_hom_restr.to_fun ϕ f S ∈ ⨆ i ∈ S, (ϕ i).range := + pi_hom_restr.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, @@ -233,7 +240,7 @@ end -- The subgroup version of pi_hom_restr.mrange lemma pi_hom_restr.range (S : finset I) : - (pi_hom_restr.hom ϕ S).range = ⨆ i ∈ S, (ϕ i).range := + (pi_hom_restr.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, @@ -249,26 +256,26 @@ include hfin namespace pi_hom -- The subgroup version of pi_hom.mrange -lemma range : (pi_hom.hom ϕ).range = ⨆ i : I, (ϕ i).range := -by { show (pi_hom_restr.hom ϕ finset.univ).range = _, simp [pi_hom_restr.range] } +lemma range : (pi_hom.hom ϕ hcomm).range = ⨆ i : I, (ϕ i).range := +by { show (pi_hom_restr.hom ϕ hcomm finset.univ).range = _, simp [pi_hom_restr.range] } lemma injective_of_independent (hind : complete_lattice.independent (λ i, (ϕ i).range)) (hinj : ∀ i, function.injective (ϕ i)) : - function.injective (pi_hom.hom ϕ):= + function.injective (pi_hom.hom ϕ hcomm):= begin apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, intro f, - show pi_hom_restr.to_fun ϕ f finset.univ = 1 → f = 1, - suffices : pi_hom_restr.to_fun ϕ f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), + show pi_hom_restr.to_fun ϕ hcomm f finset.univ = 1 → f = 1, + suffices : pi_hom_restr.to_fun ϕ hcomm f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), induction (finset.univ : finset I) using finset.induction_on with i S hnmem ih, { simp }, { intro heq1, simp only [ pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, have hnmem' : i ∉ (S : set I), by simpa, - have heq1' : ϕ i (f i) = 1 ∧ pi_hom_restr.to_fun ϕ f S = 1, + have heq1' : ϕ i (f i) = 1 ∧ pi_hom_restr.to_fun ϕ hcomm f S = 1, { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, { apply pi_hom_restr.to_fun_in_sup_range, }, }, @@ -281,6 +288,7 @@ begin { exact (ih _ h), } } end +variable (hcomm) lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := @@ -288,7 +296,7 @@ begin rintros i f ⟨hxi, hxp⟩, dsimp at hxi hxp, rw [supr_subtype', ← @pi_hom.range _ _ _ _ _ _ _ _] at hxp, rotate, apply_instance, - { split, intros _ _ hj, apply hcomm.1, exact hj ∘ subtype.ext }, + { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, cases hxp with g hgf, cases hxi with g' hg'f, have hxi : order_of f ∣ fintype.card (H i), { rw ← hg'f, exact (map_order _ _).trans order_of_dvd_card_univ }, @@ -308,8 +316,8 @@ namespace subgroup_pi_hom section subgroup_pi_hom -- We have an family of subgroups -parameters {G : Type*} [group G] -parameters {I : Type*} [hfin : fintype I] [hdec : decidable_eq I] {H : I → subgroup G} +variables {G : Type*} [group G] +variables {I : Type*} [hfin : fintype I] [hdec : decidable_eq I] {H : I → subgroup G} -- Elements of `Π (i : I), H i` are called `f` and `g` here variables (f g : Π (i : I), H i) @@ -317,28 +325,30 @@ variables (f g : Π (i : I), H i) section commuting_subgroups -- We assume that the elements of different subgroups commute -parameters (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) +variables (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) include hcomm -lemma hcomm_subtype : - fact (∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y)) := -fact.mk (by { rintros i j hne ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy }) +lemma hcomm_subtype (i j : I) (hne : i ≠ j) : + ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y) := +by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } include hfin /-- The canonical homomorphism from a pi group of subgroups -/ def hom : (Π (i : I), H i) →* G := - let _ := hcomm_subtype in by exactI pi_hom.hom (λ i, (H i).subtype) + pi_hom.hom (λ i, (H i).subtype) (hcomm_subtype hcomm) + +variable {hcomm} @[simp] -lemma hom_single (i : I) (y : H i): hom (monoid_hom.single _ i y) = y := +lemma hom_single (i : I) (y : H i): hom hcomm (monoid_hom.single _ i y) = y := by apply pi_hom.hom_single -lemma range : hom.range = ⨆ i : I, H i := +lemma range : (hom hcomm).range = ⨆ i : I, H i := by simp [hom, pi_hom.range] lemma injective_of_independent (hind : complete_lattice.independent H) : - function.injective hom := + function.injective (hom hcomm) := begin apply pi_hom.injective_of_independent, { simpa using hind }, @@ -349,8 +359,8 @@ lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent H := begin - haveI := hcomm_subtype hcomm, - simpa using pi_hom.independent_range_of_coprime_order (λ i, (H i).subtype) hcoprime, + simpa using + pi_hom.independent_range_of_coprime_order (λ i, (H i).subtype) (hcomm_subtype hcomm) hcoprime, end end commuting_subgroups From 8da6ad5d5e3aaa84c0bcab0e7001678fd3b8c264 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 14 Feb 2022 11:28:13 +0100 Subject: [PATCH 029/116] Avoid simp followed by tauto --- src/group_theory/pi_hom.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 03e5d799452e7..652aac69042c7 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -116,8 +116,8 @@ begin induction S using finset.induction_on with j S hnmem' ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], - have hij : i ≠ j, by {simp at hnmem, tauto}, - have hiS : i ∉ S, by {simp at hnmem, tauto}, + have hij : i ≠ j, by {rintro rfl, apply hnmem, exact finset.mem_insert_self i S}, + have hiS : i ∉ S, by {rintro h, apply hnmem, exact finset.mem_insert_of_mem h}, calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ hcomm f S : M)) = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ hcomm f S : by rw ← mul_assoc ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ hcomm f S : by { congr' 1, apply hcomm _ _ hij } From dfcc95612aa919e4f832c60b84f5e9879a1d5f39 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 14 Feb 2022 11:32:26 +0100 Subject: [PATCH 030/116] Update src/group_theory/pi_hom.lean --- src/group_theory/pi_hom.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/pi_hom.lean index 652aac69042c7..f3f452ea024da 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/pi_hom.lean @@ -10,7 +10,7 @@ import data.fintype.card /-! # Canonical homomorphism from a pi monoid -This file defines txinput set-prop 11 "libinput Scroll Method Enabled" 0, 0, 1 #he construction of the canoncial homomorphism from a family of monoids. +This file defines the construction of the canoncial homomorphism from a family of monoids. Given a family of morphisms `ϕ i : N i →* M` for each `i : I` where elements in the image of different morphism commute, we obtain a canoncial morphism From dfccab318bd29f817608cbc634155965036d87b8 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 15 Feb 2022 13:26:18 +0100 Subject: [PATCH 031/116] Say monoid_hom.noncomm_pi_coprod --- .../{pi_hom.lean => nocomm_pi_coprod.lean} | 139 ++++++++++-------- 1 file changed, 74 insertions(+), 65 deletions(-) rename src/group_theory/{pi_hom.lean => nocomm_pi_coprod.lean} (68%) diff --git a/src/group_theory/pi_hom.lean b/src/group_theory/nocomm_pi_coprod.lean similarity index 68% rename from src/group_theory/pi_hom.lean rename to src/group_theory/nocomm_pi_coprod.lean index f3f452ea024da..4a01a0907a3bb 100644 --- a/src/group_theory/pi_hom.lean +++ b/src/group_theory/nocomm_pi_coprod.lean @@ -8,34 +8,37 @@ import data.finset.noncomm_prod import data.fintype.card /-! -# Canonical homomorphism from a pi monoid +# Canonical homomorphism from a finite family of monoids This file defines the construction of the canoncial homomorphism from a family of monoids. Given a family of morphisms `ϕ i : N i →* M` for each `i : I` where elements in the image of different morphism commute, we obtain a canoncial morphism -`pi_hom.hom : (Π (i : I), N i) →* M` that coincides with `ϕ` +`monoid_hom.noncomm_pi_coprod : (Π (i : I), N i) →* M` that coincides with `ϕ` ## Main definitions -* `pi_hom.hom : (Π (i : I), N i) →* M` is the main homomorphism -* `pi_hom_restr.hom (S : finset S) : (Π (i : I), N i) →* M` is the homomorphism restricted to the - set `S`, and mostly of internal interest in this file, to allow inductive proofs. -* `subgroup_pi_hom.hom : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` and - the subgroup embedding. +* `monoid_hom.noncomm_pi_coprod : (Π (i : I), N i) →* M` is the main homomorphism +* `noncomm_pi_coprod_on.hom (S : finset S) : (Π (i : I), N i) →* M` is the homomorphism + restricted to the set `S`, and mostly of internal interest in this file, to allow inductive + proofs (and thus in its own namespace). +* `subgroup.noncomm_pi_coprod : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` + and the subgroup embedding. ## Main theorems -* `pi_hom.hom` conicides with `ϕ i` when restricted to `N i` -* `pi_hom.mrange`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).mrange` -* `pi_hom.range`: The range of `pi_hom.hom` is `⨆ (i : I), (ϕ i).range` -* The range of `subgroup_pi_hom.hom` is `⨆ (i : I), H i` -* `pi_hom.injective_of_independent`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` - are injective and the ranges of the `ϕ` are independent. -* `pi_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges are - independent. -* `independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, they are - independent. +* `monoid_hom.noncomm_pi_coprod` conicides with `ϕ i` when restricted to `N i` +* `monoid_hom.noncomm_pi_coprod_mrange`: The range of `monoid_hom.noncomm_pi_coprod` is + `⨆ (i : I), (ϕ i).mrange` +* `monoid_hom.noncomm_pi_coprod_range`: The range of `monoid_hom.noncomm_pi_coprod` is + `⨆ (i : I), (ϕ i).range` +* `subgroup.noncomm_pi_coprod_range`: The range of `subgroup.noncomm_pi_coprod` is `⨆ (i : I), H i`. +* `monoid_hom.injective_noncomm_pi_coprod_of_independent`: in the case of groups, `pi_hom.hom` is + injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. +* `monid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges + of the `ϕ` are independent. +* `subgroup.independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, + they are independent. -/ @@ -84,7 +87,7 @@ include hcomm -- We use `f` and `g` to denote elements of `Π (i : I), N i` variables (f g : Π (i : I), N i) -namespace pi_hom_restr +namespace noncomm_pi_coprod_on -- In this section, we restrict the hom to a set S variables (S : finset I) @@ -154,7 +157,8 @@ end variable (hcomm) -/-- The canonical homomorphism from a pi group, restricted to a subset -/ +/-- The canonical homomorphism from a family of monoids, restricted to a subset of the index space. +-/ def hom : (Π (i : I), N i) →* M := { to_fun := λ f, to_fun ϕ hcomm f S, map_one' := to_fun_one ϕ _, @@ -189,27 +193,31 @@ begin simp [hmem], } end -end pi_hom_restr +end noncomm_pi_coprod_on -namespace pi_hom +namespace monoid_hom include hfin variable (hcomm) -/-- The canonical homomorphism from a pi group -/ -def hom : (Π (i : I), N i) →* M := pi_hom_restr.hom ϕ hcomm finset.univ +/-- The canonical homomorphism from a family of monoids. -/ +def noncomm_pi_coprod : (Π (i : I), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ variable {hcomm} @[simp] -lemma hom_single (i : I) (y : N i): hom ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := -by { show pi_hom_restr.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +lemma noncomm_pi_coprod_single (i : I) (y : N i): + noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := +by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } -lemma mrange : (hom ϕ hcomm).mrange = ⨆ i : I, (ϕ i).mrange := -by { show (pi_hom_restr.hom ϕ hcomm finset.univ).mrange = _, simp [pi_hom_restr.mrange] } +lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : I, (ϕ i).mrange := +begin + show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).mrange = _, + simp [noncomm_pi_coprod_on.mrange], +end -end pi_hom +end monoid_hom end family_of_monoids @@ -225,26 +233,26 @@ include hcomm -- We use `f` and `g` to denote elements of `Π (i : I), H i` variables (f g : Π (i : I), H i) --- The subgroup version of to_fun_in_sup_mrange -lemma pi_hom_restr.to_fun_in_sup_range (S : finset I) : - pi_hom_restr.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := +-- The subgroup version of `noncomm_pi_coprod_on.to_fun_in_sup_mrange` +lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset I) : + noncomm_pi_coprod_on.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin induction S using finset.induction_on with i S hnmem ih, { simp, }, - { simp only [pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem], + { simp only [noncomm_pi_coprod_on.to_fun_insert_of_not_mem _ _ _ _ hnmem], refine (subgroup.mul_mem _ _ _), { apply subgroup.mem_Sup_of_mem, { use i }, { simp, }, }, { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).range) _ ih, exact λ i, finset.mem_insert_of_mem } } end --- The subgroup version of pi_hom_restr.mrange -lemma pi_hom_restr.range (S : finset I) : - (pi_hom_restr.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := +-- The subgroup version of `noncomm_pi_coprod_on.mrange` +lemma noncomm_pi_coprod_on.range (S : finset I) : + (noncomm_pi_coprod_on.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := begin apply le_antisymm, { rintro x ⟨f, rfl⟩, - exact (pi_hom_restr.to_fun_in_sup_range ϕ f S), }, + exact (noncomm_pi_coprod_on.to_fun_in_sup_range ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, use (monoid_hom.single _ i y), @@ -253,32 +261,36 @@ end include hfin -namespace pi_hom +namespace monoid_hom --- The subgroup version of pi_hom.mrange -lemma range : (pi_hom.hom ϕ hcomm).range = ⨆ i : I, (ϕ i).range := -by { show (pi_hom_restr.hom ϕ hcomm finset.univ).range = _, simp [pi_hom_restr.range] } +-- The subgroup version of `noncomm_pi_coprod_on_mrange` +lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : I, (ϕ i).range := +begin + show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).range = _, + simp [noncomm_pi_coprod_on.range] +end -lemma injective_of_independent +lemma injective_noncomm_pi_coprod_of_independent (hind : complete_lattice.independent (λ i, (ϕ i).range)) (hinj : ∀ i, function.injective (ϕ i)) : - function.injective (pi_hom.hom ϕ hcomm):= + function.injective (noncomm_pi_coprod ϕ hcomm):= begin apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, intro f, - show pi_hom_restr.to_fun ϕ hcomm f finset.univ = 1 → f = 1, - suffices : pi_hom_restr.to_fun ϕ hcomm f finset.univ = 1 → (∀ (i : I), i ∈ finset.univ → f i = 1), + show noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → f = 1, + suffices : noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → + (∀ (i : I), i ∈ finset.univ → f i = 1), by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), induction (finset.univ : finset I) using finset.induction_on with i S hnmem ih, { simp }, { intro heq1, - simp only [ pi_hom_restr.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, + simp only [ noncomm_pi_coprod_on.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, have hnmem' : i ∉ (S : set I), by simpa, - have heq1' : ϕ i (f i) = 1 ∧ pi_hom_restr.to_fun ϕ hcomm f S = 1, + have heq1' : ϕ i (f i) = 1 ∧ noncomm_pi_coprod_on.to_fun ϕ hcomm f S = 1, { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, - { apply pi_hom_restr.to_fun_in_sup_range, }, }, + { apply noncomm_pi_coprod_on.to_fun_in_sup_range, }, }, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, @@ -294,7 +306,7 @@ lemma independent_range_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent (λ i, (ϕ i).range) := begin rintros i f ⟨hxi, hxp⟩, dsimp at hxi hxp, - rw [supr_subtype', ← @pi_hom.range _ _ _ _ _ _ _ _] at hxp, + rw [supr_subtype', ← @noncomm_pi_coprod_range _ _ _ _ _ _ _ _] at hxp, rotate, apply_instance, { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, cases hxp with g hgf, cases hxi with g' hg'f, @@ -307,13 +319,11 @@ begin apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, end -end pi_hom +end monoid_hom end family_of_groups -namespace subgroup_pi_hom - -section subgroup_pi_hom +namespace subgroup -- We have an family of subgroups variables {G : Type*} [group G] @@ -335,22 +345,23 @@ by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } include hfin /-- The canonical homomorphism from a pi group of subgroups -/ -def hom : (Π (i : I), H i) →* G := - pi_hom.hom (λ i, (H i).subtype) (hcomm_subtype hcomm) +def noncomm_pi_coprod : (Π (i : I), H i) →* G := + monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (hcomm_subtype hcomm) variable {hcomm} @[simp] -lemma hom_single (i : I) (y : H i): hom hcomm (monoid_hom.single _ i y) = y := -by apply pi_hom.hom_single +lemma noncomm_pi_coprod_single (i : I) (y : H i) : + noncomm_pi_coprod hcomm (monoid_hom.single _ i y) = y := +by apply monoid_hom.noncomm_pi_coprod_single -lemma range : (hom hcomm).range = ⨆ i : I, H i := -by simp [hom, pi_hom.range] +lemma range : (noncomm_pi_coprod hcomm).range = ⨆ i : I, H i := +by simp [noncomm_pi_coprod, monoid_hom.noncomm_pi_coprod_range] lemma injective_of_independent (hind : complete_lattice.independent H) : - function.injective (hom hcomm) := + function.injective (noncomm_pi_coprod hcomm) := begin - apply pi_hom.injective_of_independent, + apply monoid_hom.injective_noncomm_pi_coprod_of_independent, { simpa using hind }, { intro i, exact subtype.coe_injective } end @@ -359,12 +370,10 @@ lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent H := begin - simpa using - pi_hom.independent_range_of_coprime_order (λ i, (H i).subtype) (hcomm_subtype hcomm) hcoprime, + simpa using monoid_hom.independent_range_of_coprime_order + (λ i, (H i).subtype) (hcomm_subtype hcomm) hcoprime, end end commuting_subgroups -end subgroup_pi_hom - -end subgroup_pi_hom +end subgroup From 81ecc8c25d8afa604446c1b88634f1b81fcda25e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 15 Feb 2022 14:25:49 +0100 Subject: [PATCH 032/116] Add to_additive --- src/group_theory/nocomm_pi_coprod.lean | 38 ++++++++++++++++++++------ 1 file changed, 29 insertions(+), 9 deletions(-) diff --git a/src/group_theory/nocomm_pi_coprod.lean b/src/group_theory/nocomm_pi_coprod.lean index 4a01a0907a3bb..2053fc1b263c0 100644 --- a/src/group_theory/nocomm_pi_coprod.lean +++ b/src/group_theory/nocomm_pi_coprod.lean @@ -35,7 +35,7 @@ image of different morphism commute, we obtain a canoncial morphism * `subgroup.noncomm_pi_coprod_range`: The range of `subgroup.noncomm_pi_coprod` is `⨆ (i : I), H i`. * `monoid_hom.injective_noncomm_pi_coprod_of_independent`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. -* `monid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges +* `monoid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges of the `ϕ` are independent. * `subgroup.independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, they are independent. @@ -52,6 +52,7 @@ lemma coprime_prod_left finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) -- I think it's worth keeping it and moving to appropriate file +@[to_additive] lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := begin @@ -93,26 +94,28 @@ namespace noncomm_pi_coprod_on variables (S : finset I) /-- The underlying function of `pi_hom_restr.hom` -/ +@[to_additive add_to_fun] def to_fun (S : finset I) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } } variable {hcomm} -@[simp] +@[simp, to_additive add_to_fun_empty] lemma to_fun_empty : to_fun ϕ hcomm f ∅ = 1 := finset.noncomm_prod_empty _ _ -@[simp] +@[simp, to_additive add_to_fun_insert_of_not_mem] lemma to_fun_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : to_fun ϕ hcomm f (insert i S) = ϕ i (f i) * to_fun ϕ hcomm f S := finset.noncomm_prod_insert_of_not_mem _ _ _ _ h -@[simp] +@[simp, to_additive add_to_fun_zero] lemma to_fun_one : to_fun ϕ hcomm 1 S = 1 := begin induction S using finset.cons_induction_on with i S hnmem ih, { simp }, { simp [ih, hnmem], } end +@[to_additive add_to_fun_commutes] lemma to_fun_commutes (i : I) (hnmem : i ∉ S) : commute (ϕ i (g i)) (to_fun ϕ hcomm f S) := begin @@ -129,7 +132,7 @@ begin ... = (ϕ j (f j) * to_fun ϕ hcomm f S) * ϕ i (g i) : by rw ← mul_assoc } end -@[simp] +@[simp, to_additive add_to_fun_add] lemma to_fun_mul (S : finset I) : to_fun ϕ hcomm (f * g) S = to_fun ϕ hcomm f S * to_fun ϕ hcomm g S := begin @@ -143,6 +146,7 @@ begin exact (to_fun_commutes _ _ _ S i hnmem), } end +@[to_additive add_to_fun_in_sup_mrange] lemma to_fun_in_sup_mrange (S : finset I) : to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin @@ -159,6 +163,7 @@ variable (hcomm) /-- The canonical homomorphism from a family of monoids, restricted to a subset of the index space. -/ +@[to_additive add_hom] def hom : (Π (i : I), N i) →* M := { to_fun := λ f, to_fun ϕ hcomm f S, map_one' := to_fun_one ϕ _, @@ -166,6 +171,7 @@ def hom : (Π (i : I), N i) →* M := variable {hcomm} +@[to_additive add_to_fun_single] lemma to_fun_single (i : I) (y : N i) (S : finset I) : to_fun ϕ hcomm (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := begin @@ -178,10 +184,11 @@ begin simpa [h] using ih, } } end -@[simp] +@[simp, to_additive add_hom_single] lemma hom_single (i : I) (y : N i): hom ϕ hcomm S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ _ +@[to_additive add_mrange] lemma mrange : (hom ϕ hcomm S).mrange = ⨆ i ∈ S, (ϕ i).mrange := begin apply le_antisymm, @@ -202,15 +209,17 @@ include hfin variable (hcomm) /-- The canonical homomorphism from a family of monoids. -/ +@[to_additive] def noncomm_pi_coprod : (Π (i : I), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ variable {hcomm} -@[simp] +@[simp, to_additive] lemma noncomm_pi_coprod_single (i : I) (y : N i): noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +@[to_additive] lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : I, (ϕ i).mrange := begin show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).mrange = _, @@ -234,6 +243,7 @@ include hcomm variables (f g : Π (i : I), H i) -- The subgroup version of `noncomm_pi_coprod_on.to_fun_in_sup_mrange` +@[to_additive noncomm_pi_coprod_on.add_to_fun_in_sup_range] lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset I) : noncomm_pi_coprod_on.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin @@ -247,6 +257,7 @@ begin end -- The subgroup version of `noncomm_pi_coprod_on.mrange` +@[to_additive noncomm_pi_coprod_on.add_range] lemma noncomm_pi_coprod_on.range (S : finset I) : (noncomm_pi_coprod_on.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := begin @@ -264,12 +275,14 @@ include hfin namespace monoid_hom -- The subgroup version of `noncomm_pi_coprod_on_mrange` +@[to_additive] lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : I, (ϕ i).range := begin show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).range = _, simp [noncomm_pi_coprod_on.range] end +@[to_additive] lemma injective_noncomm_pi_coprod_of_independent (hind : complete_lattice.independent (λ i, (ϕ i).range)) (hinj : ∀ i, function.injective (ϕ i)) : @@ -301,6 +314,8 @@ begin end variable (hcomm) + +@[to_additive] lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := @@ -338,6 +353,7 @@ section commuting_subgroups variables (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) include hcomm +@[to_additive] lemma hcomm_subtype (i j : I) (hne : i ≠ j) : ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y) := by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } @@ -345,19 +361,22 @@ by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } include hfin /-- The canonical homomorphism from a pi group of subgroups -/ +@[to_additive] def noncomm_pi_coprod : (Π (i : I), H i) →* G := monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (hcomm_subtype hcomm) variable {hcomm} -@[simp] +@[simp, to_additive] lemma noncomm_pi_coprod_single (i : I) (y : H i) : noncomm_pi_coprod hcomm (monoid_hom.single _ i y) = y := by apply monoid_hom.noncomm_pi_coprod_single +@[to_additive] lemma range : (noncomm_pi_coprod hcomm).range = ⨆ i : I, H i := by simp [noncomm_pi_coprod, monoid_hom.noncomm_pi_coprod_range] +@[to_additive] lemma injective_of_independent (hind : complete_lattice.independent H) : function.injective (noncomm_pi_coprod hcomm) := begin @@ -366,7 +385,8 @@ begin { intro i, exact subtype.coe_injective } end -lemma _root_.independent_of_coprime_order [∀ i, fintype (H i)] +@[to_additive] +lemma independent_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent H := begin From e5c3a9758a5c82a67127783e167da715f94929e7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 15 Feb 2022 14:26:34 +0100 Subject: [PATCH 033/116] Update src/group_theory/order_of_element.lean Co-authored-by: Eric Wieser --- src/group_theory/order_of_element.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 7cc1fe49360bc..0dd0b509b7aef 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -136,8 +136,8 @@ is_periodic_pt.minimal_period_dvd ((is_periodic_pt_mul_iff_pow_eq_one _).mpr h) lemma order_of_dvd_iff_pow_eq_one {n : ℕ} : order_of x ∣ n ↔ x ^ n = 1 := ⟨λ h, by rw [pow_eq_mod_order_of, nat.mod_eq_zero_of_dvd h, pow_zero], order_of_dvd_of_pow_eq_one⟩ -@[to_additive map_add_order] -lemma map_order {H : Type*} [monoid H] (ψ : G →* H) (x : G) : +@[to_additive add_order_of_map_dvd] +lemma order_of_map_dvd {H : Type*} [monoid H] (ψ : G →* H) (x : G) : order_of (ψ x) ∣ order_of x := by { apply order_of_dvd_of_pow_eq_one, rw [←map_pow, pow_order_of_eq_one], apply map_one } From f0d5a5c42ff8267eae2ae69190a2930cdaa56caf Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 15 Feb 2022 14:27:08 +0100 Subject: [PATCH 034/116] Use order_of_map_dvd --- src/group_theory/nocomm_pi_coprod.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/nocomm_pi_coprod.lean b/src/group_theory/nocomm_pi_coprod.lean index 2053fc1b263c0..ef9e833a31876 100644 --- a/src/group_theory/nocomm_pi_coprod.lean +++ b/src/group_theory/nocomm_pi_coprod.lean @@ -326,9 +326,9 @@ begin { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, cases hxp with g hgf, cases hxi with g' hg'f, have hxi : order_of f ∣ fintype.card (H i), - { rw ← hg'f, exact (map_order _ _).trans order_of_dvd_card_univ }, + { rw ← hg'f, exact (order_of_map_dvd _ _).trans order_of_dvd_card_univ }, have hxp : order_of f ∣ ∏ j : {j // j ≠ i}, fintype.card (H j), - { rw [← hgf, ← fintype.card_pi], exact (map_order _ _).trans order_of_dvd_card_univ }, + { rw [← hgf, ← fintype.card_pi], exact (order_of_map_dvd _ _).trans order_of_dvd_card_univ }, change f = 1, rw [← pow_one f, ← order_of_dvd_iff_pow_eq_one], convert ← nat.dvd_gcd hxp hxi, rw ← nat.coprime_iff_gcd_eq_one, apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, From 62da911cce430b6bb930cbcec3e26c0e911d83ed Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 16 Feb 2022 14:46:13 +0100 Subject: [PATCH 035/116] Additive docstrings --- src/group_theory/nocomm_pi_coprod.lean | 15 +++++++++------ 1 file changed, 9 insertions(+), 6 deletions(-) diff --git a/src/group_theory/nocomm_pi_coprod.lean b/src/group_theory/nocomm_pi_coprod.lean index ef9e833a31876..86bf220ef297c 100644 --- a/src/group_theory/nocomm_pi_coprod.lean +++ b/src/group_theory/nocomm_pi_coprod.lean @@ -93,8 +93,8 @@ namespace noncomm_pi_coprod_on -- In this section, we restrict the hom to a set S variables (S : finset I) -/-- The underlying function of `pi_hom_restr.hom` -/ -@[to_additive add_to_fun] +/-- The underlying function of `noncomm_pi_coprod_on.hom` -/ +@[to_additive add_to_fun "The underlying function of `noncomm_pi_coprod_on.add_hom` "] def to_fun (S : finset I) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } } @@ -163,7 +163,8 @@ variable (hcomm) /-- The canonical homomorphism from a family of monoids, restricted to a subset of the index space. -/ -@[to_additive add_hom] +@[to_additive add_hom "The canonical homomorphism from a family of additive monoids, restricted to a +subset of the index space"] def hom : (Π (i : I), N i) →* M := { to_fun := λ f, to_fun ϕ hcomm f S, map_one' := to_fun_one ϕ _, @@ -209,7 +210,7 @@ include hfin variable (hcomm) /-- The canonical homomorphism from a family of monoids. -/ -@[to_additive] +@[to_additive "The canonical homomorphism from a family of additive monoids."] def noncomm_pi_coprod : (Π (i : I), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ variable {hcomm} @@ -360,8 +361,10 @@ by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } include hfin -/-- The canonical homomorphism from a pi group of subgroups -/ -@[to_additive] +/-- The canonical homomorphism from a family of subgroups where elements from different subgroups +commute -/ +@[to_additive "The canonical homomorphism from a family of additive subgroups where elements from +different subgroups commute"] def noncomm_pi_coprod : (Π (i : I), H i) →* G := monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (hcomm_subtype hcomm) From ac06e68c62c38c1b98bab4495cbfbc872c4b5499 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 16 Feb 2022 14:48:05 +0100 Subject: [PATCH 036/116] =?UTF-8?q?s/I/=CE=B9/g?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/group_theory/nocomm_pi_coprod.lean | 110 ++++++++++++------------- 1 file changed, 55 insertions(+), 55 deletions(-) diff --git a/src/group_theory/nocomm_pi_coprod.lean b/src/group_theory/nocomm_pi_coprod.lean index 86bf220ef297c..563a439d41c5d 100644 --- a/src/group_theory/nocomm_pi_coprod.lean +++ b/src/group_theory/nocomm_pi_coprod.lean @@ -1,6 +1,6 @@ /- Copyright (c) 2022 Jocchim Breitner. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. +Released under Apache 2.0 license as described in the file LιCENSE. Authors: Joachim Breitner -/ import group_theory.order_of_element @@ -12,32 +12,32 @@ import data.fintype.card This file defines the construction of the canoncial homomorphism from a family of monoids. -Given a family of morphisms `ϕ i : N i →* M` for each `i : I` where elements in the +Given a family of morphisms `ϕ i : N i →* M` for each `i : ι` where elements in the image of different morphism commute, we obtain a canoncial morphism -`monoid_hom.noncomm_pi_coprod : (Π (i : I), N i) →* M` that coincides with `ϕ` +`monoid_hom.noncomm_pi_coprod : (Π i, N i) →* M` that coincides with `ϕ` ## Main definitions -* `monoid_hom.noncomm_pi_coprod : (Π (i : I), N i) →* M` is the main homomorphism -* `noncomm_pi_coprod_on.hom (S : finset S) : (Π (i : I), N i) →* M` is the homomorphism +* `monoid_hom.noncomm_pi_coprod : (Π i, N i) →* M` is the main homomorphism +* `noncomm_pi_coprod_on.hom (S : finset S) : (Π (i : ι), N i) →* M` is the homomorphism restricted to the set `S`, and mostly of internal interest in this file, to allow inductive proofs (and thus in its own namespace). -* `subgroup.noncomm_pi_coprod : (Π (i : I), H i) →* G` is the specialization to `H i : subgroup G` +* `subgroup.noncomm_pi_coprod : (Π i, H i) →* G` is the specialization to `H i : subgroup G` and the subgroup embedding. ## Main theorems * `monoid_hom.noncomm_pi_coprod` conicides with `ϕ i` when restricted to `N i` * `monoid_hom.noncomm_pi_coprod_mrange`: The range of `monoid_hom.noncomm_pi_coprod` is - `⨆ (i : I), (ϕ i).mrange` + `⨆ (i : ι), (ϕ i).mrange` * `monoid_hom.noncomm_pi_coprod_range`: The range of `monoid_hom.noncomm_pi_coprod` is - `⨆ (i : I), (ϕ i).range` -* `subgroup.noncomm_pi_coprod_range`: The range of `subgroup.noncomm_pi_coprod` is `⨆ (i : I), H i`. + `⨆ (i : ι), (ϕ i).range` +* `subgroup.noncomm_pi_coprod_range`: The range of `subgroup.noncomm_pi_coprod` is `⨆ (i : ι), H i`. * `monoid_hom.injective_noncomm_pi_coprod_of_independent`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. -* `monoid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges +* `monoid_hom.independent_range_of_coprime_order`: ιf the `N i` have coprime orders, then the ranges of the `ϕ` are independent. -* `subgroup.independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, +* `subgroup.independent_of_coprime_order`: ιf commuting, normal subgroups `H i` have coprime orders, they are independent. -/ @@ -46,12 +46,12 @@ open_locale big_operators open_locale classical lemma coprime_prod_left - {I : Type*} - {x : ℕ} {s : I → ℕ} {t : finset I} : - (∀ (i : I), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : I) in t, s i) x := + {ι : Type*} + {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : ι) in t, s i) x := finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) --- I think it's worth keeping it and moving to appropriate file +-- ι think it's worth keeping it and moving to appropriate file @[to_additive] lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := @@ -75,27 +75,27 @@ variables {M : Type*} [monoid M] -- We have a family of monoids -- The fintype assumption is not always used, but declared here, to keep things in order -variables {I : Type*} [hfin : fintype I] -variables {N : I → Type*} [∀ i, monoid (N i)] +variables {ι : Type*} [hfin : fintype ι] +variables {N : ι → Type*} [∀ i, monoid (N i)] -- And morphisms ϕ into G -variables (ϕ : Π (i : I), N i →* M) +variables (ϕ : Π (i : ι), N i →* M) -- We assume that the elements of different morphism commute -variables (hcomm : ∀ (i j : I), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y)) +variables (hcomm : ∀ (i j : ι), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y)) include hcomm --- We use `f` and `g` to denote elements of `Π (i : I), N i` -variables (f g : Π (i : I), N i) +-- We use `f` and `g` to denote elements of `Π (i : ι), N i` +variables (f g : Π (i : ι), N i) namespace noncomm_pi_coprod_on --- In this section, we restrict the hom to a set S -variables (S : finset I) +-- ιn this section, we restrict the hom to a set S +variables (S : finset ι) /-- The underlying function of `noncomm_pi_coprod_on.hom` -/ @[to_additive add_to_fun "The underlying function of `noncomm_pi_coprod_on.add_hom` "] -def to_fun (S : finset I) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ +def to_fun (S : finset ι) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } } variable {hcomm} @@ -104,7 +104,7 @@ variable {hcomm} lemma to_fun_empty : to_fun ϕ hcomm f ∅ = 1 := finset.noncomm_prod_empty _ _ @[simp, to_additive add_to_fun_insert_of_not_mem] -lemma to_fun_insert_of_not_mem (S : finset I) (i : I) (h : i ∉ S) : +lemma to_fun_insert_of_not_mem (S : finset ι) (i : ι) (h : i ∉ S) : to_fun ϕ hcomm f (insert i S) = ϕ i (f i) * to_fun ϕ hcomm f S := finset.noncomm_prod_insert_of_not_mem _ _ _ _ h @@ -116,7 +116,7 @@ begin end @[to_additive add_to_fun_commutes] -lemma to_fun_commutes (i : I) (hnmem : i ∉ S) : +lemma to_fun_commutes (i : ι) (hnmem : i ∉ S) : commute (ϕ i (g i)) (to_fun ϕ hcomm f S) := begin induction S using finset.induction_on with j S hnmem' ih, @@ -133,7 +133,7 @@ begin end @[simp, to_additive add_to_fun_add] -lemma to_fun_mul (S : finset I) : +lemma to_fun_mul (S : finset ι) : to_fun ϕ hcomm (f * g) S = to_fun ϕ hcomm f S * to_fun ϕ hcomm g S := begin induction S using finset.induction_on with i S hnmem ih, @@ -147,7 +147,7 @@ begin end @[to_additive add_to_fun_in_sup_mrange] -lemma to_fun_in_sup_mrange (S : finset I) : +lemma to_fun_in_sup_mrange (S : finset ι) : to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin induction S using finset.induction_on with i S hnmem ih, @@ -165,7 +165,7 @@ variable (hcomm) -/ @[to_additive add_hom "The canonical homomorphism from a family of additive monoids, restricted to a subset of the index space"] -def hom : (Π (i : I), N i) →* M := +def hom : (Π (i : ι), N i) →* M := { to_fun := λ f, to_fun ϕ hcomm f S, map_one' := to_fun_one ϕ _, map_mul' := λ f g, to_fun_mul ϕ f g S, } @@ -173,7 +173,7 @@ def hom : (Π (i : I), N i) →* M := variable {hcomm} @[to_additive add_to_fun_single] -lemma to_fun_single (i : I) (y : N i) (S : finset I) : +lemma to_fun_single (i : ι) (y : N i) (S : finset ι) : to_fun ϕ hcomm (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, @@ -186,7 +186,7 @@ begin end @[simp, to_additive add_hom_single] -lemma hom_single (i : I) (y : N i): +lemma hom_single (i : ι) (y : N i): hom ϕ hcomm S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ _ @[to_additive add_mrange] @@ -211,17 +211,17 @@ variable (hcomm) /-- The canonical homomorphism from a family of monoids. -/ @[to_additive "The canonical homomorphism from a family of additive monoids."] -def noncomm_pi_coprod : (Π (i : I), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ +def noncomm_pi_coprod : (Π (i : ι), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ variable {hcomm} @[simp, to_additive] -lemma noncomm_pi_coprod_single (i : I) (y : N i): +lemma noncomm_pi_coprod_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } @[to_additive] -lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : I, (ϕ i).mrange := +lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : ι, (ϕ i).mrange := begin show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).mrange = _, simp [noncomm_pi_coprod_on.mrange], @@ -234,18 +234,18 @@ end family_of_monoids section family_of_groups variables {G : Type*} [group G] -variables {I : Type*} [hfin : fintype I] -variables {H : I → Type*} [∀ i, group (H i)] -variables (ϕ : Π (i : I), H i →* G) -variables {hcomm : ∀ (i j : I), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)} +variables {ι : Type*} [hfin : fintype ι] +variables {H : ι → Type*} [∀ i, group (H i)] +variables (ϕ : Π (i : ι), H i →* G) +variables {hcomm : ∀ (i j : ι), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)} include hcomm --- We use `f` and `g` to denote elements of `Π (i : I), H i` -variables (f g : Π (i : I), H i) +-- We use `f` and `g` to denote elements of `Π (i : ι), H i` +variables (f g : Π (i : ι), H i) -- The subgroup version of `noncomm_pi_coprod_on.to_fun_in_sup_mrange` @[to_additive noncomm_pi_coprod_on.add_to_fun_in_sup_range] -lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset I) : +lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset ι) : noncomm_pi_coprod_on.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin induction S using finset.induction_on with i S hnmem ih, @@ -259,7 +259,7 @@ end -- The subgroup version of `noncomm_pi_coprod_on.mrange` @[to_additive noncomm_pi_coprod_on.add_range] -lemma noncomm_pi_coprod_on.range (S : finset I) : +lemma noncomm_pi_coprod_on.range (S : finset ι) : (noncomm_pi_coprod_on.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := begin apply le_antisymm, @@ -277,7 +277,7 @@ namespace monoid_hom -- The subgroup version of `noncomm_pi_coprod_on_mrange` @[to_additive] -lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : I, (ϕ i).range := +lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := begin show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).range = _, simp [noncomm_pi_coprod_on.range] @@ -294,13 +294,13 @@ begin intro f, show noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → f = 1, suffices : noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → - (∀ (i : I), i ∈ finset.univ → f i = 1), - by exact (λ h, funext (λ (i : I), this h i (finset.mem_univ i))), - induction (finset.univ : finset I) using finset.induction_on with i S hnmem ih, + (∀ (i : ι), i ∈ finset.univ → f i = 1), + by exact (λ h, funext (λ (i : ι), this h i (finset.mem_univ i))), + induction (finset.univ : finset ι) using finset.induction_on with i S hnmem ih, { simp }, { intro heq1, simp only [ noncomm_pi_coprod_on.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, - have hnmem' : i ∉ (S : set I), by simpa, + have hnmem' : i ∉ (S : set ι), by simpa, have heq1' : ϕ i (f i) = 1 ∧ noncomm_pi_coprod_on.to_fun ϕ hcomm f S = 1, { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, @@ -343,19 +343,19 @@ namespace subgroup -- We have an family of subgroups variables {G : Type*} [group G] -variables {I : Type*} [hfin : fintype I] [hdec : decidable_eq I] {H : I → subgroup G} +variables {ι : Type*} [hfin : fintype ι] [hdec : decidable_eq ι] {H : ι → subgroup G} --- Elements of `Π (i : I), H i` are called `f` and `g` here -variables (f g : Π (i : I), H i) +-- Elements of `Π (i : ι), H i` are called `f` and `g` here +variables (f g : Π (i : ι), H i) section commuting_subgroups -- We assume that the elements of different subgroups commute -variables (hcomm : ∀ (i j : I), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) +variables (hcomm : ∀ (i j : ι), i ≠ j → ∀ (x y : G), x ∈ H i → y ∈ H j → commute x y) include hcomm @[to_additive] -lemma hcomm_subtype (i j : I) (hne : i ≠ j) : +lemma hcomm_subtype (i j : ι) (hne : i ≠ j) : ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y) := by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } @@ -365,18 +365,18 @@ include hfin commute -/ @[to_additive "The canonical homomorphism from a family of additive subgroups where elements from different subgroups commute"] -def noncomm_pi_coprod : (Π (i : I), H i) →* G := +def noncomm_pi_coprod : (Π (i : ι), H i) →* G := monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (hcomm_subtype hcomm) variable {hcomm} @[simp, to_additive] -lemma noncomm_pi_coprod_single (i : I) (y : H i) : +lemma noncomm_pi_coprod_single (i : ι) (y : H i) : noncomm_pi_coprod hcomm (monoid_hom.single _ i y) = y := by apply monoid_hom.noncomm_pi_coprod_single @[to_additive] -lemma range : (noncomm_pi_coprod hcomm).range = ⨆ i : I, H i := +lemma range : (noncomm_pi_coprod hcomm).range = ⨆ i : ι, H i := by simp [noncomm_pi_coprod, monoid_hom.noncomm_pi_coprod_range] @[to_additive] From 1cf2f74c3b9de3c807bd0c101954518bc7c70722 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 16 Feb 2022 14:48:25 +0100 Subject: [PATCH 037/116] Fix typo in filename --- .../{nocomm_pi_coprod.lean => noncomm_pi_coprod.lean} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename src/group_theory/{nocomm_pi_coprod.lean => noncomm_pi_coprod.lean} (100%) diff --git a/src/group_theory/nocomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean similarity index 100% rename from src/group_theory/nocomm_pi_coprod.lean rename to src/group_theory/noncomm_pi_coprod.lean From 161d1a3ab577b5846ed4a6625fba2952dcdba310 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 16 Feb 2022 18:00:23 +0100 Subject: [PATCH 038/116] Fix very embarrasing search-and-replace mistake --- src/group_theory/noncomm_pi_coprod.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 563a439d41c5d..e84be8ca15dab 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -1,6 +1,6 @@ /- Copyright (c) 2022 Jocchim Breitner. All rights reserved. -Released under Apache 2.0 license as described in the file LιCENSE. +Released under Apache 2.0 license as described in the file LICENSE. Authors: Joachim Breitner -/ import group_theory.order_of_element @@ -35,9 +35,9 @@ image of different morphism commute, we obtain a canoncial morphism * `subgroup.noncomm_pi_coprod_range`: The range of `subgroup.noncomm_pi_coprod` is `⨆ (i : ι), H i`. * `monoid_hom.injective_noncomm_pi_coprod_of_independent`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. -* `monoid_hom.independent_range_of_coprime_order`: ιf the `N i` have coprime orders, then the ranges +* `monoid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges of the `ϕ` are independent. -* `subgroup.independent_of_coprime_order`: ιf commuting, normal subgroups `H i` have coprime orders, +* `subgroup.independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, they are independent. -/ @@ -51,7 +51,7 @@ lemma coprime_prod_left (∀ (i : ι), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : ι) in t, s i) x := finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) --- ι think it's worth keeping it and moving to appropriate file +-- I think it's worth keeping it and moving to appropriate file @[to_additive] lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := @@ -90,7 +90,7 @@ variables (f g : Π (i : ι), N i) namespace noncomm_pi_coprod_on --- ιn this section, we restrict the hom to a set S +-- In this section, we restrict the hom to a set S variables (S : finset ι) /-- The underlying function of `noncomm_pi_coprod_on.hom` -/ From 858a150bf2a57ad198a96fdd8505ccf79fbbcca8 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 17 Feb 2022 20:19:55 +0100 Subject: [PATCH 039/116] Refine some lemma names --- src/group_theory/noncomm_pi_coprod.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index e84be8ca15dab..bf3064e8033a8 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -355,7 +355,7 @@ variables (hcomm : ∀ (i j : ι), i ≠ j → ∀ (x y : G), x ∈ H i → y include hcomm @[to_additive] -lemma hcomm_subtype (i j : ι) (hne : i ≠ j) : +lemma commute_subtype_of_commute (i j : ι) (hne : i ≠ j) : ∀ (x : H i) (y : H j), commute ((H i).subtype x) ((H j).subtype y) := by { rintros ⟨x, hx⟩ ⟨y, hy⟩, exact hcomm i j hne x y hx hy } @@ -366,7 +366,7 @@ commute -/ @[to_additive "The canonical homomorphism from a family of additive subgroups where elements from different subgroups commute"] def noncomm_pi_coprod : (Π (i : ι), H i) →* G := - monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (hcomm_subtype hcomm) + monoid_hom.noncomm_pi_coprod (λ i, (H i).subtype) (commute_subtype_of_commute hcomm) variable {hcomm} @@ -376,11 +376,11 @@ lemma noncomm_pi_coprod_single (i : ι) (y : H i) : by apply monoid_hom.noncomm_pi_coprod_single @[to_additive] -lemma range : (noncomm_pi_coprod hcomm).range = ⨆ i : ι, H i := +lemma noncomm_pi_coprod_range : (noncomm_pi_coprod hcomm).range = ⨆ i : ι, H i := by simp [noncomm_pi_coprod, monoid_hom.noncomm_pi_coprod_range] @[to_additive] -lemma injective_of_independent (hind : complete_lattice.independent H) : +lemma injective_noncomm_pi_coprod_of_independent (hind : complete_lattice.independent H) : function.injective (noncomm_pi_coprod hcomm) := begin apply monoid_hom.injective_noncomm_pi_coprod_of_independent, @@ -394,7 +394,7 @@ lemma independent_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent H := begin simpa using monoid_hom.independent_range_of_coprime_order - (λ i, (H i).subtype) (hcomm_subtype hcomm) hcoprime, + (λ i, (H i).subtype) (commute_subtype_of_commute hcomm) hcoprime, end end commuting_subgroups From fd01d6561c0c0356a8372e6ee712fef2844c1d87 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 21 Feb 2022 11:45:45 +0100 Subject: [PATCH 040/116] Avoid open_locale classical --- src/group_theory/noncomm_pi_coprod.lean | 12 +++++------- 1 file changed, 5 insertions(+), 7 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index bf3064e8033a8..8eb8050939090 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -43,7 +43,6 @@ image of different morphism commute, we obtain a canoncial morphism -/ open_locale big_operators -open_locale classical lemma coprime_prod_left {ι : Type*} @@ -75,7 +74,7 @@ variables {M : Type*} [monoid M] -- We have a family of monoids -- The fintype assumption is not always used, but declared here, to keep things in order -variables {ι : Type*} [hfin : fintype ι] +variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] -- And morphisms ϕ into G @@ -234,7 +233,7 @@ end family_of_monoids section family_of_groups variables {G : Type*} [group G] -variables {ι : Type*} [hfin : fintype ι] +variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] variables {H : ι → Type*} [∀ i, group (H i)] variables (ϕ : Π (i : ι), H i →* G) variables {hcomm : ∀ (i j : ι), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)} @@ -322,9 +321,8 @@ lemma independent_range_of_coprime_order [∀ i, fintype (H i)] complete_lattice.independent (λ i, (ϕ i).range) := begin rintros i f ⟨hxi, hxp⟩, dsimp at hxi hxp, - rw [supr_subtype', ← @noncomm_pi_coprod_range _ _ _ _ _ _ _ _] at hxp, - rotate, apply_instance, - { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, + rw [supr_subtype', ← noncomm_pi_coprod_range] at hxp, + rotate, { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, cases hxp with g hgf, cases hxi with g' hg'f, have hxi : order_of f ∣ fintype.card (H i), { rw ← hg'f, exact (order_of_map_dvd _ _).trans order_of_dvd_card_univ }, @@ -343,7 +341,7 @@ namespace subgroup -- We have an family of subgroups variables {G : Type*} [group G] -variables {ι : Type*} [hfin : fintype ι] [hdec : decidable_eq ι] {H : ι → subgroup G} +variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] {H : ι → subgroup G} -- Elements of `Π (i : ι), H i` are called `f` and `g` here variables (f g : Π (i : ι), H i) From 45da268b35b3d445577c97cac58efd3dadc6978f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 21 Feb 2022 23:38:03 +0100 Subject: [PATCH 041/116] More pleasing classical linter --- src/group_theory/noncomm_pi_coprod.lean | 32 ++++++++++++++++++++++--- 1 file changed, 29 insertions(+), 3 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 8eb8050939090..d49ab857ac558 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -74,7 +74,7 @@ variables {M : Type*} [monoid M] -- We have a family of monoids -- The fintype assumption is not always used, but declared here, to keep things in order -variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] +variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] -- And morphisms ϕ into G @@ -102,14 +102,19 @@ variable {hcomm} @[simp, to_additive add_to_fun_empty] lemma to_fun_empty : to_fun ϕ hcomm f ∅ = 1 := finset.noncomm_prod_empty _ _ +include hdec + @[simp, to_additive add_to_fun_insert_of_not_mem] lemma to_fun_insert_of_not_mem (S : finset ι) (i : ι) (h : i ∉ S) : to_fun ϕ hcomm f (insert i S) = ϕ i (f i) * to_fun ϕ hcomm f S := finset.noncomm_prod_insert_of_not_mem _ _ _ _ h +omit hdec + @[simp, to_additive add_to_fun_zero] lemma to_fun_one : to_fun ϕ hcomm 1 S = 1 := begin + classical, induction S using finset.cons_induction_on with i S hnmem ih, { simp }, { simp [ih, hnmem], } end @@ -118,6 +123,7 @@ end lemma to_fun_commutes (i : ι) (hnmem : i ∉ S) : commute (ϕ i (g i)) (to_fun ϕ hcomm f S) := begin + classical, induction S using finset.induction_on with j S hnmem' ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], @@ -135,6 +141,7 @@ end lemma to_fun_mul (S : finset ι) : to_fun ϕ hcomm (f * g) S = to_fun ϕ hcomm f S * to_fun ϕ hcomm g S := begin + classical, induction S using finset.induction_on with i S hnmem ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], @@ -149,6 +156,7 @@ end lemma to_fun_in_sup_mrange (S : finset ι) : to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin + classical, induction S using finset.induction_on with i S hnmem ih, { simp, }, { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], @@ -171,6 +179,8 @@ def hom : (Π (i : ι), N i) →* M := variable {hcomm} +include hdec + @[to_additive add_to_fun_single] lemma to_fun_single (i : ι) (y : N i) (S : finset ι) : to_fun ϕ hcomm (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := @@ -188,9 +198,12 @@ end lemma hom_single (i : ι) (y : N i): hom ϕ hcomm S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ _ +omit hdec + @[to_additive add_mrange] lemma mrange : (hom ϕ hcomm S).mrange = ⨆ i ∈ S, (ϕ i).mrange := begin + classical, apply le_antisymm, { rintro x ⟨f, rfl⟩, exact (to_fun_in_sup_mrange ϕ f S), }, @@ -214,11 +227,15 @@ def noncomm_pi_coprod : (Π (i : ι), N i) →* M := noncomm_pi_coprod_on.hom ϕ variable {hcomm} +include hdec + @[simp, to_additive] lemma noncomm_pi_coprod_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +omit hdec + @[to_additive] lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : ι, (ϕ i).mrange := begin @@ -233,7 +250,7 @@ end family_of_monoids section family_of_groups variables {G : Type*} [group G] -variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] +variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {H : ι → Type*} [∀ i, group (H i)] variables (ϕ : Π (i : ι), H i →* G) variables {hcomm : ∀ (i j : ι), i ≠ j → ∀ (x : H i) (y : H j), commute (ϕ i x) (ϕ j y)} @@ -247,6 +264,7 @@ variables (f g : Π (i : ι), H i) lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset ι) : noncomm_pi_coprod_on.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin + classical, induction S using finset.induction_on with i S hnmem ih, { simp, }, { simp only [noncomm_pi_coprod_on.to_fun_insert_of_not_mem _ _ _ _ hnmem], @@ -261,6 +279,7 @@ end lemma noncomm_pi_coprod_on.range (S : finset ι) : (noncomm_pi_coprod_on.hom ϕ hcomm S).range = ⨆ i ∈ S, (ϕ i).range := begin + classical, apply le_antisymm, { rintro x ⟨f, rfl⟩, exact (noncomm_pi_coprod_on.to_fun_in_sup_range ϕ f S), }, @@ -288,6 +307,7 @@ lemma injective_noncomm_pi_coprod_of_independent (hinj : ∀ i, function.injective (ϕ i)) : function.injective (noncomm_pi_coprod ϕ hcomm):= begin + classical, apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, intro f, @@ -315,11 +335,13 @@ end variable (hcomm) + @[to_additive] lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : complete_lattice.independent (λ i, (ϕ i).range) := begin + classical, rintros i f ⟨hxi, hxp⟩, dsimp at hxi hxp, rw [supr_subtype', ← noncomm_pi_coprod_range] at hxp, rotate, { intros _ _ hj, apply hcomm, exact hj ∘ subtype.ext }, @@ -341,7 +363,7 @@ namespace subgroup -- We have an family of subgroups variables {G : Type*} [group G] -variables {ι : Type*} [decidable_eq ι] [hfin : fintype ι] {H : ι → subgroup G} +variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] {H : ι → subgroup G} -- Elements of `Π (i : ι), H i` are called `f` and `g` here variables (f g : Π (i : ι), H i) @@ -368,11 +390,15 @@ def noncomm_pi_coprod : (Π (i : ι), H i) →* G := variable {hcomm} +include hdec + @[simp, to_additive] lemma noncomm_pi_coprod_single (i : ι) (y : H i) : noncomm_pi_coprod hcomm (monoid_hom.single _ i y) = y := by apply monoid_hom.noncomm_pi_coprod_single +omit hdec + @[to_additive] lemma noncomm_pi_coprod_range : (noncomm_pi_coprod hcomm).range = ⨆ i : ι, H i := by simp [noncomm_pi_coprod, monoid_hom.noncomm_pi_coprod_range] From a22e26b277d22cf3158efcf85945bde467c450cb Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 09:31:23 +0100 Subject: [PATCH 042/116] Move coprime_prod_left to src/data/nat/gcd.lean --- src/data/nat/gcd.lean | 17 +++++++++++++++++ src/group_theory/noncomm_pi_coprod.lean | 6 ------ 2 files changed, 17 insertions(+), 6 deletions(-) diff --git a/src/data/nat/gcd.lean b/src/data/nat/gcd.lean index 4b92ffd5474f1..45acf8d4d2dd7 100644 --- a/src/data/nat/gcd.lean +++ b/src/data/nat/gcd.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import algebra.group_power.order +import algebra.big_operators.basic /-! # Definitions and properties of `gcd`, `lcm`, and `coprime` @@ -466,6 +467,22 @@ by simp [coprime] @[simp] theorem coprime_self (n : ℕ) : coprime n n ↔ n = 1 := by simp [coprime] +section big_operators + +open_locale big_operators + +lemma coprime_prod_left + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime (s i) x) → coprime (∏ (i : ι) in t, s i) x := +finset.prod_induction s (λ y, y.coprime x) (λ a b, coprime.mul) (by simp) + +lemma coprime_prod_right + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime x (s i)) → coprime x (∏ (i : ι) in t, s i) := +finset.prod_induction s (λ y, x.coprime y) (λ a b, coprime.mul_right) (by simp) + +end big_operators + lemma coprime.eq_of_mul_eq_zero {m n : ℕ} (h : m.coprime n) (hmn : m * n = 0) : m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := (nat.eq_zero_of_mul_eq_zero hmn).imp diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index d49ab857ac558..520c5109717e2 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -44,12 +44,6 @@ image of different morphism commute, we obtain a canoncial morphism open_locale big_operators -lemma coprime_prod_left - {ι : Type*} - {x : ℕ} {s : ι → ℕ} {t : finset ι} : - (∀ (i : ι), i ∈ t → nat.coprime (s i) x) → nat.coprime (∏ (i : ι) in t, s i) x := -finset.prod_induction s (λ y, y.coprime x) (λ a b, nat.coprime.mul) (by simp) - -- I think it's worth keeping it and moving to appropriate file @[to_additive] lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : From ceccaf73af8a52d528d6c68ea04b19043f76a833 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 09:32:46 +0100 Subject: [PATCH 043/116] feat(data/nat/gcd): add coprime_prod_left and coprime_prod_right --- src/data/nat/gcd.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/data/nat/gcd.lean b/src/data/nat/gcd.lean index 4b92ffd5474f1..45acf8d4d2dd7 100644 --- a/src/data/nat/gcd.lean +++ b/src/data/nat/gcd.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import algebra.group_power.order +import algebra.big_operators.basic /-! # Definitions and properties of `gcd`, `lcm`, and `coprime` @@ -466,6 +467,22 @@ by simp [coprime] @[simp] theorem coprime_self (n : ℕ) : coprime n n ↔ n = 1 := by simp [coprime] +section big_operators + +open_locale big_operators + +lemma coprime_prod_left + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime (s i) x) → coprime (∏ (i : ι) in t, s i) x := +finset.prod_induction s (λ y, y.coprime x) (λ a b, coprime.mul) (by simp) + +lemma coprime_prod_right + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime x (s i)) → coprime x (∏ (i : ι) in t, s i) := +finset.prod_induction s (λ y, x.coprime y) (λ a b, coprime.mul_right) (by simp) + +end big_operators + lemma coprime.eq_of_mul_eq_zero {m n : ℕ} (h : m.coprime n) (hmn : m * n = 0) : m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := (nat.eq_zero_of_mul_eq_zero hmn).imp From b477bd1c44a07fa20b59797682954c0ba143192d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 11:39:31 +0100 Subject: [PATCH 044/116] s/in_sup/mem_bsupr/g --- src/group_theory/noncomm_pi_coprod.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 520c5109717e2..c78a4f75447ed 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -146,8 +146,8 @@ begin exact (to_fun_commutes _ _ _ S i hnmem), } end -@[to_additive add_to_fun_in_sup_mrange] -lemma to_fun_in_sup_mrange (S : finset ι) : +@[to_additive add_to_fun_mem_bsupr_mrange] +lemma to_fun_mem_bsupr_mrange (S : finset ι) : to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := begin classical, @@ -200,7 +200,7 @@ begin classical, apply le_antisymm, { rintro x ⟨f, rfl⟩, - exact (to_fun_in_sup_mrange ϕ f S), }, + exact (to_fun_mem_bsupr_mrange ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, use (monoid_hom.single _ i y), @@ -253,9 +253,9 @@ include hcomm -- We use `f` and `g` to denote elements of `Π (i : ι), H i` variables (f g : Π (i : ι), H i) --- The subgroup version of `noncomm_pi_coprod_on.to_fun_in_sup_mrange` -@[to_additive noncomm_pi_coprod_on.add_to_fun_in_sup_range] -lemma noncomm_pi_coprod_on.to_fun_in_sup_range (S : finset ι) : +-- The subgroup version of `noncomm_pi_coprod_on.to_fun_mem_bsupr_mrange` +@[to_additive noncomm_pi_coprod_on.add_to_fun_mem_bsupr_range] +lemma noncomm_pi_coprod_on.to_fun_mem_bsupr_range (S : finset ι) : noncomm_pi_coprod_on.to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).range := begin classical, @@ -276,7 +276,7 @@ begin classical, apply le_antisymm, { rintro x ⟨f, rfl⟩, - exact (noncomm_pi_coprod_on.to_fun_in_sup_range ϕ f S), }, + exact (noncomm_pi_coprod_on.to_fun_mem_bsupr_range ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, use (monoid_hom.single _ i y), @@ -317,7 +317,7 @@ begin have heq1' : ϕ i (f i) = 1 ∧ noncomm_pi_coprod_on.to_fun ϕ hcomm f S = 1, { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, { simp, }, - { apply noncomm_pi_coprod_on.to_fun_in_sup_range, }, }, + { apply noncomm_pi_coprod_on.to_fun_mem_bsupr_range, }, }, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, From 61fdbfcc07b5381a663e6ff1f93e978875c9e23d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 11:41:52 +0100 Subject: [PATCH 045/116] Review comments --- src/group_theory/noncomm_pi_coprod.lean | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index c78a4f75447ed..0313a78ffb693 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -216,7 +216,9 @@ include hfin variable (hcomm) /-- The canonical homomorphism from a family of monoids. -/ -@[to_additive "The canonical homomorphism from a family of additive monoids."] +@[to_additive "The canonical homomorphism from a family of additive monoids. + +See also `linear_map.lsum` for a linear version without the commutativity assumption."] def noncomm_pi_coprod : (Π (i : ι), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ variable {hcomm} @@ -287,7 +289,7 @@ include hfin namespace monoid_hom --- The subgroup version of `noncomm_pi_coprod_on_mrange` +-- The subgroup version of `noncomm_pi_coprod_mrange` @[to_additive] lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := begin @@ -346,7 +348,7 @@ begin { rw [← hgf, ← fintype.card_pi], exact (order_of_map_dvd _ _).trans order_of_dvd_card_univ }, change f = 1, rw [← pow_one f, ← order_of_dvd_iff_pow_eq_one], convert ← nat.dvd_gcd hxp hxi, rw ← nat.coprime_iff_gcd_eq_one, - apply coprime_prod_left, intros j _, apply hcoprime, exact j.2, + apply nat.coprime_prod_left, intros j _, apply hcoprime, exact j.2, end end monoid_hom From 06e42bb56451fe7de34ae1b488b05c5eaf141cc4 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 12:34:46 +0100 Subject: [PATCH 046/116] Add universal property --- src/algebra/group/pi.lean | 16 ++++++++++++++++ src/group_theory/noncomm_pi_coprod.lean | 17 +++++++++++++++++ 2 files changed, 33 insertions(+) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 230247539b242..dc310884228fc 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -5,6 +5,7 @@ Authors: Simon Hudon, Patrick Massot -/ import data.pi import data.set.function +import data.set.pairwise import tactic.pi_instances import algebra.group.hom_instances @@ -235,6 +236,21 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : single i (x * y) = single i x * single i y := (mul_hom.single f i).map_mul x y +@[to_additive] +lemma pi.mul_single_commute [Π i, mul_one_class $ f i] : + pairwise (λ i j, ∀ (x : f i) (y : f j), commute (mul_single i x) (mul_single j y)) := +begin + intros i j hij x y, ext k, + by_cases h1 : i = k, { subst h1, simp [hij], }, + by_cases h2 : j = k, { subst h2, simp [hij], }, + simp [h1, h2], +end + +@[to_additive] +lemma monoid_hom.single_commute [Π i, mul_one_class $ f i] : + pairwise (λ i j, ∀ x y, commute (monoid_hom.single f i x) (monoid_hom.single f j y)) := +pi.mul_single_commute + @[to_additive update_eq_sub_add_single] lemma pi.update_eq_div_mul_single [Π i, group $ f i] (g : Π (i : I), f i) (x : f i) : function.update g i x = g / mul_single i (g i) * mul_single i x := diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 0313a78ffb693..65da88bc15bac 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -71,6 +71,11 @@ variables {M : Type*} [monoid M] variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] +@[to_additive] +lemma _root_.monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} : + f = g ↔ (∀ i x, f (monoid_hom.single N i x) = g (monoid_hom.single N i x)) := +sorry + -- And morphisms ϕ into G variables (ϕ : Π (i : ι), N i →* M) @@ -230,6 +235,18 @@ lemma noncomm_pi_coprod_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +/-- The universal property of `noncomm_pi_coprod` -/ +@[to_additive "The universal property of `noncomm_pi_coprod`"] +def noncomm_pi_coprod_equiv : + {ϕ : Π i, N i →* M // pairwise (λ i j, ∀ x y, commute (ϕ i x) (ϕ j y)) } + ≃ ((Π i, N i) →* M) := +{ to_fun := λ ϕ, noncomm_pi_coprod ϕ.1 ϕ.2, + inv_fun := λ f, + ⟨ λ i, f.comp (monoid_hom.single N i), + λ i j hij x y, commute.map (monoid_hom.single_commute i j hij x y) f ⟩, + left_inv := λ ϕ, by {ext ι x, simp }, + right_inv := λ f, monoid_hom.pi_ext.mpr (λ i x, by simp), } + omit hdec @[to_additive] From 02bf46612a914e59454e6b5aefdc04e45e3d3913 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 24 Feb 2022 15:55:28 +0100 Subject: [PATCH 047/116] More exploration --- src/group_theory/noncomm_pi_coprod.lean | 43 +++++++++++++++++++++++-- 1 file changed, 41 insertions(+), 2 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 65da88bc15bac..8d31f6fb3e336 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -72,10 +72,49 @@ variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] @[to_additive] -lemma _root_.monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} : - f = g ↔ (∀ i x, f (monoid_hom.single N i x) = g (monoid_hom.single N i x)) := +lemma _root_.finset.noncomm_prod_map {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + {M : Type*} [monoid M] (g : β →* M) : + g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) + (λ x hx y hy, commute.map (comm x hx y hy) g) := sorry +@[to_additive] +lemma _root_.finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := +sorry + +@[to_additive] +lemma _root_.finset.noncomm_prod_single [decidable_eq ι] [fintype ι] (x : Π i, N i) : + finset.univ.noncomm_prod (λ i, monoid_hom.single N i (x i)) + (λ i _ j _, by { by_cases h : i = j, { subst h, }, { exact pi.mul_single_commute i j h _ _ } }) + = x := +begin + ext i, + apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single N i (x i)) _ + (pi.eval_monoid_hom _ i)).trans, + rw (finset.insert_erase (finset.mem_univ i)).symm, + rw finset.noncomm_prod_insert_of_not_mem', + rw finset.noncomm_prod_eq_one, + { simp, }, + { intros i h, simp at h, simp [h], }, + { simp, }, +end + +@[to_additive] +lemma _root_.monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} + (h : ∀ i x, f (monoid_hom.single N i x) = g (monoid_hom.single N i x)) : + f = g := +begin + ext x, + rw ← finset.noncomm_prod_single x, + rw finset.univ.noncomm_prod_map _ _ f, + rw finset.univ.noncomm_prod_map _ _ g, + congr' 1, ext i, exact h i (x i), +end + + -- And morphisms ϕ into G variables (ϕ : Π (i : ι), N i →* M) From ce1e1c19dbfb295177323eb80d4d5c59b1e33b10 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 25 Feb 2022 00:07:18 -0500 Subject: [PATCH 048/116] kill sorries --- src/group_theory/noncomm_pi_coprod.lean | 34 +++++++++++++++++-------- 1 file changed, 24 insertions(+), 10 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 8d31f6fb3e336..d1004494e76a0 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -72,23 +72,36 @@ variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] @[to_additive] -lemma _root_.finset.noncomm_prod_map {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) +lemma finset.noncomm_prod_map {α : Type*} {β : Type*} [monoid β] + [decidable_eq α] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) {M : Type*} [monoid M] (g : β →* M) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, commute.map (comm x hx y hy) g) := -sorry +begin + revert comm, rw ← s.to_list_to_finset, intro, + iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, + convert g.map_list_prod _, rw list.comp_map, +end @[to_additive] -lemma _root_.finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) +lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] [decidable_eq α] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := -sorry +begin + revert comm, rw ← s.to_list_to_finset, intro, + rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, + revert h, simp_rw ← s.mem_to_list, + induction s.to_list, simp, + intro h, simp [ih (λ x hx, h x (or.inr hx)), h hd (or.inl rfl)], +end +/- may generalize to any element instead of `1`, where we get `pow` with exponent `card`. + maybe also provide a list version. -/ @[to_additive] -lemma _root_.finset.noncomm_prod_single [decidable_eq ι] [fintype ι] (x : Π i, N i) : +lemma finset.noncomm_prod_single [decidable_eq ι] [fintype ι] (x : Π i, N i) : finset.univ.noncomm_prod (λ i, monoid_hom.single N i (x i)) - (λ i _ j _, by { by_cases h : i = j, { subst h, }, { exact pi.mul_single_commute i j h _ _ } }) + (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) = x := begin ext i, @@ -98,12 +111,12 @@ begin rw finset.noncomm_prod_insert_of_not_mem', rw finset.noncomm_prod_eq_one, { simp, }, - { intros i h, simp at h, simp [h], }, + { intros i h, simp at h, simp [h], }, { simp, }, end @[to_additive] -lemma _root_.monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} +lemma monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} (h : ∀ i x, f (monoid_hom.single N i x) = g (monoid_hom.single N i x)) : f = g := begin @@ -112,6 +125,7 @@ begin rw finset.univ.noncomm_prod_map _ _ f, rw finset.univ.noncomm_prod_map _ _ g, congr' 1, ext i, exact h i (x i), + iterate 2 { by apply_instance }, end @@ -283,8 +297,8 @@ def noncomm_pi_coprod_equiv : inv_fun := λ f, ⟨ λ i, f.comp (monoid_hom.single N i), λ i j hij x y, commute.map (monoid_hom.single_commute i j hij x y) f ⟩, - left_inv := λ ϕ, by {ext ι x, simp }, - right_inv := λ f, monoid_hom.pi_ext.mpr (λ i x, by simp), } + left_inv := λ ϕ, by { ext, simp }, + right_inv := λ f, pi_ext (λ i x, by simp) } omit hdec From 3bf0e6f3254b3a3d63788df54e1d065a0ad93048 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 22:17:49 +0100 Subject: [PATCH 049/116] Classical linter --- src/group_theory/noncomm_pi_coprod.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index d1004494e76a0..321562d1fc85e 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -73,22 +73,24 @@ variables {N : ι → Type*} [∀ i, monoid (N i)] @[to_additive] lemma finset.noncomm_prod_map {α : Type*} {β : Type*} [monoid β] - [decidable_eq α] (s : finset α) (f : α → β) + (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) {M : Type*} [monoid M] (g : β →* M) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, commute.map (comm x hx y hy) g) := begin + classical, revert comm, rw ← s.to_list_to_finset, intro, iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, convert g.map_list_prod _, rw list.comp_map, end @[to_additive] -lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] [decidable_eq α] (s : finset α) +lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := begin + classical, revert comm, rw ← s.to_list_to_finset, intro, rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, revert h, simp_rw ← s.mem_to_list, @@ -125,7 +127,6 @@ begin rw finset.univ.noncomm_prod_map _ _ f, rw finset.univ.noncomm_prod_map _ _ g, congr' 1, ext i, exact h i (x i), - iterate 2 { by apply_instance }, end From 7e7dc82941f366eb7c701ebb0addc824c66a13c4 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 22:34:55 +0100 Subject: [PATCH 050/116] Generalize a bit --- src/group_theory/noncomm_pi_coprod.lean | 31 +++++++++++++++++++------ 1 file changed, 24 insertions(+), 7 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 321562d1fc85e..d9bf11f70f385 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -86,19 +86,36 @@ begin end @[to_additive] -lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) +lemma list.prod_const {α : Type*} [monoid α] (l : list α) + (m : α) (h : ∀ (x : α), x ∈ l → x = m) : l.prod = m ^ l.length := +begin + convert list.prod_repeat m l.length, + rw list.eq_repeat, + exact ⟨rfl, h⟩, +end + +@[to_additive] +lemma finset.noncomm_prod_const {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := + (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := begin classical, revert comm, rw ← s.to_list_to_finset, intro, rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, - revert h, simp_rw ← s.mem_to_list, - induction s.to_list, simp, - intro h, simp [ih (λ x hx, h x (or.inr hx)), h hd (or.inl rfl)], + revert h, simp_rw ← s.mem_to_list, intro, + rw list.prod_const _ m, + { simp, }, + { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x hx }, +end + +@[to_additive] +lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := +begin + rw finset.noncomm_prod_const s f comm 1 h, + exact one_pow _, end -/- may generalize to any element instead of `1`, where we get `pow` with exponent `card`. - maybe also provide a list version. -/ @[to_additive] lemma finset.noncomm_prod_single [decidable_eq ι] [fintype ι] (x : Π i, N i) : From 69d481bd50c7a4d4e40688974aa905f511f9de7c Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 22:36:10 +0100 Subject: [PATCH 051/116] Simplify --- src/group_theory/noncomm_pi_coprod.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index d9bf11f70f385..2bdc6518ec993 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -90,8 +90,7 @@ lemma list.prod_const {α : Type*} [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : l.prod = m ^ l.length := begin convert list.prod_repeat m l.length, - rw list.eq_repeat, - exact ⟨rfl, h⟩, + exact list.eq_repeat.mpr ⟨rfl, h⟩, end @[to_additive] From d90d9875597560865a34fff96b659d87a429e23d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 23:02:55 +0100 Subject: [PATCH 052/116] Move lemmas --- src/data/finset/noncomm_prod.lean | 72 ++++++++++++++++++++++++ src/data/list/prod_monoid.lean | 8 +++ src/group_theory/noncomm_pi_coprod.lean | 75 ------------------------- 3 files changed, 80 insertions(+), 75 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index be0b2177c6e3c..e8d19347d595d 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -217,6 +217,43 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := by simp [noncomm_prod, multiset.singleton_eq_cons] + +@[to_additive] +lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] + (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + {M : Type*} [monoid M] (g : β →* M) : + g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) + (λ x hx y hy, commute.map (comm x hx y hy) g) := +begin + classical, + revert comm, rw ← s.to_list_to_finset, intro, + iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, + convert g.map_list_prod _, rw list.comp_map, +end + +@[to_additive] +lemma noncomm_prod_eq_pow_of_forall_eq {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := +begin + classical, + revert comm, rw ← s.to_list_to_finset, intro, + rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, + rw list.prod_eq_pow_of_forall_eq _ m, + { simp }, + { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x (s.mem_to_list.mp hx) }, +end + +@[to_additive] +lemma noncomm_prod_eq_one_of_forall_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := +begin + rw finset.noncomm_prod_eq_pow_of_forall_eq s f comm 1 h, + exact one_pow _, +end + @[to_additive] lemma noncomm_prod_eq_prod {β : Type*} [comm_monoid β] (s : finset α) (f : α → β) : noncomm_prod s f (λ _ _ _ _, commute.all _ _) = s.prod f := begin @@ -244,4 +281,39 @@ begin list.nodup_append_of_nodup sl' tl' h] end +section finite_pi + +@[to_additive] +lemma noncomm_prod_single {ι : Type*} [fintype ι] [decidable_eq ι] + {M : ι → Type*} [∀ i, monoid (M i)] (x : Π i, M i) : + finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) + (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) + = x := +begin + ext i, + apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ + (pi.eval_monoid_hom _ i)).trans, + rw (finset.insert_erase (finset.mem_univ i)).symm, + rw finset.noncomm_prod_insert_of_not_mem', + rw finset.noncomm_prod_eq_one_of_forall_eq_one, + { simp, }, + { intros i h, simp at h, simp [h], }, + { simp, }, +end + +@[to_additive] +lemma _root_.monoid_hom.pi_ext {ι : Type*} [fintype ι] [decidable_eq ι] + {M : ι → Type*} [∀ i, monoid (M i)] {N : Type*} [monoid N] {f g : (Π i, M i) →* N} + (h : ∀ i x, f (monoid_hom.single M i x) = g (monoid_hom.single M i x)) : + f = g := +begin + ext x, + rw ← finset.noncomm_prod_single x, + rw finset.univ.noncomm_prod_map _ _ f, + rw finset.univ.noncomm_prod_map _ _ g, + congr' 1, ext i, exact h i (x i), +end + +end finite_pi + end finset diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 6e5da7765b1bb..8dd978044d981 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -26,6 +26,14 @@ begin { rw [list.repeat_succ, list.prod_cons, ih, pow_succ] } end +@[to_additive] +lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : + l.prod = m ^ l.length := +begin + convert list.prod_repeat m l.length, + exact list.eq_repeat.mpr ⟨rfl, h⟩, +end + @[to_additive] lemma prod_le_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 2bdc6518ec993..6be4838990121 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -71,81 +71,6 @@ variables {M : Type*} [monoid M] variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] -@[to_additive] -lemma finset.noncomm_prod_map {α : Type*} {β : Type*} [monoid β] - (s : finset α) (f : α → β) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - {M : Type*} [monoid M] (g : β →* M) : - g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) - (λ x hx y hy, commute.map (comm x hx y hy) g) := -begin - classical, - revert comm, rw ← s.to_list_to_finset, intro, - iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, - convert g.map_list_prod _, rw list.comp_map, -end - -@[to_additive] -lemma list.prod_const {α : Type*} [monoid α] (l : list α) - (m : α) (h : ∀ (x : α), x ∈ l → x = m) : l.prod = m ^ l.length := -begin - convert list.prod_repeat m l.length, - exact list.eq_repeat.mpr ⟨rfl, h⟩, -end - -@[to_additive] -lemma finset.noncomm_prod_const {α : Type*} {β : Type*} [monoid β] (s : finset α) - (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := -begin - classical, - revert comm, rw ← s.to_list_to_finset, intro, - rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, - revert h, simp_rw ← s.mem_to_list, intro, - rw list.prod_const _ m, - { simp, }, - { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x hx }, -end - -@[to_additive] -lemma finset.noncomm_prod_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) - (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := -begin - rw finset.noncomm_prod_const s f comm 1 h, - exact one_pow _, -end - -@[to_additive] -lemma finset.noncomm_prod_single [decidable_eq ι] [fintype ι] (x : Π i, N i) : - finset.univ.noncomm_prod (λ i, monoid_hom.single N i (x i)) - (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) - = x := -begin - ext i, - apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single N i (x i)) _ - (pi.eval_monoid_hom _ i)).trans, - rw (finset.insert_erase (finset.mem_univ i)).symm, - rw finset.noncomm_prod_insert_of_not_mem', - rw finset.noncomm_prod_eq_one, - { simp, }, - { intros i h, simp at h, simp [h], }, - { simp, }, -end - -@[to_additive] -lemma monoid_hom.pi_ext [decidable_eq ι] [fintype ι] {f g : (Π i, N i) →* M} - (h : ∀ i x, f (monoid_hom.single N i x) = g (monoid_hom.single N i x)) : - f = g := -begin - ext x, - rw ← finset.noncomm_prod_single x, - rw finset.univ.noncomm_prod_map _ _ f, - rw finset.univ.noncomm_prod_map _ _ g, - congr' 1, ext i, exact h i (x i), -end - - -- And morphisms ϕ into G variables (ϕ : Π (i : ι), N i →* M) From 4216e43e1b937dd425e26b7a764435514e26eb17 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 23:04:40 +0100 Subject: [PATCH 053/116] feat(data/list/prod_monoid): finite pi lemmas including a few helpers --- src/algebra/group/pi.lean | 16 +++++++ src/data/finset/noncomm_prod.lean | 72 +++++++++++++++++++++++++++++++ src/data/list/prod_monoid.lean | 8 ++++ 3 files changed, 96 insertions(+) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 230247539b242..dc310884228fc 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -5,6 +5,7 @@ Authors: Simon Hudon, Patrick Massot -/ import data.pi import data.set.function +import data.set.pairwise import tactic.pi_instances import algebra.group.hom_instances @@ -235,6 +236,21 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : single i (x * y) = single i x * single i y := (mul_hom.single f i).map_mul x y +@[to_additive] +lemma pi.mul_single_commute [Π i, mul_one_class $ f i] : + pairwise (λ i j, ∀ (x : f i) (y : f j), commute (mul_single i x) (mul_single j y)) := +begin + intros i j hij x y, ext k, + by_cases h1 : i = k, { subst h1, simp [hij], }, + by_cases h2 : j = k, { subst h2, simp [hij], }, + simp [h1, h2], +end + +@[to_additive] +lemma monoid_hom.single_commute [Π i, mul_one_class $ f i] : + pairwise (λ i j, ∀ x y, commute (monoid_hom.single f i x) (monoid_hom.single f j y)) := +pi.mul_single_commute + @[to_additive update_eq_sub_add_single] lemma pi.update_eq_div_mul_single [Π i, group $ f i] (g : Π (i : I), f i) (x : f i) : function.update g i x = g / mul_single i (g i) * mul_single i x := diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index be0b2177c6e3c..e8d19347d595d 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -217,6 +217,43 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := by simp [noncomm_prod, multiset.singleton_eq_cons] + +@[to_additive] +lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] + (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + {M : Type*} [monoid M] (g : β →* M) : + g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) + (λ x hx y hy, commute.map (comm x hx y hy) g) := +begin + classical, + revert comm, rw ← s.to_list_to_finset, intro, + iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, + convert g.map_list_prod _, rw list.comp_map, +end + +@[to_additive] +lemma noncomm_prod_eq_pow_of_forall_eq {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := +begin + classical, + revert comm, rw ← s.to_list_to_finset, intro, + rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, + rw list.prod_eq_pow_of_forall_eq _ m, + { simp }, + { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x (s.mem_to_list.mp hx) }, +end + +@[to_additive] +lemma noncomm_prod_eq_one_of_forall_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) + (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := +begin + rw finset.noncomm_prod_eq_pow_of_forall_eq s f comm 1 h, + exact one_pow _, +end + @[to_additive] lemma noncomm_prod_eq_prod {β : Type*} [comm_monoid β] (s : finset α) (f : α → β) : noncomm_prod s f (λ _ _ _ _, commute.all _ _) = s.prod f := begin @@ -244,4 +281,39 @@ begin list.nodup_append_of_nodup sl' tl' h] end +section finite_pi + +@[to_additive] +lemma noncomm_prod_single {ι : Type*} [fintype ι] [decidable_eq ι] + {M : ι → Type*} [∀ i, monoid (M i)] (x : Π i, M i) : + finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) + (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) + = x := +begin + ext i, + apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ + (pi.eval_monoid_hom _ i)).trans, + rw (finset.insert_erase (finset.mem_univ i)).symm, + rw finset.noncomm_prod_insert_of_not_mem', + rw finset.noncomm_prod_eq_one_of_forall_eq_one, + { simp, }, + { intros i h, simp at h, simp [h], }, + { simp, }, +end + +@[to_additive] +lemma _root_.monoid_hom.pi_ext {ι : Type*} [fintype ι] [decidable_eq ι] + {M : ι → Type*} [∀ i, monoid (M i)] {N : Type*} [monoid N] {f g : (Π i, M i) →* N} + (h : ∀ i x, f (monoid_hom.single M i x) = g (monoid_hom.single M i x)) : + f = g := +begin + ext x, + rw ← finset.noncomm_prod_single x, + rw finset.univ.noncomm_prod_map _ _ f, + rw finset.univ.noncomm_prod_map _ _ g, + congr' 1, ext i, exact h i (x i), +end + +end finite_pi + end finset diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 6e5da7765b1bb..8dd978044d981 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -26,6 +26,14 @@ begin { rw [list.repeat_succ, list.prod_cons, ih, pow_succ] } end +@[to_additive] +lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : + l.prod = m ^ l.length := +begin + convert list.prod_repeat m l.length, + exact list.eq_repeat.mpr ⟨rfl, h⟩, +end + @[to_additive] lemma prod_le_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := From 53f43973250ec2a813d8f07f1f178fae40f0fc58 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 25 Feb 2022 23:10:07 +0100 Subject: [PATCH 054/116] Update src/data/finset/noncomm_prod.lean --- src/data/finset/noncomm_prod.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index e8d19347d595d..76bd6cf423eee 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -217,7 +217,6 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := by simp [noncomm_prod, multiset.singleton_eq_cons] - @[to_additive] lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) From 6930f46be493b6092c9b30c6822328b81abfbf36 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 16:01:05 +0100 Subject: [PATCH 055/116] =?UTF-8?q?Don=E2=80=99=20use=20monoid=5Fhom.singl?= =?UTF-8?q?e=5Fcommute?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/group_theory/noncomm_pi_coprod.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 6be4838990121..d461ff8c4d57e 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -238,7 +238,7 @@ def noncomm_pi_coprod_equiv : { to_fun := λ ϕ, noncomm_pi_coprod ϕ.1 ϕ.2, inv_fun := λ f, ⟨ λ i, f.comp (monoid_hom.single N i), - λ i j hij x y, commute.map (monoid_hom.single_commute i j hij x y) f ⟩, + λ i j hij x y, commute.map (pi.mul_single_commute i j hij x y) f ⟩, left_inv := λ ϕ, by { ext, simp }, right_inv := λ f, pi_ext (λ i x, by simp) } From c7a6873bc2ad185756fd103d4529d2cd3fafb110 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 15:54:02 +0100 Subject: [PATCH 056/116] Rename to prod_le_pow_of_forall_le --- src/algebra/big_operators/multiset.lean | 4 ++-- src/algebra/big_operators/order.lean | 6 +++--- src/data/list/prod_monoid.lean | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index ad0a50629b95e..2d4155d787c28 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -267,10 +267,10 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := +lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, - simpa using list.prod_le_of_forall_le _ _ h, + simpa using list.prod_le_pow_of_forall_le _ _ h, end @[to_additive all_zero_of_le_zero_le_of_sum_eq_zero] diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 53c55b1661240..ddebee0a06095 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -170,10 +170,10 @@ calc f a = ∏ i in {a}, f i : prod_singleton.symm prod_le_prod_of_subset_of_one_le' (singleton_subset_iff.2 h) $ λ i hi _, hf i hi @[to_additive] -lemma prod_le_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : +lemma prod_le_pow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : s.prod f ≤ n ^ s.card := begin - refine (multiset.prod_le_of_forall_le (s.val.map f) n _).trans _, + refine (multiset.prod_le_pow_of_forall_le (s.val.map f) n _).trans _, { simpa using h }, { simpa } end @@ -181,7 +181,7 @@ end @[to_additive] lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_of_forall_le _ (order_dual N) _ _ _ _ h +@finset.prod_le_pow_of_forall_le _ (order_dual N) _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 8dd978044d981..1bcb1b7c957c7 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -35,7 +35,7 @@ begin end @[to_additive] -lemma prod_le_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : +lemma prod_le_pow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := begin induction l with y l IH, From ab6bf7d716d626cc7e0b6e757be18e069dd0e89d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 15:59:27 +0100 Subject: [PATCH 057/116] Docstring --- src/algebra/group/pi.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index dc310884228fc..9d92ca34bdee6 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -236,6 +236,10 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : single i (x * y) = single i x * single i y := (mul_hom.single f i).map_mul x y +/-- The injection into a pi group at different indices commutes. + +For injections of commmuting elements at the same index, see `commute.map` +-/ @[to_additive] lemma pi.mul_single_commute [Π i, mul_one_class $ f i] : pairwise (λ i j, ∀ (x : f i) (y : f j), commute (mul_single i x) (mul_single j y)) := From b0b717978bb044467f3e77c70f8917f95f570b9b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 16:04:56 +0100 Subject: [PATCH 058/116] Remove monoid_hom.single_commute --- src/algebra/group/pi.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 9d92ca34bdee6..b4533b9e45fd6 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -250,11 +250,6 @@ begin simp [h1, h2], end -@[to_additive] -lemma monoid_hom.single_commute [Π i, mul_one_class $ f i] : - pairwise (λ i j, ∀ x y, commute (monoid_hom.single f i x) (monoid_hom.single f j y)) := -pi.mul_single_commute - @[to_additive update_eq_sub_add_single] lemma pi.update_eq_div_mul_single [Π i, group $ f i] (g : Π (i : I), f i) (x : f i) : function.update g i x = g / mul_single i (g i) * mul_single i x := From 817f5569bf0152cb8df614012f6031b433c0044e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 16:05:05 +0100 Subject: [PATCH 059/116] Generalize noncomm_prod_map and its dependencies --- src/algebra/big_operators/basic.lean | 4 ++-- src/data/finset/noncomm_prod.lean | 15 ++++++++------- src/data/list/big_operators.lean | 10 +++++----- 3 files changed, 15 insertions(+), 14 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 65e28a5c943ac..b14bf66a121a3 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -119,12 +119,12 @@ g.to_monoid_hom.map_prod f s lemma ring_hom.map_list_prod [semiring β] [semiring γ] (f : β →+* γ) (l : list β) : f l.prod = (l.map f).prod := -f.to_monoid_hom.map_list_prod l +monoid_hom.map_list_prod f l lemma ring_hom.map_list_sum [non_assoc_semiring β] [non_assoc_semiring γ] (f : β →+* γ) (l : list β) : f l.sum = (l.map f).sum := -f.to_add_monoid_hom.map_list_sum l +add_monoid_hom.map_list_sum f l /-- A morphism into the opposite ring acts on the product by acting on the reversed elements -/ lemma ring_hom.unop_map_list_prod [semiring β] [semiring γ] (f : β →+* γᵐᵒᵖ) (l : list β) : diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 76bd6cf423eee..ee5b2b9b17c4e 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -221,14 +221,14 @@ by simp [noncomm_prod, multiset.singleton_eq_cons] lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - {M : Type*} [monoid M] (g : β →* M) : + {M : Type*} [monoid M] {F : Type*} [monoid_hom_class F β M] (g : F) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, commute.map (comm x hx y hy) g) := begin classical, revert comm, rw ← s.to_list_to_finset, intro, iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, - convert g.map_list_prod _, rw list.comp_map, + convert monoid_hom.map_list_prod g _, rw list.comp_map, end @[to_additive] @@ -290,13 +290,14 @@ lemma noncomm_prod_single {ι : Type*} [fintype ι] [decidable_eq ι] = x := begin ext i, - apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ - (pi.eval_monoid_hom _ i)).trans, + change pi.eval_monoid_hom _ i (finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) _) + = pi.eval_monoid_hom _ i x, + rw (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ (pi.eval_monoid_hom _ i)), rw (finset.insert_erase (finset.mem_univ i)).symm, rw finset.noncomm_prod_insert_of_not_mem', - rw finset.noncomm_prod_eq_one_of_forall_eq_one, - { simp, }, - { intros i h, simp at h, simp [h], }, + { rw finset.noncomm_prod_eq_one_of_forall_eq_one, + { simp, }, + { intros i h, simp at h, simp [h], } }, { simp, }, end diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index bc1558ff0e1ca..95e057532960e 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -52,10 +52,10 @@ lemma prod_hom_rel (l : list ι) {r : M → N → Prop} {f : ι → M} {g : ι list.rec_on l h₁ (λ a l hl, by simp only [map_cons, prod_cons, h₂ hl]) @[to_additive] -lemma prod_hom (l : list M) (f : M →* N) : +lemma prod_hom (l : list M) {F : Type*} [monoid_hom_class F M N] (f : F) : (l.map f).prod = f l.prod := -by { simp only [prod, foldl_map, f.map_one.symm], - exact l.foldl_hom _ _ _ 1 f.map_mul } +by { simp only [prod, foldl_map, (map_one f).symm], + exact l.foldl_hom _ _ _ 1 (map_mul f) } @[to_additive] lemma prod_hom₂ (l : list ι) (f : M → N → P) @@ -448,13 +448,13 @@ namespace monoid_hom variables [monoid M] [monoid N] @[to_additive] -lemma map_list_prod (f : M →* N) (l : list M) : +lemma map_list_prod {F : Type*} [monoid_hom_class F M N] (f : F) (l : list M) : f l.prod = (l.map f).prod := (l.prod_hom f).symm /-- A morphism into the opposite monoid acts on the product by acting on the reversed elements -/ lemma unop_map_list_prod (f : M →* Nᵐᵒᵖ) (l : list M) : (f l.prod).unop = (l.map (mul_opposite.unop ∘ f)).reverse.prod := -by rw [f.map_list_prod l, mul_opposite.unop_list_prod, list.map_map] +by rw [monoid_hom.map_list_prod f l, mul_opposite.unop_list_prod, list.map_map] end monoid_hom From 5ae1527d24103420006bb10aa2d93efe155ad821 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 26 Feb 2022 16:37:28 +0100 Subject: [PATCH 060/116] Line length --- src/algebra/big_operators/multiset.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 2d4155d787c28..81d278ca0828d 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -267,7 +267,8 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := +lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : + s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, simpa using list.prod_le_pow_of_forall_le _ _ h, From 98fbc9141238942502996e73f4b7ce6fde0d73ac Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 12:12:17 +0100 Subject: [PATCH 061/116] Fix fallout --- src/algebra/big_operators/pi.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/big_operators/pi.lean b/src/algebra/big_operators/pi.lean index 84c064dc6c6b3..86d7e2cea6da8 100644 --- a/src/algebra/big_operators/pi.lean +++ b/src/algebra/big_operators/pi.lean @@ -20,7 +20,7 @@ namespace pi @[to_additive] lemma list_prod_apply {α : Type*} {β : α → Type*} [Πa, monoid (β a)] (a : α) (l : list (Πa, β a)) : l.prod a = (l.map (λf:Πa, β a, f a)).prod := -(eval_monoid_hom β a).map_list_prod _ +monoid_hom.map_list_prod (eval_monoid_hom β a) _ @[to_additive] lemma multiset_prod_apply {α : Type*} {β : α → Type*} [∀a, comm_monoid (β a)] (a : α) From 5eaff25ca05f3481087897cb049215b8ed7af8af Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 12:14:37 +0100 Subject: [PATCH 062/116] More juggling of variables --- src/group_theory/noncomm_pi_coprod.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index d461ff8c4d57e..2653362ff17f9 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -230,6 +230,8 @@ lemma noncomm_pi_coprod_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +omit hcomm + /-- The universal property of `noncomm_pi_coprod` -/ @[to_additive "The universal property of `noncomm_pi_coprod`"] def noncomm_pi_coprod_equiv : @@ -244,6 +246,8 @@ def noncomm_pi_coprod_equiv : omit hdec +include hcomm + @[to_additive] lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : ι, (ϕ i).mrange := begin From d153089e1470421b8e18d14baedcb3c544ca313a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 12:27:01 +0100 Subject: [PATCH 063/116] Fix more fallout --- src/algebra/big_operators/order.lean | 4 ++-- src/algebra/polynomial/big_operators.lean | 2 +- src/combinatorics/double_counting.lean | 2 +- src/linear_algebra/matrix/polynomial.lean | 2 +- 4 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index ddebee0a06095..604d9f70e2146 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -186,7 +186,7 @@ lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : (s.bUnion f).card ≤ s.card * n := -card_bUnion_le.trans $ sum_le_of_forall_le _ _ _ h +card_bUnion_le.trans $ sum_le_pow_of_forall_le _ _ _ h variables {ι' : Type*} [decidable_eq ι'] @@ -264,7 +264,7 @@ times how many they are. -/ lemma sum_card_inter_le (h : ∀ a ∈ s, (B.filter $ (∈) a).card ≤ n) : ∑ t in B, (s ∩ t).card ≤ s.card * n := begin - refine le_trans _ (s.sum_le_of_forall_le _ _ h), + refine le_trans _ (s.sum_le_pow_of_forall_le _ _ h), simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 5bc884eb57d7e..53617ed1f3952 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -96,7 +96,7 @@ begin have h : nat_degree tl.prod ≤ n * tl.length, { refine (nat_degree_list_prod_le _).trans _, rw [←tl.length_map nat_degree, mul_comm], - refine list.sum_le_of_forall_le _ _ _, + refine list.sum_le_pow_of_forall_le _ _ _, simpa using hl' }, have hdn : nat_degree hd ≤ n := hl _ (list.mem_cons_self _ _), rcases hdn.eq_or_lt with rfl|hdn', diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 8f2eca8468c3e..d7fae74ea1a6d 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -64,7 +64,7 @@ calc _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.le_sum_of_forall_le _ _ hm ... = ∑ b in t, (s.bipartite_below r b).card : sum_card_bipartite_above_eq_sum_card_bipartite_below _ - ... ≤ _ : t.sum_le_of_forall_le _ _ hn + ... ≤ _ : t.sum_le_pow_of_forall_le _ _ hn lemma card_mul_le_card_mul' [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, n ≤ (s.bipartite_below r b).card) diff --git a/src/linear_algebra/matrix/polynomial.lean b/src/linear_algebra/matrix/polynomial.lean index 8e265d50ae826..6dbc03727939e 100644 --- a/src/linear_algebra/matrix/polynomial.lean +++ b/src/linear_algebra/matrix/polynomial.lean @@ -49,7 +49,7 @@ begin { rw [sg, units.neg_smul, one_smul, nat_degree_neg] } } ... ≤ ∑ (i : n), nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) : nat_degree_prod_le (finset.univ : finset n) (λ (i : n), (X • A.map C + B.map C) (g i) i) - ... ≤ finset.univ.card • 1 : finset.sum_le_of_forall_le _ _ 1 (λ (i : n) _, _) + ... ≤ finset.univ.card • 1 : finset.sum_le_pow_of_forall_le _ _ 1 (λ (i : n) _, _) ... ≤ fintype.card n : by simpa, calc nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) = nat_degree ((X : α[X]) * C (A (g i) i) + C (B (g i) i)) : by simp From 19035bb4947f9c6591caf43f5da0962ec4d963c0 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 12:40:20 +0100 Subject: [PATCH 064/116] More monoid_hom_class fallout --- src/algebra/module/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 1a3108e29bd26..060bced7eac2f 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -154,7 +154,7 @@ lemma module.eq_zero_of_zero_eq_one (zero_eq_one : (0 : R) = 1) : x = 0 := by rw [←one_smul R x, ←zero_eq_one, zero_smul] lemma list.sum_smul {l : list R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := -((smul_add_hom R M).flip x).map_list_sum l +add_monoid_hom.map_list_sum ((smul_add_hom R M).flip x) l lemma multiset.sum_smul {l : multiset R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := ((smul_add_hom R M).flip x).map_multiset_sum l From 699653c570e4ee2c352cdbb46c887288b49aa55a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 12:51:42 +0100 Subject: [PATCH 065/116] More monoid_hom_class fallout --- src/group_theory/submonoid/membership.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index ce8d5892f5e75..29ef3d82bff84 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -46,7 +46,7 @@ S.subtype.map_pow x n @[simp, norm_cast, to_additive] theorem coe_list_prod (l : list S) : (l.prod : M) = (l.map coe).prod := -S.subtype.map_list_prod l +monoid_hom.map_list_prod S.subtype l @[simp, norm_cast, to_additive] theorem coe_multiset_prod {M} [comm_monoid M] (S : submonoid M) (m : multiset S) : (m.prod : M) = (m.map coe).prod := From 6f1b066c8b92e595d95069417570dfe17541a9c8 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 17:46:22 +0100 Subject: [PATCH 066/116] More monoid_hom_class fallout --- src/group_theory/group_action/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index 5d1998374dde3..448f60db44fd8 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -490,7 +490,7 @@ variables [monoid α] [add_monoid β] [distrib_mul_action α β] lemma list.smul_sum {r : α} {l : list β} : r • l.sum = (l.map ((•) r)).sum := -(distrib_mul_action.to_add_monoid_hom β r).map_list_sum l +add_monoid_hom.map_list_sum (distrib_mul_action.to_add_monoid_hom β r) l /-- `smul` by a `k : M` over a ring is injective, if `k` is not a zero divisor. The general theory of such `k` is elaborated by `is_smul_regular`. From 5981a056acf4fbe61f9d8cacf1938e7b619c8e28 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 27 Feb 2022 17:48:21 +0100 Subject: [PATCH 067/116] Undo changing `map_list_prod` to `monoid_hom_class` --- src/algebra/big_operators/basic.lean | 4 ++-- src/algebra/big_operators/pi.lean | 2 +- src/algebra/module/basic.lean | 2 +- src/data/finset/noncomm_prod.lean | 9 ++++----- src/data/list/big_operators.lean | 10 +++++----- src/group_theory/group_action/basic.lean | 2 +- src/group_theory/submonoid/membership.lean | 2 +- 7 files changed, 15 insertions(+), 16 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index b14bf66a121a3..65e28a5c943ac 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -119,12 +119,12 @@ g.to_monoid_hom.map_prod f s lemma ring_hom.map_list_prod [semiring β] [semiring γ] (f : β →+* γ) (l : list β) : f l.prod = (l.map f).prod := -monoid_hom.map_list_prod f l +f.to_monoid_hom.map_list_prod l lemma ring_hom.map_list_sum [non_assoc_semiring β] [non_assoc_semiring γ] (f : β →+* γ) (l : list β) : f l.sum = (l.map f).sum := -add_monoid_hom.map_list_sum f l +f.to_add_monoid_hom.map_list_sum l /-- A morphism into the opposite ring acts on the product by acting on the reversed elements -/ lemma ring_hom.unop_map_list_prod [semiring β] [semiring γ] (f : β →+* γᵐᵒᵖ) (l : list β) : diff --git a/src/algebra/big_operators/pi.lean b/src/algebra/big_operators/pi.lean index 86d7e2cea6da8..84c064dc6c6b3 100644 --- a/src/algebra/big_operators/pi.lean +++ b/src/algebra/big_operators/pi.lean @@ -20,7 +20,7 @@ namespace pi @[to_additive] lemma list_prod_apply {α : Type*} {β : α → Type*} [Πa, monoid (β a)] (a : α) (l : list (Πa, β a)) : l.prod a = (l.map (λf:Πa, β a, f a)).prod := -monoid_hom.map_list_prod (eval_monoid_hom β a) _ +(eval_monoid_hom β a).map_list_prod _ @[to_additive] lemma multiset_prod_apply {α : Type*} {β : α → Type*} [∀a, comm_monoid (β a)] (a : α) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 060bced7eac2f..1a3108e29bd26 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -154,7 +154,7 @@ lemma module.eq_zero_of_zero_eq_one (zero_eq_one : (0 : R) = 1) : x = 0 := by rw [←one_smul R x, ←zero_eq_one, zero_smul] lemma list.sum_smul {l : list R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := -add_monoid_hom.map_list_sum ((smul_add_hom R M).flip x) l +((smul_add_hom R M).flip x).map_list_sum l lemma multiset.sum_smul {l : multiset R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := ((smul_add_hom R M).flip x).map_multiset_sum l diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index ee5b2b9b17c4e..e19b4ceae2fb1 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -221,14 +221,14 @@ by simp [noncomm_prod, multiset.singleton_eq_cons] lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - {M : Type*} [monoid M] {F : Type*} [monoid_hom_class F β M] (g : F) : + {M : Type*} [monoid M] (g : β →* M) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, commute.map (comm x hx y hy) g) := begin classical, revert comm, rw ← s.to_list_to_finset, intro, iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, - convert monoid_hom.map_list_prod g _, rw list.comp_map, + convert g.map_list_prod _, rw list.comp_map, end @[to_additive] @@ -290,9 +290,8 @@ lemma noncomm_prod_single {ι : Type*} [fintype ι] [decidable_eq ι] = x := begin ext i, - change pi.eval_monoid_hom _ i (finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) _) - = pi.eval_monoid_hom _ i x, - rw (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ (pi.eval_monoid_hom _ i)), + apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ + (pi.eval_monoid_hom _ i)).trans, rw (finset.insert_erase (finset.mem_univ i)).symm, rw finset.noncomm_prod_insert_of_not_mem', { rw finset.noncomm_prod_eq_one_of_forall_eq_one, diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index 95e057532960e..bc1558ff0e1ca 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -52,10 +52,10 @@ lemma prod_hom_rel (l : list ι) {r : M → N → Prop} {f : ι → M} {g : ι list.rec_on l h₁ (λ a l hl, by simp only [map_cons, prod_cons, h₂ hl]) @[to_additive] -lemma prod_hom (l : list M) {F : Type*} [monoid_hom_class F M N] (f : F) : +lemma prod_hom (l : list M) (f : M →* N) : (l.map f).prod = f l.prod := -by { simp only [prod, foldl_map, (map_one f).symm], - exact l.foldl_hom _ _ _ 1 (map_mul f) } +by { simp only [prod, foldl_map, f.map_one.symm], + exact l.foldl_hom _ _ _ 1 f.map_mul } @[to_additive] lemma prod_hom₂ (l : list ι) (f : M → N → P) @@ -448,13 +448,13 @@ namespace monoid_hom variables [monoid M] [monoid N] @[to_additive] -lemma map_list_prod {F : Type*} [monoid_hom_class F M N] (f : F) (l : list M) : +lemma map_list_prod (f : M →* N) (l : list M) : f l.prod = (l.map f).prod := (l.prod_hom f).symm /-- A morphism into the opposite monoid acts on the product by acting on the reversed elements -/ lemma unop_map_list_prod (f : M →* Nᵐᵒᵖ) (l : list M) : (f l.prod).unop = (l.map (mul_opposite.unop ∘ f)).reverse.prod := -by rw [monoid_hom.map_list_prod f l, mul_opposite.unop_list_prod, list.map_map] +by rw [f.map_list_prod l, mul_opposite.unop_list_prod, list.map_map] end monoid_hom diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index 448f60db44fd8..5d1998374dde3 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -490,7 +490,7 @@ variables [monoid α] [add_monoid β] [distrib_mul_action α β] lemma list.smul_sum {r : α} {l : list β} : r • l.sum = (l.map ((•) r)).sum := -add_monoid_hom.map_list_sum (distrib_mul_action.to_add_monoid_hom β r) l +(distrib_mul_action.to_add_monoid_hom β r).map_list_sum l /-- `smul` by a `k : M` over a ring is injective, if `k` is not a zero divisor. The general theory of such `k` is elaborated by `is_smul_regular`. diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index 29ef3d82bff84..ce8d5892f5e75 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -46,7 +46,7 @@ S.subtype.map_pow x n @[simp, norm_cast, to_additive] theorem coe_list_prod (l : list S) : (l.prod : M) = (l.map coe).prod := -monoid_hom.map_list_prod S.subtype l +S.subtype.map_list_prod l @[simp, norm_cast, to_additive] theorem coe_multiset_prod {M} [comm_monoid M] (S : submonoid M) (m : multiset S) : (m.prod : M) = (m.map coe).prod := From 177a8d7e7c166696c0517d85085a19b767f0512c Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 3 Mar 2022 21:35:23 +0100 Subject: [PATCH 068/116] Add docstrings --- src/data/nat/gcd.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/data/nat/gcd.lean b/src/data/nat/gcd.lean index 45acf8d4d2dd7..7a1445a4f5bc3 100644 --- a/src/data/nat/gcd.lean +++ b/src/data/nat/gcd.lean @@ -471,11 +471,13 @@ section big_operators open_locale big_operators +/-- See `is_coprime.prod_left` for the corresponding lemma about `is_coprime` -/ lemma coprime_prod_left {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : (∀ (i : ι), i ∈ t → coprime (s i) x) → coprime (∏ (i : ι) in t, s i) x := finset.prod_induction s (λ y, y.coprime x) (λ a b, coprime.mul) (by simp) +/-- See `is_coprime.prod_right` for the corresponding lemma about `is_coprime` -/ lemma coprime_prod_right {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : (∀ (i : ι), i ∈ t → coprime x (s i)) → coprime x (∏ (i : ι) in t, s i) := From fa6ce2e6c2263ff93c657f86cd3ee286c37e127a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 21:58:54 +0100 Subject: [PATCH 069/116] Apply code review --- src/algebra/group/pi.lean | 17 ++++++++++++++--- src/data/finset/noncomm_prod.lean | 30 +++++++++++++++--------------- 2 files changed, 29 insertions(+), 18 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index b4533b9e45fd6..c7971f34763d2 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -238,9 +238,10 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : /-- The injection into a pi group at different indices commutes. -For injections of commmuting elements at the same index, see `commute.map` --/ -@[to_additive] +For injections of commuting elements at the same index, see `commute.map` -/ +@[to_additive "The injection into an pi additive group at different indices commutes. + +For injections of commuting elements at the same index, see `add_commute.map`"] lemma pi.mul_single_commute [Π i, mul_one_class $ f i] : pairwise (λ i j, ∀ (x : f i) (y : f j), commute (mul_single i x) (mul_single j y)) := begin @@ -250,6 +251,16 @@ begin simp [h1, h2], end +/-- The injection into a pi group with the same values commutes. -/ +@[to_additive "The injection into a pi additive group with the same values commutes."] +lemma pi.mul_single_apply_commute [Π i, mul_one_class $ f i] (x : Π i, f i) (i j : I) : + commute (mul_single i (x i)) (mul_single j (x j)) := +begin + obtain rfl | hij := decidable.eq_or_ne i j, + { refl }, + { exact pi.mul_single_commute _ _ hij _ _, }, +end + @[to_additive update_eq_sub_add_single] lemma pi.update_eq_div_mul_single [Π i, group $ f i] (g : Π (i : I), f i) (x : f i) : function.update g i x = g / mul_single i (g i) * mul_single i x := diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index e19b4ceae2fb1..4ed0f07111616 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -27,7 +27,7 @@ the two must be equal. -/ -variables {α β : Type*} (f : α → β → β) (op : α → α → α) +variables {α β γ : Type*} (f : α → β → β) (op : α → α → α) namespace multiset @@ -174,7 +174,7 @@ end multiset namespace finset -variables [monoid β] +variables [monoid β] [monoid γ] /-- Product of a `s : finset α` mapped with `f : α → β` with `[monoid β]`, given a proof that `*` commutes on all elements `f x` for `x ∈ s`. -/ @@ -218,22 +218,21 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] by simp [noncomm_prod, multiset.singleton_eq_cons] @[to_additive] -lemma noncomm_prod_map {α : Type*} {β : Type*} [monoid β] - (s : finset α) (f : α → β) +lemma noncomm_prod_map (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - {M : Type*} [monoid M] (g : β →* M) : + {F : Type*} [monoid_hom_class F β γ] (g : F) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, commute.map (comm x hx y hy) g) := begin classical, revert comm, rw ← s.to_list_to_finset, intro, iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, - convert g.map_list_prod _, rw list.comp_map, + convert map_list_prod g _, rw list.comp_map, end @[to_additive] -lemma noncomm_prod_eq_pow_of_forall_eq {α : Type*} {β : Type*} [monoid β] (s : finset α) - (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) +lemma noncomm_prod_eq_pow_of_forall_eq (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := begin classical, @@ -245,8 +244,8 @@ begin end @[to_additive] -lemma noncomm_prod_eq_one_of_forall_eq_one {α : Type*} {β : Type*} [monoid β] (s : finset α) - (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) +lemma noncomm_prod_eq_one_of_forall_eq_one (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := begin rw finset.noncomm_prod_eq_pow_of_forall_eq s f comm 1 h, @@ -282,16 +281,18 @@ end section finite_pi +variables {ι : Type*} [fintype ι] [decidable_eq ι] {M : ι → Type*} [∀ i, monoid (M i)] +variables (x : Π i, M i) + @[to_additive] -lemma noncomm_prod_single {ι : Type*} [fintype ι] [decidable_eq ι] - {M : ι → Type*} [∀ i, monoid (M i)] (x : Π i, M i) : +lemma noncomm_prod_single : finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) = x := begin ext i, apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ - (pi.eval_monoid_hom _ i)).trans, + (pi.eval_monoid_hom M i)).trans, rw (finset.insert_erase (finset.mem_univ i)).symm, rw finset.noncomm_prod_insert_of_not_mem', { rw finset.noncomm_prod_eq_one_of_forall_eq_one, @@ -301,8 +302,7 @@ begin end @[to_additive] -lemma _root_.monoid_hom.pi_ext {ι : Type*} [fintype ι] [decidable_eq ι] - {M : ι → Type*} [∀ i, monoid (M i)] {N : Type*} [monoid N] {f g : (Π i, M i) →* N} +lemma _root_.monoid_hom.pi_ext {f g : (Π i, M i) →* γ} (h : ∀ i x, f (monoid_hom.single M i x) = g (monoid_hom.single M i x)) : f = g := begin From 72e316840381e89548c9c7fb33248300172ec747 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:01:23 +0100 Subject: [PATCH 070/116] Apply more code review --- src/data/finset/noncomm_prod.lean | 19 ++++++++----------- 1 file changed, 8 insertions(+), 11 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 4ed0f07111616..8f074a005433d 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -286,19 +286,18 @@ variables (x : Π i, M i) @[to_additive] lemma noncomm_prod_single : - finset.univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) + univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) = x := begin ext i, - apply (finset.univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ + apply (univ.noncomm_prod_map (λ i, monoid_hom.single M i (x i)) _ (pi.eval_monoid_hom M i)).trans, - rw (finset.insert_erase (finset.mem_univ i)).symm, - rw finset.noncomm_prod_insert_of_not_mem', - { rw finset.noncomm_prod_eq_one_of_forall_eq_one, - { simp, }, - { intros i h, simp at h, simp [h], } }, + rw [ ← insert_erase (mem_univ i), + noncomm_prod_insert_of_not_mem' _ _ _ _ (not_mem_erase _ _), + noncomm_prod_eq_one_of_forall_eq_one ], { simp, }, + { intros i h, simp at h, simp [h], }, end @[to_additive] @@ -307,10 +306,8 @@ lemma _root_.monoid_hom.pi_ext {f g : (Π i, M i) →* γ} f = g := begin ext x, - rw ← finset.noncomm_prod_single x, - rw finset.univ.noncomm_prod_map _ _ f, - rw finset.univ.noncomm_prod_map _ _ g, - congr' 1, ext i, exact h i (x i), + rw [← noncomm_prod_single x, univ.noncomm_prod_map, univ.noncomm_prod_map], + congr' 1 with i, exact h i (x i), end end finite_pi From cfd84be8bbc0e21355df82fe0f3cb37d3bc01bbd Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:06:01 +0100 Subject: [PATCH 071/116] Phrase lemmas wrt. pi.mul_single --- src/data/finset/noncomm_prod.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 8f074a005433d..4e6664ae75eb3 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -285,8 +285,8 @@ variables {ι : Type*} [fintype ι] [decidable_eq ι] {M : ι → Type*} [∀ i, variables (x : Π i, M i) @[to_additive] -lemma noncomm_prod_single : - univ.noncomm_prod (λ i, monoid_hom.single M i (x i)) +lemma noncomm_prod_mul_single : + univ.noncomm_prod (λ i, pi.mul_single i (x i)) (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) = x := begin @@ -302,11 +302,11 @@ end @[to_additive] lemma _root_.monoid_hom.pi_ext {f g : (Π i, M i) →* γ} - (h : ∀ i x, f (monoid_hom.single M i x) = g (monoid_hom.single M i x)) : + (h : ∀ i x, f (pi.mul_single i x) = g (pi.mul_single i x)) : f = g := begin ext x, - rw [← noncomm_prod_single x, univ.noncomm_prod_map, univ.noncomm_prod_map], + rw [← noncomm_prod_mul_single x, univ.noncomm_prod_map, univ.noncomm_prod_map], congr' 1 with i, exact h i (x i), end From 137fb1bdc43c696ee9a0939a89871c266e349dd6 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:16:54 +0100 Subject: [PATCH 072/116] Fix apply-lemmas for monoid_hom.single --- src/algebra/group/pi.lean | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index c7971f34763d2..0177726d07160 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -190,11 +190,15 @@ This is the `one_hom` version of `pi.mul_single`. -/ @[to_additive zero_hom.single "The zero-preserving homomorphism including a single value into a dependent family of values, as functions supported at a point. -This is the `zero_hom` version of `pi.single`.", simps] +This is the `zero_hom` version of `pi.single`."] def one_hom.single [Π i, has_one $ f i] (i : I) : one_hom (f i) (Π i, f i) := { to_fun := mul_single i, map_one' := mul_single_one i } +@[to_additive, simp] +lemma one_hom.single_apply [Π i, has_one $ f i] (i : I) (x : f i) : + one_hom.single f i x = mul_single i x := rfl + /-- The monoid homomorphism including a single monoid into a dependent family of additive monoids, as functions supported at a point. @@ -202,11 +206,15 @@ This is the `monoid_hom` version of `pi.mul_single`. -/ @[to_additive "The additive monoid homomorphism including a single additive monoid into a dependent family of additive monoids, as functions supported at a point. -This is the `add_monoid_hom` version of `pi.single`.", simps] +This is the `add_monoid_hom` version of `pi.single`."] def monoid_hom.single [Π i, mul_one_class $ f i] (i : I) : f i →* Π i, f i := { map_mul' := mul_single_op₂ (λ _, (*)) (λ _, one_mul _) _, .. (one_hom.single f i) } +@[to_additive, simp] +lemma monoid_hom.single_apply [Π i, mul_one_class $ f i] (i : I) (x : f i) : + monoid_hom.single f i x = mul_single i x := rfl + /-- The multiplicative homomorphism including a single `mul_zero_class` into a dependent family of `mul_zero_class`es, as functions supported at a point. From 5d9945c8c2a7714103bea507288daed449d504f3 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:17:34 +0100 Subject: [PATCH 073/116] Stick closer to pi.mul_single --- src/group_theory/noncomm_pi_coprod.lean | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 2653362ff17f9..53b448c209d2c 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -176,8 +176,8 @@ variable {hcomm} include hdec @[to_additive add_to_fun_single] -lemma to_fun_single (i : ι) (y : N i) (S : finset ι) : - to_fun ϕ hcomm (monoid_hom.single _ i y) S = if i ∈ S then ϕ i y else 1 := +lemma to_fun_mul_single (i : ι) (y : N i) (S : finset ι) : + to_fun ϕ hcomm (pi.mul_single i y) S = if i ∈ S then ϕ i y else 1 := begin induction S using finset.induction_on with j S hnmem ih, { simp, }, @@ -189,8 +189,8 @@ begin end @[simp, to_additive add_hom_single] -lemma hom_single (i : ι) (y : N i): - hom ϕ hcomm S (monoid_hom.single _ i y) = if i ∈ S then ϕ i y else 1 := to_fun_single _ _ _ _ +lemma hom_mul_single (i : ι) (y : N i): + hom ϕ hcomm S (pi.mul_single i y) = if i ∈ S then ϕ i y else 1 := to_fun_mul_single _ _ _ _ omit hdec @@ -203,7 +203,7 @@ begin exact (to_fun_mem_bsupr_mrange ϕ f S), }, { refine (bsupr_le _), rintro i hmem x ⟨y, rfl⟩, - use (monoid_hom.single _ i y), + use (pi.mul_single i y), simp [hmem], } end @@ -226,9 +226,9 @@ variable {hcomm} include hdec @[simp, to_additive] -lemma noncomm_pi_coprod_single (i : ι) (y : N i): - noncomm_pi_coprod ϕ hcomm (monoid_hom.single _ i y) = ϕ i y := -by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (monoid_hom.single _ i y) = ϕ i y, simp } +lemma noncomm_pi_coprod_mul_single (i : ι) (y : N i): + noncomm_pi_coprod ϕ hcomm (pi.mul_single i y) = ϕ i y := +by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (pi.mul_single i y) = ϕ i y, simp } omit hcomm @@ -241,7 +241,7 @@ def noncomm_pi_coprod_equiv : inv_fun := λ f, ⟨ λ i, f.comp (monoid_hom.single N i), λ i j hij x y, commute.map (pi.mul_single_commute i j hij x y) f ⟩, - left_inv := λ ϕ, by { ext, simp }, + left_inv := λ ϕ, by { ext, simp, }, right_inv := λ f, pi_ext (λ i x, by simp) } omit hdec @@ -405,9 +405,9 @@ variable {hcomm} include hdec @[simp, to_additive] -lemma noncomm_pi_coprod_single (i : ι) (y : H i) : - noncomm_pi_coprod hcomm (monoid_hom.single _ i y) = y := -by apply monoid_hom.noncomm_pi_coprod_single +lemma noncomm_pi_coprod_mul_single (i : ι) (y : H i) : + noncomm_pi_coprod hcomm (pi.mul_single i y) = y := +by apply monoid_hom.noncomm_pi_coprod_mul_single omit hdec From 0b30fdce52a8a3b8fde6bb7f34ae913dd38deb4d Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:16:54 +0100 Subject: [PATCH 074/116] Fix apply-lemmas for monoid_hom.single --- src/algebra/group/pi.lean | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index c7971f34763d2..0177726d07160 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -190,11 +190,15 @@ This is the `one_hom` version of `pi.mul_single`. -/ @[to_additive zero_hom.single "The zero-preserving homomorphism including a single value into a dependent family of values, as functions supported at a point. -This is the `zero_hom` version of `pi.single`.", simps] +This is the `zero_hom` version of `pi.single`."] def one_hom.single [Π i, has_one $ f i] (i : I) : one_hom (f i) (Π i, f i) := { to_fun := mul_single i, map_one' := mul_single_one i } +@[to_additive, simp] +lemma one_hom.single_apply [Π i, has_one $ f i] (i : I) (x : f i) : + one_hom.single f i x = mul_single i x := rfl + /-- The monoid homomorphism including a single monoid into a dependent family of additive monoids, as functions supported at a point. @@ -202,11 +206,15 @@ This is the `monoid_hom` version of `pi.mul_single`. -/ @[to_additive "The additive monoid homomorphism including a single additive monoid into a dependent family of additive monoids, as functions supported at a point. -This is the `add_monoid_hom` version of `pi.single`.", simps] +This is the `add_monoid_hom` version of `pi.single`."] def monoid_hom.single [Π i, mul_one_class $ f i] (i : I) : f i →* Π i, f i := { map_mul' := mul_single_op₂ (λ _, (*)) (λ _, one_mul _) _, .. (one_hom.single f i) } +@[to_additive, simp] +lemma monoid_hom.single_apply [Π i, mul_one_class $ f i] (i : I) (x : f i) : + monoid_hom.single f i x = mul_single i x := rfl + /-- The multiplicative homomorphism including a single `mul_zero_class` into a dependent family of `mul_zero_class`es, as functions supported at a point. From 88710b850177118b446991cd6b4cb219098b2d07 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:23:16 +0100 Subject: [PATCH 075/116] Fix fallout --- src/group_theory/subgroup/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 098a590c78ed2..4c6be80609688 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1166,7 +1166,7 @@ end @[simp] lemma single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)} (i : η) (x : f i) : - monoid_hom.single f i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) := + pi.mul_single i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) := begin split, { intros h hi, simpa using h i hi, }, From 799a4ef92eee076766cd56067053762659fac4d0 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:26:18 +0100 Subject: [PATCH 076/116] Throw in some to_additive attributes --- src/group_theory/subgroup/basic.lean | 19 ++++++++++++------- 1 file changed, 12 insertions(+), 7 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 4c6be80609688..7d26a73069551 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1163,8 +1163,8 @@ begin { intros h x hx i hi, refine h i hi ⟨_, hx, rfl⟩, } end -@[simp] -lemma single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)} +@[simp, to_additive] +lemma mul_single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)} (i : η) (x : f i) : pi.mul_single i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) := begin @@ -1176,7 +1176,8 @@ begin { simp [heq, one_mem], }, } end -lemma pi_mem_of_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } +@[to_additive] +lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : x ∈ H := begin @@ -1203,19 +1204,23 @@ begin { apply h2, simp, } } end -lemma pi_mem_of_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) +@[to_additive] +lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := -pi_mem_of_single_mem_aux finset.univ x (by simp) (λ i _, h i) +pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) /-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ +@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the +additive groups."] lemma pi_le_iff [decidable_eq η] [fintype η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : pi univ H ≤ J ↔ (∀ i : η, map (monoid_hom.single f i) (H i) ≤ J) := begin split, { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, - { exact λ h x hx, pi_mem_of_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } + { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } end +@[to_additive] lemma pi_eq_bot_iff (H : Π i, subgroup (f i)) : pi set.univ H = ⊥ ↔ ∀ i, H i = ⊥ := begin @@ -1224,7 +1229,7 @@ begin split, { intros h i x hx, have : monoid_hom.single f i x = 1 := - h (monoid_hom.single f i x) ((single_mem_pi i x).mpr (λ _, hx)), + h (monoid_hom.single f i x) ((mul_single_mem_pi i x).mpr (λ _, hx)), simpa using congr_fun this i, }, { exact λ h x hx, funext (λ i, h _ _ (hx i trivial)), }, end From aa370488d3fb9619bcf57f262eb62d2987b99454 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:34:35 +0100 Subject: [PATCH 077/116] lint --- src/group_theory/subgroup/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 7d26a73069551..2a90cd1f480a0 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1205,8 +1205,8 @@ begin end @[to_additive] -lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) - (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := +lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i)} + (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) /-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ From 0f37afe758b477da3e83ef83dc3f4e8c7c22489e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 09:35:19 +0100 Subject: [PATCH 078/116] feat(data/list/prod_monoid): add prod_eq_pow_of_forall_eq and rename `prod_le_of_forall_le` to `prod_le_pow_of_forall_le` --- src/algebra/big_operators/multiset.lean | 5 +++-- src/algebra/big_operators/order.lean | 10 +++++----- src/algebra/polynomial/big_operators.lean | 2 +- src/combinatorics/double_counting.lean | 2 +- src/data/list/prod_monoid.lean | 10 +++++++++- src/linear_algebra/matrix/polynomial.lean | 2 +- 6 files changed, 20 insertions(+), 11 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 29d9a2886166e..cbe0b829ffc27 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -268,10 +268,11 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := +lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : + s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, - simpa using list.prod_le_of_forall_le _ _ h, + simpa using list.prod_le_pow_of_forall_le _ _ h, end @[to_additive all_zero_of_le_zero_le_of_sum_eq_zero] diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 53c55b1661240..604d9f70e2146 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -170,10 +170,10 @@ calc f a = ∏ i in {a}, f i : prod_singleton.symm prod_le_prod_of_subset_of_one_le' (singleton_subset_iff.2 h) $ λ i hi _, hf i hi @[to_additive] -lemma prod_le_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : +lemma prod_le_pow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : s.prod f ≤ n ^ s.card := begin - refine (multiset.prod_le_of_forall_le (s.val.map f) n _).trans _, + refine (multiset.prod_le_pow_of_forall_le (s.val.map f) n _).trans _, { simpa using h }, { simpa } end @@ -181,12 +181,12 @@ end @[to_additive] lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_of_forall_le _ (order_dual N) _ _ _ _ h +@finset.prod_le_pow_of_forall_le _ (order_dual N) _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : (s.bUnion f).card ≤ s.card * n := -card_bUnion_le.trans $ sum_le_of_forall_le _ _ _ h +card_bUnion_le.trans $ sum_le_pow_of_forall_le _ _ _ h variables {ι' : Type*} [decidable_eq ι'] @@ -264,7 +264,7 @@ times how many they are. -/ lemma sum_card_inter_le (h : ∀ a ∈ s, (B.filter $ (∈) a).card ≤ n) : ∑ t in B, (s ∩ t).card ≤ s.card * n := begin - refine le_trans _ (s.sum_le_of_forall_le _ _ h), + refine le_trans _ (s.sum_le_pow_of_forall_le _ _ h), simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 5eb150eb43f56..4d1339d43479e 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -96,7 +96,7 @@ begin have h : nat_degree tl.prod ≤ n * tl.length, { refine (nat_degree_list_prod_le _).trans _, rw [←tl.length_map nat_degree, mul_comm], - refine list.sum_le_of_forall_le _ _ _, + refine list.sum_le_pow_of_forall_le _ _ _, simpa using hl' }, have hdn : nat_degree hd ≤ n := hl _ (list.mem_cons_self _ _), rcases hdn.eq_or_lt with rfl|hdn', diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 8f2eca8468c3e..d7fae74ea1a6d 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -64,7 +64,7 @@ calc _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.le_sum_of_forall_le _ _ hm ... = ∑ b in t, (s.bipartite_below r b).card : sum_card_bipartite_above_eq_sum_card_bipartite_below _ - ... ≤ _ : t.sum_le_of_forall_le _ _ hn + ... ≤ _ : t.sum_le_pow_of_forall_le _ _ hn lemma card_mul_le_card_mul' [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, n ≤ (s.bipartite_below r b).card) diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 6e5da7765b1bb..1bcb1b7c957c7 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -27,7 +27,15 @@ begin end @[to_additive] -lemma prod_le_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : +lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : + l.prod = m ^ l.length := +begin + convert list.prod_repeat m l.length, + exact list.eq_repeat.mpr ⟨rfl, h⟩, +end + +@[to_additive] +lemma prod_le_pow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := begin induction l with y l IH, diff --git a/src/linear_algebra/matrix/polynomial.lean b/src/linear_algebra/matrix/polynomial.lean index 8e265d50ae826..6dbc03727939e 100644 --- a/src/linear_algebra/matrix/polynomial.lean +++ b/src/linear_algebra/matrix/polynomial.lean @@ -49,7 +49,7 @@ begin { rw [sg, units.neg_smul, one_smul, nat_degree_neg] } } ... ≤ ∑ (i : n), nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) : nat_degree_prod_le (finset.univ : finset n) (λ (i : n), (X • A.map C + B.map C) (g i) i) - ... ≤ finset.univ.card • 1 : finset.sum_le_of_forall_le _ _ 1 (λ (i : n) _, _) + ... ≤ finset.univ.card • 1 : finset.sum_le_pow_of_forall_le _ _ 1 (λ (i : n) _, _) ... ≤ fintype.card n : by simpa, calc nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) = nat_degree ((X : α[X]) * C (A (g i) i) + C (B (g i) i)) : by simp From 47542b668b9f57887909e3ea28481979212d27b0 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 09:40:19 +0100 Subject: [PATCH 079/116] fix(src/algebra/group/pi): Fix apply-simp-lemmas for monoid_hom.single so that the simp-normal form really is `pi.mul_single`. While adjusting related lemmas in `group_theory.subgroup.basic`, add a few missing `to_additive` attributes. --- src/algebra/group/pi.lean | 12 ++++++++++-- src/group_theory/subgroup/basic.lean | 21 +++++++++++++-------- 2 files changed, 23 insertions(+), 10 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 230247539b242..2e5c1d2c7628b 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -189,11 +189,15 @@ This is the `one_hom` version of `pi.mul_single`. -/ @[to_additive zero_hom.single "The zero-preserving homomorphism including a single value into a dependent family of values, as functions supported at a point. -This is the `zero_hom` version of `pi.single`.", simps] +This is the `zero_hom` version of `pi.single`."] def one_hom.single [Π i, has_one $ f i] (i : I) : one_hom (f i) (Π i, f i) := { to_fun := mul_single i, map_one' := mul_single_one i } +@[to_additive, simp] +lemma one_hom.single_apply [Π i, has_one $ f i] (i : I) (x : f i) : + one_hom.single f i x = mul_single i x := rfl + /-- The monoid homomorphism including a single monoid into a dependent family of additive monoids, as functions supported at a point. @@ -201,11 +205,15 @@ This is the `monoid_hom` version of `pi.mul_single`. -/ @[to_additive "The additive monoid homomorphism including a single additive monoid into a dependent family of additive monoids, as functions supported at a point. -This is the `add_monoid_hom` version of `pi.single`.", simps] +This is the `add_monoid_hom` version of `pi.single`."] def monoid_hom.single [Π i, mul_one_class $ f i] (i : I) : f i →* Π i, f i := { map_mul' := mul_single_op₂ (λ _, (*)) (λ _, one_mul _) _, .. (one_hom.single f i) } +@[to_additive, simp] +lemma monoid_hom.single_apply [Π i, mul_one_class $ f i] (i : I) (x : f i) : + monoid_hom.single f i x = mul_single i x := rfl + /-- The multiplicative homomorphism including a single `mul_zero_class` into a dependent family of `mul_zero_class`es, as functions supported at a point. diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 098a590c78ed2..7d26a73069551 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1163,10 +1163,10 @@ begin { intros h x hx i hi, refine h i hi ⟨_, hx, rfl⟩, } end -@[simp] -lemma single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)} +@[simp, to_additive] +lemma mul_single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)} (i : η) (x : f i) : - monoid_hom.single f i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) := + pi.mul_single i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) := begin split, { intros h hi, simpa using h i hi, }, @@ -1176,7 +1176,8 @@ begin { simp [heq, one_mem], }, } end -lemma pi_mem_of_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } +@[to_additive] +lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : x ∈ H := begin @@ -1203,19 +1204,23 @@ begin { apply h2, simp, } } end -lemma pi_mem_of_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) +@[to_additive] +lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := -pi_mem_of_single_mem_aux finset.univ x (by simp) (λ i _, h i) +pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) /-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ +@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the +additive groups."] lemma pi_le_iff [decidable_eq η] [fintype η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : pi univ H ≤ J ↔ (∀ i : η, map (monoid_hom.single f i) (H i) ≤ J) := begin split, { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, - { exact λ h x hx, pi_mem_of_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } + { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } end +@[to_additive] lemma pi_eq_bot_iff (H : Π i, subgroup (f i)) : pi set.univ H = ⊥ ↔ ∀ i, H i = ⊥ := begin @@ -1224,7 +1229,7 @@ begin split, { intros h i x hx, have : monoid_hom.single f i x = 1 := - h (monoid_hom.single f i x) ((single_mem_pi i x).mpr (λ _, hx)), + h (monoid_hom.single f i x) ((mul_single_mem_pi i x).mpr (λ _, hx)), simpa using congr_fun this i, }, { exact λ h x hx, funext (λ i, h _ _ (hx i trivial)), }, end From cbad4370de4db4442820a2abee60e9dbb87168aa Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 5 Mar 2022 22:34:35 +0100 Subject: [PATCH 080/116] lint --- src/group_theory/subgroup/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 7d26a73069551..2a90cd1f480a0 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1205,8 +1205,8 @@ begin end @[to_additive] -lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i) - (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := +lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i)} + (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) /-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ From 2905a601cb11eff994f7da6c82565c3acbfbc9ff Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 11:42:26 +0100 Subject: [PATCH 081/116] Use npow, not pow, in name, so that to_additive does the right thing --- src/algebra/big_operators/multiset.lean | 4 ++-- src/algebra/big_operators/order.lean | 10 +++++----- src/algebra/polynomial/big_operators.lean | 2 +- src/combinatorics/double_counting.lean | 2 +- src/data/list/prod_monoid.lean | 4 ++-- src/linear_algebra/matrix/polynomial.lean | 2 +- 6 files changed, 12 insertions(+), 12 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index cbe0b829ffc27..46f0e10602e1f 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -268,11 +268,11 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : +lemma prod_le_npow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, - simpa using list.prod_le_pow_of_forall_le _ _ h, + simpa using list.prod_le_npow_of_forall_le _ _ h, end @[to_additive all_zero_of_le_zero_le_of_sum_eq_zero] diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 604d9f70e2146..68255b72edf35 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -170,10 +170,10 @@ calc f a = ∏ i in {a}, f i : prod_singleton.symm prod_le_prod_of_subset_of_one_le' (singleton_subset_iff.2 h) $ λ i hi _, hf i hi @[to_additive] -lemma prod_le_pow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : +lemma prod_le_npow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : s.prod f ≤ n ^ s.card := begin - refine (multiset.prod_le_pow_of_forall_le (s.val.map f) n _).trans _, + refine (multiset.prod_le_npow_of_forall_le (s.val.map f) n _).trans _, { simpa using h }, { simpa } end @@ -181,12 +181,12 @@ end @[to_additive] lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_pow_of_forall_le _ (order_dual N) _ _ _ _ h +@finset.prod_le_npow_of_forall_le _ (order_dual N) _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : (s.bUnion f).card ≤ s.card * n := -card_bUnion_le.trans $ sum_le_pow_of_forall_le _ _ _ h +card_bUnion_le.trans $ sum_le_nsmul_of_forall_le _ _ _ h variables {ι' : Type*} [decidable_eq ι'] @@ -264,7 +264,7 @@ times how many they are. -/ lemma sum_card_inter_le (h : ∀ a ∈ s, (B.filter $ (∈) a).card ≤ n) : ∑ t in B, (s ∩ t).card ≤ s.card * n := begin - refine le_trans _ (s.sum_le_pow_of_forall_le _ _ h), + refine le_trans _ (s.sum_le_nsmul_of_forall_le _ _ h), simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 4d1339d43479e..10e858e0911ea 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -96,7 +96,7 @@ begin have h : nat_degree tl.prod ≤ n * tl.length, { refine (nat_degree_list_prod_le _).trans _, rw [←tl.length_map nat_degree, mul_comm], - refine list.sum_le_pow_of_forall_le _ _ _, + refine list.sum_le_nsmul_of_forall_le _ _ _, simpa using hl' }, have hdn : nat_degree hd ≤ n := hl _ (list.mem_cons_self _ _), rcases hdn.eq_or_lt with rfl|hdn', diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index d7fae74ea1a6d..3f9c30fc5d4e5 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -64,7 +64,7 @@ calc _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.le_sum_of_forall_le _ _ hm ... = ∑ b in t, (s.bipartite_below r b).card : sum_card_bipartite_above_eq_sum_card_bipartite_below _ - ... ≤ _ : t.sum_le_pow_of_forall_le _ _ hn + ... ≤ _ : t.sum_le_nsmul_of_forall_le _ _ hn lemma card_mul_le_card_mul' [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, n ≤ (s.bipartite_below r b).card) diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 1bcb1b7c957c7..a86b4fabca2f5 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -27,7 +27,7 @@ begin end @[to_additive] -lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : +lemma prod_eq_npow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : l.prod = m ^ l.length := begin convert list.prod_repeat m l.length, @@ -35,7 +35,7 @@ begin end @[to_additive] -lemma prod_le_pow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : +lemma prod_le_npow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := begin induction l with y l IH, diff --git a/src/linear_algebra/matrix/polynomial.lean b/src/linear_algebra/matrix/polynomial.lean index 6dbc03727939e..9341b6e3ecf5a 100644 --- a/src/linear_algebra/matrix/polynomial.lean +++ b/src/linear_algebra/matrix/polynomial.lean @@ -49,7 +49,7 @@ begin { rw [sg, units.neg_smul, one_smul, nat_degree_neg] } } ... ≤ ∑ (i : n), nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) : nat_degree_prod_le (finset.univ : finset n) (λ (i : n), (X • A.map C + B.map C) (g i) i) - ... ≤ finset.univ.card • 1 : finset.sum_le_pow_of_forall_le _ _ 1 (λ (i : n) _, _) + ... ≤ finset.univ.card • 1 : finset.sum_le_nsmul_of_forall_le _ _ 1 (λ (i : n) _, _) ... ≤ fintype.card n : by simpa, calc nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) = nat_degree ((X : α[X]) * C (A (g i) i) + C (B (g i) i)) : by simp From 73a9e24d88a8a1f7f2319778033bd8b378249507 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 11:45:28 +0100 Subject: [PATCH 082/116] Rename le_prod_of_forall_le to npow_le_prod_of_forall_le --- src/algebra/big_operators/order.lean | 4 ++-- src/combinatorics/double_counting.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 68255b72edf35..9b77f0df9c3bb 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -179,7 +179,7 @@ begin end @[to_additive] -lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : +lemma npow_le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := @finset.prod_le_npow_of_forall_le _ (order_dual N) _ _ _ _ h @@ -281,7 +281,7 @@ times how many they are. -/ lemma le_sum_card_inter (h : ∀ a ∈ s, n ≤ (B.filter $ (∈) a).card) : s.card * n ≤ ∑ t in B, (s ∩ t).card := begin - apply (s.le_sum_of_forall_le _ _ h).trans, + apply (s.nsmul_le_sum_of_forall_le _ _ h).trans, simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 3f9c30fc5d4e5..55cf33ade699a 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -61,7 +61,7 @@ lemma card_mul_le_card_mul [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, (s.bipartite_below r b).card ≤ n) : s.card * m ≤ t.card * n := calc - _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.le_sum_of_forall_le _ _ hm + _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.nsmul_le_sum_of_forall_le _ _ hm ... = ∑ b in t, (s.bipartite_below r b).card : sum_card_bipartite_above_eq_sum_card_bipartite_below _ ... ≤ _ : t.sum_le_nsmul_of_forall_le _ _ hn From c242aaa74596fe173fb48bc8166597071f5a129f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 11:50:49 +0100 Subject: [PATCH 083/116] =?UTF-8?q?Use=20`=E2=88=80=20(x=20=E2=88=88=20l),?= =?UTF-8?q?`=20syntax?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/data/list/prod_monoid.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index a86b4fabca2f5..d1f100a0771a4 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -27,7 +27,7 @@ begin end @[to_additive] -lemma prod_eq_npow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x : α), x ∈ l → x = m) : +lemma prod_eq_npow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x ∈ l), x = m) : l.prod = m ^ l.length := begin convert list.prod_repeat m l.length, From 95b0ec9050dc568afdf464e8883a3c8523e1be31 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:08:17 +0100 Subject: [PATCH 084/116] Try to teach to_additive about `pow` naming convention --- src/algebra/group/to_additive.lean | 1 + src/algebra/group_power/basic.lean | 14 +++++++------- src/algebra/group_power/order.lean | 6 +++--- src/algebra/iterate_hom.lean | 2 +- 4 files changed, 12 insertions(+), 11 deletions(-) diff --git a/src/algebra/group/to_additive.lean b/src/algebra/group/to_additive.lean index 4c80f6719729d..eb23929d0fe1f 100644 --- a/src/algebra/group/to_additive.lean +++ b/src/algebra/group/to_additive.lean @@ -213,6 +213,7 @@ meta def tr : bool → list string → list string | is_comm ("one" :: s) := add_comm_prefix is_comm "zero" :: tr ff s | is_comm ("prod" :: s) := add_comm_prefix is_comm "sum" :: tr ff s | is_comm ("finprod" :: s) := add_comm_prefix is_comm "finsum" :: tr ff s +| is_comm ("pow" :: s) := add_comm_prefix is_comm "nsmul" :: tr ff s | is_comm ("npow" :: s) := add_comm_prefix is_comm "nsmul" :: tr ff s | is_comm ("zpow" :: s) := add_comm_prefix is_comm "zsmul" :: tr ff s | is_comm ("monoid" :: s) := ("add_" ++ add_comm_prefix is_comm "monoid") :: tr ff s diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 80ef6e1995db3..0845a981c61cf 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -55,10 +55,10 @@ by rw [pow_succ, pow_one] alias pow_two ← sq -@[to_additive nsmul_add_comm'] +@[to_additive] theorem pow_mul_comm' (a : M) (n : ℕ) : a^n * a = a * a^n := commute.pow_self a n -@[to_additive add_nsmul] +@[to_additive] theorem pow_add (a : M) (m n : ℕ) : a^(m + n) = a^m * a^n := by induction n with n ih; [rw [nat.add_zero, pow_zero, mul_one], rw [pow_succ', ← mul_assoc, ← ih, ← pow_succ', nat.add_assoc]] @@ -110,7 +110,7 @@ theorem pow_bit0 (a : M) (n : ℕ) : a ^ bit0 n = a^n * a^n := pow_add _ _ _ theorem pow_bit1 (a : M) (n : ℕ) : a ^ bit1 n = a^n * a^n * a := by rw [bit1, pow_succ', pow_bit0] -@[to_additive nsmul_add_comm] +@[to_additive] theorem pow_mul_comm (a : M) (m n : ℕ) : a^m * a^n = a^n * a^m := commute.pow_pow_self a m n @@ -143,7 +143,7 @@ theorem mul_pow (a b : M) (n : ℕ) : (a * b)^n = a^n * b^n := /-- The `n`th power map on a commutative monoid for a natural `n`, considered as a morphism of monoids. -/ -@[to_additive nsmul_add_monoid_hom "Multiplication by a natural `n` on a commutative additive +@[to_additive "Multiplication by a natural `n` on a commutative additive monoid, considered as a morphism of additive monoids.", simps] def pow_monoid_hom (n : ℕ) : M →* M := { to_fun := (^ n), @@ -195,20 +195,20 @@ open int section nat -@[simp, to_additive neg_nsmul] theorem inv_pow (a : G) (n : ℕ) : (a⁻¹)^n = (a^n)⁻¹ := +@[simp, to_additive] theorem inv_pow (a : G) (n : ℕ) : (a⁻¹)^n = (a^n)⁻¹ := begin induction n with n ih, { rw [pow_zero, pow_zero, one_inv] }, { rw [pow_succ', pow_succ, ih, mul_inv_rev] } end -@[to_additive nsmul_sub] -- rename to sub_nsmul? +@[to_additive] -- rename to sub_nsmul? theorem pow_sub (a : G) {m n : ℕ} (h : n ≤ m) : a^(m - n) = a^m * (a^n)⁻¹ := have h1 : m - n + n = m, from tsub_add_cancel_of_le h, have h2 : a^(m - n) * a^n = a^m, by rw [←pow_add, h1], eq_mul_inv_of_mul_eq h2 -@[to_additive nsmul_neg_comm] +@[to_additive] theorem pow_inv_comm (a : G) (m n : ℕ) : (a⁻¹)^m * a^n = a^n * (a⁻¹)^m := (commute.refl a).inv_left.pow_pow m n diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index ea4df89d71f84..9d5125e6143ac 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -81,7 +81,7 @@ variables [monoid M] [linear_order M] [covariant_class M M (*) (≤)] lemma one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x := ⟨le_imp_le_of_lt_imp_lt $ λ h, pow_lt_one' h hn, λ h, one_le_pow_of_one_le' h n⟩ -@[to_additive nsmul_nonpos_iff] +@[to_additive] lemma pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := @one_le_pow_iff (order_dual M) _ _ _ _ _ hn @@ -89,11 +89,11 @@ lemma pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 lemma one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := lt_iff_lt_of_le_iff_le (pow_le_one_iff hn) -@[to_additive nsmul_neg_iff] +@[to_additive] lemma pow_lt_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n < 1 ↔ x < 1 := lt_iff_lt_of_le_iff_le (one_le_pow_iff hn) -@[to_additive nsmul_eq_zero_iff] +@[to_additive] lemma pow_eq_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n = 1 ↔ x = 1 := by simp only [le_antisymm_iff, pow_le_one_iff hn, one_le_pow_iff hn] diff --git a/src/algebra/iterate_hom.lean b/src/algebra/iterate_hom.lean index 1c56297248b1c..f5bf0610a201e 100644 --- a/src/algebra/iterate_hom.lean +++ b/src/algebra/iterate_hom.lean @@ -86,7 +86,7 @@ theorem iterate_map_smul (f : M →+ M) (n m : ℕ) (x : M) : f^[n] (m • x) = m • (f^[n] x) := f.to_multiplicative.iterate_map_pow n x m -attribute [to_additive iterate_map_smul, to_additive_reorder 5] monoid_hom.iterate_map_pow +attribute [to_additive, to_additive_reorder 5] monoid_hom.iterate_map_pow theorem iterate_map_zsmul (f : G →+ G) (n : ℕ) (m : ℤ) (x : G) : f^[n] (m • x) = m • (f^[n] x) := From 5d4660bcf674f6f26062db114629b9378b5f3362 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:12:55 +0100 Subject: [PATCH 085/116] Explicit name to avoid name clashes --- src/algebra/group_power/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 0845a981c61cf..74c03055fd763 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -58,7 +58,7 @@ alias pow_two ← sq @[to_additive] theorem pow_mul_comm' (a : M) (n : ℕ) : a^n * a = a * a^n := commute.pow_self a n -@[to_additive] +@[to_additive add_nsmul] theorem pow_add (a : M) (m n : ℕ) : a^(m + n) = a^m * a^n := by induction n with n ih; [rw [nat.add_zero, pow_zero, mul_one], rw [pow_succ', ← mul_assoc, ← ih, ← pow_succ', nat.add_assoc]] From 10fdc85cb308d21d3c332bbd48a5df3c11e55e06 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:18:38 +0100 Subject: [PATCH 086/116] More names are correct now --- src/algebra/big_operators/multiset.lean | 4 ++-- src/algebra/group/pi.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 29d9a2886166e..ef6c8014b3e64 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -70,7 +70,7 @@ lemma prod_nsmul (m : multiset α) : ∀ (n : ℕ), (n • m).prod = m.prod ^ n @[simp, to_additive] lemma prod_repeat (a : α) (n : ℕ) : (repeat a n).prod = a ^ n := by simp [repeat, list.prod_repeat] -@[to_additive nsmul_count] +@[to_additive] lemma pow_count [decidable_eq α] (a : α) : a ^ s.count a = (s.filter (eq a)).prod := by rw [filter_eq, prod_repeat] @@ -105,7 +105,7 @@ lemma prod_map_one : prod (m.map (λ i, (1 : α))) = 1 := by rw [map_const, prod lemma prod_map_mul : (m.map $ λ i, f i * g i).prod = (m.map f).prod * (m.map g).prod := m.prod_hom₂ (*) mul_mul_mul_comm (mul_one _) _ _ -@[to_additive sum_map_nsmul] +@[to_additive] lemma prod_map_pow {n : ℕ} : (m.map $ λ i, f i ^ n).prod = (m.map f).prod ^ n := m.prod_hom' (pow_monoid_hom n : α →* α) f diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 230247539b242..7ff520fd5cf27 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -43,7 +43,7 @@ by refine_struct { one := (1 : Π i, f i), mul := (*), npow := λ n x i, (x i) ^ tactic.pi_instance_derive_field -- the attributes are intentionally out of order. `smul_apply` proves `nsmul_apply`. -@[to_additive nsmul_apply, simp] +@[to_additive, simp] lemma pow_apply [∀ i, monoid $ f i] (n : ℕ) : (x^n) i = (x i)^n := rfl @[to_additive] From c2d6ddeb0a22717c22a199bc5b27290dbecca5e3 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:23:24 +0100 Subject: [PATCH 087/116] Even more names are correct now --- src/algebra/big_operators/basic.lean | 2 +- src/algebra/pointwise.lean | 4 ++-- src/group_theory/submonoid/membership.lean | 4 ++-- src/topology/continuous_function/algebra.lean | 12 ++++++------ 4 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 0e85b0e3d601a..2100e97bae1f5 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1035,7 +1035,7 @@ by rw [prod_insert has, card_insert_of_not_mem has, pow_succ, ih]) @[to_additive] lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp -@[to_additive sum_nsmul] +@[to_additive] lemma prod_pow (s : finset α) (n : ℕ) (f : α → β) : ∏ x in s, f x ^ n = (∏ x in s, f x) ^ n := by haveI := classical.dec_eq α; exact diff --git a/src/algebra/pointwise.lean b/src/algebra/pointwise.lean index 2749052e0d2f0..08e38481fd92f 100644 --- a/src/algebra/pointwise.lean +++ b/src/algebra/pointwise.lean @@ -250,7 +250,7 @@ protected def comm_monoid [comm_monoid α] : comm_monoid (set α) := localized "attribute [instance] set.mul_one_class set.add_zero_class set.semigroup set.add_semigroup set.monoid set.add_monoid set.comm_monoid set.add_comm_monoid" in pointwise -@[to_additive nsmul_mem_nsmul] +@[to_additive] lemma pow_mem_pow [monoid α] (ha : a ∈ s) (n : ℕ) : a ^ n ∈ s ^ n := begin @@ -261,7 +261,7 @@ begin exact set.mul_mem_mul ha ih }, end -@[to_additive empty_nsmul] +@[to_additive] lemma empty_pow [monoid α] (n : ℕ) (hn : n ≠ 0) : (∅ : set α) ^ n = ∅ := by rw [← tsub_add_cancel_of_le (nat.succ_le_of_lt $ nat.pos_of_ne_zero hn), pow_succ, empty_mul] diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index 1008973e6e592..4438386faddac 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -40,7 +40,7 @@ namespace submonoid section assoc variables [monoid M] (S : submonoid M) -@[simp, norm_cast, to_additive coe_nsmul] theorem coe_pow (x : S) (n : ℕ) : +@[simp, norm_cast, to_additive] theorem coe_pow (x : S) (n : ℕ) : ↑(x ^ n) = (x ^ n : M) := S.subtype.map_pow x n @@ -78,7 +78,7 @@ lemma prod_mem {M : Type*} [comm_monoid M] (S : submonoid M) ∏ c in t, f c ∈ S := S.multiset_prod_mem (t.1.map f) $ λ x hx, let ⟨i, hi, hix⟩ := multiset.mem_map.1 hx in hix ▸ h i hi -@[to_additive nsmul_mem] lemma pow_mem {x : M} (hx : x ∈ S) (n : ℕ) : x ^ n ∈ S := +@[to_additive] lemma pow_mem {x : M} (hx : x ∈ S) (n : ℕ) : x ^ n ∈ S := by simpa only [coe_pow] using ((⟨x, hx⟩ : S) ^ n).coe_prop end assoc diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 2ab9803d848f6..fd21f1185d866 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -68,11 +68,11 @@ instance has_nsmul [add_monoid β] [has_continuous_add β] : has_scalar ℕ C(α instance has_pow [monoid β] [has_continuous_mul β] : has_pow C(α, β) ℕ := ⟨λ f n, ⟨f ^ n, f.continuous.pow n⟩⟩ -@[norm_cast, to_additive coe_nsmul] +@[norm_cast, to_additive] lemma coe_pow [monoid β] [has_continuous_mul β] (f : C(α, β)) (n : ℕ) : ⇑(f ^ n) = f ^ n := rfl - --- don't make `coe_nsmul` simp as the linter complains it's redundant WRT `coe_smul` + +-- don't make `coe_nsmul` simp as the linter complains it's redundant WRT `coe_smul` attribute [simp] coe_pow @[to_additive nsmul_comp] lemma pow_comp [monoid γ] [has_continuous_mul γ] @@ -80,7 +80,7 @@ attribute [simp] coe_pow (f^n).comp g = (f.comp g)^n := rfl --- don't make `nsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` +-- don't make `nsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` attribute [simp] pow_comp @[to_additive] @@ -122,7 +122,7 @@ lemma coe_zpow [group β] [topological_group β] (f : C(α, β)) (z : ℤ) : ⇑(f ^ z) = f ^ z := rfl --- don't make `coe_zsmul` simp as the linter complains it's redundant WRT `coe_smul` +-- don't make `coe_zsmul` simp as the linter complains it's redundant WRT `coe_smul` attribute [simp] coe_zpow @[to_additive] @@ -130,7 +130,7 @@ lemma zpow_comp [group γ] [topological_group γ] (f : C(β, γ)) (z : ℤ) (g : (f^z).comp g = (f.comp g)^z := rfl --- don't make `zsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` +-- don't make `zsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` attribute [simp] zpow_comp end continuous_map From 30ca72289757d4ee48ec91fe8aece5dea17e1620 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:29:02 +0100 Subject: [PATCH 088/116] Even more names are correct now (order_of_element) --- src/group_theory/order_of_element.lean | 14 +++++++------- src/topology/algebra/monoid.lean | 6 +++--- 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 2e7a63fcf5bd3..06571d2160c5a 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -43,7 +43,7 @@ variables [monoid G] [add_monoid A] section is_of_fin_order -@[to_additive is_periodic_pt_add_iff_nsmul_eq_zero] +@[to_additive] lemma is_periodic_pt_mul_iff_pow_eq_one (x : G) : is_periodic_pt ((*) x) n 1 ↔ x ^ n = 1 := by rw [is_periodic_pt, is_fixed_pt, mul_left_iterate, mul_one] @@ -154,7 +154,7 @@ is_periodic_pt.minimal_period_dvd ((is_periodic_pt_mul_iff_pow_eq_one _).mpr h) lemma order_of_dvd_iff_pow_eq_one {n : ℕ} : order_of x ∣ n ↔ x ^ n = 1 := ⟨λ h, by rw [pow_eq_mod_order_of, nat.mod_eq_zero_of_dvd h, pow_zero], order_of_dvd_of_pow_eq_one⟩ -@[to_additive exists_nsmul_eq_self_of_coprime] +@[to_additive] lemma exists_pow_eq_self_of_coprime (h : n.coprime (order_of x)) : ∃ m : ℕ, (x ^ n) ^ m = x := begin @@ -285,7 +285,7 @@ end monoid_add_monoid section cancel_monoid variables [left_cancel_monoid G] (x y) -@[to_additive nsmul_injective_aux] +@[to_additive] lemma pow_injective_aux (h : n ≤ m) (hm : m < order_of x) (eq : x ^ n = x ^ m) : n = m := by_contradiction $ assume ne : n ≠ m, @@ -362,7 +362,7 @@ begin { simp [pow_succ, IH] } } end -@[to_additive nsmul_inj_mod] +@[to_additive] lemma pow_inj_mod {n m : ℕ} : x ^ n = x ^ m ↔ n % order_of x = m % order_of x := begin @@ -403,7 +403,7 @@ section finite_cancel_monoid variables [left_cancel_monoid G] [add_left_cancel_monoid A] -- TODO: Use this to show that a finite left cancellative monoid is a group. -@[to_additive exists_nsmul_eq_zero] +@[to_additive] lemma exists_pow_eq_one (x : G) : is_of_fin_order x := begin refine (is_of_fin_order_iff_pow_eq_one _).mpr _, @@ -612,7 +612,7 @@ end let ⟨m, hm⟩ := @order_of_dvd_card_univ _ x _ _ in by simp [hm, pow_mul, pow_order_of_eq_one] -@[to_additive nsmul_eq_mod_card] lemma pow_eq_mod_card (n : ℕ) : +@[to_additive] lemma pow_eq_mod_card (n : ℕ) : x ^ n = x ^ (n % fintype.card G) := by rw [pow_eq_mod_order_of, ←nat.mod_mod_of_dvd n order_of_dvd_card_univ, ← pow_eq_mod_order_of] @@ -623,7 +623,7 @@ by rw [zpow_eq_mod_order_of, ← int.mod_mod_of_dvd n (int.coe_nat_dvd.2 order_o ← zpow_eq_mod_order_of] /-- If `gcd(|G|,n)=1` then the `n`th power map is a bijection -/ -@[to_additive nsmul_coprime "If `gcd(|G|,n)=1` then the smul by `n` is a bijection", simps] +@[to_additive "If `gcd(|G|,n)=1` then the smul by `n` is a bijection", simps] def pow_coprime (h : nat.coprime (fintype.card G) n) : G ≃ G := { to_fun := λ g, g ^ n, inv_fun := λ g, g ^ (nat.gcd_b (fintype.card G) n), diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 4b35f696ca66b..636869c9e8253 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -372,7 +372,7 @@ lemma continuous_list_prod {f : ι → X → M} (l : list ι) continuous_iff_continuous_at.2 $ assume x, tendsto_list_prod l $ assume c hc, continuous_iff_continuous_at.1 (h c hc) x -@[continuity, to_additive continuous_nsmul] +@[continuity, to_additive] lemma continuous_pow : ∀ n : ℕ, continuous (λ a : M, a ^ n) | 0 := by simpa using continuous_const | (k+1) := by { simp only [pow_succ], exact continuous_id.mul (continuous_pow _) } @@ -385,11 +385,11 @@ lemma continuous.pow {f : X → M} (h : continuous f) (n : ℕ) : continuous (λ b, (f b) ^ n) := (continuous_pow n).comp h -@[to_additive continuous_on_nsmul] +@[to_additive] lemma continuous_on_pow {s : set M} (n : ℕ) : continuous_on (λ x, x ^ n) s := (continuous_pow n).continuous_on -@[to_additive continuous_at_nsmul] +@[to_additive] lemma continuous_at_pow (x : M) (n : ℕ) : continuous_at (λ x, x ^ n) x := (continuous_pow n).continuous_at From ca6c5319d054530aedf992727185952b1cfe930f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:37:15 +0100 Subject: [PATCH 089/116] Even more names are correct now (exponent.lean) --- src/group_theory/exponent.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/group_theory/exponent.lean b/src/group_theory/exponent.lean index 513b10ab3d7d2..b042adb542108 100644 --- a/src/group_theory/exponent.lean +++ b/src/group_theory/exponent.lean @@ -89,7 +89,7 @@ begin { simp_rw [exponent, dif_neg h, pow_zero] } end -@[to_additive nsmul_eq_mod_exponent] +@[to_additive] lemma pow_eq_mod_exponent {n : ℕ} (g : G): g ^ n = g ^ (n % exponent G) := calc g ^ n = g ^ (n % exponent G + exponent G * (n / exponent G)) : by rw [nat.mod_add_div] ... = g ^ (n % exponent G) : by simp [pow_add, pow_mul, pow_exponent_eq_one] @@ -138,7 +138,7 @@ order_of_dvd_of_pow_eq_one $ pow_exponent_eq_one g variable (G) -@[to_additive exponent_dvd_of_forall_nsmul_eq_zero] +@[to_additive] lemma exponent_dvd_of_forall_pow_eq_one (G) [monoid G] (n : ℕ) (hG : ∀ g : G, g ^ n = 1) : exponent G ∣ n := begin From b076f28cd1cb9a07f3a2f892c5d56544f3784659 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 12:51:50 +0100 Subject: [PATCH 090/116] Even more names are correct now (topology/continuous_function/algebra.lean) --- src/topology/continuous_function/algebra.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index fd21f1185d866..c6a2eb3125577 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -64,7 +64,7 @@ lemma coe_one [has_one β] : ⇑(1 : C(α, β)) = 1 := rfl instance has_nsmul [add_monoid β] [has_continuous_add β] : has_scalar ℕ C(α, β) := ⟨λ n f, ⟨n • f, f.continuous.nsmul n⟩⟩ -@[to_additive has_nsmul] +@[to_additive] instance has_pow [monoid β] [has_continuous_mul β] : has_pow C(α, β) ℕ := ⟨λ f n, ⟨f ^ n, f.continuous.pow n⟩⟩ @@ -75,7 +75,7 @@ lemma coe_pow [monoid β] [has_continuous_mul β] (f : C(α, β)) (n : ℕ) : -- don't make `coe_nsmul` simp as the linter complains it's redundant WRT `coe_smul` attribute [simp] coe_pow -@[to_additive nsmul_comp] lemma pow_comp [monoid γ] [has_continuous_mul γ] +@[to_additive] lemma pow_comp [monoid γ] [has_continuous_mul γ] (f : C(β, γ)) (n : ℕ) (g : C(α, β)) : (f^n).comp g = (f.comp g)^n := rfl From b5893351902b80e23407a615fc87cb193e84f635 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 13:26:19 +0100 Subject: [PATCH 091/116] Say pow, not npow --- src/algebra/big_operators/multiset.lean | 4 ++-- src/algebra/big_operators/order.lean | 8 ++++---- src/data/list/prod_monoid.lean | 4 ++-- 3 files changed, 8 insertions(+), 8 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 46f0e10602e1f..cbe0b829ffc27 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -268,11 +268,11 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_npow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : +lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, - simpa using list.prod_le_npow_of_forall_le _ _ h, + simpa using list.prod_le_pow_of_forall_le _ _ h, end @[to_additive all_zero_of_le_zero_le_of_sum_eq_zero] diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 9b77f0df9c3bb..456ff3cf7a65c 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -170,18 +170,18 @@ calc f a = ∏ i in {a}, f i : prod_singleton.symm prod_le_prod_of_subset_of_one_le' (singleton_subset_iff.2 h) $ λ i hi _, hf i hi @[to_additive] -lemma prod_le_npow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : +lemma prod_le_pow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : s.prod f ≤ n ^ s.card := begin - refine (multiset.prod_le_npow_of_forall_le (s.val.map f) n _).trans _, + refine (multiset.prod_le_pow_of_forall_le (s.val.map f) n _).trans _, { simpa using h }, { simpa } end @[to_additive] -lemma npow_le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : +lemma pow_le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_npow_of_forall_le _ (order_dual N) _ _ _ _ h +@finset.prod_le_pow_of_forall_le _ (order_dual N) _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index d1f100a0771a4..59d7843c7cb33 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -27,7 +27,7 @@ begin end @[to_additive] -lemma prod_eq_npow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x ∈ l), x = m) : +lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x ∈ l), x = m) : l.prod = m ^ l.length := begin convert list.prod_repeat m l.length, @@ -35,7 +35,7 @@ begin end @[to_additive] -lemma prod_le_npow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : +lemma prod_le_pow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := begin induction l with y l IH, From 65c37ca16d64988f02777b8e69b5b48f8139497f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 15:38:46 +0100 Subject: [PATCH 092/116] Fix sum_divisors_prime_pow --- src/number_theory/divisors.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index ce8259d18d372..c2fcf48645ec3 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -376,7 +376,7 @@ lemma prod_proper_divisors_prime_pow {α : Type*} [comm_monoid α] {k p : ℕ} { (h : p.prime) : ∏ x in (p ^ k).proper_divisors, f x = ∏ x in range k, f (p ^ x) := by simp [h, proper_divisors_prime_pow] -@[simp, to_additive] +@[simp, to_additive sum_divisors_prime_pow] lemma prod_divisors_prime_pow {α : Type*} [comm_monoid α] {k p : ℕ} {f : ℕ → α} (h : p.prime) : ∏ x in (p ^ k).divisors, f x = ∏ x in range (k + 1), f (p ^ x) := by simp [h, divisors_prime_pow] From 6123b28866095eb5c3956059bf7e46177ce64ca1 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 6 Mar 2022 15:39:12 +0100 Subject: [PATCH 093/116] More correct names --- src/algebra/group/hom.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/group/hom.lean b/src/algebra/group/hom.lean index 39019309d8879..027bf776079da 100644 --- a/src/algebra/group/hom.lean +++ b/src/algebra/group/hom.lean @@ -305,7 +305,7 @@ by rw [map_mul, map_inv] (f : F) (x y : G) : f (x / y) = f x / f y := by rw [div_eq_mul_inv, div_eq_mul_inv, map_mul_inv] -@[simp, to_additive map_nsmul] theorem map_pow [monoid G] [monoid H] [monoid_hom_class F G H] +@[simp, to_additive] theorem map_pow [monoid G] [monoid H] [monoid_hom_class F G H] (f : F) (a : G) : ∀ (n : ℕ), f (a ^ n) = (f a) ^ n | 0 := by rw [pow_zero, pow_zero, map_one] From 48f3938c1fc5e34d0bbecd1e3dca16d077332b24 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 21:54:07 +0100 Subject: [PATCH 094/116] Big refactoring, more lemmas about noncomm_prod, no more noncomm_pi_coprod_on --- src/data/finset/noncomm_prod.lean | 142 +++++++++++++- src/group_theory/noncomm_pi_coprod.lean | 240 +++++++----------------- 2 files changed, 203 insertions(+), 179 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 4e6664ae75eb3..dc9c4302d9931 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -5,6 +5,8 @@ Authors: Yakov Pechersky -/ import algebra.big_operators.basic +import group_theory.submonoid.basic +import group_theory.subgroup.basic /-! # Products (respectively, sums) over a finset or a multiset. @@ -183,6 +185,12 @@ given a proof that `+` commutes on all elements `f x` for `x ∈ s`."] def noncomm_prod (s : finset α) (f : α → β) (comm : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) : β := (s.1.map f).noncomm_prod (by simpa [multiset.mem_map, ←finset.mem_def] using comm) +@[to_additive] +lemma noncomm_prod_congr + {s₁ s₂ : finset α} (h : s₁ = s₂) (f : α → β) + (comm : ∀ (x ∈ s₁) (y ∈ s₁), commute (f x) (f y)) : + noncomm_prod s₁ f comm = noncomm_prod s₂ f (by simpa [h] using comm) := by subst h + @[simp, to_additive] lemma noncomm_prod_to_finset [decidable_eq α] (l : list α) (f : α → β) (comm : ∀ (x ∈ l.to_finset) (y ∈ l.to_finset), commute (f x) (f y)) (hl : l.nodup) : @@ -252,6 +260,49 @@ begin exact one_pow _, end + +@[to_additive] +lemma noncomm_prod_commute (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (y : β) (h : ∀ (x : α), x ∈ s → commute y (f x)) : commute y (s.noncomm_prod f comm) := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { simp }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + exact commute.mul_right (h x (finset.mem_insert_self _ _)) ih, } +end + +@[to_additive] +lemma noncomm_prod_mem_submonoid (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : submonoid β) (h : ∀ (x : α), x ∈ s → f x ∈ m) : s.noncomm_prod f comm ∈ m := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { exact submonoid.one_mem m, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + apply submonoid.mul_mem _ (h x (finset.mem_insert_self _ _)) ih, } +end + +@[to_additive] +lemma noncomm_prod_mem_subgroup {β : Type*} [group β] (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : subgroup β) (h : ∀ (x : α), x ∈ s → f x ∈ m) : s.noncomm_prod f comm ∈ m := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { exact subgroup.one_mem m, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + apply subgroup.mul_mem _ (h x (finset.mem_insert_self _ _)) ih, } +end + @[to_additive] lemma noncomm_prod_eq_prod {β : Type*} [comm_monoid β] (s : finset α) (f : α → β) : noncomm_prod s f (λ _ _ _ _, commute.all _ _) = s.prod f := begin @@ -279,13 +330,100 @@ begin list.nodup_append_of_nodup sl' tl' h] end +/- The non-commutative version of `finset.prod_mul_distrib` -/ +@[to_additive /-" The non-commutative version of `finset.sum_add_distrib` "-/] +lemma noncomm_prod_mul_distrib [decidable_eq α] {s : finset α} + (f : α → β) (g : α → β) + (comm_fg : ∀ (x ∈ s) (y ∈ s), commute (f x * g x) (f y * g y)) + (comm_fg' : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (f x) (g y)) + (comm_ff : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) + (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) : + noncomm_prod s (f * g) comm_fg + = noncomm_prod s f comm_ff * noncomm_prod s g comm_gg := +begin + induction s using finset.induction_on with x s hnmem ih, + { simp, }, + { + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + rw pi.mul_apply, + specialize ih + (λ x hx y hy, comm_fg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy hne, comm_fg' x (mem_insert_of_mem hx) y (mem_insert_of_mem hy) hne) + (λ x hx y hy, comm_ff x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy, comm_gg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)), + rw ih, + simp only [mul_assoc], + congr' 1, + simp only [← mul_assoc], + congr' 1, + apply noncomm_prod_commute, + intros y hy, + have : y ≠ x, by {rintro rfl, contradiction}, + exact (comm_fg' y (mem_insert_of_mem hy) x (mem_insert_self x s) this).symm, + } +end + +-- I think it's worth keeping it and moving to appropriate file +@[to_additive] +lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : + disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := +begin + split, + { intros hdis x y hx hy heq, + obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), + have hy := H₂.inv_mem_iff.mp hy, + have : x ∈ H₁ ⊓ H₂, by { simp, cc }, + rw [hdis.eq_bot, subgroup.mem_bot] at this, + subst this, + simp }, + { rintros h x ⟨hx1, hx2⟩, + obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, + exact rfl, }, +end + +/-- `finset.noncomm_prod` is injective in `f` if `f` maps into independent subgroups. -/ +@[to_additive "`finset.noncomm_sum` is injective in `f` if `f` maps into independent subgroups"] +lemma eq_one_of_noncomm_prod_eq_one_of_independent {β : Type*} [group β] + (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (H : α → subgroup β) (hind : complete_lattice.independent H) (hmem : ∀ x, f x ∈ H x) + (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := +begin + classical, + revert heq1, + induction s using finset.induction_on with i s hnmem ih, + { simp, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)), + have hmem_bsupr: s.noncomm_prod f + (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) ∈ + ⨆ (i : α) (_x : i ∈ (s : set α)), H i, + { refine noncomm_prod_mem_subgroup s f _ _ _, + intros x hx, + suffices : H x ≤ ⨆ (i : α) (_ : i ∈ (s : set α)), H i, by apply this (hmem x), + exact le_bsupr x hx, }, + intro heq1, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, + have hnmem' : i ∉ (s : set α), by simpa, + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := + mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') (hmem i) hmem_bsupr heq1, + rcases heq1' with ⟨ heq1i, heq1S ⟩, + specialize ih heq1S, + intros i h, + simp only [finset.mem_insert] at h, + rcases h with ⟨rfl | _⟩, + { exact heq1i }, + { exact (ih _ h), } } +end + section finite_pi variables {ι : Type*} [fintype ι] [decidable_eq ι] {M : ι → Type*} [∀ i, monoid (M i)] variables (x : Π i, M i) + @[to_additive] -lemma noncomm_prod_mul_single : +lemma noncomm_prod_mul_single_eval : univ.noncomm_prod (λ i, pi.mul_single i (x i)) (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) = x := @@ -306,7 +444,7 @@ lemma _root_.monoid_hom.pi_ext {f g : (Π i, M i) →* γ} f = g := begin ext x, - rw [← noncomm_prod_mul_single x, univ.noncomm_prod_map, univ.noncomm_prod_map], + rw [← noncomm_prod_mul_single_eval x, univ.noncomm_prod_map, univ.noncomm_prod_map], congr' 1 with i, exact h i (x i), end diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 53b448c209d2c..233bef9e079c4 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -19,9 +19,6 @@ image of different morphism commute, we obtain a canoncial morphism ## Main definitions * `monoid_hom.noncomm_pi_coprod : (Π i, N i) →* M` is the main homomorphism -* `noncomm_pi_coprod_on.hom (S : finset S) : (Π (i : ι), N i) →* M` is the homomorphism - restricted to the set `S`, and mostly of internal interest in this file, to allow inductive - proofs (and thus in its own namespace). * `subgroup.noncomm_pi_coprod : (Π i, H i) →* G` is the specialization to `H i : subgroup G` and the subgroup embedding. @@ -44,24 +41,6 @@ image of different morphism commute, we obtain a canoncial morphism open_locale big_operators --- I think it's worth keeping it and moving to appropriate file -@[to_additive] -lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : - disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := -begin - split, - { intros hdis x y hx hy heq, - obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), - have hy := H₂.inv_mem_iff.mp hy, - have : x ∈ H₁ ⊓ H₂, by { simp, cc }, - rw [hdis.eq_bot, subgroup.mem_bot] at this, - subst this, - simp }, - { rintros h x ⟨hx1, hx2⟩, - obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, - exact rfl, }, -end - section family_of_monoids variables {M : Type*} [monoid M] @@ -81,134 +60,6 @@ include hcomm -- We use `f` and `g` to denote elements of `Π (i : ι), N i` variables (f g : Π (i : ι), N i) -namespace noncomm_pi_coprod_on - --- In this section, we restrict the hom to a set S -variables (S : finset ι) - -/-- The underlying function of `noncomm_pi_coprod_on.hom` -/ -@[to_additive add_to_fun "The underlying function of `noncomm_pi_coprod_on.add_hom` "] -def to_fun (S : finset ι) : M := finset.noncomm_prod S (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } } - -variable {hcomm} - -@[simp, to_additive add_to_fun_empty] -lemma to_fun_empty : to_fun ϕ hcomm f ∅ = 1 := finset.noncomm_prod_empty _ _ - -include hdec - -@[simp, to_additive add_to_fun_insert_of_not_mem] -lemma to_fun_insert_of_not_mem (S : finset ι) (i : ι) (h : i ∉ S) : - to_fun ϕ hcomm f (insert i S) = ϕ i (f i) * to_fun ϕ hcomm f S := -finset.noncomm_prod_insert_of_not_mem _ _ _ _ h - -omit hdec - -@[simp, to_additive add_to_fun_zero] -lemma to_fun_one : to_fun ϕ hcomm 1 S = 1 := -begin - classical, - induction S using finset.cons_induction_on with i S hnmem ih, - { simp }, { simp [ih, hnmem], } -end - -@[to_additive add_to_fun_commutes] -lemma to_fun_commutes (i : ι) (hnmem : i ∉ S) : - commute (ϕ i (g i)) (to_fun ϕ hcomm f S) := -begin - classical, - induction S using finset.induction_on with j S hnmem' ih, - { simp, }, - { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem'], - have hij : i ≠ j, by {rintro rfl, apply hnmem, exact finset.mem_insert_self i S}, - have hiS : i ∉ S, by {rintro h, apply hnmem, exact finset.mem_insert_of_mem h}, - calc ϕ i (g i) * (ϕ j (f j) * (to_fun ϕ hcomm f S : M)) - = (ϕ i (g i) * ϕ j (f j)) * to_fun ϕ hcomm f S : by rw ← mul_assoc - ... = (ϕ j (f j) * ϕ i (g i)) * to_fun ϕ hcomm f S : by { congr' 1, apply hcomm _ _ hij } - ... = ϕ j (f j) * (ϕ i (g i) * to_fun ϕ hcomm f S) : by rw mul_assoc - ... = ϕ j (f j) * (to_fun ϕ hcomm f S * ϕ i (g i)) : by { congr' 1, apply ih hiS } - ... = (ϕ j (f j) * to_fun ϕ hcomm f S) * ϕ i (g i) : by rw ← mul_assoc } -end - -@[simp, to_additive add_to_fun_add] -lemma to_fun_mul (S : finset ι) : - to_fun ϕ hcomm (f * g) S = to_fun ϕ hcomm f S * to_fun ϕ hcomm g S := -begin - classical, - induction S using finset.induction_on with i S hnmem ih, - { simp, }, - { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], - rw ih, clear ih, - simp only [pi.mul_apply, map_mul], - repeat { rw mul_assoc }, congr' 1, - repeat { rw ← mul_assoc }, congr' 1, - exact (to_fun_commutes _ _ _ S i hnmem), } -end - -@[to_additive add_to_fun_mem_bsupr_mrange] -lemma to_fun_mem_bsupr_mrange (S : finset ι) : - to_fun ϕ hcomm f S ∈ ⨆ i ∈ S, (ϕ i).mrange := -begin - classical, - induction S using finset.induction_on with i S hnmem ih, - { simp, }, - { simp only [to_fun_insert_of_not_mem _ _ _ _ hnmem], - refine (submonoid.mul_mem _ _ _), - { apply submonoid.mem_Sup_of_mem, { use i }, { simp, }, }, - { refine @bsupr_le_bsupr' _ _ _ _ _ _ (λ i, (ϕ i).mrange) _ ih, - exact λ i, finset.mem_insert_of_mem } } -end - -variable (hcomm) - -/-- The canonical homomorphism from a family of monoids, restricted to a subset of the index space. --/ -@[to_additive add_hom "The canonical homomorphism from a family of additive monoids, restricted to a -subset of the index space"] -def hom : (Π (i : ι), N i) →* M := -{ to_fun := λ f, to_fun ϕ hcomm f S, - map_one' := to_fun_one ϕ _, - map_mul' := λ f g, to_fun_mul ϕ f g S, } - -variable {hcomm} - -include hdec - -@[to_additive add_to_fun_single] -lemma to_fun_mul_single (i : ι) (y : N i) (S : finset ι) : - to_fun ϕ hcomm (pi.mul_single i y) S = if i ∈ S then ϕ i y else 1 := -begin - induction S using finset.induction_on with j S hnmem ih, - { simp, }, - { simp only [ to_fun_insert_of_not_mem _ _ _ _ hnmem], - by_cases (i = j), - { subst h, rw ih, simp [hnmem], }, - { change i ≠ j at h, - simpa [h] using ih, } } -end - -@[simp, to_additive add_hom_single] -lemma hom_mul_single (i : ι) (y : N i): - hom ϕ hcomm S (pi.mul_single i y) = if i ∈ S then ϕ i y else 1 := to_fun_mul_single _ _ _ _ - -omit hdec - -@[to_additive add_mrange] -lemma mrange : (hom ϕ hcomm S).mrange = ⨆ i ∈ S, (ϕ i).mrange := -begin - classical, - apply le_antisymm, - { rintro x ⟨f, rfl⟩, - exact (to_fun_mem_bsupr_mrange ϕ f S), }, - { refine (bsupr_le _), - rintro i hmem x ⟨y, rfl⟩, - use (pi.mul_single i y), - simp [hmem], } -end - -end noncomm_pi_coprod_on - namespace monoid_hom include hfin @@ -219,7 +70,23 @@ variable (hcomm) @[to_additive "The canonical homomorphism from a family of additive monoids. See also `linear_map.lsum` for a linear version without the commutativity assumption."] -def noncomm_pi_coprod : (Π (i : ι), N i) →* M := noncomm_pi_coprod_on.hom ϕ hcomm finset.univ +def noncomm_pi_coprod : (Π (i : ι), N i) →* M := +{ to_fun := λ f, finset.univ.noncomm_prod (λ i, ϕ i (f i)) $ + by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } }, + map_one' := finset.noncomm_prod_eq_one_of_forall_eq_one _ _ _(by simp), + map_mul' := λ f g, + begin + classical, + convert finset.noncomm_prod_mul_distrib (λ i, ϕ i (f i)) (λ i, ϕ i (g i)) _ _ _ _, + { ext i, simp, }, + { rintros i - j -, + by_cases h : i = j, + { subst h, }, + { simp only [← monoid_hom.map_mul], + exact hcomm _ _ h _ _ } }, + { rintros i - j - h, + exact hcomm _ _ h _ _ }, + end } variable {hcomm} @@ -228,7 +95,16 @@ include hdec @[simp, to_additive] lemma noncomm_pi_coprod_mul_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (pi.mul_single i y) = ϕ i y := -by { show noncomm_pi_coprod_on.hom ϕ hcomm finset.univ (pi.mul_single i y) = ϕ i y, simp } +begin + change finset.univ.noncomm_prod (λ j, ϕ j (pi.mul_single i y j)) _ = ϕ i y, + have h : finset.univ = insert i (finset.univ.erase i) := + (finset.insert_erase (finset.mem_univ i)).symm, + rw finset.noncomm_prod_congr h, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ (finset.not_mem_erase i finset.univ), + rw finset.noncomm_prod_eq_one_of_forall_eq_one, + { simp, }, + { intros j hj, simp only [finset.mem_erase] at hj, simp [hj], }, +end omit hcomm @@ -251,8 +127,20 @@ include hcomm @[to_additive] lemma noncomm_pi_coprod_mrange : (noncomm_pi_coprod ϕ hcomm).mrange = ⨆ i : ι, (ϕ i).mrange := begin - show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).mrange = _, - simp [noncomm_pi_coprod_on.mrange], + classical, + apply le_antisymm, + { rintro x ⟨f, rfl⟩, + refine finset.noncomm_prod_mem_submonoid _ _ _ _ _, + intros i hi, + apply submonoid.mem_Sup_of_mem, { use i }, + simp, }, + { have : ∀ i, i ∈ finset.univ → (ϕ i).mrange ≤ (noncomm_pi_coprod ϕ hcomm).mrange, + { rintro i hmem x ⟨y, rfl⟩, + refine ⟨pi.mul_single i y, noncomm_pi_coprod_mul_single _ _ _⟩, }, + -- Why does this bsupr_le not work? + exact bsupr_le this, + -- exact @bsupr_le (submonoid M) ι submonoid.complete_lattice (noncomm_pi_coprod ϕ hcomm).mrange (λ i, i ∈ finset.univ) (λ i _, (ϕ i).mrange) this , + }, end end monoid_hom @@ -271,6 +159,7 @@ include hcomm -- We use `f` and `g` to denote elements of `Π (i : ι), H i` variables (f g : Π (i : ι), H i) + /- -- The subgroup version of `noncomm_pi_coprod_on.to_fun_mem_bsupr_mrange` @[to_additive noncomm_pi_coprod_on.add_to_fun_mem_bsupr_range] lemma noncomm_pi_coprod_on.to_fun_mem_bsupr_range (S : finset ι) : @@ -300,6 +189,7 @@ begin use (monoid_hom.single _ i y), simp [hmem], } end + -/ include hfin @@ -309,8 +199,18 @@ namespace monoid_hom @[to_additive] lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := begin - show (noncomm_pi_coprod_on.hom ϕ hcomm finset.univ).range = _, - simp [noncomm_pi_coprod_on.range] + apply le_antisymm, + { rintro x ⟨f, rfl⟩, + refine finset.noncomm_prod_mem_subgroup _ _ _ _ _, + intros i hi, + apply subgroup.mem_Sup_of_mem, { use i }, + simp, }, + sorry;{ + refine bsupr_le _, + rintro i hmem x ⟨y, rfl⟩, + use (pi.mul_single i y), + simp [hmem], + } end @[to_additive] @@ -322,32 +222,18 @@ begin classical, apply (monoid_hom.ker_eq_bot_iff _).mp, apply eq_bot_iff.mpr, - intro f, - show noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → f = 1, - suffices : noncomm_pi_coprod_on.to_fun ϕ hcomm f finset.univ = 1 → - (∀ (i : ι), i ∈ finset.univ → f i = 1), - by exact (λ h, funext (λ (i : ι), this h i (finset.mem_univ i))), - induction (finset.univ : finset ι) using finset.induction_on with i S hnmem ih, - { simp }, - { intro heq1, - simp only [ noncomm_pi_coprod_on.to_fun_insert_of_not_mem _ _ _ _ hnmem] at heq1, - have hnmem' : i ∉ (S : set ι), by simpa, - have heq1' : ϕ i (f i) = 1 ∧ noncomm_pi_coprod_on.to_fun ϕ hcomm f S = 1, - { apply mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') _ _ heq1, - { simp, }, - { apply noncomm_pi_coprod_on.to_fun_mem_bsupr_range, }, }, - rcases heq1' with ⟨ heq1i, heq1S ⟩, - specialize ih heq1S, - intros i h, - simp only [finset.mem_insert] at h, - rcases h with ⟨rfl | _⟩, - { apply hinj i, simpa, }, - { exact (ih _ h), } } + intros f heq1, + change finset.univ.noncomm_prod (λ i, ϕ i (f i)) _ = 1 at heq1, + change f = 1, + have : ∀ i, i ∈ finset.univ → ϕ i (f i) = 1 := + finset.eq_one_of_noncomm_prod_eq_one_of_independent _ _ _ _ hind (by simp) heq1, + ext i, + apply hinj, + simp [this i (finset.mem_univ i)], end variable (hcomm) - @[to_additive] lemma independent_range_of_coprime_order [∀ i, fintype (H i)] (hcoprime : ∀ i j, i ≠ j → nat.coprime (fintype.card (H i)) (fintype.card (H j))) : From 9a823712dfca31bad94dbcaf537d2292fbad8d95 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 22:01:41 +0100 Subject: [PATCH 095/116] supr_le, not bsupr_le! --- src/group_theory/noncomm_pi_coprod.lean | 20 +++++++------------- 1 file changed, 7 insertions(+), 13 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 233bef9e079c4..d0902dda644e1 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -134,13 +134,9 @@ begin intros i hi, apply submonoid.mem_Sup_of_mem, { use i }, simp, }, - { have : ∀ i, i ∈ finset.univ → (ϕ i).mrange ≤ (noncomm_pi_coprod ϕ hcomm).mrange, - { rintro i hmem x ⟨y, rfl⟩, - refine ⟨pi.mul_single i y, noncomm_pi_coprod_mul_single _ _ _⟩, }, - -- Why does this bsupr_le not work? - exact bsupr_le this, - -- exact @bsupr_le (submonoid M) ι submonoid.complete_lattice (noncomm_pi_coprod ϕ hcomm).mrange (λ i, i ∈ finset.univ) (λ i _, (ϕ i).mrange) this , - }, + { refine supr_le _, + rintro i x ⟨y, rfl⟩, + refine ⟨pi.mul_single i y, noncomm_pi_coprod_mul_single _ _ _⟩, }, end end monoid_hom @@ -199,18 +195,16 @@ namespace monoid_hom @[to_additive] lemma noncomm_pi_coprod_range : (noncomm_pi_coprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := begin + classical, apply le_antisymm, { rintro x ⟨f, rfl⟩, refine finset.noncomm_prod_mem_subgroup _ _ _ _ _, intros i hi, apply subgroup.mem_Sup_of_mem, { use i }, simp, }, - sorry;{ - refine bsupr_le _, - rintro i hmem x ⟨y, rfl⟩, - use (pi.mul_single i y), - simp [hmem], - } + { refine supr_le _, + rintro i x ⟨y, rfl⟩, + refine ⟨pi.mul_single i y, noncomm_pi_coprod_mul_single _ _ _⟩, }, end @[to_additive] From 40276be72664509253636363bb52668b70f815cb Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 22:08:24 +0100 Subject: [PATCH 096/116] Weaken lemma --- src/data/finset/noncomm_prod.lean | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index dc9c4302d9931..6e25f1da13fc7 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -387,26 +387,27 @@ end @[to_additive "`finset.noncomm_sum` is injective in `f` if `f` maps into independent subgroups"] lemma eq_one_of_noncomm_prod_eq_one_of_independent {β : Type*} [group β] (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (H : α → subgroup β) (hind : complete_lattice.independent H) (hmem : ∀ x, f x ∈ H x) + (γ : α → subgroup β) (hind : complete_lattice.independent γ) (hmem : ∀ (x ∈ s), f x ∈ γ x) (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := begin classical, revert heq1, induction s using finset.induction_on with i s hnmem ih, { simp, }, - { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)), + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, hmem x (mem_insert_of_mem hx)), have hmem_bsupr: s.noncomm_prod f (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) ∈ - ⨆ (i : α) (_x : i ∈ (s : set α)), H i, + ⨆ (i : α) (_x : i ∈ (s : set α)), γ i, { refine noncomm_prod_mem_subgroup s f _ _ _, intros x hx, - suffices : H x ≤ ⨆ (i : α) (_ : i ∈ (s : set α)), H i, by apply this (hmem x), - exact le_bsupr x hx, }, + have : γ x ≤ ⨆ (i ∈ (s : set α)), γ i := le_bsupr x hx, + exact this (hmem x (mem_insert_of_mem hx)), }, intro heq1, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, have hnmem' : i ∉ (s : set α), by simpa, have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := - mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') (hmem i) hmem_bsupr heq1, + mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, From 923f7f8a9ef91e2f54f9bf24fd6d013dc106c503 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 22:09:58 +0100 Subject: [PATCH 097/116] Make lint a bit happy --- src/data/finset/noncomm_prod.lean | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 6e25f1da13fc7..c16785e96c873 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -343,8 +343,7 @@ lemma noncomm_prod_mul_distrib [decidable_eq α] {s : finset α} begin induction s using finset.induction_on with x s hnmem ih, { simp, }, - { - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + { rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, rw pi.mul_apply, @@ -361,8 +360,7 @@ begin apply noncomm_prod_commute, intros y hy, have : y ≠ x, by {rintro rfl, contradiction}, - exact (comm_fg' y (mem_insert_of_mem hy) x (mem_insert_self x s) this).symm, - } + exact (comm_fg' y (mem_insert_of_mem hy) x (mem_insert_self x s) this).symm, } end -- I think it's worth keeping it and moving to appropriate file @@ -406,8 +404,8 @@ begin intro heq1, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, have hnmem' : i ∉ (s : set α), by simpa, - have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := - mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := mul_eq_one_iff_disjoint.mp + (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, intros i h, From 406600960ed125a5e6d07af8ef29cf8204f6a070 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 22:58:17 +0100 Subject: [PATCH 098/116] Reorder some hypotheses --- src/data/finset/noncomm_prod.lean | 23 ++++++++++------------- 1 file changed, 10 insertions(+), 13 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index c16785e96c873..4b35bc08fea40 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -334,24 +334,21 @@ end @[to_additive /-" The non-commutative version of `finset.sum_add_distrib` "-/] lemma noncomm_prod_mul_distrib [decidable_eq α] {s : finset α} (f : α → β) (g : α → β) - (comm_fg : ∀ (x ∈ s) (y ∈ s), commute (f x * g x) (f y * g y)) - (comm_fg' : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (f x) (g y)) + (comm_fgfg : ∀ (x ∈ s) (y ∈ s), commute (f x * g x) (f y * g y)) (comm_ff : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) - (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) : - noncomm_prod s (f * g) comm_fg + (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) + (comm_gf : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (g x) (f y)) : + noncomm_prod s (f * g) comm_fgfg = noncomm_prod s f comm_ff * noncomm_prod s g comm_gg := begin induction s using finset.induction_on with x s hnmem ih, { simp, }, - { rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, - rw pi.mul_apply, + { simp only [finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, pi.mul_apply], specialize ih - (λ x hx y hy, comm_fg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) - (λ x hx y hy hne, comm_fg' x (mem_insert_of_mem hx) y (mem_insert_of_mem hy) hne) + (λ x hx y hy, comm_fgfg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) (λ x hx y hy, comm_ff x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) - (λ x hx y hy, comm_gg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)), + (λ x hx y hy, comm_gg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy hne, comm_gf x (mem_insert_of_mem hx) y (mem_insert_of_mem hy) hne), rw ih, simp only [mul_assoc], congr' 1, @@ -359,8 +356,8 @@ begin congr' 1, apply noncomm_prod_commute, intros y hy, - have : y ≠ x, by {rintro rfl, contradiction}, - exact (comm_fg' y (mem_insert_of_mem hy) x (mem_insert_self x s) this).symm, } + have : x ≠ y, by {rintro rfl, contradiction}, + exact comm_gf x (mem_insert_self x s) y (mem_insert_of_mem hy) this, } end -- I think it's worth keeping it and moving to appropriate file From c9932f36b1db8401a42816f043921f64a5e05bf9 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:11:01 +0100 Subject: [PATCH 099/116] Moving some lemmas --- src/data/finset/noncomm_prod.lean | 24 +++--------------------- src/group_theory/noncomm_pi_coprod.lean | 9 +++------ src/group_theory/subgroup/basic.lean | 17 +++++++++++++++++ 3 files changed, 23 insertions(+), 27 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 4b35bc08fea40..b950fcb858983 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -313,7 +313,7 @@ begin end /- The non-commutative version of `finset.prod_union` -/ -@[to_additive /-" The non-commutative version of `finset.sum_union` "-/] +@[to_additive "The non-commutative version of `finset.sum_union`"] lemma noncomm_prod_union_of_disjoint [decidable_eq α] {s t : finset α} (h : disjoint s t) (f : α → β) (comm : ∀ (x ∈ s ∪ t) (y ∈ s ∪ t), commute (f x) (f y)) @@ -331,7 +331,7 @@ begin end /- The non-commutative version of `finset.prod_mul_distrib` -/ -@[to_additive /-" The non-commutative version of `finset.sum_add_distrib` "-/] +@[to_additive "The non-commutative version of `finset.sum_add_distrib`"] lemma noncomm_prod_mul_distrib [decidable_eq α] {s : finset α} (f : α → β) (g : α → β) (comm_fgfg : ∀ (x ∈ s) (y ∈ s), commute (f x * g x) (f y * g y)) @@ -360,24 +360,6 @@ begin exact comm_gf x (mem_insert_self x s) y (mem_insert_of_mem hy) this, } end --- I think it's worth keeping it and moving to appropriate file -@[to_additive] -lemma mul_eq_one_iff_disjoint {G : Type*} [group G] {H₁ H₂ : subgroup G} : - disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := -begin - split, - { intros hdis x y hx hy heq, - obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), - have hy := H₂.inv_mem_iff.mp hy, - have : x ∈ H₁ ⊓ H₂, by { simp, cc }, - rw [hdis.eq_bot, subgroup.mem_bot] at this, - subst this, - simp }, - { rintros h x ⟨hx1, hx2⟩, - obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, - exact rfl, }, -end - /-- `finset.noncomm_prod` is injective in `f` if `f` maps into independent subgroups. -/ @[to_additive "`finset.noncomm_sum` is injective in `f` if `f` maps into independent subgroups"] lemma eq_one_of_noncomm_prod_eq_one_of_independent {β : Type*} [group β] @@ -401,7 +383,7 @@ begin intro heq1, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, have hnmem' : i ∉ (s : set α), by simpa, - have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := mul_eq_one_iff_disjoint.mp + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.mul_eq_one_iff_disjoint.mp (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index d0902dda644e1..2470b8aec6e48 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -47,7 +47,7 @@ variables {M : Type*} [monoid M] -- We have a family of monoids -- The fintype assumption is not always used, but declared here, to keep things in order -variables {ι : Type*} [hdec : decidable_eq ι] [hfin : fintype ι] +variables {ι : Type*} [hdec : decidable_eq ι] [fintype ι] variables {N : ι → Type*} [∀ i, monoid (N i)] -- And morphisms ϕ into G @@ -62,10 +62,6 @@ variables (f g : Π (i : ι), N i) namespace monoid_hom -include hfin - -variable (hcomm) - /-- The canonical homomorphism from a family of monoids. -/ @[to_additive "The canonical homomorphism from a family of additive monoids. @@ -101,8 +97,9 @@ begin (finset.insert_erase (finset.mem_univ i)).symm, rw finset.noncomm_prod_congr h, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ (finset.not_mem_erase i finset.univ), + rw pi.mul_single_eq_same, rw finset.noncomm_prod_eq_one_of_forall_eq_one, - { simp, }, + { exact mul_one _, }, { intros j hj, simp only [finset.mem_erase] at hj, simp [hj], }, end diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index d818fc0bb0354..6ea346572ad93 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2716,6 +2716,23 @@ end end subgroup_normal +@[to_additive] +lemma mul_eq_one_iff_disjoint {H₁ H₂ : subgroup G} : + disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := +begin + split, + { intros hdis x y hx hy heq, + obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), + have hy := H₂.inv_mem_iff.mp hy, + have : x ∈ H₁ ⊓ H₂, by { simp, cc }, + rw [hdis.eq_bot, subgroup.mem_bot] at this, + subst this, + simp }, + { rintros h x ⟨hx1, hx2⟩, + obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, + exact rfl, }, +end + end subgroup namespace is_conj From 4e78f24b93f5be282ccd89f288fea446cb6728eb Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:13:35 +0100 Subject: [PATCH 100/116] Rename lemma --- src/data/finset/noncomm_prod.lean | 2 +- src/group_theory/subgroup/basic.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index b950fcb858983..85f834ca28fcf 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -383,7 +383,7 @@ begin intro heq1, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, have hnmem' : i ∉ (s : set α), by simpa, - have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.mul_eq_one_iff_disjoint.mp + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 6ea346572ad93..fea09abce30e8 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2717,7 +2717,7 @@ end end subgroup_normal @[to_additive] -lemma mul_eq_one_iff_disjoint {H₁ H₂ : subgroup G} : +lemma disjoint_iff_mul_eq_one {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := begin split, From 06a48a1d7feeddb6746f267eadc851a27f26a2dd Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:14:10 +0100 Subject: [PATCH 101/116] feat(group_theory/subgroup/basic): disjoint_iff_mul_eq_one --- src/group_theory/subgroup/basic.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index d818fc0bb0354..fea09abce30e8 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2716,6 +2716,23 @@ end end subgroup_normal +@[to_additive] +lemma disjoint_iff_mul_eq_one {H₁ H₂ : subgroup G} : + disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x * y = 1 → x = 1 ∧ y = 1 := +begin + split, + { intros hdis x y hx hy heq, + obtain rfl : y = x⁻¹ := symm (inv_eq_iff_mul_eq_one.mpr heq), + have hy := H₂.inv_mem_iff.mp hy, + have : x ∈ H₁ ⊓ H₂, by { simp, cc }, + rw [hdis.eq_bot, subgroup.mem_bot] at this, + subst this, + simp }, + { rintros h x ⟨hx1, hx2⟩, + obtain rfl : x = 1 := (h hx1 (H₂.inv_mem hx2) (mul_inv_self x)).1, + exact rfl, }, +end + end subgroup namespace is_conj From 048a979b7cba3111d98793619f62500fb8b74c35 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:17:42 +0100 Subject: [PATCH 102/116] feat(data/finset/noncomm_prod): assorted lemmas * `noncomm_prod_congr` * `noncomm_prod_commute` * `noncomm_prod_mem_submonoid` * `noncomm_prod_mem_subgroup` * `noncomm_prod_mul_distrib` * `eq_one_of_noncomm_prod_eq_one_of_independent` --- src/data/finset/noncomm_prod.lean | 117 +++++++++++++++++++++++++++++- 1 file changed, 116 insertions(+), 1 deletion(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index be0b2177c6e3c..ac456c0959412 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -5,6 +5,8 @@ Authors: Yakov Pechersky -/ import algebra.big_operators.basic +import group_theory.submonoid.basic +import group_theory.subgroup.basic /-! # Products (respectively, sums) over a finset or a multiset. @@ -183,6 +185,12 @@ given a proof that `+` commutes on all elements `f x` for `x ∈ s`."] def noncomm_prod (s : finset α) (f : α → β) (comm : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) : β := (s.1.map f).noncomm_prod (by simpa [multiset.mem_map, ←finset.mem_def] using comm) +@[to_additive] +lemma noncomm_prod_congr + {s₁ s₂ : finset α} (h : s₁ = s₂) (f : α → β) + (comm : ∀ (x ∈ s₁) (y ∈ s₁), commute (f x) (f y)) : + noncomm_prod s₁ f comm = noncomm_prod s₂ f (by simpa [h] using comm) := by subst h + @[simp, to_additive] lemma noncomm_prod_to_finset [decidable_eq α] (l : list α) (f : α → β) (comm : ∀ (x ∈ l.to_finset) (y ∈ l.to_finset), commute (f x) (f y)) (hl : l.nodup) : @@ -217,6 +225,49 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := by simp [noncomm_prod, multiset.singleton_eq_cons] + +@[to_additive] +lemma noncomm_prod_commute (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (y : β) (h : ∀ (x : α), x ∈ s → commute y (f x)) : commute y (s.noncomm_prod f comm) := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { simp }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + exact commute.mul_right (h x (finset.mem_insert_self _ _)) ih, } +end + +@[to_additive] +lemma noncomm_prod_mem_submonoid (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : submonoid β) (h : ∀ (x : α), x ∈ s → f x ∈ m) : s.noncomm_prod f comm ∈ m := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { exact submonoid.one_mem m, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + apply submonoid.mul_mem _ (h x (finset.mem_insert_self _ _)) ih, } +end + +@[to_additive] +lemma noncomm_prod_mem_subgroup {β : Type*} [group β] (s : finset α) (f : α → β) + (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (m : subgroup β) (h : ∀ (x : α), x ∈ s → f x ∈ m) : s.noncomm_prod f comm ∈ m := +begin + classical, + induction s using finset.induction_on with x s hnmem ih, + { exact subgroup.one_mem m, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, h x (mem_insert_of_mem hx)), + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, + apply subgroup.mul_mem _ (h x (finset.mem_insert_self _ _)) ih, } +end + @[to_additive] lemma noncomm_prod_eq_prod {β : Type*} [comm_monoid β] (s : finset α) (f : α → β) : noncomm_prod s f (λ _ _ _ _, commute.all _ _) = s.prod f := begin @@ -227,7 +278,7 @@ begin end /- The non-commutative version of `finset.prod_union` -/ -@[to_additive /-" The non-commutative version of `finset.sum_union` "-/] +@[to_additive "The non-commutative version of `finset.sum_union`"] lemma noncomm_prod_union_of_disjoint [decidable_eq α] {s t : finset α} (h : disjoint s t) (f : α → β) (comm : ∀ (x ∈ s ∪ t) (y ∈ s ∪ t), commute (f x) (f y)) @@ -244,4 +295,68 @@ begin list.nodup_append_of_nodup sl' tl' h] end +/- The non-commutative version of `finset.prod_mul_distrib` -/ +@[to_additive "The non-commutative version of `finset.sum_add_distrib`"] +lemma noncomm_prod_mul_distrib [decidable_eq α] {s : finset α} + (f : α → β) (g : α → β) + (comm_fgfg : ∀ (x ∈ s) (y ∈ s), commute (f x * g x) (f y * g y)) + (comm_ff : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) + (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) + (comm_gf : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (g x) (f y)) : + noncomm_prod s (f * g) comm_fgfg + = noncomm_prod s f comm_ff * noncomm_prod s g comm_gg := +begin + induction s using finset.induction_on with x s hnmem ih, + { simp, }, + { simp only [finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem, pi.mul_apply], + specialize ih + (λ x hx y hy, comm_fgfg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy, comm_ff x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy, comm_gg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx y hy hne, comm_gf x (mem_insert_of_mem hx) y (mem_insert_of_mem hy) hne), + rw ih, + simp only [mul_assoc], + congr' 1, + simp only [← mul_assoc], + congr' 1, + apply noncomm_prod_commute, + intros y hy, + have : x ≠ y, by {rintro rfl, contradiction}, + exact comm_gf x (mem_insert_self x s) y (mem_insert_of_mem hy) this, } +end + +/-- `finset.noncomm_prod` is injective in `f` if `f` maps into independent subgroups. -/ +@[to_additive "`finset.noncomm_sum` is injective in `f` if `f` maps into independent subgroups"] +lemma eq_one_of_noncomm_prod_eq_one_of_independent {β : Type*} [group β] + (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) + (γ : α → subgroup β) (hind : complete_lattice.independent γ) (hmem : ∀ (x ∈ s), f x ∈ γ x) + (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := +begin + classical, + revert heq1, + induction s using finset.induction_on with i s hnmem ih, + { simp, }, + { specialize ih (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) + (λ x hx, hmem x (mem_insert_of_mem hx)), + have hmem_bsupr: s.noncomm_prod f + (λ x hx y hy, comm x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) ∈ + ⨆ (i : α) (_x : i ∈ (s : set α)), γ i, + { refine noncomm_prod_mem_subgroup s f _ _ _, + intros x hx, + have : γ x ≤ ⨆ (i ∈ (s : set α)), γ i := le_bsupr x hx, + exact this (hmem x (mem_insert_of_mem hx)), }, + intro heq1, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, + have hnmem' : i ∉ (s : set α), by simpa, + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.mul_eq_one_iff_disjoint.mp + (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, + rcases heq1' with ⟨ heq1i, heq1S ⟩, + specialize ih heq1S, + intros i h, + simp only [finset.mem_insert] at h, + rcases h with ⟨rfl | _⟩, + { exact heq1i }, + { exact (ih _ h), } } +end + end finset From 96db2d5f1cd1066b4a4d0841705e61b775402926 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:22:51 +0100 Subject: [PATCH 103/116] Update src/data/finset/noncomm_prod.lean --- src/data/finset/noncomm_prod.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index ac456c0959412..83ea1a7ea9bde 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -348,7 +348,7 @@ begin intro heq1, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, have hnmem' : i ∉ (s : set α), by simpa, - have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.mul_eq_one_iff_disjoint.mp + have heq1' : f i = 1 ∧ s.noncomm_prod f _ = 1 := subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') (hmem i (mem_insert_self _ _)) hmem_bsupr heq1, rcases heq1' with ⟨ heq1i, heq1S ⟩, specialize ih heq1S, From 15d7b4a606dd9fd67b9f67abcbfdaa363ae45864 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 7 Mar 2022 23:41:26 +0100 Subject: [PATCH 104/116] Fix merge mistake --- src/data/finset/noncomm_prod.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 49f42c560beba..78b52f7893732 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -225,7 +225,6 @@ by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := by simp [noncomm_prod, multiset.singleton_eq_cons] -<<<<<<< HEAD @[to_additive] lemma noncomm_prod_map (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) From e6f2b3bdc459a410286f4c5ad2fcba1b6c1dd273 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 8 Mar 2022 12:10:50 +0100 Subject: [PATCH 105/116] Avoid the need for finset.noncomm_prod_congr --- src/group_theory/noncomm_pi_coprod.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 2470b8aec6e48..7ae41fa69549e 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -93,10 +93,12 @@ lemma noncomm_pi_coprod_mul_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (pi.mul_single i y) = ϕ i y := begin change finset.univ.noncomm_prod (λ j, ϕ j (pi.mul_single i y j)) _ = ϕ i y, - have h : finset.univ = insert i (finset.univ.erase i) := - (finset.insert_erase (finset.mem_univ i)).symm, - rw finset.noncomm_prod_congr h, - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ (finset.not_mem_erase i finset.univ), + obtain ⟨s, h⟩ : ∃ (s : finset ι), finset.univ = insert i (s.erase i) := ⟨_, + (finset.insert_erase (finset.mem_univ i)).symm⟩, + -- NB: need simp_rw, not rw, because of the dependent argument + -- NB: but then need to hide the finset.univ on the RHS of h, else simp_rw does not work + simp_rw h, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ (finset.not_mem_erase i _), rw pi.mul_single_eq_same, rw finset.noncomm_prod_eq_one_of_forall_eq_one, { exact mul_one _, }, From 8f678af60f1f4ab05014ecbb06262e2c58644c52 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 8 Mar 2022 12:15:36 +0100 Subject: [PATCH 106/116] Simplify further --- src/group_theory/noncomm_pi_coprod.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 7ae41fa69549e..454d2cee41234 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -93,11 +93,7 @@ lemma noncomm_pi_coprod_mul_single (i : ι) (y : N i): noncomm_pi_coprod ϕ hcomm (pi.mul_single i y) = ϕ i y := begin change finset.univ.noncomm_prod (λ j, ϕ j (pi.mul_single i y j)) _ = ϕ i y, - obtain ⟨s, h⟩ : ∃ (s : finset ι), finset.univ = insert i (s.erase i) := ⟨_, - (finset.insert_erase (finset.mem_univ i)).symm⟩, - -- NB: need simp_rw, not rw, because of the dependent argument - -- NB: but then need to hide the finset.univ on the RHS of h, else simp_rw does not work - simp_rw h, + simp only [←finset.insert_erase (finset.mem_univ i)] {single_pass := tt}, rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ (finset.not_mem_erase i _), rw pi.mul_single_eq_same, rw finset.noncomm_prod_eq_one_of_forall_eq_one, From b4c4c3a5dd303f56ee7684b6a0abbf62312d852b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 11 Mar 2022 12:05:05 +0100 Subject: [PATCH 107/116] Undo renaming, now happening in #12589 --- src/algebra/big_operators/multiset.lean | 5 ++--- src/algebra/big_operators/order.lean | 14 +++++++------- src/algebra/polynomial/big_operators.lean | 2 +- src/combinatorics/double_counting.lean | 4 ++-- src/data/list/prod_monoid.lean | 2 +- src/linear_algebra/matrix/polynomial.lean | 2 +- 6 files changed, 14 insertions(+), 15 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index f35c83d0359f0..ef6c8014b3e64 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -268,11 +268,10 @@ lemma single_le_prod : (∀ x ∈ s, (1 : α) ≤ x) → ∀ x ∈ s, x ≤ s.pr quotient.induction_on s $ λ l hl x hx, by simpa using list.single_le_prod hl x hx @[to_additive] -lemma prod_le_pow_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : - s.prod ≤ n ^ s.card := +lemma prod_le_of_forall_le (s : multiset α) (n : α) (h : ∀ x ∈ s, x ≤ n) : s.prod ≤ n ^ s.card := begin induction s using quotient.induction_on, - simpa using list.prod_le_pow_of_forall_le _ _ h, + simpa using list.prod_le_of_forall_le _ _ h, end @[to_additive all_zero_of_le_zero_le_of_sum_eq_zero] diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 456ff3cf7a65c..53c55b1661240 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -170,23 +170,23 @@ calc f a = ∏ i in {a}, f i : prod_singleton.symm prod_le_prod_of_subset_of_one_le' (singleton_subset_iff.2 h) $ λ i hi _, hf i hi @[to_additive] -lemma prod_le_pow_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : +lemma prod_le_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, f x ≤ n) : s.prod f ≤ n ^ s.card := begin - refine (multiset.prod_le_pow_of_forall_le (s.val.map f) n _).trans _, + refine (multiset.prod_le_of_forall_le (s.val.map f) n _).trans _, { simpa using h }, { simpa } end @[to_additive] -lemma pow_le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : +lemma le_prod_of_forall_le (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_pow_of_forall_le _ (order_dual N) _ _ _ _ h +@finset.prod_le_of_forall_le _ (order_dual N) _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : (s.bUnion f).card ≤ s.card * n := -card_bUnion_le.trans $ sum_le_nsmul_of_forall_le _ _ _ h +card_bUnion_le.trans $ sum_le_of_forall_le _ _ _ h variables {ι' : Type*} [decidable_eq ι'] @@ -264,7 +264,7 @@ times how many they are. -/ lemma sum_card_inter_le (h : ∀ a ∈ s, (B.filter $ (∈) a).card ≤ n) : ∑ t in B, (s ∩ t).card ≤ s.card * n := begin - refine le_trans _ (s.sum_le_nsmul_of_forall_le _ _ h), + refine le_trans _ (s.sum_le_of_forall_le _ _ h), simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end @@ -281,7 +281,7 @@ times how many they are. -/ lemma le_sum_card_inter (h : ∀ a ∈ s, n ≤ (B.filter $ (∈) a).card) : s.card * n ≤ ∑ t in B, (s ∩ t).card := begin - apply (s.nsmul_le_sum_of_forall_le _ _ h).trans, + apply (s.le_sum_of_forall_le _ _ h).trans, simp_rw [←filter_mem_eq_inter, card_eq_sum_ones, sum_filter], exact sum_comm.le, end diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 10e858e0911ea..5eb150eb43f56 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -96,7 +96,7 @@ begin have h : nat_degree tl.prod ≤ n * tl.length, { refine (nat_degree_list_prod_le _).trans _, rw [←tl.length_map nat_degree, mul_comm], - refine list.sum_le_nsmul_of_forall_le _ _ _, + refine list.sum_le_of_forall_le _ _ _, simpa using hl' }, have hdn : nat_degree hd ≤ n := hl _ (list.mem_cons_self _ _), rcases hdn.eq_or_lt with rfl|hdn', diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 55cf33ade699a..8f2eca8468c3e 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -61,10 +61,10 @@ lemma card_mul_le_card_mul [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, (s.bipartite_below r b).card ≤ n) : s.card * m ≤ t.card * n := calc - _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.nsmul_le_sum_of_forall_le _ _ hm + _ ≤ ∑ a in s, (t.bipartite_above r a).card : s.le_sum_of_forall_le _ _ hm ... = ∑ b in t, (s.bipartite_below r b).card : sum_card_bipartite_above_eq_sum_card_bipartite_below _ - ... ≤ _ : t.sum_le_nsmul_of_forall_le _ _ hn + ... ≤ _ : t.sum_le_of_forall_le _ _ hn lemma card_mul_le_card_mul' [Π a b, decidable (r a b)] (hn : ∀ b ∈ t, n ≤ (s.bipartite_below r b).card) diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index b08e4f5c2d2a5..197b19ca9a772 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -36,7 +36,7 @@ begin end @[to_additive] -lemma prod_le_pow_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : +lemma prod_le_of_forall_le [ordered_comm_monoid α] (l : list α) (n : α) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := begin induction l with y l IH, diff --git a/src/linear_algebra/matrix/polynomial.lean b/src/linear_algebra/matrix/polynomial.lean index 9341b6e3ecf5a..8e265d50ae826 100644 --- a/src/linear_algebra/matrix/polynomial.lean +++ b/src/linear_algebra/matrix/polynomial.lean @@ -49,7 +49,7 @@ begin { rw [sg, units.neg_smul, one_smul, nat_degree_neg] } } ... ≤ ∑ (i : n), nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) : nat_degree_prod_le (finset.univ : finset n) (λ (i : n), (X • A.map C + B.map C) (g i) i) - ... ≤ finset.univ.card • 1 : finset.sum_le_nsmul_of_forall_le _ _ 1 (λ (i : n) _, _) + ... ≤ finset.univ.card • 1 : finset.sum_le_of_forall_le _ _ 1 (λ (i : n) _, _) ... ≤ fintype.card n : by simpa, calc nat_degree (((X : α[X]) • A.map C + B.map C) (g i) i) = nat_degree ((X : α[X]) * C (A (g i) i) + C (B (g i) i)) : by simp From c5dd22f2f7f667c6910419c9ab53fed7866dd45b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 11 Mar 2022 12:06:48 +0100 Subject: [PATCH 108/116] Rename the new lemma --- src/data/list/prod_monoid.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/list/prod_monoid.lean b/src/data/list/prod_monoid.lean index 197b19ca9a772..3495328547234 100644 --- a/src/data/list/prod_monoid.lean +++ b/src/data/list/prod_monoid.lean @@ -27,8 +27,8 @@ begin { rw [list.repeat_succ, list.prod_cons, ih, pow_succ] } end -@[to_additive] -lemma prod_eq_pow_of_forall_eq [monoid α] (l : list α) (m : α) (h : ∀ (x ∈ l), x = m) : +@[to_additive sum_eq_card_nsmul] +lemma prod_eq_pow_card [monoid α] (l : list α) (m : α) (h : ∀ (x ∈ l), x = m) : l.prod = m ^ l.length := begin convert list.prod_repeat m l.length, From 804d7d883649dfdb862f8ef3e38408311a17ceff Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 11 Mar 2022 22:22:49 +0100 Subject: [PATCH 109/116] Rename lemmas according to latest fashion --- src/data/finset/noncomm_prod.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 7158de98de7dc..f661948cd088b 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -248,15 +248,15 @@ begin convert map_list_prod g _, rw list.comp_map, end -@[to_additive] -lemma noncomm_prod_eq_pow_of_forall_eq (s : finset α) (f : α → β) +@[to_additive noncomm_sum_eq_card_nsmul] +lemma noncomm_prod_eq_pow_card (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := begin classical, revert comm, rw ← s.to_list_to_finset, intro, rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, - rw list.prod_eq_pow_of_forall_eq _ m, + rw list.prod_eq_pow_card _ m, { simp }, { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x (s.mem_to_list.mp hx) }, end @@ -266,7 +266,7 @@ lemma noncomm_prod_eq_one_of_forall_eq_one (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := begin - rw finset.noncomm_prod_eq_pow_of_forall_eq s f comm 1 h, + rw finset.noncomm_prod_eq_pow_card s f comm 1 h, exact one_pow _, end From d80f84409f68474bc4ca888763f2418267e994ce Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Mar 2022 10:31:50 +0100 Subject: [PATCH 110/116] Multisetify noncomm_prod_map --- src/data/finset/noncomm_prod.lean | 32 +++++++++++++++++++++++-------- 1 file changed, 24 insertions(+), 8 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index f661948cd088b..85dcfdb1c4d98 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -106,7 +106,7 @@ begin end omit assoc -variables [monoid α] +variables [monoid α] [monoid β] /-- Product of a `s : multiset α` with `[monoid α]`, given a proof that `*` commutes on all elements `x ∈ s`. -/ @@ -163,6 +163,27 @@ begin simp [hy] } } } end +@[to_additive, protected] +lemma nocomm_prod_map_aux (s : multiset α) + (comm : ∀ (x ∈ s) (y ∈ s), commute x y) + {F : Type*} [monoid_hom_class F α β] (f : F) : + ∀ (x ∈ s.map f) (y ∈ s.map f), commute x y := +begin + simp only [multiset.mem_map], + rintros _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩, + exact (comm _ hx _ hy).map f, +end + +@[to_additive] +lemma noncomm_prod_map (s : multiset α) + (comm : ∀ (x ∈ s) (y ∈ s), commute x y) + {F : Type*} [monoid_hom_class F α β] (f : F) : + f (s.noncomm_prod comm) = (s.map f).noncomm_prod (nocomm_prod_map_aux s comm f) := +begin + induction s using quotient.induction_on, + simpa using map_list_prod f _, +end + @[to_additive] lemma noncomm_prod_eq_prod {α : Type*} [comm_monoid α] (s : multiset α) : noncomm_prod s (λ _ _ _ _, commute.all _ _) = prod s := begin @@ -240,13 +261,8 @@ lemma noncomm_prod_map (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) {F : Type*} [monoid_hom_class F β γ] (g : F) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) - (λ x hx y hy, commute.map (comm x hx y hy) g) := -begin - classical, - revert comm, rw ← s.to_list_to_finset, intro, - iterate 2 { rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list }, - convert map_list_prod g _, rw list.comp_map, -end + (λ x hx y hy, (comm x hx y hy).map g) := +by simp [noncomm_prod, multiset.noncomm_prod_map], @[to_additive noncomm_sum_eq_card_nsmul] lemma noncomm_prod_eq_pow_card (s : finset α) (f : α → β) From 4961f66fda35e5ffd2863c2bd50f05e42596b5d8 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Mar 2022 10:41:42 +0100 Subject: [PATCH 111/116] Multisetify noncomm_prod_eq_pow_card --- src/data/finset/noncomm_prod.lean | 23 ++++++++++++++--------- 1 file changed, 14 insertions(+), 9 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 85dcfdb1c4d98..ecedb181b8b17 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -175,8 +175,7 @@ begin end @[to_additive] -lemma noncomm_prod_map (s : multiset α) - (comm : ∀ (x ∈ s) (y ∈ s), commute x y) +lemma noncomm_prod_map (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s), commute x y) {F : Type*} [monoid_hom_class F α β] (f : F) : f (s.noncomm_prod comm) = (s.map f).noncomm_prod (nocomm_prod_map_aux s comm f) := begin @@ -184,6 +183,15 @@ begin simpa using map_list_prod f _, end +@[to_additive noncomm_sum_eq_card_nsmul] +lemma noncomm_prod_eq_pow_card (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s), commute x y) + (m : α) (h : ∀ (x ∈ s), x = m) : s.noncomm_prod comm = m ^ s.card := +begin + induction s using quotient.induction_on, + simp only [quot_mk_to_coe, noncomm_prod_coe, coe_card, mem_coe] at *, + exact list.prod_eq_pow_card _ m h, +end + @[to_additive] lemma noncomm_prod_eq_prod {α : Type*} [comm_monoid α] (s : multiset α) : noncomm_prod s (λ _ _ _ _, commute.all _ _) = prod s := begin @@ -262,19 +270,16 @@ lemma noncomm_prod_map (s : finset α) (f : α → β) {F : Type*} [monoid_hom_class F β γ] (g : F) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) (λ x hx y hy, (comm x hx y hy).map g) := -by simp [noncomm_prod, multiset.noncomm_prod_map], +by simp [noncomm_prod, multiset.noncomm_prod_map] @[to_additive noncomm_sum_eq_card_nsmul] lemma noncomm_prod_eq_pow_card (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := begin - classical, - revert comm, rw ← s.to_list_to_finset, intro, - rw finset.noncomm_prod_to_finset _ _ _ s.nodup_to_list, - rw list.prod_eq_pow_card _ m, - { simp }, - { intro _, rw list.mem_map, rintros ⟨x, hx, rfl⟩, apply h x (s.mem_to_list.mp hx) }, + rw [noncomm_prod, multiset.noncomm_prod_eq_pow_card _ _ m], + simp only [finset.card_def, multiset.card_map], + simpa using h, end @[to_additive] From ab6f6395f442e39ae708d682fa3d4e22bfe3ee8f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Mar 2022 10:42:28 +0100 Subject: [PATCH 112/116] Golf --- src/data/finset/noncomm_prod.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index ecedb181b8b17..3ea408d4f5dc0 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -286,10 +286,7 @@ end lemma noncomm_prod_eq_one_of_forall_eq_one (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := -begin - rw finset.noncomm_prod_eq_pow_card s f comm 1 h, - exact one_pow _, -end +by rw [noncomm_prod_eq_pow_card s f comm 1 h, one_pow] @[to_additive] lemma noncomm_prod_commute (s : finset α) (f : α → β) From 700f30989d60680d89c5867e7091b64ddfb0cafe Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Mar 2022 10:44:27 +0100 Subject: [PATCH 113/116] Avoid rather specialized lemma --- src/data/finset/noncomm_prod.lean | 9 ++------- 1 file changed, 2 insertions(+), 7 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 3ea408d4f5dc0..9966552e4a3ae 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -282,12 +282,6 @@ begin simpa using h, end -@[to_additive] -lemma noncomm_prod_eq_one_of_forall_eq_one (s : finset α) (f : α → β) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (h : ∀ (x : α), x ∈ s → f x = 1) : s.noncomm_prod f comm = 1 := -by rw [noncomm_prod_eq_pow_card s f comm 1 h, one_pow] - @[to_additive] lemma noncomm_prod_commute (s : finset α) (f : α → β) (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) @@ -387,7 +381,8 @@ begin (pi.eval_monoid_hom M i)).trans, rw [ ← insert_erase (mem_univ i), noncomm_prod_insert_of_not_mem' _ _ _ _ (not_mem_erase _ _), - noncomm_prod_eq_one_of_forall_eq_one ], + noncomm_prod_eq_pow_card, + one_pow], { simp, }, { intros i h, simp at h, simp [h], }, end From 92719f45bf240b404c8ed7c8ac307fdbfa337315 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sat, 12 Mar 2022 21:28:41 +0100 Subject: [PATCH 114/116] Use pi.mul_single_apply_commute --- src/data/finset/noncomm_prod.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 9966552e4a3ae..dcc5f05b87bcf 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -372,8 +372,7 @@ variables (x : Π i, M i) @[to_additive] lemma noncomm_prod_mul_single : - univ.noncomm_prod (λ i, pi.mul_single i (x i)) - (λ i _ j _, by { by_cases h : i = j, { rw h }, { apply pi.mul_single_commute i j h } }) + univ.noncomm_prod (λ i, pi.mul_single i (x i)) (λ i _ j _, pi.mul_single_apply_commute x i j) = x := begin ext i, From bf98c6095eee201a3ed1170f9673f3ae09d6db22 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 14 Mar 2022 09:16:49 +0100 Subject: [PATCH 115/116] Apply suggestions from code review Co-authored-by: Johan Commelin --- src/algebra/group/pi.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index b299141674c4a..b754de48ff589 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -247,7 +247,7 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : /-- The injection into a pi group at different indices commutes. For injections of commuting elements at the same index, see `commute.map` -/ -@[to_additive "The injection into an pi additive group at different indices commutes. +@[to_additive "The injection into an additive pi group at different indices commutes. For injections of commuting elements at the same index, see `add_commute.map`"] lemma pi.mul_single_commute [Π i, mul_one_class $ f i] : @@ -260,7 +260,7 @@ begin end /-- The injection into a pi group with the same values commutes. -/ -@[to_additive "The injection into a pi additive group with the same values commutes."] +@[to_additive "The injection into an additive pi group with the same values commutes."] lemma pi.mul_single_apply_commute [Π i, mul_one_class $ f i] (x : Π i, f i) (i j : I) : commute (mul_single i (x i)) (mul_single j (x j)) := begin From 91973158db336365017d0f958aff97d802aad60a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 15 Mar 2022 10:36:16 +0100 Subject: [PATCH 116/116] fix partly embarrassing typos Co-authored-by: Scott Morrison --- src/group_theory/noncomm_pi_coprod.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 6bd5d6c264da7..4149cb880d908 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -1,5 +1,5 @@ /- -Copyright (c) 2022 Jocchim Breitner. All rights reserved. +Copyright (c) 2022 Joachim Breitner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joachim Breitner -/ @@ -10,10 +10,10 @@ import data.fintype.card /-! # Canonical homomorphism from a finite family of monoids -This file defines the construction of the canoncial homomorphism from a family of monoids. +This file defines the construction of the canonical homomorphism from a family of monoids. Given a family of morphisms `ϕ i : N i →* M` for each `i : ι` where elements in the -image of different morphism commute, we obtain a canoncial morphism +images of different morphisms commute, we obtain a canonical morphism `monoid_hom.noncomm_pi_coprod : (Π i, N i) →* M` that coincides with `ϕ` ## Main definitions @@ -24,7 +24,7 @@ image of different morphism commute, we obtain a canoncial morphism ## Main theorems -* `monoid_hom.noncomm_pi_coprod` conicides with `ϕ i` when restricted to `N i` +* `monoid_hom.noncomm_pi_coprod` coincides with `ϕ i` when restricted to `N i` * `monoid_hom.noncomm_pi_coprod_mrange`: The range of `monoid_hom.noncomm_pi_coprod` is `⨆ (i : ι), (ϕ i).mrange` * `monoid_hom.noncomm_pi_coprod_range`: The range of `monoid_hom.noncomm_pi_coprod` is @@ -34,7 +34,7 @@ image of different morphism commute, we obtain a canoncial morphism injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. * `monoid_hom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges of the `ϕ` are independent. -* `subgroup.independent_of_coprime_order`: If commuting, normal subgroups `H i` have coprime orders, +* `subgroup.independent_of_coprime_order`: If commuting normal subgroups `H i` have coprime orders, they are independent. -/