From d17ef230b88237d72224ec2ebeee6e0f54611b4a Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 23 Jan 2023 21:16:38 +0100 Subject: [PATCH 01/34] feat: port GroupTheory.Subgroup.Basic From 853951ab19880e7d380a71595661ffb4dec09a44 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 23 Jan 2023 21:16:38 +0100 Subject: [PATCH 02/34] Initial file copy from mathport --- Mathlib/GroupTheory/Subgroup/Basic.lean | 3693 +++++++++++++++++++++++ 1 file changed, 3693 insertions(+) create mode 100644 Mathlib/GroupTheory/Subgroup/Basic.lean diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean new file mode 100644 index 0000000000000..b8816a5d6a3ec --- /dev/null +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -0,0 +1,3693 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying + +! This file was ported from Lean 3 source module group_theory.subgroup.basic +! leanprover-community/mathlib commit 1f0096e6caa61e9c849ec2adbd227e960e9dff58 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Algebra.Group.Conj +import Mathbin.Algebra.Module.Basic +import Mathbin.Algebra.Order.Group.InjSurj +import Mathbin.Data.Countable.Basic +import Mathbin.GroupTheory.Submonoid.Centralizer +import Mathbin.Logic.Encodable.Basic +import Mathbin.Order.Atoms +import Mathbin.Tactic.ApplyFun + +/-! +# Subgroups + +This file defines multiplicative and additive subgroups as an extension of submonoids, in a bundled +form (unbundled subgroups are in `deprecated/subgroups.lean`). + +We prove subgroups of a group form a complete lattice, and results about images and preimages of +subgroups under group homomorphisms. The bundled subgroups use bundled monoid homomorphisms. + +There are also theorems about the subgroups generated by an element or a subset of a group, +defined both inductively and as the infimum of the set of subgroups containing a given +element/subset. + +Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and inspiration. + +## Main definitions + +Notation used here: + +- `G N` are `group`s + +- `A` is an `add_group` + +- `H K` are `subgroup`s of `G` or `add_subgroup`s of `A` + +- `x` is an element of type `G` or type `A` + +- `f g : N →* G` are group homomorphisms + +- `s k` are sets of elements of type `G` + +Definitions in the file: + +* `subgroup G` : the type of subgroups of a group `G` + +* `add_subgroup A` : the type of subgroups of an additive group `A` + +* `complete_lattice (subgroup G)` : the subgroups of `G` form a complete lattice + +* `subgroup.closure k` : the minimal subgroup that includes the set `k` + +* `subgroup.subtype` : the natural group homomorphism from a subgroup of group `G` to `G` + +* `subgroup.gi` : `closure` forms a Galois insertion with the coercion to set + +* `subgroup.comap H f` : the preimage of a subgroup `H` along the group homomorphism `f` is also a + subgroup + +* `subgroup.map f H` : the image of a subgroup `H` along the group homomorphism `f` is also a + subgroup + +* `subgroup.prod H K` : the product of subgroups `H`, `K` of groups `G`, `N` respectively, `H × K` + is a subgroup of `G × N` + +* `monoid_hom.range f` : the range of the group homomorphism `f` is a subgroup + +* `monoid_hom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` + such that `f x = 1` + +* `monoid_hom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that + `f x = g x` form a subgroup of `G` + +## Implementation notes + +Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as +membership of a subgroup's underlying set. + +## Tags +subgroup, subgroups +-/ + + +open Function + +variable {G G' : Type _} [Group G] [Group G'] + +variable {A : Type _} [AddGroup A] + +section SubgroupClass + +/-- `inv_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ +class InvMemClass (S G : Type _) [Inv G] [SetLike S G] : Prop where + inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s +#align inv_mem_class InvMemClass + +export InvMemClass (inv_mem) + +/-- `neg_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ +class NegMemClass (S G : Type _) [Neg G] [SetLike S G] : Prop where + neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s +#align neg_mem_class NegMemClass + +export NegMemClass (neg_mem) + +/-- `subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ +class SubgroupClass (S G : Type _) [DivInvMonoid G] [SetLike S G] extends SubmonoidClass S G, + InvMemClass S G : Prop +#align subgroup_class SubgroupClass + +/-- `add_subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are +additive subgroups of `G`. -/ +class AddSubgroupClass (S G : Type _) [SubNegMonoid G] [SetLike S G] extends AddSubmonoidClass S G, + NegMemClass S G : Prop +#align add_subgroup_class AddSubgroupClass + +attribute [to_additive] InvMemClass SubgroupClass + +variable {M S : Type _} [DivInvMonoid M] [SetLike S M] [hSM : SubgroupClass S M] {H K : S} + +include hSM + +/-- A subgroup is closed under division. -/ +@[to_additive "An additive subgroup is closed under subtraction."] +theorem div_mem {x y : M} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := by + rw [div_eq_mul_inv] <;> exact mul_mem hx (inv_mem hy) +#align div_mem div_mem +#align sub_mem sub_mem + +@[to_additive] +theorem zpow_mem {x : M} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K + | (n : ℕ) => by + rw [zpow_ofNat] + exact pow_mem hx n + | -[n+1] => by + rw [zpow_negSucc] + exact inv_mem (pow_mem hx n.succ) +#align zpow_mem zpow_mem +#align zsmul_mem zsmul_mem + +omit hSM + +variable [SetLike S G] [hSG : SubgroupClass S G] + +include hSG + +@[simp, to_additive] +theorem inv_mem_iff {x : G} : x⁻¹ ∈ H ↔ x ∈ H := + ⟨fun h => inv_inv x ▸ inv_mem h, inv_mem⟩ +#align inv_mem_iff inv_mem_iff +#align neg_mem_iff neg_mem_iff + +@[to_additive] +theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := by + rw [← inv_mem_iff, div_eq_mul_inv, div_eq_mul_inv, mul_inv_rev, inv_inv] +#align div_mem_comm_iff div_mem_comm_iff +#align sub_mem_comm_iff sub_mem_comm_iff + +@[simp, to_additive] +theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ P x⁻¹) ↔ ∃ x ∈ H, P x := by + constructor <;> + · rintro ⟨x, x_in, hx⟩ + exact ⟨x⁻¹, inv_mem x_in, by simp [hx]⟩ +#align exists_inv_mem_iff_exists_mem exists_inv_mem_iff_exists_mem +#align exists_neg_mem_iff_exists_mem exists_neg_mem_iff_exists_mem + +@[to_additive] +theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := + ⟨fun hba => by simpa using mul_mem hba (inv_mem h), fun hb => mul_mem hb h⟩ +#align mul_mem_cancel_right mul_mem_cancel_right +#align add_mem_cancel_right add_mem_cancel_right + +@[to_additive] +theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := + ⟨fun hab => by simpa using mul_mem (inv_mem h) hab, mul_mem h⟩ +#align mul_mem_cancel_left mul_mem_cancel_left +#align add_mem_cancel_left add_mem_cancel_left + +namespace SubgroupClass + +omit hSG + +include hSM + +/-- A subgroup of a group inherits an inverse. -/ +@[to_additive "An additive subgroup of a `add_group` inherits an inverse."] +instance hasInv : Inv H := + ⟨fun a => ⟨a⁻¹, inv_mem a.2⟩⟩ +#align subgroup_class.has_inv SubgroupClass.hasInv +#align add_subgroup_class.has_neg AddSubgroupClass.hasNeg + +/-- A subgroup of a group inherits a division -/ +@[to_additive "An additive subgroup of an `add_group` inherits a subtraction."] +instance hasDiv : Div H := + ⟨fun a b => ⟨a / b, div_mem a.2 b.2⟩⟩ +#align subgroup_class.has_div SubgroupClass.hasDiv +#align add_subgroup_class.has_sub AddSubgroupClass.hasSub + +omit hSM + +/-- An additive subgroup of an `add_group` inherits an integer scaling. -/ +instance AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] + {H : S} : SMul ℤ H := + ⟨fun n a => ⟨n • a, zsmul_mem a.2 n⟩⟩ +#align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul + +include hSM + +/-- A subgroup of a group inherits an integer power. -/ +@[to_additive] +instance hasZpow : Pow H ℤ := + ⟨fun a n => ⟨a ^ n, zpow_mem a.2 n⟩⟩ +#align subgroup_class.has_zpow SubgroupClass.hasZpow +#align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul + +@[simp, norm_cast, to_additive] +theorem coe_inv (x : H) : ↑(x⁻¹ : H) = (x⁻¹ : M) := + rfl +#align subgroup_class.coe_inv SubgroupClass.coe_inv +#align add_subgroup_class.coe_neg AddSubgroupClass.coe_neg + +@[simp, norm_cast, to_additive] +theorem coe_div (x y : H) : (↑(x / y) : M) = ↑x / ↑y := + rfl +#align subgroup_class.coe_div SubgroupClass.coe_div +#align add_subgroup_class.coe_sub AddSubgroupClass.coe_sub + +omit hSM + +variable (H) + +include hSG + +-- Prefer subclasses of `group` over subclasses of `subgroup_class`. +/-- A subgroup of a group inherits a group structure. -/ +@[to_additive "An additive subgroup of an `add_group` inherits an `add_group` structure."] +instance (priority := 75) toGroup : Group H := + Subtype.coe_injective.Group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup_class.to_group SubgroupClass.toGroup +#align add_subgroup_class.to_add_group AddSubgroupClass.toAddGroup + +omit hSG + +-- Prefer subclasses of `comm_group` over subclasses of `subgroup_class`. +/-- A subgroup of a `comm_group` is a `comm_group`. -/ +@[to_additive "An additive subgroup of an `add_comm_group` is an `add_comm_group`."] +instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [SubgroupClass S G] : + CommGroup H := + Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup_class.to_comm_group SubgroupClass.toCommGroup +#align add_subgroup_class.to_add_comm_group AddSubgroupClass.toAddCommGroup + +-- Prefer subclasses of `group` over subclasses of `subgroup_class`. +/-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/ +@[to_additive "An additive subgroup of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] +instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [SetLike S G] + [SubgroupClass S G] : OrderedCommGroup H := + Subtype.coe_injective.OrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup_class.to_ordered_comm_group SubgroupClass.toOrderedCommGroup +#align add_subgroup_class.to_ordered_add_comm_group AddSubgroupClass.toOrderedAddCommGroup + +-- Prefer subclasses of `group` over subclasses of `subgroup_class`. +/-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ +@[to_additive + "An additive subgroup of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] +instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] + [SetLike S G] [SubgroupClass S G] : LinearOrderedCommGroup H := + Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl +#align subgroup_class.to_linear_ordered_comm_group SubgroupClass.toLinearOrderedCommGroup +#align add_subgroup_class.to_linear_ordered_add_comm_group AddSubgroupClass.toLinearOrderedAddCommGroup + +include hSG + +/-- The natural group hom from a subgroup of group `G` to `G`. -/ +@[to_additive "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] +def subtype : H →* G := + ⟨coe, rfl, fun _ _ => rfl⟩ +#align subgroup_class.subtype SubgroupClass.subtype +#align add_subgroup_class.subtype AddSubgroupClass.subtype + +@[simp, to_additive] +theorem coeSubtype : (subtype H : H → G) = coe := + rfl +#align subgroup_class.coe_subtype SubgroupClass.coeSubtype +#align add_subgroup_class.coe_subtype AddSubgroupClass.coeSubtype + +variable {H} + +@[simp, norm_cast, to_additive coe_smul] +theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := + (subtype H : H →* G).map_pow _ _ +#align subgroup_class.coe_pow SubgroupClass.coe_pow +#align add_subgroup_class.coe_smul AddSubgroupClass.coe_smul + +@[simp, norm_cast, to_additive] +theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := + (subtype H : H →* G).map_zpow _ _ +#align subgroup_class.coe_zpow SubgroupClass.coe_zpow +#align add_subgroup_class.coe_zsmul AddSubgroupClass.coe_zsmul + +/-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ +@[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] +def inclusion {H K : S} (h : H ≤ K) : H →* K := + MonoidHom.mk' (fun x => ⟨x, h x.Prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl +#align subgroup_class.inclusion SubgroupClass.inclusion +#align add_subgroup_class.inclusion AddSubgroupClass.inclusion + +@[simp, to_additive] +theorem inclusion_self (x : H) : inclusion le_rfl x = x := + by + cases x + rfl +#align subgroup_class.inclusion_self SubgroupClass.inclusion_self +#align add_subgroup_class.inclusion_self AddSubgroupClass.inclusion_self + +@[simp, to_additive] +theorem inclusion_mk {h : H ≤ K} (x : G) (hx : x ∈ H) : inclusion h ⟨x, hx⟩ = ⟨x, h hx⟩ := + rfl +#align subgroup_class.inclusion_mk SubgroupClass.inclusion_mk +#align add_subgroup_class.inclusion_mk AddSubgroupClass.inclusion_mk + +@[to_additive] +theorem inclusion_right (h : H ≤ K) (x : K) (hx : (x : G) ∈ H) : inclusion h ⟨x, hx⟩ = x := + by + cases x + rfl +#align subgroup_class.inclusion_right SubgroupClass.inclusion_right +#align add_subgroup_class.inclusion_right AddSubgroupClass.inclusion_right + +@[simp] +theorem inclusion_inclusion {L : S} (hHK : H ≤ K) (hKL : K ≤ L) (x : H) : + inclusion hKL (inclusion hHK x) = inclusion (hHK.trans hKL) x := + by + cases x + rfl +#align subgroup_class.inclusion_inclusion SubgroupClass.inclusion_inclusion + +@[simp, to_additive] +theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := + by + cases a + simp only [inclusion, [anonymous], MonoidHom.mk'_apply] +#align subgroup_class.coe_inclusion SubgroupClass.coe_inclusion +#align add_subgroup_class.coe_inclusion AddSubgroupClass.coe_inclusion + +@[simp, to_additive] +theorem subtype_comp_inclusion {H K : S} (hH : H ≤ K) : + (subtype K).comp (inclusion hH) = subtype H := + by + ext + simp only [MonoidHom.comp_apply, coeSubtype, coe_inclusion] +#align subgroup_class.subtype_comp_inclusion SubgroupClass.subtype_comp_inclusion +#align add_subgroup_class.subtype_comp_inclusion AddSubgroupClass.subtype_comp_inclusion + +end SubgroupClass + +end SubgroupClass + +/-- A subgroup of a group `G` is a subset containing 1, closed under multiplication +and closed under multiplicative inverse. -/ +structure Subgroup (G : Type _) [Group G] extends Submonoid G where + inv_mem' {x} : x ∈ carrier → x⁻¹ ∈ carrier +#align subgroup Subgroup + +/-- An additive subgroup of an additive group `G` is a subset containing 0, closed +under addition and additive inverse. -/ +structure AddSubgroup (G : Type _) [AddGroup G] extends AddSubmonoid G where + neg_mem' {x} : x ∈ carrier → -x ∈ carrier +#align add_subgroup AddSubgroup + +attribute [to_additive] Subgroup + +attribute [to_additive AddSubgroup.toAddSubmonoid] Subgroup.toSubmonoid + +/-- Reinterpret a `subgroup` as a `submonoid`. -/ +add_decl_doc Subgroup.toSubmonoid + +/-- Reinterpret an `add_subgroup` as an `add_submonoid`. -/ +add_decl_doc AddSubgroup.toAddSubmonoid + +namespace Subgroup + +@[to_additive] +instance : SetLike (Subgroup G) G where + coe := Subgroup.carrier + coe_injective' p q h := by cases p <;> cases q <;> congr + +@[to_additive] +instance : SubgroupClass (Subgroup G) G + where + mul_mem := Subgroup.mul_mem' + one_mem := Subgroup.one_mem' + inv_mem := Subgroup.inv_mem' + +@[simp, to_additive] +theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := + Iff.rfl +#align subgroup.mem_carrier Subgroup.mem_carrier +#align add_subgroup.mem_carrier AddSubgroup.mem_carrier + +@[simp, to_additive] +theorem mem_mk {s : Set G} {x : G} (h_one) (h_mul) (h_inv) : x ∈ mk s h_one h_mul h_inv ↔ x ∈ s := + Iff.rfl +#align subgroup.mem_mk Subgroup.mem_mk +#align add_subgroup.mem_mk AddSubgroup.mem_mk + +@[simp, to_additive] +theorem coe_set_mk {s : Set G} (h_one) (h_mul) (h_inv) : (mk s h_one h_mul h_inv : Set G) = s := + rfl +#align subgroup.coe_set_mk Subgroup.coe_set_mk +#align add_subgroup.coe_set_mk AddSubgroup.coe_set_mk + +@[simp, to_additive] +theorem mk_le_mk {s t : Set G} (h_one) (h_mul) (h_inv) (h_one') (h_mul') (h_inv') : + mk s h_one h_mul h_inv ≤ mk t h_one' h_mul' h_inv' ↔ s ⊆ t := + Iff.rfl +#align subgroup.mk_le_mk Subgroup.mk_le_mk +#align add_subgroup.mk_le_mk AddSubgroup.mk_le_mk + +/-- See Note [custom simps projection] -/ +@[to_additive "See Note [custom simps projection]"] +def Simps.coe (S : Subgroup G) : Set G := + S +#align subgroup.simps.coe Subgroup.Simps.coe +#align add_subgroup.simps.coe AddSubgroup.Simps.coe + +initialize_simps_projections Subgroup (carrier → coe) + +initialize_simps_projections AddSubgroup (carrier → coe) + +@[simp, to_additive] +theorem coe_toSubmonoid (K : Subgroup G) : (K.toSubmonoid : Set G) = K := + rfl +#align subgroup.coe_to_submonoid Subgroup.coe_toSubmonoid +#align add_subgroup.coe_to_add_submonoid AddSubgroup.coe_to_add_submonoid + +@[simp, to_additive] +theorem mem_toSubmonoid (K : Subgroup G) (x : G) : x ∈ K.toSubmonoid ↔ x ∈ K := + Iff.rfl +#align subgroup.mem_to_submonoid Subgroup.mem_toSubmonoid +#align add_subgroup.mem_to_add_submonoid AddSubgroup.mem_to_add_submonoid + +@[to_additive] +theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subgroup G → Submonoid G) := + fun p q h => SetLike.ext'_iff.2 (show _ from SetLike.ext'_iff.1 h) +#align subgroup.to_submonoid_injective Subgroup.toSubmonoid_injective +#align add_subgroup.to_add_submonoid_injective AddSubgroup.to_add_submonoid_injective + +@[simp, to_additive] +theorem toSubmonoid_eq {p q : Subgroup G} : p.toSubmonoid = q.toSubmonoid ↔ p = q := + toSubmonoid_injective.eq_iff +#align subgroup.to_submonoid_eq Subgroup.toSubmonoid_eq +#align add_subgroup.to_add_submonoid_eq AddSubgroup.to_add_submonoid_eq + +@[to_additive, mono] +theorem toSubmonoid_strictMono : StrictMono (toSubmonoid : Subgroup G → Submonoid G) := fun _ _ => + id +#align subgroup.to_submonoid_strict_mono Subgroup.toSubmonoid_strictMono +#align add_subgroup.to_add_submonoid_strict_mono AddSubgroup.to_add_submonoid_strictMono + +attribute [mono] AddSubgroup.to_add_submonoid_strictMono + +@[to_additive, mono] +theorem toSubmonoid_mono : Monotone (toSubmonoid : Subgroup G → Submonoid G) := + toSubmonoid_strictMono.Monotone +#align subgroup.to_submonoid_mono Subgroup.toSubmonoid_mono +#align add_subgroup.to_add_submonoid_mono AddSubgroup.to_add_submonoid_mono + +attribute [mono] AddSubgroup.to_add_submonoid_mono + +@[simp, to_additive] +theorem toSubmonoid_le {p q : Subgroup G} : p.toSubmonoid ≤ q.toSubmonoid ↔ p ≤ q := + Iff.rfl +#align subgroup.to_submonoid_le Subgroup.toSubmonoid_le +#align add_subgroup.to_add_submonoid_le AddSubgroup.to_add_submonoid_le + +end Subgroup + +/-! +### Conversion to/from `additive`/`multiplicative` +-/ + + +section mul_add + +/-- Supgroups of a group `G` are isomorphic to additive subgroups of `additive G`. -/ +@[simps] +def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) + where + toFun S := { S.toSubmonoid.toAddSubmonoid with neg_mem' := fun _ => S.inv_mem' } + invFun S := { S.toAddSubmonoid.toSubmonoid' with inv_mem' := fun _ => S.neg_mem' } + left_inv x := by cases x <;> rfl + right_inv x := by cases x <;> rfl + map_rel_iff' a b := Iff.rfl +#align subgroup.to_add_subgroup Subgroup.toAddSubgroup + +/-- Additive subgroup of an additive group `additive G` are isomorphic to subgroup of `G`. -/ +abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := + Subgroup.toAddSubgroup.symm +#align add_subgroup.to_subgroup' AddSubgroup.toSubgroup' + +/-- Additive supgroups of an additive group `A` are isomorphic to subgroups of `multiplicative A`. +-/ +@[simps] +def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) + where + toFun S := { S.toAddSubmonoid.toSubmonoid with inv_mem' := fun _ => S.neg_mem' } + invFun S := { S.toSubmonoid.toAddSubmonoid' with neg_mem' := fun _ => S.inv_mem' } + left_inv x := by cases x <;> rfl + right_inv x := by cases x <;> rfl + map_rel_iff' a b := Iff.rfl +#align add_subgroup.to_subgroup AddSubgroup.toSubgroup + +/-- Subgroups of an additive group `multiplicative A` are isomorphic to additive subgroups of `A`. +-/ +abbrev Subgroup.toAddSubgroup' : Subgroup (Multiplicative A) ≃o AddSubgroup A := + AddSubgroup.toSubgroup.symm +#align subgroup.to_add_subgroup' Subgroup.toAddSubgroup' + +end mul_add + +namespace Subgroup + +variable (H K : Subgroup G) + +/-- Copy of a subgroup with a new `carrier` equal to the old one. Useful to fix definitional +equalities.-/ +@[to_additive + "Copy of an additive subgroup with a new `carrier` equal to the old one.\nUseful to fix definitional equalities"] +protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G + where + carrier := s + one_mem' := hs.symm ▸ K.one_mem' + mul_mem' _ _ := hs.symm ▸ K.mul_mem' + inv_mem' _ := hs.symm ▸ K.inv_mem' +#align subgroup.copy Subgroup.copy +#align add_subgroup.copy AddSubgroup.copy + +@[simp, to_additive] +theorem coe_copy (K : Subgroup G) (s : Set G) (hs : s = ↑K) : (K.copy s hs : Set G) = s := + rfl +#align subgroup.coe_copy Subgroup.coe_copy +#align add_subgroup.coe_copy AddSubgroup.coe_copy + +@[to_additive] +theorem copy_eq (K : Subgroup G) (s : Set G) (hs : s = ↑K) : K.copy s hs = K := + SetLike.coe_injective hs +#align subgroup.copy_eq Subgroup.copy_eq +#align add_subgroup.copy_eq AddSubgroup.copy_eq + +/-- Two subgroups are equal if they have the same elements. -/ +@[ext, to_additive "Two `add_subgroup`s are equal if they have the same elements."] +theorem ext {H K : Subgroup G} (h : ∀ x, x ∈ H ↔ x ∈ K) : H = K := + SetLike.ext h +#align subgroup.ext Subgroup.ext +#align add_subgroup.ext AddSubgroup.ext + +/-- A subgroup contains the group's 1. -/ +@[to_additive "An `add_subgroup` contains the group's 0."] +protected theorem one_mem : (1 : G) ∈ H := + one_mem _ +#align subgroup.one_mem Subgroup.one_mem +#align add_subgroup.zero_mem AddSubgroup.zero_mem + +/-- A subgroup is closed under multiplication. -/ +@[to_additive "An `add_subgroup` is closed under addition."] +protected theorem mul_mem {x y : G} : x ∈ H → y ∈ H → x * y ∈ H := + mul_mem +#align subgroup.mul_mem Subgroup.mul_mem +#align add_subgroup.add_mem AddSubgroup.add_mem + +/-- A subgroup is closed under inverse. -/ +@[to_additive "An `add_subgroup` is closed under inverse."] +protected theorem inv_mem {x : G} : x ∈ H → x⁻¹ ∈ H := + inv_mem +#align subgroup.inv_mem Subgroup.inv_mem +#align add_subgroup.neg_mem AddSubgroup.neg_mem + +/-- A subgroup is closed under division. -/ +@[to_additive "An `add_subgroup` is closed under subtraction."] +protected theorem div_mem {x y : G} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := + div_mem hx hy +#align subgroup.div_mem Subgroup.div_mem +#align add_subgroup.sub_mem AddSubgroup.sub_mem + +@[to_additive] +protected theorem inv_mem_iff {x : G} : x⁻¹ ∈ H ↔ x ∈ H := + inv_mem_iff +#align subgroup.inv_mem_iff Subgroup.inv_mem_iff +#align add_subgroup.neg_mem_iff AddSubgroup.neg_mem_iff + +@[to_additive] +protected theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := + div_mem_comm_iff +#align subgroup.div_mem_comm_iff Subgroup.div_mem_comm_iff +#align add_subgroup.sub_mem_comm_iff AddSubgroup.sub_mem_comm_iff + +@[to_additive] +protected theorem exists_inv_mem_iff_exists_mem (K : Subgroup G) {P : G → Prop} : + (∃ x : G, x ∈ K ∧ P x⁻¹) ↔ ∃ x ∈ K, P x := + exists_inv_mem_iff_exists_mem +#align subgroup.exists_inv_mem_iff_exists_mem Subgroup.exists_inv_mem_iff_exists_mem +#align add_subgroup.exists_neg_mem_iff_exists_mem AddSubgroup.exists_neg_mem_iff_exists_mem + +@[to_additive] +protected theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := + mul_mem_cancel_right h +#align subgroup.mul_mem_cancel_right Subgroup.mul_mem_cancel_right +#align add_subgroup.add_mem_cancel_right AddSubgroup.add_mem_cancel_right + +@[to_additive] +protected theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := + mul_mem_cancel_left h +#align subgroup.mul_mem_cancel_left Subgroup.mul_mem_cancel_left +#align add_subgroup.add_mem_cancel_left AddSubgroup.add_mem_cancel_left + +@[to_additive AddSubgroup.nsmul_mem] +protected theorem pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := + pow_mem hx +#align subgroup.pow_mem Subgroup.pow_mem +#align add_subgroup.nsmul_mem AddSubgroup.nsmul_mem + +@[to_additive] +protected theorem zpow_mem {x : G} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K := + zpow_mem hx +#align subgroup.zpow_mem Subgroup.zpow_mem +#align add_subgroup.zsmul_mem AddSubgroup.zsmul_mem + +/- ./././Mathport/Syntax/Translate/Basic.lean:632:2: warning: expanding binder collection (x y «expr ∈ » s) -/ +/-- Construct a subgroup from a nonempty set that is closed under division. -/ +@[to_additive "Construct a subgroup from a nonempty set that is closed under subtraction"] +def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y ∈ s), x * y⁻¹ ∈ s) : + Subgroup G := + have one_mem : (1 : G) ∈ s := by + let ⟨x, hx⟩ := hsn + simpa using hs x hx x hx + have inv_mem : ∀ x, x ∈ s → x⁻¹ ∈ s := fun x hx => by simpa using hs 1 one_mem x hx + { carrier := s + one_mem' := one_mem + inv_mem' := inv_mem + mul_mem' := fun x y hx hy => by simpa using hs x hx y⁻¹ (inv_mem y hy) } +#align subgroup.of_div Subgroup.ofDiv +#align add_subgroup.of_sub AddSubgroup.ofSub + +/-- A subgroup of a group inherits a multiplication. -/ +@[to_additive "An `add_subgroup` of an `add_group` inherits an addition."] +instance hasMul : Mul H := + H.toSubmonoid.HasMul +#align subgroup.has_mul Subgroup.hasMul +#align add_subgroup.has_add AddSubgroup.hasAdd + +/-- A subgroup of a group inherits a 1. -/ +@[to_additive "An `add_subgroup` of an `add_group` inherits a zero."] +instance hasOne : One H := + H.toSubmonoid.HasOne +#align subgroup.has_one Subgroup.hasOne +#align add_subgroup.has_zero AddSubgroup.hasZero + +/-- A subgroup of a group inherits an inverse. -/ +@[to_additive "A `add_subgroup` of a `add_group` inherits an inverse."] +instance hasInv : Inv H := + ⟨fun a => ⟨a⁻¹, H.inv_mem a.2⟩⟩ +#align subgroup.has_inv Subgroup.hasInv +#align add_subgroup.has_neg AddSubgroup.hasNeg + +/-- A subgroup of a group inherits a division -/ +@[to_additive "An `add_subgroup` of an `add_group` inherits a subtraction."] +instance hasDiv : Div H := + ⟨fun a b => ⟨a / b, H.div_mem a.2 b.2⟩⟩ +#align subgroup.has_div Subgroup.hasDiv +#align add_subgroup.has_sub AddSubgroup.hasSub + +/-- An `add_subgroup` of an `add_group` inherits a natural scaling. -/ +instance AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := + ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ +#align add_subgroup.has_nsmul AddSubgroup.hasNsmul + +/-- A subgroup of a group inherits a natural power -/ +@[to_additive] +instance hasNpow : Pow H ℕ := + ⟨fun a n => ⟨a ^ n, H.pow_mem a.2 n⟩⟩ +#align subgroup.has_npow Subgroup.hasNpow +#align add_subgroup.has_nsmul AddSubgroup.hasNsmul + +/-- An `add_subgroup` of an `add_group` inherits an integer scaling. -/ +instance AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := + ⟨fun n a => ⟨n • a, H.zsmul_mem a.2 n⟩⟩ +#align add_subgroup.has_zsmul AddSubgroup.hasZsmul + +/-- A subgroup of a group inherits an integer power -/ +@[to_additive] +instance hasZpow : Pow H ℤ := + ⟨fun a n => ⟨a ^ n, H.zpow_mem a.2 n⟩⟩ +#align subgroup.has_zpow Subgroup.hasZpow +#align add_subgroup.has_zsmul AddSubgroup.hasZsmul + +@[simp, norm_cast, to_additive] +theorem coe_mul (x y : H) : (↑(x * y) : G) = ↑x * ↑y := + rfl +#align subgroup.coe_mul Subgroup.coe_mul +#align add_subgroup.coe_add AddSubgroup.coe_add + +@[simp, norm_cast, to_additive] +theorem coe_one : ((1 : H) : G) = 1 := + rfl +#align subgroup.coe_one Subgroup.coe_one +#align add_subgroup.coe_zero AddSubgroup.coe_zero + +@[simp, norm_cast, to_additive] +theorem coe_inv (x : H) : ↑(x⁻¹ : H) = (x⁻¹ : G) := + rfl +#align subgroup.coe_inv Subgroup.coe_inv +#align add_subgroup.coe_neg AddSubgroup.coe_neg + +@[simp, norm_cast, to_additive] +theorem coe_div (x y : H) : (↑(x / y) : G) = ↑x / ↑y := + rfl +#align subgroup.coe_div Subgroup.coe_div +#align add_subgroup.coe_sub AddSubgroup.coe_sub + +@[simp, norm_cast, to_additive] +theorem coe_mk (x : G) (hx : x ∈ H) : ((⟨x, hx⟩ : H) : G) = x := + rfl +#align subgroup.coe_mk Subgroup.coe_mk +#align add_subgroup.coe_mk AddSubgroup.coe_mk + +@[simp, norm_cast, to_additive] +theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := + rfl +#align subgroup.coe_pow Subgroup.coe_pow +#align add_subgroup.coe_nsmul AddSubgroup.coe_nsmul + +@[simp, norm_cast, to_additive] +theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := + rfl +#align subgroup.coe_zpow Subgroup.coe_zpow +#align add_subgroup.coe_zsmul AddSubgroup.coe_zsmul + +@[simp, to_additive] +theorem mk_eq_one_iff {g : G} {h} : (⟨g, h⟩ : H) = 1 ↔ g = 1 := + show (⟨g, h⟩ : H) = (⟨1, H.one_mem⟩ : H) ↔ g = 1 by simp +#align subgroup.mk_eq_one_iff Subgroup.mk_eq_one_iff +#align add_subgroup.mk_eq_zero_iff AddSubgroup.mk_eq_zero_iff + +/-- A subgroup of a group inherits a group structure. -/ +@[to_additive "An `add_subgroup` of an `add_group` inherits an `add_group` structure."] +instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := + Subtype.coe_injective.Group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup.to_group Subgroup.toGroup +#align add_subgroup.to_add_group AddSubgroup.toAddGroup + +/-- A subgroup of a `comm_group` is a `comm_group`. -/ +@[to_additive "An `add_subgroup` of an `add_comm_group` is an `add_comm_group`."] +instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H := + Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup.to_comm_group Subgroup.toCommGroup +#align add_subgroup.to_add_comm_group AddSubgroup.toAddCommGroup + +/-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/ +@[to_additive "An `add_subgroup` of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] +instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : + OrderedCommGroup H := + Subtype.coe_injective.OrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) fun _ _ => rfl +#align subgroup.to_ordered_comm_group Subgroup.toOrderedCommGroup +#align add_subgroup.to_ordered_add_comm_group AddSubgroup.toOrderedAddCommGroup + +/-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ +@[to_additive + "An `add_subgroup` of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] +instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : Subgroup G) : + LinearOrderedCommGroup H := + Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl +#align subgroup.to_linear_ordered_comm_group Subgroup.toLinearOrderedCommGroup +#align add_subgroup.to_linear_ordered_add_comm_group AddSubgroup.toLinearOrderedAddCommGroup + +/-- The natural group hom from a subgroup of group `G` to `G`. -/ +@[to_additive "The natural group hom from an `add_subgroup` of `add_group` `G` to `G`."] +def subtype : H →* G := + ⟨coe, rfl, fun _ _ => rfl⟩ +#align subgroup.subtype Subgroup.subtype +#align add_subgroup.subtype AddSubgroup.subtype + +@[simp, to_additive] +theorem coeSubtype : ⇑H.Subtype = coe := + rfl +#align subgroup.coe_subtype Subgroup.coeSubtype +#align add_subgroup.coe_subtype AddSubgroup.coeSubtype + +@[to_additive] +theorem subtype_injective : Function.Injective (subtype H) := + Subtype.coe_injective +#align subgroup.subtype_injective Subgroup.subtype_injective +#align add_subgroup.subtype_injective AddSubgroup.subtype_injective + +/-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ +@[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] +def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := + MonoidHom.mk' (fun x => ⟨x, h x.Prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl +#align subgroup.inclusion Subgroup.inclusion +#align add_subgroup.inclusion AddSubgroup.inclusion + +@[simp, to_additive] +theorem coe_inclusion {H K : Subgroup G} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := + by + cases a + simp only [inclusion, coe_mk, MonoidHom.mk'_apply] +#align subgroup.coe_inclusion Subgroup.coe_inclusion +#align add_subgroup.coe_inclusion AddSubgroup.coe_inclusion + +@[to_additive] +theorem inclusion_injective {H K : Subgroup G} (h : H ≤ K) : Function.Injective <| inclusion h := + Set.inclusion_injective h +#align subgroup.inclusion_injective Subgroup.inclusion_injective +#align add_subgroup.inclusion_injective AddSubgroup.inclusion_injective + +@[simp, to_additive] +theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : + K.Subtype.comp (inclusion hH) = H.Subtype := + rfl +#align subgroup.subtype_comp_inclusion Subgroup.subtype_comp_inclusion +#align add_subgroup.subtype_comp_inclusion AddSubgroup.subtype_comp_inclusion + +/-- The subgroup `G` of the group `G`. -/ +@[to_additive "The `add_subgroup G` of the `add_group G`."] +instance : Top (Subgroup G) := + ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ _ => Set.mem_univ _ }⟩ + +/-- The top subgroup is isomorphic to the group. + +This is the group version of `submonoid.top_equiv`. -/ +@[to_additive + "The top additive subgroup is isomorphic to the additive group.\n\nThis is the additive group version of `add_submonoid.top_equiv`.", + simps] +def topEquiv : (⊤ : Subgroup G) ≃* G := + Submonoid.topEquiv +#align subgroup.top_equiv Subgroup.topEquiv +#align add_subgroup.top_equiv AddSubgroup.topEquiv + +/-- The trivial subgroup `{1}` of an group `G`. -/ +@[to_additive "The trivial `add_subgroup` `{0}` of an `add_group` `G`."] +instance : Bot (Subgroup G) := + ⟨{ (⊥ : Submonoid G) with inv_mem' := fun _ => by simp [*] }⟩ + +@[to_additive] +instance : Inhabited (Subgroup G) := + ⟨⊥⟩ + +@[simp, to_additive] +theorem mem_bot {x : G} : x ∈ (⊥ : Subgroup G) ↔ x = 1 := + Iff.rfl +#align subgroup.mem_bot Subgroup.mem_bot +#align add_subgroup.mem_bot AddSubgroup.mem_bot + +@[simp, to_additive] +theorem mem_top (x : G) : x ∈ (⊤ : Subgroup G) := + Set.mem_univ x +#align subgroup.mem_top Subgroup.mem_top +#align add_subgroup.mem_top AddSubgroup.mem_top + +@[simp, to_additive] +theorem coe_top : ((⊤ : Subgroup G) : Set G) = Set.univ := + rfl +#align subgroup.coe_top Subgroup.coe_top +#align add_subgroup.coe_top AddSubgroup.coe_top + +@[simp, to_additive] +theorem coe_bot : ((⊥ : Subgroup G) : Set G) = {1} := + rfl +#align subgroup.coe_bot Subgroup.coe_bot +#align add_subgroup.coe_bot AddSubgroup.coe_bot + +@[to_additive] +instance : Unique (⊥ : Subgroup G) := + ⟨⟨1⟩, fun g => Subtype.ext g.2⟩ + +@[simp, to_additive] +theorem top_toSubmonoid : (⊤ : Subgroup G).toSubmonoid = ⊤ := + rfl +#align subgroup.top_to_submonoid Subgroup.top_toSubmonoid +#align add_subgroup.top_to_add_submonoid AddSubgroup.top_to_add_submonoid + +@[simp, to_additive] +theorem bot_toSubmonoid : (⊥ : Subgroup G).toSubmonoid = ⊥ := + rfl +#align subgroup.bot_to_submonoid Subgroup.bot_toSubmonoid +#align add_subgroup.bot_to_add_submonoid AddSubgroup.bot_to_add_submonoid + +@[to_additive] +theorem eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := + toSubmonoid_injective.eq_iff.symm.trans <| Submonoid.eq_bot_iff_forall _ +#align subgroup.eq_bot_iff_forall Subgroup.eq_bot_iff_forall +#align add_subgroup.eq_bot_iff_forall AddSubgroup.eq_bot_iff_forall + +@[to_additive] +theorem eq_bot_of_subsingleton [Subsingleton H] : H = ⊥ := + by + rw [Subgroup.eq_bot_iff_forall] + intro y hy + rw [← Subgroup.coe_mk H y hy, Subsingleton.elim (⟨y, hy⟩ : H) 1, Subgroup.coe_one] +#align subgroup.eq_bot_of_subsingleton Subgroup.eq_bot_of_subsingleton +#align add_subgroup.eq_bot_of_subsingleton AddSubgroup.eq_bot_of_subsingleton + +@[to_additive] +theorem coe_eq_univ {H : Subgroup G} : (H : Set G) = Set.univ ↔ H = ⊤ := + (SetLike.ext'_iff.trans (by rfl)).symm +#align subgroup.coe_eq_univ Subgroup.coe_eq_univ +#align add_subgroup.coe_eq_univ AddSubgroup.coe_eq_univ + +@[to_additive] +theorem coe_eq_singleton {H : Subgroup G} : (∃ g : G, (H : Set G) = {g}) ↔ H = ⊥ := + ⟨fun ⟨g, hg⟩ => + haveI : Subsingleton (H : Set G) := by + rw [hg] + infer_instance + H.eq_bot_of_subsingleton, + fun h => ⟨1, SetLike.ext'_iff.mp h⟩⟩ +#align subgroup.coe_eq_singleton Subgroup.coe_eq_singleton +#align add_subgroup.coe_eq_singleton AddSubgroup.coe_eq_singleton + +@[to_additive] +theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x ∈ H, x ≠ (1 : G) := + Subtype.nontrivial_iff_exists_ne (fun x => x ∈ H) (1 : H) +#align subgroup.nontrivial_iff_exists_ne_one Subgroup.nontrivial_iff_exists_ne_one +#align add_subgroup.nontrivial_iff_exists_ne_zero AddSubgroup.nontrivial_iff_exists_ne_zero + +/-- A subgroup is either the trivial subgroup or nontrivial. -/ +@[to_additive "A subgroup is either the trivial subgroup or nontrivial."] +theorem bot_or_nontrivial (H : Subgroup G) : H = ⊥ ∨ Nontrivial H := by + classical + by_cases h : ∀ x ∈ H, x = (1 : G) + · left + exact H.eq_bot_iff_forall.mpr h + · right + simp only [not_forall] at h + simpa only [nontrivial_iff_exists_ne_one] +#align subgroup.bot_or_nontrivial Subgroup.bot_or_nontrivial +#align add_subgroup.bot_or_nontrivial AddSubgroup.bot_or_nontrivial + +/-- A subgroup is either the trivial subgroup or contains a non-identity element. -/ +@[to_additive "A subgroup is either the trivial subgroup or contains a nonzero element."] +theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ (1 : G) := + by + convert H.bot_or_nontrivial + rw [nontrivial_iff_exists_ne_one] +#align subgroup.bot_or_exists_ne_one Subgroup.bot_or_exists_ne_one +#align add_subgroup.bot_or_exists_ne_zero AddSubgroup.bot_or_exists_ne_zero + +/-- The inf of two subgroups is their intersection. -/ +@[to_additive "The inf of two `add_subgroups`s is their intersection."] +instance : HasInf (Subgroup G) := + ⟨fun H₁ H₂ => + { H₁.toSubmonoid ⊓ H₂.toSubmonoid with + inv_mem' := fun _ ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ + +@[simp, to_additive] +theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = p ∩ p' := + rfl +#align subgroup.coe_inf Subgroup.coe_inf +#align add_subgroup.coe_inf AddSubgroup.coe_inf + +@[simp, to_additive] +theorem mem_inf {p p' : Subgroup G} {x : G} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' := + Iff.rfl +#align subgroup.mem_inf Subgroup.mem_inf +#align add_subgroup.mem_inf AddSubgroup.mem_inf + +@[to_additive] +instance : InfSet (Subgroup G) := + ⟨fun s => + { (⨅ S ∈ s, Subgroup.toSubmonoid S).copy (⋂ S ∈ s, ↑S) (by simp) with + inv_mem' := fun x hx => + Set.mem_binterᵢ fun i h => i.inv_mem (by apply Set.mem_interᵢ₂.1 hx i h) }⟩ + +@[simp, norm_cast, to_additive] +theorem coe_infₛ (H : Set (Subgroup G)) : ((infₛ H : Subgroup G) : Set G) = ⋂ s ∈ H, ↑s := + rfl +#align subgroup.coe_Inf Subgroup.coe_infₛ +#align add_subgroup.coe_Inf AddSubgroup.coe_Inf + +@[simp, to_additive] +theorem mem_infₛ {S : Set (Subgroup G)} {x : G} : x ∈ infₛ S ↔ ∀ p ∈ S, x ∈ p := + Set.mem_interᵢ₂ +#align subgroup.mem_Inf Subgroup.mem_infₛ +#align add_subgroup.mem_Inf AddSubgroup.mem_Inf + +@[to_additive] +theorem mem_infᵢ {ι : Sort _} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by + simp only [infᵢ, mem_Inf, Set.forall_range_iff] +#align subgroup.mem_infi Subgroup.mem_infᵢ +#align add_subgroup.mem_infi AddSubgroup.mem_infi + +@[simp, norm_cast, to_additive] +theorem coe_infᵢ {ι : Sort _} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by + simp only [infᵢ, coe_Inf, Set.binterᵢ_range] +#align subgroup.coe_infi Subgroup.coe_infᵢ +#align add_subgroup.coe_infi AddSubgroup.coe_infi + +/-- Subgroups of a group form a complete lattice. -/ +@[to_additive "The `add_subgroup`s of an `add_group` form a complete lattice."] +instance : CompleteLattice (Subgroup G) := + { + completeLatticeOfInf (Subgroup G) fun s => + IsGLB.of_image (fun H K => show (H : Set G) ≤ K ↔ H ≤ K from SetLike.coe_subset_coe) + isGLB_binfᵢ with + bot := ⊥ + bot_le := fun S x hx => (mem_bot.1 hx).symm ▸ S.one_mem + top := ⊤ + le_top := fun S x hx => mem_top x + inf := (· ⊓ ·) + le_inf := fun a b c ha hb x hx => ⟨ha hx, hb hx⟩ + inf_le_left := fun a b x => And.left + inf_le_right := fun a b x => And.right } + +@[to_additive] +theorem mem_sup_left {S T : Subgroup G} : ∀ {x : G}, x ∈ S → x ∈ S ⊔ T := + show S ≤ S ⊔ T from le_sup_left +#align subgroup.mem_sup_left Subgroup.mem_sup_left +#align add_subgroup.mem_sup_left AddSubgroup.mem_sup_left + +@[to_additive] +theorem mem_sup_right {S T : Subgroup G} : ∀ {x : G}, x ∈ T → x ∈ S ⊔ T := + show T ≤ S ⊔ T from le_sup_right +#align subgroup.mem_sup_right Subgroup.mem_sup_right +#align add_subgroup.mem_sup_right AddSubgroup.mem_sup_right + +@[to_additive] +theorem mul_mem_sup {S T : Subgroup G} {x y : G} (hx : x ∈ S) (hy : y ∈ T) : x * y ∈ S ⊔ T := + (S ⊔ T).mul_mem (mem_sup_left hx) (mem_sup_right hy) +#align subgroup.mul_mem_sup Subgroup.mul_mem_sup +#align add_subgroup.add_mem_sup AddSubgroup.add_mem_sup + +@[to_additive] +theorem mem_supᵢ_of_mem {ι : Sort _} {S : ι → Subgroup G} (i : ι) : + ∀ {x : G}, x ∈ S i → x ∈ supᵢ S := + show S i ≤ supᵢ S from le_supᵢ _ _ +#align subgroup.mem_supr_of_mem Subgroup.mem_supᵢ_of_mem +#align add_subgroup.mem_supr_of_mem AddSubgroup.mem_supᵢ_of_mem + +@[to_additive] +theorem mem_supₛ_of_mem {S : Set (Subgroup G)} {s : Subgroup G} (hs : s ∈ S) : + ∀ {x : G}, x ∈ s → x ∈ supₛ S := + show s ≤ supₛ S from le_supₛ hs +#align subgroup.mem_Sup_of_mem Subgroup.mem_supₛ_of_mem +#align add_subgroup.mem_Sup_of_mem AddSubgroup.mem_supₛ_of_mem + +@[simp, to_additive] +theorem subsingleton_iff : Subsingleton (Subgroup G) ↔ Subsingleton G := + ⟨fun h => + ⟨fun x y => + have : ∀ i : G, i = 1 := fun i => + mem_bot.mp <| Subsingleton.elim (⊤ : Subgroup G) ⊥ ▸ mem_top i + (this x).trans (this y).symm⟩, + fun h => ⟨fun x y => Subgroup.ext fun i => Subsingleton.elim 1 i ▸ by simp [Subgroup.one_mem]⟩⟩ +#align subgroup.subsingleton_iff Subgroup.subsingleton_iff +#align add_subgroup.subsingleton_iff AddSubgroup.subsingleton_iff + +@[simp, to_additive] +theorem nontrivial_iff : Nontrivial (Subgroup G) ↔ Nontrivial G := + not_iff_not.mp + ((not_nontrivial_iff_subsingleton.trans subsingleton_iff).trans + not_nontrivial_iff_subsingleton.symm) +#align subgroup.nontrivial_iff Subgroup.nontrivial_iff +#align add_subgroup.nontrivial_iff AddSubgroup.nontrivial_iff + +@[to_additive] +instance [Subsingleton G] : Unique (Subgroup G) := + ⟨⟨⊥⟩, fun a => @Subsingleton.elim _ (subsingleton_iff.mpr ‹_›) a _⟩ + +@[to_additive] +instance [Nontrivial G] : Nontrivial (Subgroup G) := + nontrivial_iff.mpr ‹_› + +@[to_additive] +theorem eq_top_iff' : H = ⊤ ↔ ∀ x : G, x ∈ H := + eq_top_iff.trans ⟨fun h m => h <| mem_top m, fun h m _ => h m⟩ +#align subgroup.eq_top_iff' Subgroup.eq_top_iff' +#align add_subgroup.eq_top_iff' AddSubgroup.eq_top_iff' + +/-- The `subgroup` generated by a set. -/ +@[to_additive "The `add_subgroup` generated by a set"] +def closure (k : Set G) : Subgroup G := + infₛ { K | k ⊆ K } +#align subgroup.closure Subgroup.closure +#align add_subgroup.closure AddSubgroup.closure + +variable {k : Set G} + +@[to_additive] +theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K → x ∈ K := + mem_Inf +#align subgroup.mem_closure Subgroup.mem_closure +#align add_subgroup.mem_closure AddSubgroup.mem_closure + +/-- The subgroup generated by a set includes the set. -/ +@[simp, to_additive "The `add_subgroup` generated by a set includes the set."] +theorem subset_closure : k ⊆ closure k := fun x hx => mem_closure.2 fun K hK => hK hx +#align subgroup.subset_closure Subgroup.subset_closure +#align add_subgroup.subset_closure AddSubgroup.subset_closure + +@[to_additive] +theorem not_mem_of_not_mem_closure {P : G} (hP : P ∉ closure k) : P ∉ k := fun h => + hP (subset_closure h) +#align subgroup.not_mem_of_not_mem_closure Subgroup.not_mem_of_not_mem_closure +#align add_subgroup.not_mem_of_not_mem_closure AddSubgroup.not_mem_of_not_mem_closure + +open Set + +/-- A subgroup `K` includes `closure k` if and only if it includes `k`. -/ +@[simp, to_additive "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] +theorem closure_le : closure k ≤ K ↔ k ⊆ K := + ⟨Subset.trans subset_closure, fun h => infₛ_le h⟩ +#align subgroup.closure_le Subgroup.closure_le +#align add_subgroup.closure_le AddSubgroup.closure_le + +@[to_additive] +theorem closure_eq_of_le (h₁ : k ⊆ K) (h₂ : K ≤ closure k) : closure k = K := + le_antisymm ((closure_le <| K).2 h₁) h₂ +#align subgroup.closure_eq_of_le Subgroup.closure_eq_of_le +#align add_subgroup.closure_eq_of_le AddSubgroup.closure_eq_of_le + +/-- An induction principle for closure membership. If `p` holds for `1` and all elements of `k`, and +is preserved under multiplication and inverse, then `p` holds for all elements of the closure +of `k`. -/ +@[elab_as_elim, + to_additive + "An induction principle for additive closure membership. If `p`\nholds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` holds\nfor all elements of the additive closure of `k`."] +theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x ∈ k, p x) (H1 : p 1) + (Hmul : ∀ x y, p x → p y → p (x * y)) (Hinv : ∀ x, p x → p x⁻¹) : p x := + (@closure_le _ _ ⟨p, Hmul, H1, Hinv⟩ _).2 Hk h +#align subgroup.closure_induction Subgroup.closure_induction +#align add_subgroup.closure_induction AddSubgroup.closure_induction + +/-- A dependent version of `subgroup.closure_induction`. -/ +@[elab_as_elim, to_additive "A dependent version of `add_subgroup.closure_induction`. "] +theorem closure_induction' {p : ∀ x, x ∈ closure k → Prop} + (Hs : ∀ (x) (h : x ∈ k), p x (subset_closure h)) (H1 : p 1 (one_mem _)) + (Hmul : ∀ x hx y hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) + (Hinv : ∀ x hx, p x hx → p x⁻¹ (inv_mem hx)) {x} (hx : x ∈ closure k) : p x hx := + by + refine' Exists.elim _ fun (hx : x ∈ closure k) (hc : p x hx) => hc + exact + closure_induction hx (fun x hx => ⟨_, Hs x hx⟩) ⟨_, H1⟩ + (fun x y ⟨hx', hx⟩ ⟨hy', hy⟩ => ⟨_, Hmul _ _ _ _ hx hy⟩) fun x ⟨hx', hx⟩ => ⟨_, Hinv _ _ hx⟩ +#align subgroup.closure_induction' Subgroup.closure_induction' +#align add_subgroup.closure_induction' AddSubgroup.closure_induction' + +/-- An induction principle for closure membership for predicates with two arguments. -/ +@[elab_as_elim, + to_additive + "An induction principle for additive closure membership, for\npredicates with two arguments."] +theorem closure_induction₂ {p : G → G → Prop} {x} {y : G} (hx : x ∈ closure k) (hy : y ∈ closure k) + (Hk : ∀ x ∈ k, ∀ y ∈ k, p x y) (H1_left : ∀ x, p 1 x) (H1_right : ∀ x, p x 1) + (Hmul_left : ∀ x₁ x₂ y, p x₁ y → p x₂ y → p (x₁ * x₂) y) + (Hmul_right : ∀ x y₁ y₂, p x y₁ → p x y₂ → p x (y₁ * y₂)) (Hinv_left : ∀ x y, p x y → p x⁻¹ y) + (Hinv_right : ∀ x y, p x y → p x y⁻¹) : p x y := + closure_induction hx + (fun x xk => closure_induction hy (Hk x xk) (H1_right x) (Hmul_right x) (Hinv_right x)) + (H1_left y) (fun z z' => Hmul_left z z' y) fun z => Hinv_left z y +#align subgroup.closure_induction₂ Subgroup.closure_induction₂ +#align add_subgroup.closure_induction₂ AddSubgroup.closure_induction₂ + +@[simp, to_additive] +theorem closure_closure_coe_preimage {k : Set G} : closure ((coe : closure k → G) ⁻¹' k) = ⊤ := + eq_top_iff.2 fun x => + Subtype.recOn x fun x hx _ => + by + refine' closure_induction' (fun g hg => _) _ (fun g₁ g₂ hg₁ hg₂ => _) (fun g hg => _) hx + · exact subset_closure hg + · exact one_mem _ + · exact mul_mem + · exact inv_mem +#align subgroup.closure_closure_coe_preimage Subgroup.closure_closure_coe_preimage +#align add_subgroup.closure_closure_coe_preimage AddSubgroup.closure_closure_coe_preimage + +/-- If all the elements of a set `s` commute, then `closure s` is a commutative group. -/ +@[to_additive + "If all the elements of a set `s` commute, then `closure s` is an additive\ncommutative group."] +def closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * y = y * x) : + CommGroup (closure k) := + { (closure k).toGroup with + mul_comm := fun x y => by + ext + simp only [Subgroup.coe_mul] + refine' + closure_induction₂ x.prop y.prop hcomm (fun x => by simp only [mul_one, one_mul]) + (fun x => by simp only [mul_one, one_mul]) + (fun x y z h₁ h₂ => by rw [mul_assoc, h₂, ← mul_assoc, h₁, mul_assoc]) + (fun x y z h₁ h₂ => by rw [← mul_assoc, h₁, mul_assoc, h₂, ← mul_assoc]) + (fun x y h => by + rw [inv_mul_eq_iff_eq_mul, ← mul_assoc, h, mul_assoc, mul_inv_self, mul_one]) + fun x y h => by + rw [mul_inv_eq_iff_eq_mul, mul_assoc, h, ← mul_assoc, inv_mul_self, one_mul] } +#align subgroup.closure_comm_group_of_comm Subgroup.closureCommGroupOfComm +#align add_subgroup.closure_add_comm_group_of_comm AddSubgroup.closureAddCommGroupOfComm + +variable (G) + +/-- `closure` forms a Galois insertion with the coercion to set. -/ +@[to_additive "`closure` forms a Galois insertion with the coercion to set."] +protected def gi : GaloisInsertion (@closure G _) coe + where + choice s _ := closure s + gc s t := @closure_le _ _ t s + le_l_u s := subset_closure + choice_eq s h := rfl +#align subgroup.gi Subgroup.gi +#align add_subgroup.gi AddSubgroup.gi + +variable {G} + +/-- Subgroup closure of a set is monotone in its argument: if `h ⊆ k`, +then `closure h ≤ closure k`. -/ +@[to_additive + "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`,\nthen `closure h ≤ closure k`"] +theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k := + (Subgroup.gi G).gc.monotone_l h' +#align subgroup.closure_mono Subgroup.closure_mono +#align add_subgroup.closure_mono AddSubgroup.closure_mono + +/-- Closure of a subgroup `K` equals `K`. -/ +@[simp, to_additive "Additive closure of an additive subgroup `K` equals `K`"] +theorem closure_eq : closure (K : Set G) = K := + (Subgroup.gi G).l_u_eq K +#align subgroup.closure_eq Subgroup.closure_eq +#align add_subgroup.closure_eq AddSubgroup.closure_eq + +@[simp, to_additive] +theorem closure_empty : closure (∅ : Set G) = ⊥ := + (Subgroup.gi G).gc.l_bot +#align subgroup.closure_empty Subgroup.closure_empty +#align add_subgroup.closure_empty AddSubgroup.closure_empty + +@[simp, to_additive] +theorem closure_univ : closure (univ : Set G) = ⊤ := + @coe_top G _ ▸ closure_eq ⊤ +#align subgroup.closure_univ Subgroup.closure_univ +#align add_subgroup.closure_univ AddSubgroup.closure_univ + +@[to_additive] +theorem closure_union (s t : Set G) : closure (s ∪ t) = closure s ⊔ closure t := + (Subgroup.gi G).gc.l_sup +#align subgroup.closure_union Subgroup.closure_union +#align add_subgroup.closure_union AddSubgroup.closure_union + +@[to_additive] +theorem closure_unionᵢ {ι} (s : ι → Set G) : closure (⋃ i, s i) = ⨆ i, closure (s i) := + (Subgroup.gi G).gc.l_supr +#align subgroup.closure_Union Subgroup.closure_unionᵢ +#align add_subgroup.closure_Union AddSubgroup.closure_unionᵢ + +@[to_additive] +theorem closure_eq_bot_iff (G : Type _) [Group G] (S : Set G) : closure S = ⊥ ↔ S ⊆ {1} := + by + rw [← le_bot_iff] + exact closure_le _ +#align subgroup.closure_eq_bot_iff Subgroup.closure_eq_bot_iff +#align add_subgroup.closure_eq_bot_iff AddSubgroup.closure_eq_bot_iff + +@[to_additive] +theorem supᵢ_eq_closure {ι : Sort _} (p : ι → Subgroup G) : + (⨆ i, p i) = closure (⋃ i, (p i : Set G)) := by simp_rw [closure_Union, closure_eq] +#align subgroup.supr_eq_closure Subgroup.supᵢ_eq_closure +#align add_subgroup.supr_eq_closure AddSubgroup.supᵢ_eq_closure + +/-- The subgroup generated by an element of a group equals the set of integer number powers of + the element. -/ +@[to_additive + "The `add_subgroup` generated by an element of an `add_group` equals the set of\nnatural number multiples of the element."] +theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := + by + refine' + ⟨fun hy => closure_induction hy _ _ _ _, fun ⟨n, hn⟩ => + hn ▸ zpow_mem (subset_closure <| mem_singleton x) n⟩ + · intro y hy + rw [eq_of_mem_singleton hy] + exact ⟨1, zpow_one x⟩ + · exact ⟨0, zpow_zero x⟩ + · rintro _ _ ⟨n, rfl⟩ ⟨m, rfl⟩ + exact ⟨n + m, zpow_add x n m⟩ + rintro _ ⟨n, rfl⟩ + exact ⟨-n, zpow_neg x n⟩ +#align subgroup.mem_closure_singleton Subgroup.mem_closure_singleton +#align add_subgroup.mem_closure_singleton AddSubgroup.mem_closure_singleton + +@[to_additive] +theorem closure_singleton_one : closure ({1} : Set G) = ⊥ := by + simp [eq_bot_iff_forall, mem_closure_singleton] +#align subgroup.closure_singleton_one Subgroup.closure_singleton_one +#align add_subgroup.closure_singleton_zero AddSubgroup.closure_singleton_zero + +@[to_additive] +theorem le_closure_toSubmonoid (S : Set G) : Submonoid.closure S ≤ (closure S).toSubmonoid := + Submonoid.closure_le.2 subset_closure +#align subgroup.le_closure_to_submonoid Subgroup.le_closure_toSubmonoid +#align add_subgroup.le_closure_to_add_submonoid AddSubgroup.le_closure_to_add_submonoid + +@[to_additive] +theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = ⊤) : + closure S = ⊤ := + (eq_top_iff' _).2 fun x => le_closure_toSubmonoid _ <| h.symm ▸ trivial +#align subgroup.closure_eq_top_of_mclosure_eq_top Subgroup.closure_eq_top_of_mclosure_eq_top +#align add_subgroup.closure_eq_top_of_mclosure_eq_top AddSubgroup.closure_eq_top_of_mclosure_eq_top + +@[to_additive] +theorem mem_supᵢ_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} (hK : Directed (· ≤ ·) K) + {x : G} : x ∈ (supᵢ K : Subgroup G) ↔ ∃ i, x ∈ K i := + by + refine' ⟨_, fun ⟨i, hi⟩ => (SetLike.le_def.1 <| le_supᵢ K i) hi⟩ + suffices x ∈ closure (⋃ i, (K i : Set G)) → ∃ i, x ∈ K i by + simpa only [closure_Union, closure_eq (K _)] using this + refine' fun hx => closure_induction hx (fun _ => mem_Union.1) _ _ _ + · exact hι.elim fun i => ⟨i, (K i).one_mem⟩ + · rintro x y ⟨i, hi⟩ ⟨j, hj⟩ + rcases hK i j with ⟨k, hki, hkj⟩ + exact ⟨k, mul_mem (hki hi) (hkj hj)⟩ + rintro _ ⟨i, hi⟩ + exact ⟨i, inv_mem hi⟩ +#align subgroup.mem_supr_of_directed Subgroup.mem_supᵢ_of_directed +#align add_subgroup.mem_supr_of_directed AddSubgroup.mem_supᵢ_of_directed + +@[to_additive] +theorem coe_supᵢ_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : Directed (· ≤ ·) S) : + ((⨆ i, S i : Subgroup G) : Set G) = ⋃ i, ↑(S i) := + Set.ext fun x => by simp [mem_supr_of_directed hS] +#align subgroup.coe_supr_of_directed Subgroup.coe_supᵢ_of_directed +#align add_subgroup.coe_supr_of_directed AddSubgroup.coe_supᵢ_of_directed + +@[to_additive] +theorem mem_supₛ_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) + {x : G} : x ∈ supₛ K ↔ ∃ s ∈ K, x ∈ s := + by + haveI : Nonempty K := Kne.to_subtype + simp only [supₛ_eq_supᵢ', mem_supr_of_directed hK.directed_coe, SetCoe.exists, Subtype.coe_mk] +#align subgroup.mem_Sup_of_directed_on Subgroup.mem_supₛ_of_directedOn +#align add_subgroup.mem_Sup_of_directed_on AddSubgroup.mem_supₛ_of_directedOn + +variable {N : Type _} [Group N] {P : Type _} [Group P] + +/-- The preimage of a subgroup along a monoid homomorphism is a subgroup. -/ +@[to_additive + "The preimage of an `add_subgroup` along an `add_monoid` homomorphism\nis an `add_subgroup`."] +def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := + { H.toSubmonoid.comap f with + carrier := f ⁻¹' H + inv_mem' := fun a ha => show f a⁻¹ ∈ H by rw [f.map_inv] <;> exact H.inv_mem ha } +#align subgroup.comap Subgroup.comap +#align add_subgroup.comap AddSubgroup.comap + +@[simp, to_additive] +theorem coe_comap (K : Subgroup N) (f : G →* N) : (K.comap f : Set G) = f ⁻¹' K := + rfl +#align subgroup.coe_comap Subgroup.coe_comap +#align add_subgroup.coe_comap AddSubgroup.coe_comap + +@[simp, to_additive] +theorem mem_comap {K : Subgroup N} {f : G →* N} {x : G} : x ∈ K.comap f ↔ f x ∈ K := + Iff.rfl +#align subgroup.mem_comap Subgroup.mem_comap +#align add_subgroup.mem_comap AddSubgroup.mem_comap + +@[to_additive] +theorem comap_mono {f : G →* N} {K K' : Subgroup N} : K ≤ K' → comap f K ≤ comap f K' := + preimage_mono +#align subgroup.comap_mono Subgroup.comap_mono +#align add_subgroup.comap_mono AddSubgroup.comap_mono + +@[to_additive] +theorem comap_comap (K : Subgroup P) (g : N →* P) (f : G →* N) : + (K.comap g).comap f = K.comap (g.comp f) := + rfl +#align subgroup.comap_comap Subgroup.comap_comap +#align add_subgroup.comap_comap AddSubgroup.comap_comap + +@[simp, to_additive] +theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := + by + ext + rfl +#align subgroup.comap_id Subgroup.comap_id +#align add_subgroup.comap_id AddSubgroup.comap_id + +/-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ +@[to_additive + "The image of an `add_subgroup` along an `add_monoid` homomorphism\nis an `add_subgroup`."] +def map (f : G →* N) (H : Subgroup G) : Subgroup N := + { H.toSubmonoid.map f with + carrier := f '' H + inv_mem' := by + rintro _ ⟨x, hx, rfl⟩ + exact ⟨x⁻¹, H.inv_mem hx, f.map_inv x⟩ } +#align subgroup.map Subgroup.map +#align add_subgroup.map AddSubgroup.map + +@[simp, to_additive] +theorem coe_map (f : G →* N) (K : Subgroup G) : (K.map f : Set N) = f '' K := + rfl +#align subgroup.coe_map Subgroup.coe_map +#align add_subgroup.coe_map AddSubgroup.coe_map + +@[simp, to_additive] +theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := + mem_image_iff_bex +#align subgroup.mem_map Subgroup.mem_map +#align add_subgroup.mem_map AddSubgroup.mem_map + +@[to_additive] +theorem mem_map_of_mem (f : G →* N) {K : Subgroup G} {x : G} (hx : x ∈ K) : f x ∈ K.map f := + mem_image_of_mem f hx +#align subgroup.mem_map_of_mem Subgroup.mem_map_of_mem +#align add_subgroup.mem_map_of_mem AddSubgroup.mem_map_of_mem + +@[to_additive] +theorem apply_coe_mem_map (f : G →* N) (K : Subgroup G) (x : K) : f x ∈ K.map f := + mem_map_of_mem f x.Prop +#align subgroup.apply_coe_mem_map Subgroup.apply_coe_mem_map +#align add_subgroup.apply_coe_mem_map AddSubgroup.apply_coe_mem_map + +@[to_additive] +theorem map_mono {f : G →* N} {K K' : Subgroup G} : K ≤ K' → map f K ≤ map f K' := + image_subset _ +#align subgroup.map_mono Subgroup.map_mono +#align add_subgroup.map_mono AddSubgroup.map_mono + +@[simp, to_additive] +theorem map_id : K.map (MonoidHom.id G) = K := + SetLike.coe_injective <| image_id _ +#align subgroup.map_id Subgroup.map_id +#align add_subgroup.map_id AddSubgroup.map_id + +@[to_additive] +theorem map_map (g : N →* P) (f : G →* N) : (K.map f).map g = K.map (g.comp f) := + SetLike.coe_injective <| image_image _ _ _ +#align subgroup.map_map Subgroup.map_map +#align add_subgroup.map_map AddSubgroup.map_map + +@[simp, to_additive] +theorem map_one_eq_bot : K.map (1 : G →* N) = ⊥ := + eq_bot_iff.mpr <| by + rintro x ⟨y, _, rfl⟩ + simp +#align subgroup.map_one_eq_bot Subgroup.map_one_eq_bot +#align add_subgroup.map_zero_eq_bot AddSubgroup.map_zero_eq_bot + +@[to_additive] +theorem mem_map_equiv {f : G ≃* N} {K : Subgroup G} {x : N} : + x ∈ K.map f.toMonoidHom ↔ f.symm x ∈ K := + @Set.mem_image_equiv _ _ (↑K) f.toEquiv x +#align subgroup.mem_map_equiv Subgroup.mem_map_equiv +#align add_subgroup.mem_map_equiv AddSubgroup.mem_map_equiv + +@[to_additive] +theorem mem_map_iff_mem {f : G →* N} (hf : Function.Injective f) {K : Subgroup G} {x : G} : + f x ∈ K.map f ↔ x ∈ K := + hf.mem_set_image +#align subgroup.mem_map_iff_mem Subgroup.mem_map_iff_mem +#align add_subgroup.mem_map_iff_mem AddSubgroup.mem_map_iff_mem + +@[to_additive] +theorem map_equiv_eq_comap_symm (f : G ≃* N) (K : Subgroup G) : + K.map f.toMonoidHom = K.comap f.symm.toMonoidHom := + SetLike.coe_injective (f.toEquiv.image_eq_preimage K) +#align subgroup.map_equiv_eq_comap_symm Subgroup.map_equiv_eq_comap_symm +#align add_subgroup.map_equiv_eq_comap_symm AddSubgroup.map_equiv_eq_comap_symm + +@[to_additive] +theorem comap_equiv_eq_map_symm (f : N ≃* G) (K : Subgroup G) : + K.comap f.toMonoidHom = K.map f.symm.toMonoidHom := + (map_equiv_eq_comap_symm f.symm K).symm +#align subgroup.comap_equiv_eq_map_symm Subgroup.comap_equiv_eq_map_symm +#align add_subgroup.comap_equiv_eq_map_symm AddSubgroup.comap_equiv_eq_map_symm + +@[to_additive] +theorem map_symm_eq_iff_map_eq {H : Subgroup N} {e : G ≃* N} : H.map ↑e.symm = K ↔ K.map ↑e = H := + by + constructor <;> rintro rfl + · + rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.symm_trans_self, + MulEquiv.coe_monoidHom_refl, map_id] + · + rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.self_trans_symm, + MulEquiv.coe_monoidHom_refl, map_id] +#align subgroup.map_symm_eq_iff_map_eq Subgroup.map_symm_eq_iff_map_eq +#align add_subgroup.map_symm_eq_iff_map_eq AddSubgroup.map_symm_eq_iff_map_eq + +@[to_additive] +theorem map_le_iff_le_comap {f : G →* N} {K : Subgroup G} {H : Subgroup N} : + K.map f ≤ H ↔ K ≤ H.comap f := + image_subset_iff +#align subgroup.map_le_iff_le_comap Subgroup.map_le_iff_le_comap +#align add_subgroup.map_le_iff_le_comap AddSubgroup.map_le_iff_le_comap + +@[to_additive] +theorem gc_map_comap (f : G →* N) : GaloisConnection (map f) (comap f) := fun _ _ => + map_le_iff_le_comap +#align subgroup.gc_map_comap Subgroup.gc_map_comap +#align add_subgroup.gc_map_comap AddSubgroup.gc_map_comap + +@[to_additive] +theorem map_sup (H K : Subgroup G) (f : G →* N) : (H ⊔ K).map f = H.map f ⊔ K.map f := + (gc_map_comap f).l_sup +#align subgroup.map_sup Subgroup.map_sup +#align add_subgroup.map_sup AddSubgroup.map_sup + +@[to_additive] +theorem map_supᵢ {ι : Sort _} (f : G →* N) (s : ι → Subgroup G) : + (supᵢ s).map f = ⨆ i, (s i).map f := + (gc_map_comap f).l_supr +#align subgroup.map_supr Subgroup.map_supᵢ +#align add_subgroup.map_supr AddSubgroup.map_supᵢ + +@[to_additive] +theorem comap_sup_comap_le (H K : Subgroup N) (f : G →* N) : + comap f H ⊔ comap f K ≤ comap f (H ⊔ K) := + Monotone.le_map_sup (fun _ _ => comap_mono) H K +#align subgroup.comap_sup_comap_le Subgroup.comap_sup_comap_le +#align add_subgroup.comap_sup_comap_le AddSubgroup.comap_sup_comap_le + +@[to_additive] +theorem supᵢ_comap_le {ι : Sort _} (f : G →* N) (s : ι → Subgroup N) : + (⨆ i, (s i).comap f) ≤ (supᵢ s).comap f := + Monotone.le_map_supᵢ fun _ _ => comap_mono +#align subgroup.supr_comap_le Subgroup.supᵢ_comap_le +#align add_subgroup.supr_comap_le AddSubgroup.supᵢ_comap_le + +@[to_additive] +theorem comap_inf (H K : Subgroup N) (f : G →* N) : (H ⊓ K).comap f = H.comap f ⊓ K.comap f := + (gc_map_comap f).u_inf +#align subgroup.comap_inf Subgroup.comap_inf +#align add_subgroup.comap_inf AddSubgroup.comap_inf + +@[to_additive] +theorem comap_infᵢ {ι : Sort _} (f : G →* N) (s : ι → Subgroup N) : + (infᵢ s).comap f = ⨅ i, (s i).comap f := + (gc_map_comap f).u_infi +#align subgroup.comap_infi Subgroup.comap_infᵢ +#align add_subgroup.comap_infi AddSubgroup.comap_infi + +@[to_additive] +theorem map_inf_le (H K : Subgroup G) (f : G →* N) : map f (H ⊓ K) ≤ map f H ⊓ map f K := + le_inf (map_mono inf_le_left) (map_mono inf_le_right) +#align subgroup.map_inf_le Subgroup.map_inf_le +#align add_subgroup.map_inf_le AddSubgroup.map_inf_le + +@[to_additive] +theorem map_inf_eq (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : + map f (H ⊓ K) = map f H ⊓ map f K := + by + rw [← SetLike.coe_set_eq] + simp [Set.image_inter hf] +#align subgroup.map_inf_eq Subgroup.map_inf_eq +#align add_subgroup.map_inf_eq AddSubgroup.map_inf_eq + +@[simp, to_additive] +theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := + (gc_map_comap f).l_bot +#align subgroup.map_bot Subgroup.map_bot +#align add_subgroup.map_bot AddSubgroup.map_bot + +@[simp, to_additive] +theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := + by + rw [eq_top_iff] + intro x hx + obtain ⟨y, hy⟩ := h x + exact ⟨y, trivial, hy⟩ +#align subgroup.map_top_of_surjective Subgroup.map_top_of_surjective +#align add_subgroup.map_top_of_surjective AddSubgroup.map_top_of_surjective + +@[simp, to_additive] +theorem comap_top (f : G →* N) : (⊤ : Subgroup N).comap f = ⊤ := + (gc_map_comap f).u_top +#align subgroup.comap_top Subgroup.comap_top +#align add_subgroup.comap_top AddSubgroup.comap_top + +/-- For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`. -/ +@[to_additive "For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`."] +def subgroupOf (H K : Subgroup G) : Subgroup K := + H.comap K.Subtype +#align subgroup.subgroup_of Subgroup.subgroupOf +#align add_subgroup.add_subgroup_of AddSubgroup.addSubgroupOf + +/-- If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`. -/ +@[to_additive "If `H ≤ K`, then `H` as a subgroup of `K` is isomorphic to `H`.", simps] +def subgroupOfEquivOfLe {G : Type _} [Group G] {H K : Subgroup G} (h : H ≤ K) : H.subgroupOf K ≃* H + where + toFun g := ⟨g.1, g.2⟩ + invFun g := ⟨⟨g.1, h g.2⟩, g.2⟩ + left_inv g := Subtype.ext (Subtype.ext rfl) + right_inv g := Subtype.ext rfl + map_mul' g h := rfl +#align subgroup.subgroup_of_equiv_of_le Subgroup.subgroupOfEquivOfLe +#align add_subgroup.add_subgroup_of_equiv_of_le AddSubgroup.addSubgroupOfEquivOfLe + +@[simp, to_additive] +theorem comap_subtype (H K : Subgroup G) : H.comap K.Subtype = H.subgroupOf K := + rfl +#align subgroup.comap_subtype Subgroup.comap_subtype +#align add_subgroup.comap_subtype AddSubgroup.comap_subtype + +@[simp, to_additive] +theorem comap_inclusion_subgroupOf {K₁ K₂ : Subgroup G} (h : K₁ ≤ K₂) (H : Subgroup G) : + (H.subgroupOf K₂).comap (inclusion h) = H.subgroupOf K₁ := + rfl +#align subgroup.comap_inclusion_subgroup_of Subgroup.comap_inclusion_subgroupOf +#align add_subgroup.comap_inclusion_add_subgroup_of AddSubgroup.comap_inclusion_add_subgroupOf + +@[to_additive] +theorem coe_subgroupOf (H K : Subgroup G) : (H.subgroupOf K : Set K) = K.Subtype ⁻¹' H := + rfl +#align subgroup.coe_subgroup_of Subgroup.coe_subgroupOf +#align add_subgroup.coe_add_subgroup_of AddSubgroup.coe_add_subgroupOf + +@[to_additive] +theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h : G) ∈ H := + Iff.rfl +#align subgroup.mem_subgroup_of Subgroup.mem_subgroupOf +#align add_subgroup.mem_add_subgroup_of AddSubgroup.mem_add_subgroupOf + +@[simp, to_additive] +theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.Subtype = H ⊓ K := + SetLike.ext' <| Subtype.image_preimage_coe _ _ +#align subgroup.subgroup_of_map_subtype Subgroup.subgroupOf_map_subtype +#align add_subgroup.add_subgroup_of_map_subtype AddSubgroup.add_subgroupOf_map_subtype + +@[simp, to_additive] +theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := + Eq.symm (Subgroup.ext fun g => Subtype.ext_iff) +#align subgroup.bot_subgroup_of Subgroup.bot_subgroupOf +#align add_subgroup.bot_add_subgroup_of AddSubgroup.bot_add_subgroupOf + +@[simp, to_additive] +theorem top_subgroupOf : (⊤ : Subgroup G).subgroupOf H = ⊤ := + rfl +#align subgroup.top_subgroup_of Subgroup.top_subgroupOf +#align add_subgroup.top_add_subgroup_of AddSubgroup.top_add_subgroupOf + +@[to_additive] +theorem subgroupOf_bot_eq_bot : H.subgroupOf ⊥ = ⊥ := + Subsingleton.elim _ _ +#align subgroup.subgroup_of_bot_eq_bot Subgroup.subgroupOf_bot_eq_bot +#align add_subgroup.add_subgroup_of_bot_eq_bot AddSubgroup.add_subgroupOf_bot_eq_bot + +@[to_additive] +theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := + Subsingleton.elim _ _ +#align subgroup.subgroup_of_bot_eq_top Subgroup.subgroupOf_bot_eq_top +#align add_subgroup.add_subgroup_of_bot_eq_top AddSubgroup.add_subgroupOf_bot_eq_top + +@[simp, to_additive] +theorem subgroupOf_self : H.subgroupOf H = ⊤ := + top_unique fun g hg => g.2 +#align subgroup.subgroup_of_self Subgroup.subgroupOf_self +#align add_subgroup.add_subgroup_of_self AddSubgroup.add_subgroupOf_self + +@[simp, to_additive] +theorem subgroupOf_inj {H₁ H₂ K : Subgroup G} : + H₁.subgroupOf K = H₂.subgroupOf K ↔ H₁ ⊓ K = H₂ ⊓ K := by + simpa only [SetLike.ext_iff, mem_inf, mem_subgroup_of, and_congr_left_iff] using Subtype.forall +#align subgroup.subgroup_of_inj Subgroup.subgroupOf_inj +#align add_subgroup.add_subgroup_of_inj AddSubgroup.add_subgroupOf_inj + +@[simp, to_additive] +theorem inf_subgroupOf_right (H K : Subgroup G) : (H ⊓ K).subgroupOf K = H.subgroupOf K := + subgroupOf_inj.2 inf_right_idem +#align subgroup.inf_subgroup_of_right Subgroup.inf_subgroupOf_right +#align add_subgroup.inf_add_subgroup_of_right AddSubgroup.inf_add_subgroupOf_right + +@[simp, to_additive] +theorem inf_subgroupOf_left (H K : Subgroup G) : (K ⊓ H).subgroupOf K = H.subgroupOf K := by + rw [inf_comm, inf_subgroup_of_right] +#align subgroup.inf_subgroup_of_left Subgroup.inf_subgroupOf_left +#align add_subgroup.inf_add_subgroup_of_left AddSubgroup.inf_add_subgroupOf_left + +@[simp, to_additive] +theorem subgroupOf_eq_bot {H K : Subgroup G} : H.subgroupOf K = ⊥ ↔ Disjoint H K := by + rw [disjoint_iff, ← bot_subgroup_of, subgroup_of_inj, bot_inf_eq] +#align subgroup.subgroup_of_eq_bot Subgroup.subgroupOf_eq_bot +#align add_subgroup.add_subgroup_of_eq_bot AddSubgroup.add_subgroup_of_eq_bot + +@[simp, to_additive] +theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H := by + rw [← top_subgroup_of, subgroup_of_inj, top_inf_eq, inf_eq_right] +#align subgroup.subgroup_of_eq_top Subgroup.subgroupOf_eq_top +#align add_subgroup.add_subgroup_of_eq_top AddSubgroup.add_subgroupOf_eq_top + +/-- Given `subgroup`s `H`, `K` of groups `G`, `N` respectively, `H × K` as a subgroup of `G × N`. -/ +@[to_additive Prod + "Given `add_subgroup`s `H`, `K` of `add_group`s `A`, `B` respectively, `H × K`\nas an `add_subgroup` of `A × B`."] +def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := + { Submonoid.prod H.toSubmonoid K.toSubmonoid with + inv_mem' := fun _ hx => ⟨H.inv_mem' hx.1, K.inv_mem' hx.2⟩ } +#align subgroup.prod Subgroup.prod +#align add_subgroup.prod AddSubgroup.prod + +/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ +@[to_additive coe_prod] +theorem coe_prod (H : Subgroup G) (K : Subgroup N) : (H.Prod K : Set (G × N)) = H ×ˢ K := + rfl +#align subgroup.coe_prod Subgroup.coe_prod +#align add_subgroup.coe_prod AddSubgroup.coe_prod + +@[to_additive mem_prod] +theorem mem_prod {H : Subgroup G} {K : Subgroup N} {p : G × N} : p ∈ H.Prod K ↔ p.1 ∈ H ∧ p.2 ∈ K := + Iff.rfl +#align subgroup.mem_prod Subgroup.mem_prod +#align add_subgroup.mem_prod AddSubgroup.mem_prod + +@[to_additive prod_mono] +theorem prod_mono : ((· ≤ ·) ⇒ (· ≤ ·) ⇒ (· ≤ ·)) (@prod G _ N _) (@prod G _ N _) := + fun s s' hs t t' ht => Set.prod_mono hs ht +#align subgroup.prod_mono Subgroup.prod_mono +#align add_subgroup.prod_mono AddSubgroup.prod_mono + +@[to_additive prod_mono_right] +theorem prod_mono_right (K : Subgroup G) : Monotone fun t : Subgroup N => K.Prod t := + prod_mono (le_refl K) +#align subgroup.prod_mono_right Subgroup.prod_mono_right +#align add_subgroup.prod_mono_right AddSubgroup.prod_mono_right + +@[to_additive prod_mono_left] +theorem prod_mono_left (H : Subgroup N) : Monotone fun K : Subgroup G => K.Prod H := fun s₁ s₂ hs => + prod_mono hs (le_refl H) +#align subgroup.prod_mono_left Subgroup.prod_mono_left +#align add_subgroup.prod_mono_left AddSubgroup.prod_mono_left + +@[to_additive prod_top] +theorem prod_top (K : Subgroup G) : K.Prod (⊤ : Subgroup N) = K.comap (MonoidHom.fst G N) := + ext fun x => by simp [mem_prod, MonoidHom.coe_fst] +#align subgroup.prod_top Subgroup.prod_top +#align add_subgroup.prod_top AddSubgroup.prod_top + +@[to_additive top_prod] +theorem top_prod (H : Subgroup N) : (⊤ : Subgroup G).Prod H = H.comap (MonoidHom.snd G N) := + ext fun x => by simp [mem_prod, MonoidHom.coe_snd] +#align subgroup.top_prod Subgroup.top_prod +#align add_subgroup.top_prod AddSubgroup.top_prod + +@[simp, to_additive top_prod_top] +theorem top_prod_top : (⊤ : Subgroup G).Prod (⊤ : Subgroup N) = ⊤ := + (top_prod _).trans <| comap_top _ +#align subgroup.top_prod_top Subgroup.top_prod_top +#align add_subgroup.top_prod_top AddSubgroup.top_prod_top + +@[to_additive] +theorem bot_prod_bot : (⊥ : Subgroup G).Prod (⊥ : Subgroup N) = ⊥ := + SetLike.coe_injective <| by simp [coe_prod, Prod.one_eq_mk] +#align subgroup.bot_prod_bot Subgroup.bot_prod_bot +#align add_subgroup.bot_sum_bot AddSubgroup.bot_sum_bot + +@[to_additive le_prod_iff] +theorem le_prod_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : + J ≤ H.Prod K ↔ map (MonoidHom.fst G N) J ≤ H ∧ map (MonoidHom.snd G N) J ≤ K := by + simpa only [← Subgroup.toSubmonoid_le] using Submonoid.le_prod_iff +#align subgroup.le_prod_iff Subgroup.le_prod_iff +#align add_subgroup.le_prod_iff AddSubgroup.le_prod_iff + +@[to_additive prod_le_iff] +theorem prod_le_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : + H.Prod K ≤ J ↔ map (MonoidHom.inl G N) H ≤ J ∧ map (MonoidHom.inr G N) K ≤ J := by + simpa only [← Subgroup.toSubmonoid_le] using Submonoid.prod_le_iff +#align subgroup.prod_le_iff Subgroup.prod_le_iff +#align add_subgroup.prod_le_iff AddSubgroup.prod_le_iff + +@[simp, to_additive prod_eq_bot_iff] +theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.Prod K = ⊥ ↔ H = ⊥ ∧ K = ⊥ := by + simpa only [← Subgroup.toSubmonoid_eq] using Submonoid.prod_eq_bot_iff +#align subgroup.prod_eq_bot_iff Subgroup.prod_eq_bot_iff +#align add_subgroup.prod_eq_bot_iff AddSubgroup.prod_eq_bot_iff + +/-- Product of subgroups is isomorphic to their product as groups. -/ +@[to_additive prod_equiv + "Product of additive subgroups is isomorphic to their product\nas additive groups"] +def prodEquiv (H : Subgroup G) (K : Subgroup N) : H.Prod K ≃* H × K := + { Equiv.Set.prod ↑H ↑K with map_mul' := fun x y => rfl } +#align subgroup.prod_equiv Subgroup.prodEquiv +#align add_subgroup.prod_equiv AddSubgroup.prodEquiv + +section Pi + +variable {η : Type _} {f : η → Type _} + +-- defined here and not in group_theory.submonoid.operations to have access to algebra.group.pi +/-- A version of `set.pi` for submonoids. Given an index set `I` and a family of submodules +`s : Π i, submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that +`f i` belongs to `pi I s` whenever `i ∈ I`. -/ +@[to_additive + " A version of `set.pi` for `add_submonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, add_submonoid f i`, `pi I s` is the `add_submonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] +def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : + Submonoid (∀ i, f i) where + carrier := I.pi fun i => (s i).carrier + one_mem' i _ := (s i).one_mem + mul_mem' p q hp hq i hI := (s i).mul_mem (hp i hI) (hq i hI) +#align submonoid.pi Submonoid.pi +#align add_submonoid.pi AddSubmonoid.pi + +variable [∀ i, Group (f i)] + +/-- A version of `set.pi` for subgroups. Given an index set `I` and a family of submodules +`s : Π i, subgroup f i`, `pi I s` is the subgroup of dependent functions `f : Π i, f i` such that +`f i` belongs to `pi I s` whenever `i ∈ I`. -/ +@[to_additive + " A version of `set.pi` for `add_subgroup`s. Given an index set `I` and a family\nof submodules `s : Π i, add_subgroup f i`, `pi I s` is the `add_subgroup` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] +def pi (I : Set η) (H : ∀ i, Subgroup (f i)) : Subgroup (∀ i, f i) := + { Submonoid.pi I fun i => (H i).toSubmonoid with + inv_mem' := fun p hp i hI => (H i).inv_mem (hp i hI) } +#align subgroup.pi Subgroup.pi +#align add_subgroup.pi AddSubgroup.pi + +@[to_additive] +theorem coe_pi (I : Set η) (H : ∀ i, Subgroup (f i)) : + (pi I H : Set (∀ i, f i)) = Set.pi I fun i => (H i : Set (f i)) := + rfl +#align subgroup.coe_pi Subgroup.coe_pi +#align add_subgroup.coe_pi AddSubgroup.coe_pi + +@[to_additive] +theorem mem_pi (I : Set η) {H : ∀ i, Subgroup (f i)} {p : ∀ i, f i} : + p ∈ pi I H ↔ ∀ i : η, i ∈ I → p i ∈ H i := + Iff.rfl +#align subgroup.mem_pi Subgroup.mem_pi +#align add_subgroup.mem_pi AddSubgroup.mem_pi + +@[to_additive] +theorem pi_top (I : Set η) : (pi I fun i => (⊤ : Subgroup (f i))) = ⊤ := + ext fun x => by simp [mem_pi] +#align subgroup.pi_top Subgroup.pi_top +#align add_subgroup.pi_top AddSubgroup.pi_top + +@[to_additive] +theorem pi_empty (H : ∀ i, Subgroup (f i)) : pi ∅ H = ⊤ := + ext fun x => by simp [mem_pi] +#align subgroup.pi_empty Subgroup.pi_empty +#align add_subgroup.pi_empty AddSubgroup.pi_empty + +@[to_additive] +theorem pi_bot : (pi Set.univ fun i => (⊥ : Subgroup (f i))) = ⊥ := + (eq_bot_iff_forall _).mpr fun p hp => + by + simp only [mem_pi, mem_bot] at * + ext j + exact hp j trivial +#align subgroup.pi_bot Subgroup.pi_bot +#align add_subgroup.pi_bot AddSubgroup.pi_bot + +@[to_additive] +theorem le_pi_iff {I : Set η} {H : ∀ i, Subgroup (f i)} {J : Subgroup (∀ i, f i)} : + J ≤ pi I H ↔ ∀ i : η, i ∈ I → map (Pi.evalMonoidHom f i) J ≤ H i := + by + constructor + · intro h i hi + rintro _ ⟨x, hx, rfl⟩ + exact (h hx) _ hi + · intro h x hx i hi + refine' h i hi ⟨_, hx, rfl⟩ +#align subgroup.le_pi_iff Subgroup.le_pi_iff +#align add_subgroup.le_pi_iff AddSubgroup.le_pi_iff + +@[simp, to_additive] +theorem mulSingle_mem_pi [DecidableEq η] {I : Set η} {H : ∀ i, Subgroup (f i)} (i : η) (x : f i) : + Pi.mulSingle i x ∈ pi I H ↔ i ∈ I → x ∈ H i := + by + constructor + · intro h hi + simpa using h i hi + · intro h j hj + by_cases heq : j = i + · subst HEq + simpa using h hj + · simp [HEq, one_mem] +#align subgroup.mul_single_mem_pi Subgroup.mulSingle_mem_pi +#align add_subgroup.single_mem_pi AddSubgroup.single_mem_pi + +@[to_additive] +theorem pi_eq_bot_iff (H : ∀ i, Subgroup (f i)) : pi Set.univ H = ⊥ ↔ ∀ i, H i = ⊥ := by + classical + simp only [eq_bot_iff_forall] + constructor + · intro h i x hx + have : MonoidHom.single f i x = 1 := + h (MonoidHom.single f i x) ((mul_single_mem_pi i x).mpr fun _ => hx) + simpa using congr_fun this i + · exact fun h x hx => funext fun i => h _ _ (hx i trivial) +#align subgroup.pi_eq_bot_iff Subgroup.pi_eq_bot_iff +#align add_subgroup.pi_eq_bot_iff AddSubgroup.pi_eq_bot_iff + +end Pi + +/-- A subgroup is normal if whenever `n ∈ H`, then `g * n * g⁻¹ ∈ H` for every `g : G` -/ +structure Normal : Prop where + conj_mem : ∀ n, n ∈ H → ∀ g : G, g * n * g⁻¹ ∈ H +#align subgroup.normal Subgroup.Normal + +attribute [class] normal + +end Subgroup + +namespace AddSubgroup + +/- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: #[`conj_mem] [] -/ +/-- An add_subgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ +structure Normal (H : AddSubgroup A) : Prop where + conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H +#align add_subgroup.normal AddSubgroup.Normal + +attribute [to_additive AddSubgroup.Normal] Subgroup.Normal + +attribute [class] normal + +end AddSubgroup + +namespace Subgroup + +variable {H K : Subgroup G} + +@[to_additive] +instance (priority := 100) normal_of_comm {G : Type _} [CommGroup G] (H : Subgroup G) : H.Normal := + ⟨by simp [mul_comm, mul_left_comm]⟩ +#align subgroup.normal_of_comm Subgroup.normal_of_comm +#align add_subgroup.normal_of_comm AddSubgroup.normal_of_comm + +namespace Normal + +variable (nH : H.Normal) + +@[to_additive] +theorem mem_comm {a b : G} (h : a * b ∈ H) : b * a ∈ H := + by + have : a⁻¹ * (a * b) * a⁻¹⁻¹ ∈ H := nH.conj_mem (a * b) h a⁻¹ + simpa +#align subgroup.normal.mem_comm Subgroup.Normal.mem_comm +#align add_subgroup.normal.mem_comm AddSubgroup.Normal.mem_comm + +@[to_additive] +theorem mem_comm_iff {a b : G} : a * b ∈ H ↔ b * a ∈ H := + ⟨nH.mem_comm, nH.mem_comm⟩ +#align subgroup.normal.mem_comm_iff Subgroup.Normal.mem_comm_iff +#align add_subgroup.normal.mem_comm_iff AddSubgroup.Normal.mem_comm_iff + +end Normal + +variable (H) + +/-- A subgroup is characteristic if it is fixed by all automorphisms. + Several equivalent conditions are provided by lemmas of the form `characteristic.iff...` -/ +structure Characteristic : Prop where + fixed : ∀ ϕ : G ≃* G, H.comap ϕ.toMonoidHom = H +#align subgroup.characteristic Subgroup.Characteristic + +attribute [class] characteristic + +instance (priority := 100) normal_of_characteristic [h : H.Characteristic] : H.Normal := + ⟨fun a ha b => (SetLike.ext_iff.mp (h.fixed (MulAut.conj b)) a).mpr ha⟩ +#align subgroup.normal_of_characteristic Subgroup.normal_of_characteristic + +end Subgroup + +namespace AddSubgroup + +variable (H : AddSubgroup A) + +/-- A add_subgroup is characteristic if it is fixed by all automorphisms. + Several equivalent conditions are provided by lemmas of the form `characteristic.iff...` -/ +structure Characteristic : Prop where + fixed : ∀ ϕ : A ≃+ A, H.comap ϕ.toAddMonoidHom = H +#align add_subgroup.characteristic AddSubgroup.Characteristic + +attribute [to_additive AddSubgroup.Characteristic] Subgroup.Characteristic + +attribute [class] characteristic + +instance (priority := 100) normal_of_characteristic [h : H.Characteristic] : H.Normal := + ⟨fun a ha b => (SetLike.ext_iff.mp (h.fixed (AddAut.conj b)) a).mpr ha⟩ +#align add_subgroup.normal_of_characteristic AddSubgroup.normal_of_characteristic + +end AddSubgroup + +namespace Subgroup + +variable {H K : Subgroup G} + +@[to_additive] +theorem characteristic_iff_comap_eq : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.comap ϕ.toMonoidHom = H := + ⟨Characteristic.fixed, Characteristic.mk⟩ +#align subgroup.characteristic_iff_comap_eq Subgroup.characteristic_iff_comap_eq +#align add_subgroup.characteristic_iff_comap_eq AddSubgroup.characteristic_iff_comap_eq + +@[to_additive] +theorem characteristic_iff_comap_le : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.comap ϕ.toMonoidHom ≤ H := + characteristic_iff_comap_eq.trans + ⟨fun h ϕ => le_of_eq (h ϕ), fun h ϕ => + le_antisymm (h ϕ) fun g hg => h ϕ.symm ((congr_arg (· ∈ H) (ϕ.symm_apply_apply g)).mpr hg)⟩ +#align subgroup.characteristic_iff_comap_le Subgroup.characteristic_iff_comap_le +#align add_subgroup.characteristic_iff_comap_le AddSubgroup.characteristic_iff_comap_le + +@[to_additive] +theorem characteristic_iff_le_comap : H.Characteristic ↔ ∀ ϕ : G ≃* G, H ≤ H.comap ϕ.toMonoidHom := + characteristic_iff_comap_eq.trans + ⟨fun h ϕ => ge_of_eq (h ϕ), fun h ϕ => + le_antisymm (fun g hg => (congr_arg (· ∈ H) (ϕ.symm_apply_apply g)).mp (h ϕ.symm hg)) (h ϕ)⟩ +#align subgroup.characteristic_iff_le_comap Subgroup.characteristic_iff_le_comap +#align add_subgroup.characteristic_iff_le_comap AddSubgroup.characteristic_iff_le_comap + +@[to_additive] +theorem characteristic_iff_map_eq : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom = H := + by + simp_rw [map_equiv_eq_comap_symm] + exact characteristic_iff_comap_eq.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ +#align subgroup.characteristic_iff_map_eq Subgroup.characteristic_iff_map_eq +#align add_subgroup.characteristic_iff_map_eq AddSubgroup.characteristic_iff_map_eq + +@[to_additive] +theorem characteristic_iff_map_le : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom ≤ H := + by + simp_rw [map_equiv_eq_comap_symm] + exact characteristic_iff_comap_le.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ +#align subgroup.characteristic_iff_map_le Subgroup.characteristic_iff_map_le +#align add_subgroup.characteristic_iff_map_le AddSubgroup.characteristic_iff_map_le + +@[to_additive] +theorem characteristic_iff_le_map : H.Characteristic ↔ ∀ ϕ : G ≃* G, H ≤ H.map ϕ.toMonoidHom := + by + simp_rw [map_equiv_eq_comap_symm] + exact characteristic_iff_le_comap.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ +#align subgroup.characteristic_iff_le_map Subgroup.characteristic_iff_le_map +#align add_subgroup.characteristic_iff_le_map AddSubgroup.characteristic_iff_le_map + +@[to_additive] +instance botCharacteristic : Characteristic (⊥ : Subgroup G) := + characteristic_iff_le_map.mpr fun ϕ => bot_le +#align subgroup.bot_characteristic Subgroup.botCharacteristic +#align add_subgroup.bot_characteristic AddSubgroup.bot_characteristic + +@[to_additive] +instance topCharacteristic : Characteristic (⊤ : Subgroup G) := + characteristic_iff_map_le.mpr fun ϕ => le_top +#align subgroup.top_characteristic Subgroup.topCharacteristic +#align add_subgroup.top_characteristic AddSubgroup.top_characteristic + +variable (G) + +/-- The center of a group `G` is the set of elements that commute with everything in `G` -/ +@[to_additive + "The center of an additive group `G` is the set of elements that commute with\neverything in `G`"] +def center : Subgroup G := + { Submonoid.center G with + carrier := Set.center G + inv_mem' := fun a => Set.inv_mem_center } +#align subgroup.center Subgroup.center +#align add_subgroup.center AddSubgroup.center + +@[to_additive] +theorem coe_center : ↑(center G) = Set.center G := + rfl +#align subgroup.coe_center Subgroup.coe_center +#align add_subgroup.coe_center AddSubgroup.coe_center + +@[simp, to_additive] +theorem center_toSubmonoid : (center G).toSubmonoid = Submonoid.center G := + rfl +#align subgroup.center_to_submonoid Subgroup.center_toSubmonoid +#align add_subgroup.center_to_add_submonoid AddSubgroup.center_to_add_submonoid + +variable {G} + +@[to_additive] +theorem mem_center_iff {z : G} : z ∈ center G ↔ ∀ g, g * z = z * g := + Iff.rfl +#align subgroup.mem_center_iff Subgroup.mem_center_iff +#align add_subgroup.mem_center_iff AddSubgroup.mem_center_iff + +instance decidableMemCenter (z : G) [Decidable (∀ g, g * z = z * g)] : Decidable (z ∈ center G) := + decidable_of_iff' _ mem_center_iff +#align subgroup.decidable_mem_center Subgroup.decidableMemCenter + +@[to_additive] +instance centerCharacteristic : (center G).Characteristic := + by + refine' characteristic_iff_comap_le.mpr fun ϕ g hg h => _ + rw [← ϕ.injective.eq_iff, ϕ.map_mul, ϕ.map_mul] + exact hg (ϕ h) +#align subgroup.center_characteristic Subgroup.centerCharacteristic +#align add_subgroup.center_characteristic AddSubgroup.center_characteristic + +theorem CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := + by + rw [eq_top_iff'] + intro x y + exact mul_comm y x +#align comm_group.center_eq_top CommGroup.center_eq_top + +/-- A group is commutative if the center is the whole group -/ +def Group.commGroupOfCenterEqTop (h : center G = ⊤) : CommGroup G := + { (_ : Group G) with + mul_comm := by + rw [eq_top_iff'] at h + intro x y + exact h y x } +#align group.comm_group_of_center_eq_top Group.commGroupOfCenterEqTop + +variable {G} (H) + +section Normalizer + +/-- 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."] +def normalizer : Subgroup G + where + carrier := { g : G | ∀ n, n ∈ H ↔ g * n * g⁻¹ ∈ H } + one_mem' := by simp + mul_mem' a b (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) (hb : ∀ n, n ∈ H ↔ b * n * b⁻¹ ∈ H) n := + by + rw [hb, ha] + simp [mul_assoc] + inv_mem' a (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) n := + by + rw [ha (a⁻¹ * n * a⁻¹⁻¹)] + simp [mul_assoc] +#align subgroup.normalizer Subgroup.normalizer +#align add_subgroup.normalizer AddSubgroup.normalizer + +-- variant for sets. +-- TODO should this replace `normalizer`? +/-- The `set_normalizer` of `S` is the subgroup of `G` whose elements satisfy `g*S*g⁻¹=S` -/ +@[to_additive + "The `set_normalizer` of `S` is the subgroup of `G` whose elements satisfy\n`g+S-g=S`."] +def setNormalizer (S : Set G) : Subgroup G + where + carrier := { g : G | ∀ n, n ∈ S ↔ g * n * g⁻¹ ∈ S } + one_mem' := by simp + mul_mem' a b (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) (hb : ∀ n, n ∈ S ↔ b * n * b⁻¹ ∈ S) n := + by + rw [hb, ha] + simp [mul_assoc] + inv_mem' a (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n := + by + rw [ha (a⁻¹ * n * a⁻¹⁻¹)] + simp [mul_assoc] +#align subgroup.set_normalizer Subgroup.setNormalizer +#align add_subgroup.set_normalizer AddSubgroup.setNormalizer + +variable {H} + +@[to_additive] +theorem mem_normalizer_iff {g : G} : g ∈ H.normalizer ↔ ∀ h, h ∈ H ↔ g * h * g⁻¹ ∈ H := + Iff.rfl +#align subgroup.mem_normalizer_iff Subgroup.mem_normalizer_iff +#align add_subgroup.mem_normalizer_iff AddSubgroup.mem_normalizer_iff + +@[to_additive] +theorem mem_normalizer_iff'' {g : G} : g ∈ H.normalizer ↔ ∀ h : G, h ∈ H ↔ g⁻¹ * h * g ∈ H := by + rw [← inv_mem_iff, mem_normalizer_iff, inv_inv] +#align subgroup.mem_normalizer_iff'' Subgroup.mem_normalizer_iff'' +#align add_subgroup.mem_normalizer_iff'' AddSubgroup.mem_normalizer_iff'' + +@[to_additive] +theorem mem_normalizer_iff' {g : G} : g ∈ H.normalizer ↔ ∀ n, n * g ∈ H ↔ g * n ∈ H := + ⟨fun h n => by rw [h, mul_assoc, mul_inv_cancel_right], fun h n => by + rw [mul_assoc, ← h, inv_mul_cancel_right]⟩ +#align subgroup.mem_normalizer_iff' Subgroup.mem_normalizer_iff' +#align add_subgroup.mem_normalizer_iff' AddSubgroup.mem_normalizer_iff' + +@[to_additive] +theorem le_normalizer : H ≤ normalizer H := fun x xH n => by + rw [H.mul_mem_cancel_right (H.inv_mem xH), H.mul_mem_cancel_left xH] +#align subgroup.le_normalizer Subgroup.le_normalizer +#align add_subgroup.le_normalizer AddSubgroup.le_normalizer + +@[to_additive] +instance (priority := 100) normal_in_normalizer : (H.subgroupOf H.normalizer).Normal := + ⟨fun x xH g => by simpa using (g.2 x).1 xH⟩ +#align subgroup.normal_in_normalizer Subgroup.normal_in_normalizer +#align add_subgroup.normal_in_normalizer AddSubgroup.normal_in_normalizer + +@[to_additive] +theorem normalizer_eq_top : H.normalizer = ⊤ ↔ H.Normal := + eq_top_iff.trans + ⟨fun h => ⟨fun a ha b => (h (mem_top b) a).mp ha⟩, fun h a ha b => + ⟨fun hb => h.conj_mem b hb a, fun hb => by rwa [h.mem_comm_iff, inv_mul_cancel_left] at hb⟩⟩ +#align subgroup.normalizer_eq_top Subgroup.normalizer_eq_top +#align add_subgroup.normalizer_eq_top AddSubgroup.normalizer_eq_top + +@[to_additive] +theorem center_le_normalizer : center G ≤ H.normalizer := fun x hx y => by + simp [← mem_center_iff.mp hx y, mul_assoc] +#align subgroup.center_le_normalizer Subgroup.center_le_normalizer +#align add_subgroup.center_le_normalizer AddSubgroup.center_le_normalizer + +open Classical + +@[to_additive] +theorem le_normalizer_of_normal [hK : (H.subgroupOf K).Normal] (HK : H ≤ K) : K ≤ H.normalizer := + fun x hx y => + ⟨fun yH => hK.conj_mem ⟨y, HK yH⟩ yH ⟨x, hx⟩, fun yH => by + simpa [mem_subgroup_of, mul_assoc] using + hK.conj_mem ⟨x * y * x⁻¹, HK yH⟩ yH ⟨x⁻¹, K.inv_mem hx⟩⟩ +#align subgroup.le_normalizer_of_normal Subgroup.le_normalizer_of_normal +#align add_subgroup.le_normalizer_of_normal AddSubgroup.le_normalizer_of_normal + +variable {N : Type _} [Group N] + +/-- The preimage of the normalizer is contained in the normalizer of the preimage. -/ +@[to_additive "The preimage of the normalizer is contained in the normalizer of the preimage."] +theorem le_normalizer_comap (f : N →* G) : H.normalizer.comap f ≤ (H.comap f).normalizer := fun x => + by + simp only [mem_normalizer_iff, mem_comap] + intro h n + simp [h (f n)] +#align subgroup.le_normalizer_comap Subgroup.le_normalizer_comap +#align add_subgroup.le_normalizer_comap AddSubgroup.le_normalizer_comap + +/-- The image of the normalizer is contained in the normalizer of the image. -/ +@[to_additive "The image of the normalizer is contained in the normalizer of the image."] +theorem le_normalizer_map (f : G →* N) : H.normalizer.map f ≤ (H.map f).normalizer := fun _ => + by + simp only [and_imp, exists_prop, mem_map, exists_imp, mem_normalizer_iff] + rintro x hx rfl n + constructor + · rintro ⟨y, hy, rfl⟩ + use x * y * x⁻¹, (hx y).1 hy + simp + · rintro ⟨y, hyH, hy⟩ + use x⁻¹ * y * x + rw [hx] + simp [hy, hyH, mul_assoc] +#align subgroup.le_normalizer_map Subgroup.le_normalizer_map +#align add_subgroup.le_normalizer_map AddSubgroup.le_normalizer_map + +variable (G) + +/-- Every proper subgroup `H` of `G` is a proper normal subgroup of the normalizer of `H` in `G`. -/ +def NormalizerCondition := + ∀ H : Subgroup G, H < ⊤ → H < normalizer H +#align normalizer_condition NormalizerCondition + +variable {G} + +/-- Alternative phrasing of the normalizer condition: Only the full group is self-normalizing. +This may be easier to work with, as it avoids inequalities and negations. -/ +theorem normalizerCondition_iff_only_full_group_self_normalizing : + NormalizerCondition G ↔ ∀ H : Subgroup G, H.normalizer = H → H = ⊤ := + by + apply forall_congr'; intro H + simp only [lt_iff_le_and_ne, le_normalizer, true_and_iff, le_top, Ne.def] + tauto +#align normalizer_condition_iff_only_full_group_self_normalizing normalizerCondition_iff_only_full_group_self_normalizing + +variable (H) + +/-- In a group that satisifes the normalizer condition, every maximal subgroup is normal -/ +theorem NormalizerCondition.normal_of_coatom (hnc : NormalizerCondition G) (hmax : IsCoatom H) : + H.Normal := + normalizer_eq_top.mp (hmax.2 _ (hnc H (lt_top_iff_ne_top.mpr hmax.1))) +#align subgroup.normalizer_condition.normal_of_coatom Subgroup.NormalizerCondition.normal_of_coatom + +end Normalizer + +section Centralizer + +/-- The `centralizer` of `H` is the subgroup of `g : G` commuting with every `h : H`. -/ +@[to_additive + "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with\nevery `h : H`."] +def centralizer : Subgroup G := + { Submonoid.centralizer ↑H with + carrier := Set.centralizer H + inv_mem' := fun g => Set.inv_mem_centralizer } +#align subgroup.centralizer Subgroup.centralizer +#align add_subgroup.centralizer AddSubgroup.centralizer + +variable {H} + +@[to_additive] +theorem mem_centralizer_iff {g : G} : g ∈ H.centralizer ↔ ∀ h ∈ H, h * g = g * h := + Iff.rfl +#align subgroup.mem_centralizer_iff Subgroup.mem_centralizer_iff +#align add_subgroup.mem_centralizer_iff AddSubgroup.mem_centralizer_iff + +@[to_additive] +theorem mem_centralizer_iff_commutator_eq_one {g : G} : + g ∈ H.centralizer ↔ ∀ h ∈ H, h * g * h⁻¹ * g⁻¹ = 1 := by + simp only [mem_centralizer_iff, mul_inv_eq_iff_eq_mul, one_mul] +#align subgroup.mem_centralizer_iff_commutator_eq_one Subgroup.mem_centralizer_iff_commutator_eq_one +#align add_subgroup.mem_centralizer_iff_commutator_eq_zero AddSubgroup.mem_centralizer_iff_commutator_eq_zero + +@[to_additive] +theorem centralizer_top : centralizer ⊤ = center G := + SetLike.ext' (Set.centralizer_univ G) +#align subgroup.centralizer_top Subgroup.centralizer_top +#align add_subgroup.centralizer_top AddSubgroup.centralizer_top + +@[to_additive] +theorem le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := + ⟨fun h x hx y hy => (h hy x hx).symm, fun h x hx y hy => (h hy x hx).symm⟩ +#align subgroup.le_centralizer_iff Subgroup.le_centralizer_iff +#align add_subgroup.le_centralizer_iff AddSubgroup.le_centralizer_iff + +@[to_additive] +theorem centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := + Submonoid.centralizer_le h +#align subgroup.centralizer_le Subgroup.centralizer_le +#align add_subgroup.centralizer_le AddSubgroup.centralizer_le + +@[to_additive] +instance Subgroup.Centralizer.characteristic [hH : H.Characteristic] : + H.centralizer.Characteristic := + by + refine' subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.Injective _ + rw [map_mul, map_mul] + exact hg (ϕ h) (subgroup.characteristic_iff_le_comap.mp hH ϕ hh) +#align subgroup.subgroup.centralizer.characteristic Subgroup.Subgroup.Centralizer.characteristic +#align add_subgroup.subgroup.centralizer.characteristic AddSubgroup.Subgroup.Centralizer.characteristic + +end Centralizer + +/-- Commutivity of a subgroup -/ +structure IsCommutative : Prop where + is_comm : IsCommutative H (· * ·) +#align subgroup.is_commutative Subgroup.IsCommutative + +attribute [class] IsCommutative + +/-- Commutivity of an additive subgroup -/ +structure AddSubgroup.IsCommutative (H : AddSubgroup A) : Prop where + is_comm : IsCommutative H (· + ·) +#align add_subgroup.is_commutative AddSubgroup.IsCommutative + +attribute [to_additive AddSubgroup.IsCommutative] Subgroup.IsCommutative + +attribute [class] AddSubgroup.IsCommutative + +/-- A commutative subgroup is commutative. -/ +@[to_additive "A commutative subgroup is commutative."] +instance IsCommutative.commGroup [h : H.IsCommutative] : CommGroup H := + { H.toGroup with mul_comm := h.is_comm.comm } +#align subgroup.is_commutative.comm_group Subgroup.IsCommutative.commGroup +#align add_subgroup.is_commutative.add_comm_group AddSubgroup.IsCommutative.addCommGroup + +instance center.isCommutative : (center G).IsCommutative := + ⟨⟨fun a b => Subtype.ext (b.2 a)⟩⟩ +#align subgroup.center.is_commutative Subgroup.center.isCommutative + +@[to_additive] +instance map_isCommutative (f : G →* G') [H.IsCommutative] : (H.map f).IsCommutative := + ⟨⟨by + rintro ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩ + rw [Subtype.ext_iff, coe_mul, coe_mul, Subtype.coe_mk, Subtype.coe_mk, ← map_mul, ← map_mul] + exact congr_arg f (subtype.ext_iff.mp (mul_comm ⟨a, ha⟩ ⟨b, hb⟩))⟩⟩ +#align subgroup.map_is_commutative Subgroup.map_isCommutative +#align add_subgroup.map_is_commutative AddSubgroup.map_isCommutative + +@[to_additive] +theorem comap_injective_isCommutative {f : G' →* G} (hf : Injective f) [H.IsCommutative] : + (H.comap f).IsCommutative := + ⟨⟨fun a b => + Subtype.ext + (by + have := mul_comm (⟨f a, a.2⟩ : H) (⟨f b, b.2⟩ : H) + rwa [Subtype.ext_iff, coe_mul, coe_mul, coe_mk, coe_mk, ← map_mul, ← map_mul, + hf.eq_iff] at this)⟩⟩ +#align subgroup.comap_injective_is_commutative Subgroup.comap_injective_isCommutative +#align add_subgroup.comap_injective_is_commutative AddSubgroup.comap_injective_isCommutative + +@[to_additive] +instance subgroupOf_isCommutative [H.IsCommutative] : (H.subgroupOf K).IsCommutative := + H.comap_injective_is_commutative Subtype.coe_injective +#align subgroup.subgroup_of_is_commutative Subgroup.subgroupOf_isCommutative +#align add_subgroup.add_subgroup_of_is_commutative AddSubgroup.add_subgroupOf_isCommutative + +@[to_additive] +theorem le_centralizer_iff_isCommutative : K ≤ K.centralizer ↔ K.IsCommutative := + ⟨fun h => ⟨⟨fun x y => Subtype.ext (h y.2 x x.2)⟩⟩, fun h x hx y hy => + congr_arg coe (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ +#align subgroup.le_centralizer_iff_is_commutative Subgroup.le_centralizer_iff_isCommutative +#align add_subgroup.le_centralizer_iff_is_commutative AddSubgroup.le_centralizer_iff_isCommutative + +@[to_additive] +theorem le_centralizer [h : H.IsCommutative] : H ≤ H.centralizer := + le_centralizer_iff_isCommutative.mpr h +#align subgroup.le_centralizer Subgroup.le_centralizer +#align add_subgroup.le_centralizer AddSubgroup.le_centralizer + +end Subgroup + +namespace Group + +variable {s : Set G} + +/-- Given a set `s`, `conjugates_of_set s` is the set of all conjugates of +the elements of `s`. -/ +def conjugatesOfSet (s : Set G) : Set G := + ⋃ a ∈ s, conjugatesOf a +#align group.conjugates_of_set Group.conjugatesOfSet + +theorem mem_conjugatesOfSet_iff {x : G} : x ∈ conjugatesOfSet s ↔ ∃ a ∈ s, IsConj a x := + Set.mem_unionᵢ₂ +#align group.mem_conjugates_of_set_iff Group.mem_conjugatesOfSet_iff + +theorem subset_conjugatesOfSet : s ⊆ conjugatesOfSet s := fun (x : G) (h : x ∈ s) => + mem_conjugatesOfSet_iff.2 ⟨x, h, IsConj.refl _⟩ +#align group.subset_conjugates_of_set Group.subset_conjugatesOfSet + +theorem conjugatesOfSet_mono {s t : Set G} (h : s ⊆ t) : conjugatesOfSet s ⊆ conjugatesOfSet t := + Set.bunionᵢ_subset_bunionᵢ_left h +#align group.conjugates_of_set_mono Group.conjugatesOfSet_mono + +theorem conjugates_subset_normal {N : Subgroup G} [tn : N.Normal] {a : G} (h : a ∈ N) : + conjugatesOf a ⊆ N := by + rintro a hc + obtain ⟨c, rfl⟩ := isConj_iff.1 hc + exact tn.conj_mem a h c +#align group.conjugates_subset_normal Group.conjugates_subset_normal + +theorem conjugatesOfSet_subset {s : Set G} {N : Subgroup G} [N.Normal] (h : s ⊆ N) : + conjugatesOfSet s ⊆ N := + Set.unionᵢ₂_subset fun x H => conjugates_subset_normal (h H) +#align group.conjugates_of_set_subset Group.conjugatesOfSet_subset + +/-- The set of conjugates of `s` is closed under conjugation. -/ +theorem conj_mem_conjugatesOfSet {x c : G} : + x ∈ conjugatesOfSet s → c * x * c⁻¹ ∈ conjugatesOfSet s := fun H => + by + rcases mem_conjugates_of_set_iff.1 H with ⟨a, h₁, h₂⟩ + exact mem_conjugates_of_set_iff.2 ⟨a, h₁, h₂.trans (isConj_iff.2 ⟨c, rfl⟩)⟩ +#align group.conj_mem_conjugates_of_set Group.conj_mem_conjugatesOfSet + +end Group + +namespace Subgroup + +open Group + +variable {s : Set G} + +/-- The normal closure of a set `s` is the subgroup closure of all the conjugates of +elements of `s`. It is the smallest normal subgroup containing `s`. -/ +def normalClosure (s : Set G) : Subgroup G := + closure (conjugatesOfSet s) +#align subgroup.normal_closure Subgroup.normalClosure + +theorem conjugatesOfSet_subset_normalClosure : conjugatesOfSet s ⊆ normalClosure s := + subset_closure +#align subgroup.conjugates_of_set_subset_normal_closure Subgroup.conjugatesOfSet_subset_normalClosure + +theorem subset_normalClosure : s ⊆ normalClosure s := + Set.Subset.trans subset_conjugatesOfSet conjugatesOfSet_subset_normalClosure +#align subgroup.subset_normal_closure Subgroup.subset_normalClosure + +theorem le_normalClosure {H : Subgroup G} : H ≤ normalClosure ↑H := fun _ h => + subset_normalClosure h +#align subgroup.le_normal_closure Subgroup.le_normalClosure + +/-- The normal closure of `s` is a normal subgroup. -/ +instance normalClosure_normal : (normalClosure s).Normal := + ⟨fun n h g => + by + refine' Subgroup.closure_induction h (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ + · exact conjugates_of_set_subset_normal_closure (conj_mem_conjugates_of_set hx) + · simpa using (normal_closure s).one_mem + · rw [← conj_mul] + exact mul_mem ihx ihy + · rw [← conj_inv] + exact inv_mem ihx⟩ +#align subgroup.normal_closure_normal Subgroup.normalClosure_normal + +/-- The normal closure of `s` is the smallest normal subgroup containing `s`. -/ +theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : normalClosure s ≤ N := + by + intro a w + refine' closure_induction w (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ + · exact conjugates_of_set_subset h hx + · exact one_mem _ + · exact mul_mem ihx ihy + · exact inv_mem ihx +#align subgroup.normal_closure_le_normal Subgroup.normalClosure_le_normal + +theorem normalClosure_subset_iff {N : Subgroup G} [N.Normal] : s ⊆ N ↔ normalClosure s ≤ N := + ⟨normalClosure_le_normal, Set.Subset.trans subset_normalClosure⟩ +#align subgroup.normal_closure_subset_iff Subgroup.normalClosure_subset_iff + +theorem normalClosure_mono {s t : Set G} (h : s ⊆ t) : normalClosure s ≤ normalClosure t := + normalClosure_le_normal (Set.Subset.trans h subset_normalClosure) +#align subgroup.normal_closure_mono Subgroup.normalClosure_mono + +theorem normalClosure_eq_infᵢ : + normalClosure s = ⨅ (N : Subgroup G) (_ : Normal N) (hs : s ⊆ N), N := + le_antisymm (le_infᵢ fun N => le_infᵢ fun hN => le_infᵢ normal_closure_le_normal) + (infᵢ_le_of_le (normalClosure s) + (infᵢ_le_of_le (by infer_instance) (infᵢ_le_of_le subset_normalClosure le_rfl))) +#align subgroup.normal_closure_eq_infi Subgroup.normalClosure_eq_infᵢ + +@[simp] +theorem normalClosure_eq_self (H : Subgroup G) [H.Normal] : normalClosure ↑H = H := + le_antisymm (normalClosure_le_normal rfl.Subset) le_normalClosure +#align subgroup.normal_closure_eq_self Subgroup.normalClosure_eq_self + +@[simp] +theorem normalClosure_idempotent : normalClosure ↑(normalClosure s) = normalClosure s := + normalClosure_eq_self _ +#align subgroup.normal_closure_idempotent Subgroup.normalClosure_idempotent + +theorem closure_le_normalClosure {s : Set G} : closure s ≤ normalClosure s := by + simp only [subset_normal_closure, closure_le] +#align subgroup.closure_le_normal_closure Subgroup.closure_le_normalClosure + +@[simp] +theorem normalClosure_closure_eq_normalClosure {s : Set G} : + normalClosure ↑(closure s) = normalClosure s := + le_antisymm (normalClosure_le_normal closure_le_normalClosure) (normalClosure_mono subset_closure) +#align subgroup.normal_closure_closure_eq_normal_closure Subgroup.normalClosure_closure_eq_normalClosure + +/-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, +as shown by `subgroup.normal_core_eq_supr`. -/ +def normalCore (H : Subgroup G) : Subgroup G + where + carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } + one_mem' a := by rw [mul_one, mul_inv_self] <;> exact H.one_mem + inv_mem' a h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) + mul_mem' a b ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) +#align subgroup.normal_core Subgroup.normalCore + +theorem normalCore_le (H : Subgroup G) : H.normalCore ≤ H := fun a h => + by + rw [← mul_one a, ← inv_one, ← one_mul a] + exact h 1 +#align subgroup.normal_core_le Subgroup.normalCore_le + +instance normalCore_normal (H : Subgroup G) : H.normalCore.Normal := + ⟨fun a h b c => by + rw [mul_assoc, mul_assoc, ← mul_inv_rev, ← mul_assoc, ← mul_assoc] <;> exact h (c * b)⟩ +#align subgroup.normal_core_normal Subgroup.normalCore_normal + +theorem normal_le_normalCore {H : Subgroup G} {N : Subgroup G} [hN : N.Normal] : + N ≤ H.normalCore ↔ N ≤ H := + ⟨ge_trans H.normal_core_le, fun h_le n hn g => h_le (hN.conj_mem n hn g)⟩ +#align subgroup.normal_le_normal_core Subgroup.normal_le_normalCore + +theorem normalCore_mono {H K : Subgroup G} (h : H ≤ K) : H.normalCore ≤ K.normalCore := + normal_le_normalCore.mpr (H.normal_core_le.trans h) +#align subgroup.normal_core_mono Subgroup.normalCore_mono + +theorem normalCore_eq_supᵢ (H : Subgroup G) : + H.normalCore = ⨆ (N : Subgroup G) (_ : Normal N) (hs : N ≤ H), N := + le_antisymm + (le_supᵢ_of_le H.normalCore + (le_supᵢ_of_le H.normal_core_normal (le_supᵢ_of_le H.normal_core_le le_rfl))) + (supᵢ_le fun N => supᵢ_le fun hN => supᵢ_le normal_le_normal_core.mpr) +#align subgroup.normal_core_eq_supr Subgroup.normalCore_eq_supᵢ + +@[simp] +theorem normalCore_eq_self (H : Subgroup G) [H.Normal] : H.normalCore = H := + le_antisymm H.normal_core_le (normal_le_normalCore.mpr le_rfl) +#align subgroup.normal_core_eq_self Subgroup.normalCore_eq_self + +@[simp] +theorem normalCore_idempotent (H : Subgroup G) : H.normalCore.normalCore = H.normalCore := + H.normalCore.normal_core_eq_self +#align subgroup.normal_core_idempotent Subgroup.normalCore_idempotent + +end Subgroup + +namespace MonoidHom + +variable {N : Type _} {P : Type _} [Group N] [Group P] (K : Subgroup G) + +open Subgroup + +/-- The range of a monoid homomorphism from a group is a subgroup. -/ +@[to_additive "The range of an `add_monoid_hom` from an `add_group` is an `add_subgroup`."] +def range (f : G →* N) : Subgroup N := + Subgroup.copy ((⊤ : Subgroup G).map f) (Set.range f) (by simp [Set.ext_iff]) +#align monoid_hom.range MonoidHom.range +#align add_monoid_hom.range AddMonoidHom.range + +@[simp, to_additive] +theorem coe_range (f : G →* N) : (f.range : Set N) = Set.range f := + rfl +#align monoid_hom.coe_range MonoidHom.coe_range +#align add_monoid_hom.coe_range AddMonoidHom.coe_range + +@[simp, to_additive] +theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := + Iff.rfl +#align monoid_hom.mem_range MonoidHom.mem_range +#align add_monoid_hom.mem_range AddMonoidHom.mem_range + +@[to_additive] +theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext <;> simp +#align monoid_hom.range_eq_map MonoidHom.range_eq_map +#align add_monoid_hom.range_eq_map AddMonoidHom.range_eq_map + +@[simp, to_additive] +theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by + simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, Subtype.coe_mk, + iff_self_iff, forall_const] +#align monoid_hom.restrict_range MonoidHom.restrict_range +#align add_monoid_hom.restrict_range AddMonoidHom.restrict_range + +/-- The canonical surjective group homomorphism `G →* f(G)` induced by a group +homomorphism `G →* N`. -/ +@[to_additive + "The canonical surjective `add_group` homomorphism `G →+ f(G)` induced by a group\nhomomorphism `G →+ N`."] +def rangeRestrict (f : G →* N) : G →* f.range := + codRestrict f _ fun x => ⟨x, rfl⟩ +#align monoid_hom.range_restrict MonoidHom.rangeRestrict +#align add_monoid_hom.range_restrict AddMonoidHom.rangeRestrict + +@[simp, to_additive] +theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.range_restrict g : N) = f g := + rfl +#align monoid_hom.coe_range_restrict MonoidHom.coe_rangeRestrict +#align add_monoid_hom.coe_range_restrict AddMonoidHom.coe_rangeRestrict + +@[to_additive] +theorem coe_comp_rangeRestrict (f : G →* N) : + (coe : f.range → N) ∘ (⇑f.range_restrict : G → f.range) = f := + rfl +#align monoid_hom.coe_comp_range_restrict MonoidHom.coe_comp_rangeRestrict +#align add_monoid_hom.coe_comp_range_restrict AddMonoidHom.coe_comp_rangeRestrict + +@[to_additive] +theorem subtype_comp_rangeRestrict (f : G →* N) : f.range.Subtype.comp f.range_restrict = f := + ext <| f.coe_range_restrict +#align monoid_hom.subtype_comp_range_restrict MonoidHom.subtype_comp_rangeRestrict +#align add_monoid_hom.subtype_comp_range_restrict AddMonoidHom.subtype_comp_rangeRestrict + +@[to_additive] +theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.range_restrict := + fun ⟨_, g, rfl⟩ => ⟨g, rfl⟩ +#align monoid_hom.range_restrict_surjective MonoidHom.rangeRestrict_surjective +#align add_monoid_hom.range_restrict_surjective AddMonoidHom.rangeRestrict_surjective + +@[to_additive] +theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by + rw [range_eq_map, range_eq_map] <;> exact (⊤ : Subgroup G).map_map g f +#align monoid_hom.map_range MonoidHom.map_range +#align add_monoid_hom.map_range AddMonoidHom.map_range + +@[to_additive] +theorem range_top_iff_surjective {N} [Group N] {f : G →* N} : + f.range = (⊤ : Subgroup N) ↔ Function.Surjective f := + SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_range, coe_top]) Set.range_iff_surjective +#align monoid_hom.range_top_iff_surjective MonoidHom.range_top_iff_surjective +#align add_monoid_hom.range_top_iff_surjective AddMonoidHom.range_top_iff_surjective + +/-- The range of a surjective monoid homomorphism is the whole of the codomain. -/ +@[to_additive "The range of a surjective `add_monoid` homomorphism is the whole of the codomain."] +theorem range_top_of_surjective {N} [Group N] (f : G →* N) (hf : Function.Surjective f) : + f.range = (⊤ : Subgroup N) := + range_top_iff_surjective.2 hf +#align monoid_hom.range_top_of_surjective MonoidHom.range_top_of_surjective +#align add_monoid_hom.range_top_of_surjective AddMonoidHom.range_top_of_surjective + +@[simp, to_additive] +theorem range_one : (1 : G →* N).range = ⊥ := + SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x +#align monoid_hom.range_one MonoidHom.range_one +#align add_monoid_hom.range_zero AddMonoidHom.range_zero + +@[simp, to_additive] +theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := + by + rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] + ext + simp +#align subgroup.subtype_range Subgroup.subtype_range +#align add_subgroup.subtype_range AddSubgroup.subtype_range + +@[simp, to_additive] +theorem Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : + (inclusion h_le).range = H.subgroupOf K := + Subgroup.ext fun g => Set.ext_iff.mp (Set.range_inclusion h_le) g +#align subgroup.inclusion_range Subgroup.inclusion_range +#align add_subgroup.inclusion_range AddSubgroup.inclusion_range + +@[to_additive] +theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type _} [Group G₁] [Group G₂] {K : Subgroup G₂} + (f : G₁ →* G₂) (h : f.range ≤ K) : + f.range.subgroupOf K = (f.codRestrict K fun x => h ⟨x, rfl⟩).range := + by + ext k + refine' exists_congr _ + simp [Subtype.ext_iff] +#align monoid_hom.subgroup_of_range_eq_of_le MonoidHom.subgroupOf_range_eq_of_le +#align add_monoid_hom.add_subgroup_of_range_eq_of_le AddMonoidHom.add_subgroup_of_range_eq_of_le + +/-- Computable alternative to `monoid_hom.of_injective`. -/ +@[to_additive "Computable alternative to `add_monoid_hom.of_injective`."] +def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : G ≃* f.range := + { f.range_restrict with + toFun := f.range_restrict + invFun := g ∘ f.range.Subtype + left_inv := h + right_inv := by + rintro ⟨x, y, rfl⟩ + apply Subtype.ext + rw [coe_range_restrict, Function.comp_apply, Subgroup.coeSubtype, Subtype.coe_mk, h] } +#align monoid_hom.of_left_inverse MonoidHom.ofLeftInverse +#align add_monoid_hom.of_left_inverse AddMonoidHom.ofLeftInverse + +@[simp, to_additive] +theorem ofLeftInverse_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) (x : G) : + ↑(ofLeftInverse h x) = f x := + rfl +#align monoid_hom.of_left_inverse_apply MonoidHom.ofLeftInverse_apply +#align add_monoid_hom.of_left_inverse_apply AddMonoidHom.ofLeftInverse_apply + +@[simp, to_additive] +theorem ofLeftInverse_symm_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) + (x : f.range) : (ofLeftInverse h).symm x = g x := + rfl +#align monoid_hom.of_left_inverse_symm_apply MonoidHom.ofLeftInverse_symm_apply +#align add_monoid_hom.of_left_inverse_symm_apply AddMonoidHom.ofLeftInverse_symm_apply + +/-- The range of an injective group homomorphism is isomorphic to its domain. -/ +@[to_additive "The range of an injective additive group homomorphism is isomorphic to its\ndomain."] +noncomputable def ofInjective {f : G →* N} (hf : Function.Injective f) : G ≃* f.range := + MulEquiv.ofBijective (f.codRestrict f.range fun x => ⟨x, rfl⟩) + ⟨fun x y h => hf (Subtype.ext_iff.mp h), + by + rintro ⟨x, y, rfl⟩ + exact ⟨y, rfl⟩⟩ +#align monoid_hom.of_injective MonoidHom.ofInjective +#align add_monoid_hom.of_injective AddMonoidHom.ofInjective + +@[to_additive] +theorem ofInjective_apply {f : G →* N} (hf : Function.Injective f) {x : G} : + ↑(ofInjective hf x) = f x := + rfl +#align monoid_hom.of_injective_apply MonoidHom.ofInjective_apply +#align add_monoid_hom.of_injective_apply AddMonoidHom.ofInjective_apply + +section Ker + +variable {M : Type _} [MulOneClass M] + +/-- The multiplicative kernel of a monoid homomorphism is the subgroup of elements `x : G` such that +`f x = 1` -/ +@[to_additive + "The additive kernel of an `add_monoid` homomorphism is the `add_subgroup` of elements\nsuch that `f x = 0`"] +def ker (f : G →* M) : Subgroup G := + { f.mker with + inv_mem' := fun x (hx : f x = 1) => + calc + f x⁻¹ = f x * f x⁻¹ := by rw [hx, one_mul] + _ = 1 := by rw [← map_mul, mul_inv_self, map_one] + } +#align monoid_hom.ker MonoidHom.ker +#align add_monoid_hom.ker AddMonoidHom.ker + +@[to_additive] +theorem mem_ker (f : G →* M) {x : G} : x ∈ f.ker ↔ f x = 1 := + Iff.rfl +#align monoid_hom.mem_ker MonoidHom.mem_ker +#align add_monoid_hom.mem_ker AddMonoidHom.mem_ker + +@[to_additive] +theorem coe_ker (f : G →* M) : (f.ker : Set G) = (f : G → M) ⁻¹' {1} := + rfl +#align monoid_hom.coe_ker MonoidHom.coe_ker +#align add_monoid_hom.coe_ker AddMonoidHom.coe_ker + +@[simp, to_additive] +theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := + by + ext x + simp [mem_ker, Units.ext_iff] +#align monoid_hom.ker_to_hom_units MonoidHom.ker_toHomUnits +#align add_monoid_hom.ker_to_hom_add_units AddMonoidHom.ker_to_hom_add_units + +@[to_additive] +theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := + by + constructor <;> intro h + · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_self, map_one] + · rw [← one_mul x, ← mul_inv_self y, mul_assoc, map_mul, f.mem_ker.1 h, mul_one] +#align monoid_hom.eq_iff MonoidHom.eq_iff +#align add_monoid_hom.eq_iff AddMonoidHom.eq_iff + +@[to_additive] +instance decidableMemKer [DecidableEq M] (f : G →* M) : DecidablePred (· ∈ f.ker) := fun x => + decidable_of_iff (f x = 1) f.mem_ker +#align monoid_hom.decidable_mem_ker MonoidHom.decidableMemKer +#align add_monoid_hom.decidable_mem_ker AddMonoidHom.decidableMemKer + +@[to_additive] +theorem comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker := + rfl +#align monoid_hom.comap_ker MonoidHom.comap_ker +#align add_monoid_hom.comap_ker AddMonoidHom.comap_ker + +@[simp, to_additive] +theorem comap_bot (f : G →* N) : (⊥ : Subgroup N).comap f = f.ker := + rfl +#align monoid_hom.comap_bot MonoidHom.comap_bot +#align add_monoid_hom.comap_bot AddMonoidHom.comap_bot + +@[simp, to_additive] +theorem ker_restrict (f : G →* N) : (f.restrict K).ker = f.ker.subgroupOf K := + rfl +#align monoid_hom.ker_restrict MonoidHom.ker_restrict +#align add_monoid_hom.ker_restrict AddMonoidHom.ker_restrict + +@[simp, to_additive] +theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s : S) + (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := + SetLike.ext fun x => Subtype.ext_iff +#align monoid_hom.ker_cod_restrict MonoidHom.ker_codRestrict +#align add_monoid_hom.ker_cod_restrict AddMonoidHom.ker_cod_restrict + +@[simp, to_additive] +theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := + ker_codRestrict _ _ _ +#align monoid_hom.ker_range_restrict MonoidHom.ker_rangeRestrict +#align add_monoid_hom.ker_range_restrict AddMonoidHom.ker_rangeRestrict + +@[simp, to_additive] +theorem ker_one : (1 : G →* M).ker = ⊤ := + SetLike.ext fun x => eq_self_iff_true _ +#align monoid_hom.ker_one MonoidHom.ker_one +#align add_monoid_hom.ker_zero AddMonoidHom.ker_zero + +@[simp, to_additive] +theorem ker_id : (MonoidHom.id G).ker = ⊥ := + rfl +#align monoid_hom.ker_id MonoidHom.ker_id +#align add_monoid_hom.ker_id AddMonoidHom.ker_id + +@[to_additive] +theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := + ⟨fun h x y hxy => by rwa [eq_iff, h, mem_bot, inv_mul_eq_one, eq_comm] at hxy, fun h => + bot_unique fun x hx => h (hx.trans f.map_one.symm)⟩ +#align monoid_hom.ker_eq_bot_iff MonoidHom.ker_eq_bot_iff +#align add_monoid_hom.ker_eq_bot_iff AddMonoidHom.ker_eq_bot_iff + +@[simp, to_additive] +theorem Subgroup.ker_subtype (H : Subgroup G) : H.Subtype.ker = ⊥ := + H.Subtype.ker_eq_bot_iff.mpr Subtype.coe_injective +#align subgroup.ker_subtype Subgroup.ker_subtype +#align add_subgroup.ker_subtype AddSubgroup.ker_subtype + +@[simp, to_additive] +theorem Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := + (inclusion h).ker_eq_bot_iff.mpr (Set.inclusion_injective h) +#align subgroup.ker_inclusion Subgroup.ker_inclusion +#align add_subgroup.ker_inclusion AddSubgroup.ker_inclusion + +@[to_additive] +theorem prodMap_comap_prod {G' : Type _} {N' : Type _} [Group G'] [Group N'] (f : G →* N) + (g : G' →* N') (S : Subgroup N) (S' : Subgroup N') : + (S.Prod S').comap (prodMap f g) = (S.comap f).Prod (S'.comap g) := + SetLike.coe_injective <| Set.preimage_prod_map_prod f g _ _ +#align monoid_hom.prod_map_comap_prod MonoidHom.prodMap_comap_prod +#align add_monoid_hom.sum_map_comap_sum AddMonoidHom.sum_map_comap_sum + +@[to_additive] +theorem ker_prodMap {G' : Type _} {N' : Type _} [Group G'] [Group N'] (f : G →* N) (g : G' →* N') : + (prodMap f g).ker = f.ker.Prod g.ker := by + rw [← comap_bot, ← comap_bot, ← comap_bot, ← prod_map_comap_prod, bot_prod_bot] +#align monoid_hom.ker_prod_map MonoidHom.ker_prodMap +#align add_monoid_hom.ker_sum_map AddMonoidHom.ker_sum_map + +@[to_additive] +instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := + ⟨fun x hx y => by + rw [mem_ker, map_mul, map_mul, f.mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_self y)]⟩ +#align monoid_hom.normal_ker MonoidHom.normal_ker +#align add_monoid_hom.normal_ker AddMonoidHom.normal_ker + +end Ker + +section EqLocus + +variable {M : Type _} [Monoid M] + +/-- The subgroup of elements `x : G` such that `f x = g x` -/ +@[to_additive "The additive subgroup of elements `x : G` such that `f x = g x`"] +def eqLocus (f g : G →* M) : Subgroup G := + { eqLocusM f g with inv_mem' := fun x => eq_on_inv f g } +#align monoid_hom.eq_locus MonoidHom.eqLocus +#align add_monoid_hom.eq_locus AddMonoidHom.eqLocus + +@[simp, to_additive] +theorem eqLocus_same (f : G →* N) : f.eqLocus f = ⊤ := + SetLike.ext fun _ => eq_self_iff_true _ +#align monoid_hom.eq_locus_same MonoidHom.eqLocus_same +#align add_monoid_hom.eq_locus_same AddMonoidHom.eqLocus_same + +/-- If two monoid homomorphisms are equal on a set, then they are equal on its subgroup closure. -/ +@[to_additive + "If two monoid homomorphisms are equal on a set, then they are equal on its subgroup\nclosure."] +theorem eqOn_closure {f g : G →* M} {s : Set G} (h : Set.EqOn f g s) : Set.EqOn f g (closure s) := + show closure s ≤ f.eqLocus g from (closure_le _).2 h +#align monoid_hom.eq_on_closure MonoidHom.eqOn_closure +#align add_monoid_hom.eq_on_closure AddMonoidHom.eqOn_closure + +@[to_additive] +theorem eq_of_eqOn_top {f g : G →* M} (h : Set.EqOn f g (⊤ : Subgroup G)) : f = g := + ext fun x => h trivial +#align monoid_hom.eq_of_eq_on_top MonoidHom.eq_of_eqOn_top +#align add_monoid_hom.eq_of_eq_on_top AddMonoidHom.eq_of_eqOn_top + +@[to_additive] +theorem eq_of_eqOn_dense {s : Set G} (hs : closure s = ⊤) {f g : G →* M} (h : s.EqOn f g) : f = g := + eq_of_eq_on_top <| hs ▸ eqOn_closure h +#align monoid_hom.eq_of_eq_on_dense MonoidHom.eq_of_eqOn_dense +#align add_monoid_hom.eq_of_eq_on_dense AddMonoidHom.eq_of_eqOn_dense + +end EqLocus + +@[to_additive] +theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) ≤ (closure s).comap f := + (closure_le _).2 fun x hx => by rw [SetLike.mem_coe, mem_comap] <;> exact subset_closure hx +#align monoid_hom.closure_preimage_le MonoidHom.closure_preimage_le +#align add_monoid_hom.closure_preimage_le AddMonoidHom.closure_preimage_le + +/-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup +generated by the image of the set. -/ +@[to_additive + "The image under an `add_monoid` hom of the `add_subgroup` generated by a set equals\nthe `add_subgroup` generated by the image of the set."] +theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := + Set.image_preimage.l_comm_of_u_comm (Subgroup.gc_map_comap f) (Subgroup.gi N).gc + (Subgroup.gi G).gc fun t => rfl +#align monoid_hom.map_closure MonoidHom.map_closure +#align add_monoid_hom.map_closure AddMonoidHom.map_closure + +end MonoidHom + +namespace Subgroup + +variable {N : Type _} [Group N] (H : Subgroup G) + +@[to_additive] +theorem Normal.map {H : Subgroup G} (h : H.Normal) (f : G →* N) (hf : Function.Surjective f) : + (H.map f).Normal := + by + rw [← normalizer_eq_top, ← top_le_iff, ← f.range_top_of_surjective hf, f.range_eq_map, ← + normalizer_eq_top.2 h] + exact le_normalizer_map _ +#align subgroup.normal.map Subgroup.Normal.map +#align add_subgroup.normal.map AddSubgroup.Normal.map + +@[to_additive] +theorem map_eq_bot_iff {f : G →* N} : H.map f = ⊥ ↔ H ≤ f.ker := + (gc_map_comap f).l_eq_bot +#align subgroup.map_eq_bot_iff Subgroup.map_eq_bot_iff +#align add_subgroup.map_eq_bot_iff AddSubgroup.map_eq_bot_iff + +@[to_additive] +theorem map_eq_bot_iff_of_injective {f : G →* N} (hf : Function.Injective f) : + H.map f = ⊥ ↔ H = ⊥ := by rw [map_eq_bot_iff, f.ker_eq_bot_iff.mpr hf, le_bot_iff] +#align subgroup.map_eq_bot_iff_of_injective Subgroup.map_eq_bot_iff_of_injective +#align add_subgroup.map_eq_bot_iff_of_injective AddSubgroup.map_eq_bot_iff_of_injective + +end Subgroup + +namespace Subgroup + +open MonoidHom + +variable {N : Type _} [Group N] (f : G →* N) + +@[to_additive] +theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := + (range_eq_map f).symm ▸ map_mono le_top +#align subgroup.map_le_range Subgroup.map_le_range +#align add_subgroup.map_le_range AddSubgroup.map_le_range + +@[to_additive] +theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.Subtype ≤ H := + (K.map_le_range H.Subtype).trans (le_of_eq H.subtype_range) +#align subgroup.map_subtype_le Subgroup.map_subtype_le +#align add_subgroup.map_subtype_le AddSubgroup.map_subtype_le + +@[to_additive] +theorem ker_le_comap (H : Subgroup N) : f.ker ≤ comap f H := + comap_bot f ▸ comap_mono bot_le +#align subgroup.ker_le_comap Subgroup.ker_le_comap +#align add_subgroup.ker_le_comap AddSubgroup.ker_le_comap + +@[to_additive] +theorem map_comap_le (H : Subgroup N) : map f (comap f H) ≤ H := + (gc_map_comap f).l_u_le _ +#align subgroup.map_comap_le Subgroup.map_comap_le +#align add_subgroup.map_comap_le AddSubgroup.map_comap_le + +@[to_additive] +theorem le_comap_map (H : Subgroup G) : H ≤ comap f (map f H) := + (gc_map_comap f).le_u_l _ +#align subgroup.le_comap_map Subgroup.le_comap_map +#align add_subgroup.le_comap_map AddSubgroup.le_comap_map + +@[to_additive] +theorem map_comap_eq (H : Subgroup N) : map f (comap f H) = f.range ⊓ H := + SetLike.ext' <| by + rw [coe_map, coe_comap, Set.image_preimage_eq_inter_range, coe_inf, coe_range, Set.inter_comm] +#align subgroup.map_comap_eq Subgroup.map_comap_eq +#align add_subgroup.map_comap_eq AddSubgroup.map_comap_eq + +@[to_additive] +theorem comap_map_eq (H : Subgroup G) : comap f (map f H) = H ⊔ f.ker := + by + refine' le_antisymm _ (sup_le (le_comap_map _ _) (ker_le_comap _ _)) + intro x hx; simp only [exists_prop, mem_map, mem_comap] at hx + rcases hx with ⟨y, hy, hy'⟩ + rw [← mul_inv_cancel_left y x] + exact mul_mem_sup hy (by simp [mem_ker, hy']) +#align subgroup.comap_map_eq Subgroup.comap_map_eq +#align add_subgroup.comap_map_eq AddSubgroup.comap_map_eq + +@[to_additive] +theorem map_comap_eq_self {f : G →* N} {H : Subgroup N} (h : H ≤ f.range) : map f (comap f H) = H := + by rwa [map_comap_eq, inf_eq_right] +#align subgroup.map_comap_eq_self Subgroup.map_comap_eq_self +#align add_subgroup.map_comap_eq_self AddSubgroup.map_comap_eq_self + +@[to_additive] +theorem map_comap_eq_self_of_surjective {f : G →* N} (h : Function.Surjective f) (H : Subgroup N) : + map f (comap f H) = H := + map_comap_eq_self ((range_top_of_surjective _ h).symm ▸ le_top) +#align subgroup.map_comap_eq_self_of_surjective Subgroup.map_comap_eq_self_of_surjective +#align add_subgroup.map_comap_eq_self_of_surjective AddSubgroup.map_comap_eq_self_of_surjective + +@[to_additive] +theorem comap_le_comap_of_le_range {f : G →* N} {K L : Subgroup N} (hf : K ≤ f.range) : + K.comap f ≤ L.comap f ↔ K ≤ L := + ⟨(map_comap_eq_self hf).ge.trans ∘ map_le_iff_le_comap.mpr, comap_mono⟩ +#align subgroup.comap_le_comap_of_le_range Subgroup.comap_le_comap_of_le_range +#align add_subgroup.comap_le_comap_of_le_range AddSubgroup.comap_le_comap_of_le_range + +@[to_additive] +theorem comap_le_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : + K.comap f ≤ L.comap f ↔ K ≤ L := + comap_le_comap_of_le_range (le_top.trans (f.range_top_of_surjective hf).ge) +#align subgroup.comap_le_comap_of_surjective Subgroup.comap_le_comap_of_surjective +#align add_subgroup.comap_le_comap_of_surjective AddSubgroup.comap_le_comap_of_surjective + +@[to_additive] +theorem comap_lt_comap_of_surjective {f : G →* N} {K L : Subgroup N} (hf : Function.Surjective f) : + K.comap f < L.comap f ↔ K < L := by simp_rw [lt_iff_le_not_le, comap_le_comap_of_surjective hf] +#align subgroup.comap_lt_comap_of_surjective Subgroup.comap_lt_comap_of_surjective +#align add_subgroup.comap_lt_comap_of_surjective AddSubgroup.comap_lt_comap_of_surjective + +@[to_additive] +theorem comap_injective {f : G →* N} (h : Function.Surjective f) : Function.Injective (comap f) := + fun K L => by simp only [le_antisymm_iff, comap_le_comap_of_surjective h, imp_self] +#align subgroup.comap_injective Subgroup.comap_injective +#align add_subgroup.comap_injective AddSubgroup.comap_injective + +@[to_additive] +theorem comap_map_eq_self {f : G →* N} {H : Subgroup G} (h : f.ker ≤ H) : comap f (map f H) = H := + by rwa [comap_map_eq, sup_eq_left] +#align subgroup.comap_map_eq_self Subgroup.comap_map_eq_self +#align add_subgroup.comap_map_eq_self AddSubgroup.comap_map_eq_self + +@[to_additive] +theorem comap_map_eq_self_of_injective {f : G →* N} (h : Function.Injective f) (H : Subgroup G) : + comap f (map f H) = H := + comap_map_eq_self (((ker_eq_bot_iff _).mpr h).symm ▸ bot_le) +#align subgroup.comap_map_eq_self_of_injective Subgroup.comap_map_eq_self_of_injective +#align add_subgroup.comap_map_eq_self_of_injective AddSubgroup.comap_map_eq_self_of_injective + +@[to_additive] +theorem map_le_map_iff {f : G →* N} {H K : Subgroup G} : H.map f ≤ K.map f ↔ H ≤ K ⊔ f.ker := by + rw [map_le_iff_le_comap, comap_map_eq] +#align subgroup.map_le_map_iff Subgroup.map_le_map_iff +#align add_subgroup.map_le_map_iff AddSubgroup.map_le_map_iff + +@[to_additive] +theorem map_le_map_iff' {f : G →* N} {H K : Subgroup G} : + H.map f ≤ K.map f ↔ H ⊔ f.ker ≤ K ⊔ f.ker := by + simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true_iff] +#align subgroup.map_le_map_iff' Subgroup.map_le_map_iff' +#align add_subgroup.map_le_map_iff' AddSubgroup.map_le_map_iff' + +@[to_additive] +theorem map_eq_map_iff {f : G →* N} {H K : Subgroup G} : + H.map f = K.map f ↔ H ⊔ f.ker = K ⊔ f.ker := by simp only [le_antisymm_iff, map_le_map_iff'] +#align subgroup.map_eq_map_iff Subgroup.map_eq_map_iff +#align add_subgroup.map_eq_map_iff AddSubgroup.map_eq_map_iff + +@[to_additive] +theorem map_eq_range_iff {f : G →* N} {H : Subgroup G} : H.map f = f.range ↔ Codisjoint H f.ker := + by rw [f.range_eq_map, map_eq_map_iff, codisjoint_iff, top_sup_eq] +#align subgroup.map_eq_range_iff Subgroup.map_eq_range_iff +#align add_subgroup.map_eq_range_iff AddSubgroup.map_eq_range_iff + +@[to_additive] +theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) {H K : Subgroup G} : + H.map f ≤ K.map f ↔ H ≤ K := by rw [map_le_iff_le_comap, comap_map_eq_self_of_injective hf] +#align subgroup.map_le_map_iff_of_injective Subgroup.map_le_map_iff_of_injective +#align add_subgroup.map_le_map_iff_of_injective AddSubgroup.map_le_map_iff_of_injective + +@[simp, to_additive] +theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : + H.map G'.Subtype ≤ K.map G'.Subtype ↔ H ≤ K := + map_le_map_iff_of_injective Subtype.coe_injective +#align subgroup.map_subtype_le_map_subtype Subgroup.map_subtype_le_map_subtype +#align add_subgroup.map_subtype_le_map_subtype AddSubgroup.map_subtype_le_map_subtype + +@[to_additive] +theorem map_injective {f : G →* N} (h : Function.Injective f) : Function.Injective (map f) := + Function.LeftInverse.injective <| comap_map_eq_self_of_injective h +#align subgroup.map_injective Subgroup.map_injective +#align add_subgroup.map_injective AddSubgroup.map_injective + +@[to_additive] +theorem map_eq_comap_of_inverse {f : G →* N} {g : N →* G} (hl : Function.LeftInverse g f) + (hr : Function.RightInverse g f) (H : Subgroup G) : map f H = comap g H := + SetLike.ext' <| by rw [coe_map, coe_comap, Set.image_eq_preimage_of_inverse hl hr] +#align subgroup.map_eq_comap_of_inverse Subgroup.map_eq_comap_of_inverse +#align add_subgroup.map_eq_comap_of_inverse AddSubgroup.map_eq_comap_of_inverse + +/-- Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`. -/ +@[to_additive "Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`."] +theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ker ≤ K) + (hf : map f H = map f K) : H = K := + by + apply_fun comap f at hf + rwa [comap_map_eq, comap_map_eq, sup_of_le_left hH, sup_of_le_left hK] at hf +#align subgroup.map_injective_of_ker_le Subgroup.map_injective_of_ker_le +#align add_subgroup.map_injective_of_ker_le AddSubgroup.map_injective_of_ker_le + +@[to_additive] +theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).Subtype ⁻¹' s) = ⊤ := + by + apply map_injective (closure s).subtype_injective + rwa [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, + Set.image_preimage_eq_of_subset] + rw [coeSubtype, Subtype.range_coe_subtype] + exact subset_closure +#align subgroup.closure_preimage_eq_top Subgroup.closure_preimage_eq_top +#align add_subgroup.closure_preimage_eq_top AddSubgroup.closure_preimage_eq_top + +@[to_additive] +theorem comap_sup_eq_of_le_range {H K : Subgroup N} (hH : H ≤ f.range) (hK : K ≤ f.range) : + comap f H ⊔ comap f K = comap f (H ⊔ K) := + map_injective_of_ker_le f ((ker_le_comap f H).trans le_sup_left) (ker_le_comap f (H ⊔ K)) + (by + rw [map_comap_eq, map_sup, map_comap_eq, map_comap_eq, inf_eq_right.mpr hH, + inf_eq_right.mpr hK, inf_eq_right.mpr (sup_le hH hK)]) +#align subgroup.comap_sup_eq_of_le_range Subgroup.comap_sup_eq_of_le_range +#align add_subgroup.comap_sup_eq_of_le_range AddSubgroup.comap_sup_eq_of_le_range + +@[to_additive] +theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : + comap f H ⊔ comap f K = comap f (H ⊔ K) := + comap_sup_eq_of_le_range f (le_top.trans (ge_of_eq (f.range_top_of_surjective hf))) + (le_top.trans (ge_of_eq (f.range_top_of_surjective hf))) +#align subgroup.comap_sup_eq Subgroup.comap_sup_eq +#align add_subgroup.comap_sup_eq AddSubgroup.comap_sup_eq + +@[to_additive] +theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : + H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := + comap_sup_eq_of_le_range L.Subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) +#align subgroup.sup_subgroup_of_eq Subgroup.sup_subgroupOf_eq +#align add_subgroup.sup_add_subgroup_of_eq AddSubgroup.sup_add_subgroupOf_eq + +@[to_additive] +theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : + Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := + by + rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self] + exacts[le_sup_left, le_sup_right] +#align subgroup.codisjoint_subgroup_of_sup Subgroup.codisjoint_subgroupOf_sup +#align add_subgroup.codisjoint_add_subgroup_of_sup AddSubgroup.codisjoint_add_subgroup_of_sup + +/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, +use `mul_equiv.subgroup_map` for better definitional equalities. -/ +@[to_additive + "An additive subgroup is isomorphic to its image under an injective function. If you\nhave an isomorphism, use `add_equiv.add_subgroup_map` for better definitional equalities."] +noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) : + H ≃* H.map f := + { Equiv.Set.image f H hf with map_mul' := fun _ _ => Subtype.ext (f.map_mul _ _) } +#align subgroup.equiv_map_of_injective Subgroup.equivMapOfInjective +#align add_subgroup.equiv_map_of_injective AddSubgroup.equivMapOfInjective + +@[simp, to_additive] +theorem coe_equivMapOfInjective_apply (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) + (h : H) : (equivMapOfInjective H f hf h : N) = f h := + rfl +#align subgroup.coe_equiv_map_of_injective_apply Subgroup.coe_equivMapOfInjective_apply +#align add_subgroup.coe_equiv_map_of_injective_apply AddSubgroup.coe_equivMapOfInjective_apply + +/-- The preimage of the normalizer is equal to the normalizer of the preimage of a surjective + function. -/ +@[to_additive + "The preimage of the normalizer is equal to the normalizer of the preimage of\na surjective function."] +theorem comap_normalizer_eq_of_surjective (H : Subgroup G) {f : N →* G} + (hf : Function.Surjective f) : H.normalizer.comap f = (H.comap f).normalizer := + le_antisymm (le_normalizer_comap f) + (by + intro x hx + simp only [mem_comap, mem_normalizer_iff] at * + intro n + rcases hf n with ⟨y, rfl⟩ + simp [hx y]) +#align subgroup.comap_normalizer_eq_of_surjective Subgroup.comap_normalizer_eq_of_surjective +#align add_subgroup.comap_normalizer_eq_of_surjective AddSubgroup.comap_normalizer_eq_of_surjective + +@[to_additive] +theorem comap_normalizer_eq_of_injective_of_le_range {N : Type _} [Group N] (H : Subgroup G) + {f : N →* G} (hf : Function.Injective f) (h : H.normalizer ≤ f.range) : + comap f H.normalizer = (comap f H).normalizer := + by + apply Subgroup.map_injective hf + rw [map_comap_eq_self h] + apply le_antisymm + · refine' le_trans (le_of_eq _) (map_mono (le_normalizer_comap _)) + rw [map_comap_eq_self h] + · refine' le_trans (le_normalizer_map f) (le_of_eq _) + rw [map_comap_eq_self (le_trans le_normalizer h)] +#align subgroup.comap_normalizer_eq_of_injective_of_le_range Subgroup.comap_normalizer_eq_of_injective_of_le_range +#align add_subgroup.comap_normalizer_eq_of_injective_of_le_range AddSubgroup.comap_normalizer_eq_of_injective_of_le_range + +@[to_additive] +theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : + H.normalizer.subgroupOf N = (H.subgroupOf N).normalizer := + by + apply comap_normalizer_eq_of_injective_of_le_range + exact Subtype.coe_injective + simpa +#align subgroup.subgroup_of_normalizer_eq Subgroup.subgroupOf_normalizer_eq +#align add_subgroup.add_subgroup_of_normalizer_eq AddSubgroup.add_subgroupOf_normalizer_eq + +/-- The image of the normalizer is equal to the normalizer of the image of an isomorphism. -/ +@[to_additive + "The image of the normalizer is equal to the normalizer of the image of an\nisomorphism."] +theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : + H.normalizer.map f.toMonoidHom = (H.map f.toMonoidHom).normalizer := + by + ext x + simp only [mem_normalizer_iff, mem_map_equiv] + rw [f.to_equiv.forall_congr] + simp +#align subgroup.map_equiv_normalizer_eq Subgroup.map_equiv_normalizer_eq +#align add_subgroup.map_equiv_normalizer_eq AddSubgroup.map_equiv_normalizer_eq + +/-- The image of the normalizer is equal to the normalizer of the image of a bijective + function. -/ +@[to_additive + "The image of the normalizer is equal to the normalizer of the image of a bijective\n function."] +theorem map_normalizer_eq_of_bijective (H : Subgroup G) {f : G →* N} (hf : Function.Bijective f) : + H.normalizer.map f = (H.map f).normalizer := + map_equiv_normalizer_eq H (MulEquiv.ofBijective f hf) +#align subgroup.map_normalizer_eq_of_bijective Subgroup.map_normalizer_eq_of_bijective +#align add_subgroup.map_normalizer_eq_of_bijective AddSubgroup.map_normalizer_eq_of_bijective + +end Subgroup + +namespace MonoidHom + +variable {G₁ G₂ G₃ : Type _} [Group G₁] [Group G₂] [Group G₃] + +variable (f : G₁ →* G₂) (f_inv : G₂ → G₁) + +/-- Auxiliary definition used to define `lift_of_right_inverse` -/ +@[to_additive "Auxiliary definition used to define `lift_of_right_inverse`"] +def liftOfRightInverseAux (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) : + G₂ →* G₃ where + toFun b := g (f_inv b) + map_one' := hg (hf 1) + map_mul' := by + intro x y + rw [← g.map_mul, ← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker] + apply hg + rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one, f.map_mul] + simp only [hf _] +#align monoid_hom.lift_of_right_inverse_aux MonoidHom.liftOfRightInverseAux +#align add_monoid_hom.lift_of_right_inverse_aux AddMonoidHom.liftOfRightInverseAux + +@[simp, to_additive] +theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) + (hg : f.ker ≤ g.ker) (x : G₁) : (f.liftOfRightInverseAux f_inv hf g hg) (f x) = g x := + by + dsimp [lift_of_right_inverse_aux] + rw [← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker] + apply hg + rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one] + simp only [hf _] +#align monoid_hom.lift_of_right_inverse_aux_comp_apply MonoidHom.liftOfRightInverseAux_comp_apply +#align add_monoid_hom.lift_of_right_inverse_aux_comp_apply AddMonoidHom.liftOfRightInverseAux_comp_apply + +/-- `lift_of_right_inverse f hf g hg` is the unique group homomorphism `φ` + +* such that `φ.comp f = g` (`monoid_hom.lift_of_right_inverse_comp`), +* where `f : G₁ →+* G₂` has a right_inverse `f_inv` (`hf`), +* and `g : G₂ →+* G₃` satisfies `hg : f.ker ≤ g.ker`. + +See `monoid_hom.eq_lift_of_right_inverse` for the uniqueness lemma. + +``` + G₁. + | \ + f | \ g + | \ + v \⌟ + G₂----> G₃ + ∃!φ +``` + -/ +@[to_additive + "`lift_of_right_inverse f f_inv hf g hg` is the unique additive group homomorphism `φ`\n\n* such that `φ.comp f = g` (`add_monoid_hom.lift_of_right_inverse_comp`),\n* where `f : G₁ →+ G₂` has a right_inverse `f_inv` (`hf`),\n* and `g : G₂ →+ G₃` satisfies `hg : f.ker ≤ g.ker`.\n\nSee `add_monoid_hom.eq_lift_of_right_inverse` for the uniqueness lemma.\n\n```\n G₁.\n | \\\n f | \\ g\n | \\\n v \\⌟\n G₂----> G₃\n ∃!φ\n```"] +def liftOfRightInverse (hf : Function.RightInverse f_inv f) : + { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) + where + toFun g := f.liftOfRightInverseAux f_inv hf g.1 g.2 + invFun φ := ⟨φ.comp f, fun x hx => (mem_ker _).mpr <| by simp [(mem_ker _).mp hx]⟩ + left_inv g := by + ext + simp only [comp_apply, lift_of_right_inverse_aux_comp_apply, Subtype.coe_mk, Subtype.val_eq_coe] + right_inv φ := by + ext b + simp [lift_of_right_inverse_aux, hf b] +#align monoid_hom.lift_of_right_inverse MonoidHom.liftOfRightInverse +#align add_monoid_hom.lift_of_right_inverse AddMonoidHom.liftOfRightInverse + +/-- A non-computable version of `monoid_hom.lift_of_right_inverse` for when no computable right +inverse is available, that uses `function.surj_inv`. -/ +@[simp, + to_additive + "A non-computable version of `add_monoid_hom.lift_of_right_inverse` for when no\ncomputable right inverse is available."] +noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : + { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) := + f.liftOfRightInverse (Function.surjInv hf) (Function.rightInverse_surjInv hf) +#align monoid_hom.lift_of_surjective MonoidHom.liftOfSurjective +#align add_monoid_hom.lift_of_surjective AddMonoidHom.liftOfSurjective + +@[simp, to_additive] +theorem liftOfRightInverse_comp_apply (hf : Function.RightInverse f_inv f) + (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) (x : G₁) : + (f.liftOfRightInverse f_inv hf g) (f x) = g x := + f.lift_of_right_inverse_aux_comp_apply f_inv hf g.1 g.2 x +#align monoid_hom.lift_of_right_inverse_comp_apply MonoidHom.liftOfRightInverse_comp_apply +#align add_monoid_hom.lift_of_right_inverse_comp_apply AddMonoidHom.lift_of_right_inverse_comp_apply + +@[simp, to_additive] +theorem liftOfRightInverse_comp (hf : Function.RightInverse f_inv f) + (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) : (f.liftOfRightInverse f_inv hf g).comp f = g := + MonoidHom.ext <| f.lift_of_right_inverse_comp_apply f_inv hf g +#align monoid_hom.lift_of_right_inverse_comp MonoidHom.liftOfRightInverse_comp +#align add_monoid_hom.lift_of_right_inverse_comp AddMonoidHom.lift_of_right_inverse_comp + +@[to_additive] +theorem eq_liftOfRightInverse (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) + (hg : f.ker ≤ g.ker) (h : G₂ →* G₃) (hh : h.comp f = g) : + h = f.liftOfRightInverse f_inv hf ⟨g, hg⟩ := + by + simp_rw [← hh] + exact ((f.lift_of_right_inverse f_inv hf).apply_symm_apply _).symm +#align monoid_hom.eq_lift_of_right_inverse MonoidHom.eq_liftOfRightInverse +#align add_monoid_hom.eq_lift_of_right_inverse AddMonoidHom.eq_lift_of_right_inverse + +end MonoidHom + +variable {N : Type _} [Group N] + +-- Here `H.normal` is an explicit argument so we can use dot notation with `comap`. +@[to_additive] +theorem Subgroup.Normal.comap {H : Subgroup N} (hH : H.Normal) (f : G →* N) : (H.comap f).Normal := + ⟨fun _ => by simp (config := { contextual := true }) [Subgroup.mem_comap, hH.conj_mem]⟩ +#align subgroup.normal.comap Subgroup.Normal.comap +#align add_subgroup.normal.comap AddSubgroup.Normal.comap + +@[to_additive] +instance (priority := 100) Subgroup.normal_comap {H : Subgroup N} [nH : H.Normal] (f : G →* N) : + (H.comap f).Normal := + nH.comap _ +#align subgroup.normal_comap Subgroup.normal_comap +#align add_subgroup.normal_comap AddSubgroup.normal_comap + +-- Here `H.normal` is an explicit argument so we can use dot notation with `subgroup_of`. +@[to_additive] +theorem Subgroup.Normal.subgroupOf {H : Subgroup G} (hH : H.Normal) (K : Subgroup G) : + (H.subgroupOf K).Normal := + hH.comap _ +#align subgroup.normal.subgroup_of Subgroup.Normal.subgroupOf +#align add_subgroup.normal.add_subgroup_of AddSubgroup.Normal.add_subgroupOf + +@[to_additive] +instance (priority := 100) Subgroup.normal_subgroupOf {H N : Subgroup G} [N.Normal] : + (N.subgroupOf H).Normal := + Subgroup.normal_comap _ +#align subgroup.normal_subgroup_of Subgroup.normal_subgroupOf +#align add_subgroup.normal_add_subgroup_of AddSubgroup.normal_add_subgroupOf + +namespace MonoidHom + +/-- The `monoid_hom` from the preimage of a subgroup to itself. -/ +@[to_additive "the `add_monoid_hom` from the preimage of an additive subgroup to itself.", simps] +def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := + f.submonoidComap H'.toSubmonoid +#align monoid_hom.subgroup_comap MonoidHom.subgroupComap +#align add_monoid_hom.add_subgroup_comap AddMonoidHom.addSubgroupComap + +/-- The `monoid_hom` from a subgroup to its image. -/ +@[to_additive "the `add_monoid_hom` from an additive subgroup to its image", simps] +def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := + f.submonoidMap H.toSubmonoid +#align monoid_hom.subgroup_map MonoidHom.subgroupMap +#align add_monoid_hom.add_subgroup_map AddMonoidHom.addSubgroupMap + +@[to_additive] +theorem subgroupMap_surjective (f : G →* G') (H : Subgroup G) : + Function.Surjective (f.subgroupMap H) := + f.submonoid_map_surjective H.toSubmonoid +#align monoid_hom.subgroup_map_surjective MonoidHom.subgroupMap_surjective +#align add_monoid_hom.add_subgroup_map_surjective AddMonoidHom.add_subgroupMap_surjective + +end MonoidHom + +namespace MulEquiv + +variable {H K : Subgroup G} + +/-- Makes the identity isomorphism from a proof two subgroups of a multiplicative + group are equal. -/ +@[to_additive + "Makes the identity additive isomorphism from a proof\ntwo subgroups of an additive group are equal."] +def subgroupCongr (h : H = K) : H ≃* K := + { Equiv.setCongr <| congr_arg _ h with map_mul' := fun _ _ => rfl } +#align mul_equiv.subgroup_congr MulEquiv.subgroupCongr +#align add_equiv.add_subgroup_congr AddEquiv.addSubgroupCongr + +/-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, +use `subgroup.equiv_map_of_injective`. -/ +@[to_additive + "An additive subgroup is isomorphic to its image under an an isomorphism. If you only\nhave an injective map, use `add_subgroup.equiv_map_of_injective`."] +def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') := + MulEquiv.submonoidMap (e : G ≃* G') H.toSubmonoid +#align mul_equiv.subgroup_map MulEquiv.subgroupMap +#align add_equiv.add_subgroup_map AddEquiv.addSubgroupMap + +@[simp, to_additive] +theorem coe_subgroupMap_apply (e : G ≃* G') (H : Subgroup G) (g : H) : + ((subgroupMap e H g : H.map (e : G →* G')) : G') = e g := + rfl +#align mul_equiv.coe_subgroup_map_apply MulEquiv.coe_subgroupMap_apply +#align add_equiv.coe_add_subgroup_map_apply AddEquiv.coe_add_subgroupMap_apply + +@[simp, to_additive] +theorem subgroupMap_symm_apply (e : G ≃* G') (H : Subgroup G) (g : H.map (e : G →* G')) : + (e.subgroupMap H).symm g = ⟨e.symm g, SetLike.mem_coe.1 <| Set.mem_image_equiv.1 g.2⟩ := + rfl +#align mul_equiv.subgroup_map_symm_apply MulEquiv.subgroupMap_symm_apply +#align add_equiv.add_subgroup_map_symm_apply AddEquiv.add_subgroup_map_symm_apply + +end MulEquiv + +namespace Subgroup + +@[simp, to_additive] +theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : + H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := + by + ext + rfl +#align subgroup.equiv_map_of_injective_coe_mul_equiv Subgroup.equivMapOfInjective_coe_mulEquiv +#align add_subgroup.equiv_map_of_injective_coe_add_equiv AddSubgroup.equiv_map_of_injective_coe_add_equiv + +variable {C : Type _} [CommGroup C] {s t : Subgroup C} {x : C} + +@[to_additive] +theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := + ⟨fun h => by + rw [← closure_eq s, ← closure_eq t, ← closure_union] at h + apply closure_induction h + · rintro y (h | h) + · exact ⟨y, h, 1, t.one_mem, by simp⟩ + · exact ⟨1, s.one_mem, y, h, by simp⟩ + · exact ⟨1, s.one_mem, 1, ⟨t.one_mem, mul_one 1⟩⟩ + · rintro _ _ ⟨y₁, hy₁, z₁, hz₁, rfl⟩ ⟨y₂, hy₂, z₂, hz₂, rfl⟩ + exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc] <;> cc⟩ + · rintro _ ⟨y, hy, z, hz, rfl⟩ + exact ⟨_, inv_mem hy, _, inv_mem hz, mul_comm z y ▸ (mul_inv_rev z y).symm⟩, by + rintro ⟨y, hy, z, hz, rfl⟩ <;> exact mul_mem_sup hy hz⟩ +#align subgroup.mem_sup Subgroup.mem_sup +#align add_subgroup.mem_sup AddSubgroup.mem_sup + +@[to_additive] +theorem mem_sup' : x ∈ s ⊔ t ↔ ∃ (y : s)(z : t), (y : C) * z = x := + mem_sup.trans <| by simp only [SetLike.exists, coe_mk] +#align subgroup.mem_sup' Subgroup.mem_sup' +#align add_subgroup.mem_sup' AddSubgroup.mem_sup' + +@[to_additive] +theorem mem_closure_pair {x y z : C} : + z ∈ closure ({x, y} : Set C) ↔ ∃ m n : ℤ, x ^ m * y ^ n = z := + by + rw [← Set.singleton_union, Subgroup.closure_union, mem_sup] + simp_rw [exists_prop, mem_closure_singleton, exists_exists_eq_and] +#align subgroup.mem_closure_pair Subgroup.mem_closure_pair +#align add_subgroup.mem_closure_pair AddSubgroup.mem_closure_pair + +@[to_additive] +instance : IsModularLattice (Subgroup C) := + ⟨fun x y z xz a ha => by + rw [mem_inf, mem_sup] at ha + rcases ha with ⟨⟨b, hb, c, hc, rfl⟩, haz⟩ + rw [mem_sup] + exact ⟨b, hb, c, mem_inf.2 ⟨hc, (mul_mem_cancel_left (xz hb)).1 haz⟩, rfl⟩⟩ + +end Subgroup + +namespace Subgroup + +section SubgroupNormal + +@[to_additive] +theorem normal_subgroupOf_iff {H K : Subgroup G} (hHK : H ≤ K) : + (H.subgroupOf K).Normal ↔ ∀ h k, h ∈ H → k ∈ K → k * h * k⁻¹ ∈ H := + ⟨fun hN h k hH hK => hN.conj_mem ⟨h, hHK hH⟩ hH ⟨k, hK⟩, fun hN => + { conj_mem := fun h hm k => hN h.1 k.1 hm k.2 }⟩ +#align subgroup.normal_subgroup_of_iff Subgroup.normal_subgroupOf_iff +#align add_subgroup.normal_add_subgroup_of_iff AddSubgroup.normal_add_subgroupOf_iff + +@[to_additive] +instance prod_subgroupOf_prod_normal {H₁ K₁ : Subgroup G} {H₂ K₂ : Subgroup N} + [h₁ : (H₁.subgroupOf K₁).Normal] [h₂ : (H₂.subgroupOf K₂).Normal] : + ((H₁.Prod H₂).subgroupOf (K₁.Prod K₂)).Normal + where conj_mem n hgHK g := + ⟨h₁.conj_mem ⟨(n : G × N).fst, (mem_prod.mp n.2).1⟩ hgHK.1 + ⟨(g : G × N).fst, (mem_prod.mp g.2).1⟩, + h₂.conj_mem ⟨(n : G × N).snd, (mem_prod.mp n.2).2⟩ hgHK.2 + ⟨(g : G × N).snd, (mem_prod.mp g.2).2⟩⟩ +#align subgroup.prod_subgroup_of_prod_normal Subgroup.prod_subgroupOf_prod_normal +#align add_subgroup.sum_add_subgroup_of_sum_normal AddSubgroup.sum_add_subgroupOf_sum_normal + +@[to_additive] +instance prod_normal (H : Subgroup G) (K : Subgroup N) [hH : H.Normal] [hK : K.Normal] : + (H.Prod K).Normal + where conj_mem n hg g := + ⟨hH.conj_mem n.fst (Subgroup.mem_prod.mp hg).1 g.fst, + hK.conj_mem n.snd (Subgroup.mem_prod.mp hg).2 g.snd⟩ +#align subgroup.prod_normal Subgroup.prod_normal +#align add_subgroup.sum_normal AddSubgroup.sum_normal + +@[to_additive] +theorem inf_subgroupOf_inf_normal_of_right (A B' B : Subgroup G) (hB : B' ≤ B) + [hN : (B'.subgroupOf B).Normal] : ((A ⊓ B').subgroupOf (A ⊓ B)).Normal := + { + conj_mem := fun n hn g => + ⟨mul_mem (mul_mem (mem_inf.1 g.2).1 (mem_inf.1 n.2).1) (inv_mem (mem_inf.1 g.2).1), + (normal_subgroupOf_iff hB).mp hN n g hn.2 (mem_inf.mp g.2).2⟩ } +#align subgroup.inf_subgroup_of_inf_normal_of_right Subgroup.inf_subgroupOf_inf_normal_of_right +#align add_subgroup.inf_add_subgroup_of_inf_normal_of_right AddSubgroup.inf_add_subgroupOf_inf_normal_of_right + +@[to_additive] +theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) (hA : A' ≤ A) + [hN : (A'.subgroupOf A).Normal] : ((A' ⊓ B).subgroupOf (A ⊓ B)).Normal := + { + conj_mem := fun n hn g => + ⟨(normal_subgroupOf_iff hA).mp hN n g hn.1 (mem_inf.mp g.2).1, + mul_mem (mul_mem (mem_inf.1 g.2).2 (mem_inf.1 n.2).2) (inv_mem (mem_inf.1 g.2).2)⟩ } +#align subgroup.inf_subgroup_of_inf_normal_of_left Subgroup.inf_subgroupOf_inf_normal_of_left +#align add_subgroup.inf_add_subgroup_of_inf_normal_of_left AddSubgroup.inf_add_subgroupOf_inf_normal_of_left + +@[to_additive] +instance normal_inf_normal (H K : Subgroup G) [hH : H.Normal] [hK : K.Normal] : (H ⊓ K).Normal := + ⟨fun n hmem g => ⟨hH.conj_mem n hmem.1 g, hK.conj_mem n hmem.2 g⟩⟩ +#align subgroup.normal_inf_normal Subgroup.normal_inf_normal +#align add_subgroup.normal_inf_normal AddSubgroup.normal_inf_normal + +@[to_additive] +theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : + (A ⊔ A').subgroupOf B = A.subgroupOf B ⊔ A'.subgroupOf B := + by + refine' + map_injective_of_ker_le B.subtype (ker_le_comap _ _) + (le_trans (ker_le_comap B.subtype _) le_sup_left) _ + · simp only [subgroup_of, map_comap_eq, map_sup, subtype_range] + rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] +#align subgroup.subgroup_of_sup Subgroup.subgroupOf_sup +#align add_subgroup.add_subgroup_of_sup AddSubgroup.add_subgroupOf_sup + +@[to_additive] +theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgroupOf K).Normal] + {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := + by + have := (normal_subgroup_of_iff hK).mp hN (a * b) b h hb + rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this +#align subgroup.subgroup_normal.mem_comm Subgroup.SubgroupNormal.mem_comm +#align add_subgroup.subgroup_normal.mem_comm AddSubgroup.SubgroupNormal.mem_comm + +/-- Elements of disjoint, normal subgroups commute. -/ +@[to_additive "Elements of disjoint, normal subgroups commute."] +theorem 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 := + by + suffices x * y * x⁻¹ * y⁻¹ = 1 by + show x * y = y * x + · rw [mul_assoc, mul_eq_one_iff_eq_inv] at this + simpa + apply hdis.le_bot + constructor + · suffices x * (y * x⁻¹ * y⁻¹) ∈ H₁ by simpa [mul_assoc] + exact H₁.mul_mem hx (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 +#align subgroup.commute_of_normal_of_disjoint Subgroup.commute_of_normal_of_disjoint +#align add_subgroup.commute_of_normal_of_disjoint AddSubgroup.commute_of_normal_of_disjoint + +end SubgroupNormal + +@[to_additive] +theorem disjoint_def {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x : G}, x ∈ H₁ → x ∈ H₂ → x = 1 := + disjoint_iff_inf_le.trans <| by simp only [Disjoint, SetLike.le_def, mem_inf, mem_bot, and_imp] +#align subgroup.disjoint_def Subgroup.disjoint_def +#align add_subgroup.disjoint_def AddSubgroup.disjoint_def + +@[to_additive] +theorem disjoint_def' {H₁ H₂ : Subgroup G} : + Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := + disjoint_def.trans ⟨fun h x y hx hy hxy => h hx <| hxy.symm ▸ hy, fun h x hx hx' => h hx hx' rfl⟩ +#align subgroup.disjoint_def' Subgroup.disjoint_def' +#align add_subgroup.disjoint_def' AddSubgroup.disjoint_def' + +@[to_additive] +theorem 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 := + disjoint_def'.trans + ⟨fun h x y hx hy hxy => + let hx1 : x = 1 := h hx (H₂.inv_mem hy) (eq_inv_iff_mul_eq_one.mpr hxy) + ⟨hx1, by simpa [hx1] using hxy⟩, + fun h x y hx hy hxy => (h hx (H₂.inv_mem hy) (mul_inv_eq_one.mpr hxy)).1⟩ +#align subgroup.disjoint_iff_mul_eq_one Subgroup.disjoint_iff_mul_eq_one +#align add_subgroup.disjoint_iff_add_eq_zero AddSubgroup.disjoint_iff_add_eq_zero + +@[to_additive] +theorem mul_injective_of_disjoint {H₁ H₂ : Subgroup G} (h : Disjoint H₁ H₂) : + Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := + by + intro x y hxy + rw [← inv_mul_eq_iff_eq_mul, ← mul_assoc, ← mul_inv_eq_one, mul_assoc] at hxy + replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).Prop (x.2 * y.2⁻¹).Prop hxy + rwa [coe_mul, coe_mul, coe_inv, coe_inv, inv_mul_eq_one, mul_inv_eq_one, ← Subtype.ext_iff, ← + Subtype.ext_iff, eq_comm, ← Prod.ext_iff] at hxy +#align subgroup.mul_injective_of_disjoint Subgroup.mul_injective_of_disjoint +#align add_subgroup.add_injective_of_disjoint AddSubgroup.add_injective_of_disjoint + +end Subgroup + +namespace IsConj + +open Subgroup + +theorem normalClosure_eq_top_of {N : Subgroup G} [hn : N.Normal] {g g' : G} {hg : g ∈ N} + {hg' : g' ∈ N} (hc : IsConj g g') (ht : normalClosure ({⟨g, hg⟩} : Set N) = ⊤) : + normalClosure ({⟨g', hg'⟩} : Set N) = ⊤ := + by + obtain ⟨c, rfl⟩ := isConj_iff.1 hc + have h : ∀ x : N, (MulAut.conj c) x ∈ N := + by + rintro ⟨x, hx⟩ + exact hn.conj_mem _ hx c + have hs : Function.Surjective (((MulAut.conj c).toMonoidHom.restrict N).codRestrict _ h) := + by + rintro ⟨x, hx⟩ + refine' ⟨⟨c⁻¹ * x * c, _⟩, _⟩ + · have h := hn.conj_mem _ hx c⁻¹ + rwa [inv_inv] at h + simp only [MonoidHom.codRestrict_apply, MulEquiv.coe_toMonoidHom, MulAut.conj_apply, coe_mk, + MonoidHom.restrict_apply, Subtype.mk_eq_mk, ← mul_assoc, mul_inv_self, one_mul] + rw [mul_assoc, mul_inv_self, mul_one] + have ht' := map_mono (eq_top_iff.1 ht) + rw [← MonoidHom.range_eq_map, MonoidHom.range_top_of_surjective _ hs] at ht' + refine' eq_top_iff.2 (le_trans ht' (map_le_iff_le_comap.2 (normal_closure_le_normal _))) + rw [Set.singleton_subset_iff, SetLike.mem_coe] + simp only [MonoidHom.codRestrict_apply, MulEquiv.coe_toMonoidHom, MulAut.conj_apply, coe_mk, + MonoidHom.restrict_apply, mem_comap] + exact subset_normal_closure (Set.mem_singleton _) +#align is_conj.normal_closure_eq_top_of IsConj.normalClosure_eq_top_of + +end IsConj + +assert_not_exists Multiset + From 389e6a7ff1e505359f8c2eb3a5bda7e5b548464a Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 23 Jan 2023 21:16:38 +0100 Subject: [PATCH 03/34] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/GroupTheory/Subgroup/Basic.lean | 169 +++++++++--------------- 2 files changed, 60 insertions(+), 110 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 34842f49e3f45..1fd4341953385 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -536,6 +536,7 @@ import Mathlib.GroupTheory.GroupAction.Units import Mathlib.GroupTheory.Perm.Basic import Mathlib.GroupTheory.Perm.Support import Mathlib.GroupTheory.Perm.ViaEmbedding +import Mathlib.GroupTheory.Subgroup.Basic import Mathlib.GroupTheory.Submonoid.Basic import Mathlib.GroupTheory.Submonoid.Center import Mathlib.GroupTheory.Submonoid.Centralizer diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index b8816a5d6a3ec..487b3dd84fe72 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -8,14 +8,14 @@ Authors: Kexing Ying ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Algebra.Group.Conj -import Mathbin.Algebra.Module.Basic -import Mathbin.Algebra.Order.Group.InjSurj -import Mathbin.Data.Countable.Basic -import Mathbin.GroupTheory.Submonoid.Centralizer -import Mathbin.Logic.Encodable.Basic -import Mathbin.Order.Atoms -import Mathbin.Tactic.ApplyFun +import Mathlib.Algebra.Group.Conj +import Mathlib.Algebra.Module.Basic +import Mathlib.Algebra.Order.Group.InjSurj +import Mathlib.Data.Countable.Basic +import Mathlib.GroupTheory.Submonoid.Centralizer +import Mathlib.Logic.Encodable.Basic +import Mathlib.Order.Atoms +import Mathlib.Tactic.ApplyFun /-! # Subgroups @@ -318,8 +318,7 @@ def inclusion {H K : S} (h : H ≤ K) : H →* K := #align add_subgroup_class.inclusion AddSubgroupClass.inclusion @[simp, to_additive] -theorem inclusion_self (x : H) : inclusion le_rfl x = x := - by +theorem inclusion_self (x : H) : inclusion le_rfl x = x := by cases x rfl #align subgroup_class.inclusion_self SubgroupClass.inclusion_self @@ -332,8 +331,7 @@ theorem inclusion_mk {h : H ≤ K} (x : G) (hx : x ∈ H) : inclusion h ⟨x, hx #align add_subgroup_class.inclusion_mk AddSubgroupClass.inclusion_mk @[to_additive] -theorem inclusion_right (h : H ≤ K) (x : K) (hx : (x : G) ∈ H) : inclusion h ⟨x, hx⟩ = x := - by +theorem inclusion_right (h : H ≤ K) (x : K) (hx : (x : G) ∈ H) : inclusion h ⟨x, hx⟩ = x := by cases x rfl #align subgroup_class.inclusion_right SubgroupClass.inclusion_right @@ -341,15 +339,13 @@ theorem inclusion_right (h : H ≤ K) (x : K) (hx : (x : G) ∈ H) : inclusion h @[simp] theorem inclusion_inclusion {L : S} (hHK : H ≤ K) (hKL : K ≤ L) (x : H) : - inclusion hKL (inclusion hHK x) = inclusion (hHK.trans hKL) x := - by + inclusion hKL (inclusion hHK x) = inclusion (hHK.trans hKL) x := by cases x rfl #align subgroup_class.inclusion_inclusion SubgroupClass.inclusion_inclusion @[simp, to_additive] -theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := - by +theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := by cases a simp only [inclusion, [anonymous], MonoidHom.mk'_apply] #align subgroup_class.coe_inclusion SubgroupClass.coe_inclusion @@ -357,8 +353,7 @@ theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a @[simp, to_additive] theorem subtype_comp_inclusion {H K : S} (hH : H ≤ K) : - (subtype K).comp (inclusion hH) = subtype H := - by + (subtype K).comp (inclusion hH) = subtype H := by ext simp only [MonoidHom.comp_apply, coeSubtype, coe_inclusion] #align subgroup_class.subtype_comp_inclusion SubgroupClass.subtype_comp_inclusion @@ -816,8 +811,7 @@ def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := #align add_subgroup.inclusion AddSubgroup.inclusion @[simp, to_additive] -theorem coe_inclusion {H K : Subgroup G} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := - by +theorem coe_inclusion {H K : Subgroup G} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := by cases a simp only [inclusion, coe_mk, MonoidHom.mk'_apply] #align subgroup.coe_inclusion Subgroup.coe_inclusion @@ -908,8 +902,7 @@ theorem eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := #align add_subgroup.eq_bot_iff_forall AddSubgroup.eq_bot_iff_forall @[to_additive] -theorem eq_bot_of_subsingleton [Subsingleton H] : H = ⊥ := - by +theorem eq_bot_of_subsingleton [Subsingleton H] : H = ⊥ := by rw [Subgroup.eq_bot_iff_forall] intro y hy rw [← Subgroup.coe_mk H y hy, Subsingleton.elim (⟨y, hy⟩ : H) 1, Subgroup.coe_one] @@ -954,8 +947,7 @@ theorem bot_or_nontrivial (H : Subgroup G) : H = ⊥ ∨ Nontrivial H := by /-- A subgroup is either the trivial subgroup or contains a non-identity element. -/ @[to_additive "A subgroup is either the trivial subgroup or contains a nonzero element."] -theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ (1 : G) := - by +theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ (1 : G) := by convert H.bot_or_nontrivial rw [nontrivial_iff_exists_ne_one] #align subgroup.bot_or_exists_ne_one Subgroup.bot_or_exists_ne_one @@ -1151,8 +1143,7 @@ theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x theorem closure_induction' {p : ∀ x, x ∈ closure k → Prop} (Hs : ∀ (x) (h : x ∈ k), p x (subset_closure h)) (H1 : p 1 (one_mem _)) (Hmul : ∀ x hx y hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) - (Hinv : ∀ x hx, p x hx → p x⁻¹ (inv_mem hx)) {x} (hx : x ∈ closure k) : p x hx := - by + (Hinv : ∀ x hx, p x hx → p x⁻¹ (inv_mem hx)) {x} (hx : x ∈ closure k) : p x hx := by refine' Exists.elim _ fun (hx : x ∈ closure k) (hc : p x hx) => hc exact closure_induction hx (fun x hx => ⟨_, Hs x hx⟩) ⟨_, H1⟩ @@ -1265,8 +1256,7 @@ theorem closure_unionᵢ {ι} (s : ι → Set G) : closure (⋃ i, s i) = ⨆ i, #align add_subgroup.closure_Union AddSubgroup.closure_unionᵢ @[to_additive] -theorem closure_eq_bot_iff (G : Type _) [Group G] (S : Set G) : closure S = ⊥ ↔ S ⊆ {1} := - by +theorem closure_eq_bot_iff (G : Type _) [Group G] (S : Set G) : closure S = ⊥ ↔ S ⊆ {1} := by rw [← le_bot_iff] exact closure_le _ #align subgroup.closure_eq_bot_iff Subgroup.closure_eq_bot_iff @@ -1282,8 +1272,7 @@ theorem supᵢ_eq_closure {ι : Sort _} (p : ι → Subgroup G) : the element. -/ @[to_additive "The `add_subgroup` generated by an element of an `add_group` equals the set of\nnatural number multiples of the element."] -theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := - by +theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := by refine' ⟨fun hy => closure_induction hy _ _ _ _, fun ⟨n, hn⟩ => hn ▸ zpow_mem (subset_closure <| mem_singleton x) n⟩ @@ -1319,8 +1308,7 @@ theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = @[to_additive] theorem mem_supᵢ_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} (hK : Directed (· ≤ ·) K) - {x : G} : x ∈ (supᵢ K : Subgroup G) ↔ ∃ i, x ∈ K i := - by + {x : G} : x ∈ (supᵢ K : Subgroup G) ↔ ∃ i, x ∈ K i := by refine' ⟨_, fun ⟨i, hi⟩ => (SetLike.le_def.1 <| le_supᵢ K i) hi⟩ suffices x ∈ closure (⋃ i, (K i : Set G)) → ∃ i, x ∈ K i by simpa only [closure_Union, closure_eq (K _)] using this @@ -1343,8 +1331,7 @@ theorem coe_supᵢ_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : @[to_additive] theorem mem_supₛ_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) - {x : G} : x ∈ supₛ K ↔ ∃ s ∈ K, x ∈ s := - by + {x : G} : x ∈ supₛ K ↔ ∃ s ∈ K, x ∈ s := by haveI : Nonempty K := Kne.to_subtype simp only [supₛ_eq_supᵢ', mem_supr_of_directed hK.directed_coe, SetCoe.exists, Subtype.coe_mk] #align subgroup.mem_Sup_of_directed_on Subgroup.mem_supₛ_of_directedOn @@ -1388,8 +1375,7 @@ theorem comap_comap (K : Subgroup P) (g : N →* P) (f : G →* N) : #align add_subgroup.comap_comap AddSubgroup.comap_comap @[simp, to_additive] -theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := - by +theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by ext rfl #align subgroup.comap_id Subgroup.comap_id @@ -1559,8 +1545,7 @@ theorem map_inf_le (H K : Subgroup G) (f : G →* N) : map f (H ⊓ K) ≤ map f @[to_additive] theorem map_inf_eq (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) : - map f (H ⊓ K) = map f H ⊓ map f K := - by + map f (H ⊓ K) = map f H ⊓ map f K := by rw [← SetLike.coe_set_eq] simp [Set.image_inter hf] #align subgroup.map_inf_eq Subgroup.map_inf_eq @@ -1573,8 +1558,7 @@ theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := #align add_subgroup.map_bot AddSubgroup.map_bot @[simp, to_additive] -theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := - by +theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := by rw [eq_top_iff] intro x hx obtain ⟨y, hy⟩ := h x @@ -1860,8 +1844,7 @@ theorem pi_bot : (pi Set.univ fun i => (⊥ : Subgroup (f i))) = ⊥ := @[to_additive] theorem le_pi_iff {I : Set η} {H : ∀ i, Subgroup (f i)} {J : Subgroup (∀ i, f i)} : - J ≤ pi I H ↔ ∀ i : η, i ∈ I → map (Pi.evalMonoidHom f i) J ≤ H i := - by + J ≤ pi I H ↔ ∀ i : η, i ∈ I → map (Pi.evalMonoidHom f i) J ≤ H i := by constructor · intro h i hi rintro _ ⟨x, hx, rfl⟩ @@ -1873,8 +1856,7 @@ theorem le_pi_iff {I : Set η} {H : ∀ i, Subgroup (f i)} {J : Subgroup (∀ i, @[simp, to_additive] theorem mulSingle_mem_pi [DecidableEq η] {I : Set η} {H : ∀ i, Subgroup (f i)} (i : η) (x : f i) : - Pi.mulSingle i x ∈ pi I H ↔ i ∈ I → x ∈ H i := - by + Pi.mulSingle i x ∈ pi I H ↔ i ∈ I → x ∈ H i := by constructor · intro h hi simpa using h i hi @@ -1939,8 +1921,7 @@ namespace Normal variable (nH : H.Normal) @[to_additive] -theorem mem_comm {a b : G} (h : a * b ∈ H) : b * a ∈ H := - by +theorem mem_comm {a b : G} (h : a * b ∈ H) : b * a ∈ H := by have : a⁻¹ * (a * b) * a⁻¹⁻¹ ∈ H := nH.conj_mem (a * b) h a⁻¹ simpa #align subgroup.normal.mem_comm Subgroup.Normal.mem_comm @@ -2017,24 +1998,21 @@ theorem characteristic_iff_le_comap : H.Characteristic ↔ ∀ ϕ : G ≃* G, H #align add_subgroup.characteristic_iff_le_comap AddSubgroup.characteristic_iff_le_comap @[to_additive] -theorem characteristic_iff_map_eq : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom = H := - by +theorem characteristic_iff_map_eq : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom = H := by simp_rw [map_equiv_eq_comap_symm] exact characteristic_iff_comap_eq.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ #align subgroup.characteristic_iff_map_eq Subgroup.characteristic_iff_map_eq #align add_subgroup.characteristic_iff_map_eq AddSubgroup.characteristic_iff_map_eq @[to_additive] -theorem characteristic_iff_map_le : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom ≤ H := - by +theorem characteristic_iff_map_le : H.Characteristic ↔ ∀ ϕ : G ≃* G, H.map ϕ.toMonoidHom ≤ H := by simp_rw [map_equiv_eq_comap_symm] exact characteristic_iff_comap_le.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ #align subgroup.characteristic_iff_map_le Subgroup.characteristic_iff_map_le #align add_subgroup.characteristic_iff_map_le AddSubgroup.characteristic_iff_map_le @[to_additive] -theorem characteristic_iff_le_map : H.Characteristic ↔ ∀ ϕ : G ≃* G, H ≤ H.map ϕ.toMonoidHom := - by +theorem characteristic_iff_le_map : H.Characteristic ↔ ∀ ϕ : G ≃* G, H ≤ H.map ϕ.toMonoidHom := by simp_rw [map_equiv_eq_comap_symm] exact characteristic_iff_le_comap.trans ⟨fun h ϕ => h ϕ.symm, fun h ϕ => h ϕ.symm⟩ #align subgroup.characteristic_iff_le_map Subgroup.characteristic_iff_le_map @@ -2089,16 +2067,14 @@ instance decidableMemCenter (z : G) [Decidable (∀ g, g * z = z * g)] : Decidab #align subgroup.decidable_mem_center Subgroup.decidableMemCenter @[to_additive] -instance centerCharacteristic : (center G).Characteristic := - by +instance centerCharacteristic : (center G).Characteristic := by refine' characteristic_iff_comap_le.mpr fun ϕ g hg h => _ rw [← ϕ.injective.eq_iff, ϕ.map_mul, ϕ.map_mul] exact hg (ϕ h) #align subgroup.center_characteristic Subgroup.centerCharacteristic #align add_subgroup.center_characteristic AddSubgroup.center_characteristic -theorem CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := - by +theorem CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := by rw [eq_top_iff'] intro x y exact mul_comm y x @@ -2226,8 +2202,7 @@ theorem le_normalizer_comap (f : N →* G) : H.normalizer.comap f ≤ (H.comap f /-- The image of the normalizer is contained in the normalizer of the image. -/ @[to_additive "The image of the normalizer is contained in the normalizer of the image."] -theorem le_normalizer_map (f : G →* N) : H.normalizer.map f ≤ (H.map f).normalizer := fun _ => - by +theorem le_normalizer_map (f : G →* N) : H.normalizer.map f ≤ (H.map f).normalizer := fun _ => by simp only [and_imp, exists_prop, mem_map, exists_imp, mem_normalizer_iff] rintro x hx rfl n constructor @@ -2253,8 +2228,7 @@ variable {G} /-- Alternative phrasing of the normalizer condition: Only the full group is self-normalizing. This may be easier to work with, as it avoids inequalities and negations. -/ theorem normalizerCondition_iff_only_full_group_self_normalizing : - NormalizerCondition G ↔ ∀ H : Subgroup G, H.normalizer = H → H = ⊤ := - by + NormalizerCondition G ↔ ∀ H : Subgroup G, H.normalizer = H → H = ⊤ := by apply forall_congr'; intro H simp only [lt_iff_le_and_ne, le_normalizer, true_and_iff, le_top, Ne.def] tauto @@ -2317,8 +2291,7 @@ theorem centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := @[to_additive] instance Subgroup.Centralizer.characteristic [hH : H.Characteristic] : - H.centralizer.Characteristic := - by + H.centralizer.Characteristic := by refine' subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.Injective _ rw [map_mul, map_mul] exact hg (ϕ h) (subgroup.characteristic_iff_le_comap.mp hH ϕ hh) @@ -2432,8 +2405,7 @@ theorem conjugatesOfSet_subset {s : Set G} {N : Subgroup G} [N.Normal] (h : s /-- The set of conjugates of `s` is closed under conjugation. -/ theorem conj_mem_conjugatesOfSet {x c : G} : - x ∈ conjugatesOfSet s → c * x * c⁻¹ ∈ conjugatesOfSet s := fun H => - by + x ∈ conjugatesOfSet s → c * x * c⁻¹ ∈ conjugatesOfSet s := fun H => by rcases mem_conjugates_of_set_iff.1 H with ⟨a, h₁, h₂⟩ exact mem_conjugates_of_set_iff.2 ⟨a, h₁, h₂.trans (isConj_iff.2 ⟨c, rfl⟩)⟩ #align group.conj_mem_conjugates_of_set Group.conj_mem_conjugatesOfSet @@ -2478,8 +2450,7 @@ instance normalClosure_normal : (normalClosure s).Normal := #align subgroup.normal_closure_normal Subgroup.normalClosure_normal /-- The normal closure of `s` is the smallest normal subgroup containing `s`. -/ -theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : normalClosure s ≤ N := - by +theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : normalClosure s ≤ N := by intro a w refine' closure_induction w (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ · exact conjugates_of_set_subset h hx @@ -2533,8 +2504,7 @@ def normalCore (H : Subgroup G) : Subgroup G mul_mem' a b ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) #align subgroup.normal_core Subgroup.normalCore -theorem normalCore_le (H : Subgroup G) : H.normalCore ≤ H := fun a h => - by +theorem normalCore_le (H : Subgroup G) : H.normalCore ≤ H := fun a h => by rw [← mul_one a, ← inv_one, ← one_mul a] exact h 1 #align subgroup.normal_core_le Subgroup.normalCore_le @@ -2672,8 +2642,7 @@ theorem range_one : (1 : G →* N).range = ⊥ := #align add_monoid_hom.range_zero AddMonoidHom.range_zero @[simp, to_additive] -theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := - by +theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := by rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] ext simp @@ -2690,8 +2659,7 @@ theorem Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : @[to_additive] theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type _} [Group G₁] [Group G₂] {K : Subgroup G₂} (f : G₁ →* G₂) (h : f.range ≤ K) : - f.range.subgroupOf K = (f.codRestrict K fun x => h ⟨x, rfl⟩).range := - by + f.range.subgroupOf K = (f.codRestrict K fun x => h ⟨x, rfl⟩).range := by ext k refine' exists_congr _ simp [Subtype.ext_iff] @@ -2775,16 +2743,14 @@ theorem coe_ker (f : G →* M) : (f.ker : Set G) = (f : G → M) ⁻¹' {1} := #align add_monoid_hom.coe_ker AddMonoidHom.coe_ker @[simp, to_additive] -theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := - by +theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := by ext x simp [mem_ker, Units.ext_iff] #align monoid_hom.ker_to_hom_units MonoidHom.ker_toHomUnits #align add_monoid_hom.ker_to_hom_add_units AddMonoidHom.ker_to_hom_add_units @[to_additive] -theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := - by +theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := by constructor <;> intro h · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_self, map_one] · rw [← one_mul x, ← mul_inv_self y, mul_assoc, map_mul, f.mem_ker.1 h, mul_one] @@ -2946,8 +2912,7 @@ variable {N : Type _} [Group N] (H : Subgroup G) @[to_additive] theorem Normal.map {H : Subgroup G} (h : H.Normal) (f : G →* N) (hf : Function.Surjective f) : - (H.map f).Normal := - by + (H.map f).Normal := by rw [← normalizer_eq_top, ← top_le_iff, ← f.range_top_of_surjective hf, f.range_eq_map, ← normalizer_eq_top.2 h] exact le_normalizer_map _ @@ -3012,8 +2977,7 @@ theorem map_comap_eq (H : Subgroup N) : map f (comap f H) = f.range ⊓ H := #align add_subgroup.map_comap_eq AddSubgroup.map_comap_eq @[to_additive] -theorem comap_map_eq (H : Subgroup G) : comap f (map f H) = H ⊔ f.ker := - by +theorem comap_map_eq (H : Subgroup G) : comap f (map f H) = H ⊔ f.ker := by refine' le_antisymm _ (sup_le (le_comap_map _ _) (ker_le_comap _ _)) intro x hx; simp only [exists_prop, mem_map, mem_comap] at hx rcases hx with ⟨y, hy, hy'⟩ @@ -3128,16 +3092,14 @@ theorem map_eq_comap_of_inverse {f : G →* N} {g : N →* G} (hl : Function.Lef /-- Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`. -/ @[to_additive "Given `f(A) = f(B)`, `ker f ≤ A`, and `ker f ≤ B`, deduce that `A = B`."] theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ker ≤ K) - (hf : map f H = map f K) : H = K := - by + (hf : map f H = map f K) : H = K := by apply_fun comap f at hf rwa [comap_map_eq, comap_map_eq, sup_of_le_left hH, sup_of_le_left hK] at hf #align subgroup.map_injective_of_ker_le Subgroup.map_injective_of_ker_le #align add_subgroup.map_injective_of_ker_le AddSubgroup.map_injective_of_ker_le @[to_additive] -theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).Subtype ⁻¹' s) = ⊤ := - by +theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).Subtype ⁻¹' s) = ⊤ := by apply map_injective (closure s).subtype_injective rwa [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, Set.image_preimage_eq_of_subset] @@ -3173,8 +3135,7 @@ theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : @[to_additive] theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : - Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := - by + Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := by rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self] exacts[le_sup_left, le_sup_right] #align subgroup.codisjoint_subgroup_of_sup Subgroup.codisjoint_subgroupOf_sup @@ -3216,8 +3177,7 @@ theorem comap_normalizer_eq_of_surjective (H : Subgroup G) {f : N →* G} @[to_additive] theorem comap_normalizer_eq_of_injective_of_le_range {N : Type _} [Group N] (H : Subgroup G) {f : N →* G} (hf : Function.Injective f) (h : H.normalizer ≤ f.range) : - comap f H.normalizer = (comap f H).normalizer := - by + comap f H.normalizer = (comap f H).normalizer := by apply Subgroup.map_injective hf rw [map_comap_eq_self h] apply le_antisymm @@ -3230,8 +3190,7 @@ theorem comap_normalizer_eq_of_injective_of_le_range {N : Type _} [Group N] (H : @[to_additive] theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : - H.normalizer.subgroupOf N = (H.subgroupOf N).normalizer := - by + H.normalizer.subgroupOf N = (H.subgroupOf N).normalizer := by apply comap_normalizer_eq_of_injective_of_le_range exact Subtype.coe_injective simpa @@ -3242,8 +3201,7 @@ theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : @[to_additive "The image of the normalizer is equal to the normalizer of the image of an\nisomorphism."] theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : - H.normalizer.map f.toMonoidHom = (H.map f.toMonoidHom).normalizer := - by + H.normalizer.map f.toMonoidHom = (H.map f.toMonoidHom).normalizer := by ext x simp only [mem_normalizer_iff, mem_map_equiv] rw [f.to_equiv.forall_congr] @@ -3286,8 +3244,7 @@ def liftOfRightInverseAux (hf : Function.RightInverse f_inv f) (g : G₁ →* G @[simp, to_additive] theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) - (hg : f.ker ≤ g.ker) (x : G₁) : (f.liftOfRightInverseAux f_inv hf g hg) (f x) = g x := - by + (hg : f.ker ≤ g.ker) (x : G₁) : (f.liftOfRightInverseAux f_inv hf g hg) (f x) = g x := by dsimp [lift_of_right_inverse_aux] rw [← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker] apply hg @@ -3359,8 +3316,7 @@ theorem liftOfRightInverse_comp (hf : Function.RightInverse f_inv f) @[to_additive] theorem eq_liftOfRightInverse (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) (h : G₂ →* G₃) (hh : h.comp f = g) : - h = f.liftOfRightInverse f_inv hf ⟨g, hg⟩ := - by + h = f.liftOfRightInverse f_inv hf ⟨g, hg⟩ := by simp_rw [← hh] exact ((f.lift_of_right_inverse f_inv hf).apply_symm_apply _).symm #align monoid_hom.eq_lift_of_right_inverse MonoidHom.eq_liftOfRightInverse @@ -3466,8 +3422,7 @@ namespace Subgroup @[simp, to_additive] theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : - H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := - by + H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := by ext rfl #align subgroup.equiv_map_of_injective_coe_mul_equiv Subgroup.equivMapOfInjective_coe_mulEquiv @@ -3500,8 +3455,7 @@ theorem mem_sup' : x ∈ s ⊔ t ↔ ∃ (y : s)(z : t), (y : C) * z = x := @[to_additive] theorem mem_closure_pair {x y z : C} : - z ∈ closure ({x, y} : Set C) ↔ ∃ m n : ℤ, x ^ m * y ^ n = z := - by + z ∈ closure ({x, y} : Set C) ↔ ∃ m n : ℤ, x ^ m * y ^ n = z := by rw [← Set.singleton_union, Subgroup.closure_union, mem_sup] simp_rw [exists_prop, mem_closure_singleton, exists_exists_eq_and] #align subgroup.mem_closure_pair Subgroup.mem_closure_pair @@ -3578,8 +3532,7 @@ instance normal_inf_normal (H K : Subgroup G) [hH : H.Normal] [hK : K.Normal] : @[to_additive] theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : - (A ⊔ A').subgroupOf B = A.subgroupOf B ⊔ A'.subgroupOf B := - by + (A ⊔ A').subgroupOf B = A.subgroupOf B ⊔ A'.subgroupOf B := by refine' map_injective_of_ker_le B.subtype (ker_le_comap _ _) (le_trans (ker_le_comap B.subtype _) le_sup_left) _ @@ -3590,8 +3543,7 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : @[to_additive] theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgroupOf K).Normal] - {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := - by + {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := by have := (normal_subgroup_of_iff hK).mp hN (a * b) b h hb rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this #align subgroup.subgroup_normal.mem_comm Subgroup.SubgroupNormal.mem_comm @@ -3600,8 +3552,7 @@ theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgr /-- Elements of disjoint, normal subgroups commute. -/ @[to_additive "Elements of disjoint, normal subgroups commute."] theorem 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 := - by + (hdis : Disjoint H₁ H₂) (x y : G) (hx : x ∈ H₁) (hy : y ∈ H₂) : Commute x y := by suffices x * y * x⁻¹ * y⁻¹ = 1 by show x * y = y * x · rw [mul_assoc, mul_eq_one_iff_eq_inv] at this @@ -3644,8 +3595,7 @@ theorem disjoint_iff_mul_eq_one {H₁ H₂ : Subgroup G} : @[to_additive] theorem mul_injective_of_disjoint {H₁ H₂ : Subgroup G} (h : Disjoint H₁ H₂) : - Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := - by + Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := by intro x y hxy rw [← inv_mul_eq_iff_eq_mul, ← mul_assoc, ← mul_inv_eq_one, mul_assoc] at hxy replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).Prop (x.2 * y.2⁻¹).Prop hxy @@ -3662,8 +3612,7 @@ open Subgroup theorem normalClosure_eq_top_of {N : Subgroup G} [hn : N.Normal] {g g' : G} {hg : g ∈ N} {hg' : g' ∈ N} (hc : IsConj g g') (ht : normalClosure ({⟨g, hg⟩} : Set N) = ⊤) : - normalClosure ({⟨g', hg'⟩} : Set N) = ⊤ := - by + normalClosure ({⟨g', hg'⟩} : Set N) = ⊤ := by obtain ⟨c, rfl⟩ := isConj_iff.1 hc have h : ∀ x : N, (MulAut.conj c) x ∈ N := by From b7d6afc025097fb7f01ed0489a727130ab6c7bec Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 23 Jan 2023 21:36:21 +0100 Subject: [PATCH 04/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 57 ++++++++----------------- 1 file changed, 17 insertions(+), 40 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 487b3dd84fe72..77d78f4404c3e 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -90,6 +90,7 @@ subgroup, subgroups open Function +open Int variable {G G' : Type _} [Group G] [Group G'] @@ -126,8 +127,6 @@ attribute [to_additive] InvMemClass SubgroupClass variable {M S : Type _} [DivInvMonoid M] [SetLike S M] [hSM : SubgroupClass S M] {H K : S} -include hSM - /-- A subgroup is closed under division. -/ @[to_additive "An additive subgroup is closed under subtraction."] theorem div_mem {x y : M} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := by @@ -146,11 +145,7 @@ theorem zpow_mem {x : M} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K #align zpow_mem zpow_mem #align zsmul_mem zsmul_mem -omit hSM - -variable [SetLike S G] [hSG : SubgroupClass S G] - -include hSG +variable [SetLike S G] [SubgroupClass S G] @[simp, to_additive] theorem inv_mem_iff {x : G} : x⁻¹ ∈ H ↔ x ∈ H := @@ -186,10 +181,6 @@ theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := namespace SubgroupClass -omit hSG - -include hSM - /-- A subgroup of a group inherits an inverse. -/ @[to_additive "An additive subgroup of a `add_group` inherits an inverse."] instance hasInv : Inv H := @@ -204,58 +195,48 @@ instance hasDiv : Div H := #align subgroup_class.has_div SubgroupClass.hasDiv #align add_subgroup_class.has_sub AddSubgroupClass.hasSub -omit hSM - /-- An additive subgroup of an `add_group` inherits an integer scaling. -/ instance AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] {H : S} : SMul ℤ H := - ⟨fun n a => ⟨n • a, zsmul_mem a.2 n⟩⟩ + ⟨fun n a => ⟨n • a.1, zsmul_mem a.2 n⟩⟩ #align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul -include hSM - /-- A subgroup of a group inherits an integer power. -/ @[to_additive] instance hasZpow : Pow H ℤ := - ⟨fun a n => ⟨a ^ n, zpow_mem a.2 n⟩⟩ + ⟨fun a n => ⟨a.1 ^ n, zpow_mem a.2 n⟩⟩ #align subgroup_class.has_zpow SubgroupClass.hasZpow #align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul @[simp, norm_cast, to_additive] -theorem coe_inv (x : H) : ↑(x⁻¹ : H) = (x⁻¹ : M) := +theorem coe_inv (x : H) : (x⁻¹).1 = x.1⁻¹ := rfl #align subgroup_class.coe_inv SubgroupClass.coe_inv #align add_subgroup_class.coe_neg AddSubgroupClass.coe_neg @[simp, norm_cast, to_additive] -theorem coe_div (x y : H) : (↑(x / y) : M) = ↑x / ↑y := +theorem coe_div (x y : H) : (x / y).1 = x.1 / y.1 := rfl #align subgroup_class.coe_div SubgroupClass.coe_div #align add_subgroup_class.coe_sub AddSubgroupClass.coe_sub -omit hSM - variable (H) -include hSG - -- Prefer subclasses of `group` over subclasses of `subgroup_class`. /-- A subgroup of a group inherits a group structure. -/ @[to_additive "An additive subgroup of an `add_group` inherits an `add_group` structure."] instance (priority := 75) toGroup : Group H := - Subtype.coe_injective.Group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup_class.to_group SubgroupClass.toGroup #align add_subgroup_class.to_add_group AddSubgroupClass.toAddGroup -omit hSG - -- Prefer subclasses of `comm_group` over subclasses of `subgroup_class`. /-- A subgroup of a `comm_group` is a `comm_group`. -/ @[to_additive "An additive subgroup of an `add_comm_group` is an `add_comm_group`."] instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [SubgroupClass S G] : CommGroup H := - Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.commGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup_class.to_comm_group SubgroupClass.toCommGroup #align add_subgroup_class.to_add_comm_group AddSubgroupClass.toAddCommGroup @@ -265,7 +246,7 @@ instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [ @[to_additive "An additive subgroup of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : OrderedCommGroup H := - Subtype.coe_injective.OrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.orderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup_class.to_ordered_comm_group SubgroupClass.toOrderedCommGroup #align add_subgroup_class.to_ordered_add_comm_group AddSubgroupClass.toOrderedAddCommGroup @@ -276,17 +257,15 @@ instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [ "An additive subgroup of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : LinearOrderedCommGroup H := - Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) + Subtype.coe_injective.linearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup_class.to_linear_ordered_comm_group SubgroupClass.toLinearOrderedCommGroup #align add_subgroup_class.to_linear_ordered_add_comm_group AddSubgroupClass.toLinearOrderedAddCommGroup -include hSG - /-- The natural group hom from a subgroup of group `G` to `G`. -/ @[to_additive "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] def subtype : H →* G := - ⟨coe, rfl, fun _ _ => rfl⟩ + ⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩ #align subgroup_class.subtype SubgroupClass.subtype #align add_subgroup_class.subtype AddSubgroupClass.subtype @@ -313,7 +292,7 @@ theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] def inclusion {H K : S} (h : H ≤ K) : H →* K := - MonoidHom.mk' (fun x => ⟨x, h x.Prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl + MonoidHom.mk' (fun x => ⟨x, h x.prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl #align subgroup_class.inclusion SubgroupClass.inclusion #align add_subgroup_class.inclusion AddSubgroupClass.inclusion @@ -347,7 +326,7 @@ theorem inclusion_inclusion {L : S} (hHK : H ≤ K) (hKL : K ≤ L) (x : H) : @[simp, to_additive] theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := by cases a - simp only [inclusion, [anonymous], MonoidHom.mk'_apply] + simp only [inclusion, MonoidHom.mk'_apply] #align subgroup_class.coe_inclusion SubgroupClass.coe_inclusion #align add_subgroup_class.coe_inclusion AddSubgroupClass.coe_inclusion @@ -389,12 +368,11 @@ namespace Subgroup @[to_additive] instance : SetLike (Subgroup G) G where - coe := Subgroup.carrier + coe s := s.carrier coe_injective' p q h := by cases p <;> cases q <;> congr @[to_additive] -instance : SubgroupClass (Subgroup G) G - where +instance : SubgroupClass (Subgroup G) G where mul_mem := Subgroup.mul_mem' one_mem := Subgroup.one_mem' inv_mem := Subgroup.inv_mem' @@ -425,7 +403,7 @@ theorem mk_le_mk {s t : Set G} (h_one) (h_mul) (h_inv) (h_one') (h_mul') (h_inv' #align add_subgroup.mk_le_mk AddSubgroup.mk_le_mk /-- See Note [custom simps projection] -/ -@[to_additive "See Note [custom simps projection]"] +--@[to_additive "See Note [custom simps projection]"] def Simps.coe (S : Subgroup G) : Set G := S #align subgroup.simps.coe Subgroup.Simps.coe @@ -3638,5 +3616,4 @@ theorem normalClosure_eq_top_of {N : Subgroup G} [hn : N.Normal] {g g' : G} {hg end IsConj -assert_not_exists Multiset - +-- assert_not_exists Multiset From 0b6c65ce3bb20c65ea53285d74f1bd8c20d9570a Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 24 Jan 2023 08:36:26 +0000 Subject: [PATCH 05/34] minifix --- Mathlib/GroupTheory/Subgroup/Basic.lean | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 77d78f4404c3e..0ecbf5fcd7dea 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -169,12 +169,16 @@ theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ @[to_additive] theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := + -- Porting note: whut? why do we need this? + have : SubmonoidClass S G := SubgroupClass.toSubmonoidClass ⟨fun hba => by simpa using mul_mem hba (inv_mem h), fun hb => mul_mem hb h⟩ #align mul_mem_cancel_right mul_mem_cancel_right #align add_mem_cancel_right add_mem_cancel_right @[to_additive] theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := + -- Porting note: whut? why do we need this? + have : SubmonoidClass S G := SubgroupClass.toSubmonoidClass ⟨fun hab => by simpa using mul_mem (inv_mem h) hab, mul_mem h⟩ #align mul_mem_cancel_left mul_mem_cancel_left #align add_mem_cancel_left add_mem_cancel_left @@ -196,7 +200,7 @@ instance hasDiv : Div H := #align add_subgroup_class.has_sub AddSubgroupClass.hasSub /-- An additive subgroup of an `add_group` inherits an integer scaling. -/ -instance AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] +instance _root_.AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] {H : S} : SMul ℤ H := ⟨fun n a => ⟨n • a.1, zsmul_mem a.2 n⟩⟩ #align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul @@ -206,7 +210,7 @@ instance AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubg instance hasZpow : Pow H ℤ := ⟨fun a n => ⟨a.1 ^ n, zpow_mem a.2 n⟩⟩ #align subgroup_class.has_zpow SubgroupClass.hasZpow -#align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul +-- Porting note: additive align statement is given above @[simp, norm_cast, to_additive] theorem coe_inv (x : H) : (x⁻¹).1 = x.1⁻¹ := @@ -375,7 +379,7 @@ instance : SetLike (Subgroup G) G where instance : SubgroupClass (Subgroup G) G where mul_mem := Subgroup.mul_mem' one_mem := Subgroup.one_mem' - inv_mem := Subgroup.inv_mem' + inv_mem := Subgroup.inv_mem' _ @[simp, to_additive] theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := From b41d24992ab7045bf31cd7894c573ca832efb707 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Tue, 24 Jan 2023 20:03:38 +0100 Subject: [PATCH 06/34] give hasZsmul and hasZpow congruent types --- Mathlib/GroupTheory/Subgroup/Basic.lean | 204 ++++++++++++------------ 1 file changed, 102 insertions(+), 102 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 0ecbf5fcd7dea..93b3a6fac8187 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -130,7 +130,7 @@ variable {M S : Type _} [DivInvMonoid M] [SetLike S M] [hSM : SubgroupClass S M] /-- A subgroup is closed under division. -/ @[to_additive "An additive subgroup is closed under subtraction."] theorem div_mem {x y : M} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := by - rw [div_eq_mul_inv] <;> exact mul_mem hx (inv_mem hy) + rw [div_eq_mul_inv]; exact mul_mem hx (inv_mem hy) #align div_mem div_mem #align sub_mem sub_mem @@ -147,7 +147,7 @@ theorem zpow_mem {x : M} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K variable [SetLike S G] [SubgroupClass S G] -@[simp, to_additive] +@[to_additive (attr := simp)] theorem inv_mem_iff {x : G} : x⁻¹ ∈ H ↔ x ∈ H := ⟨fun h => inv_inv x ▸ inv_mem h, inv_mem⟩ #align inv_mem_iff inv_mem_iff @@ -159,7 +159,7 @@ theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := by #align div_mem_comm_iff div_mem_comm_iff #align sub_mem_comm_iff sub_mem_comm_iff -@[simp, to_additive] +@[to_additive (attr := simp)] theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ P x⁻¹) ↔ ∃ x ∈ H, P x := by constructor <;> · rintro ⟨x, x_in, hx⟩ @@ -170,7 +170,7 @@ theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ @[to_additive] theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := -- Porting note: whut? why do we need this? - have : SubmonoidClass S G := SubgroupClass.toSubmonoidClass + haveI : SubmonoidClass S G := SubgroupClass.toSubmonoidClass ⟨fun hba => by simpa using mul_mem hba (inv_mem h), fun hb => mul_mem hb h⟩ #align mul_mem_cancel_right mul_mem_cancel_right #align add_mem_cancel_right add_mem_cancel_right @@ -200,25 +200,25 @@ instance hasDiv : Div H := #align add_subgroup_class.has_sub AddSubgroupClass.hasSub /-- An additive subgroup of an `add_group` inherits an integer scaling. -/ -instance _root_.AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] - {H : S} : SMul ℤ H := +instance _root_.AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] + [AddSubgroupClass S M] {H : S} : SMul ℤ H := ⟨fun n a => ⟨n • a.1, zsmul_mem a.2 n⟩⟩ #align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul /-- A subgroup of a group inherits an integer power. -/ @[to_additive] -instance hasZpow : Pow H ℤ := +instance hasZpow {M S} [DivInvMonoid M] [SetLike S M] [SubgroupClass S M] {H : S} : Pow H ℤ := ⟨fun a n => ⟨a.1 ^ n, zpow_mem a.2 n⟩⟩ #align subgroup_class.has_zpow SubgroupClass.hasZpow -- Porting note: additive align statement is given above -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_inv (x : H) : (x⁻¹).1 = x.1⁻¹ := rfl #align subgroup_class.coe_inv SubgroupClass.coe_inv #align add_subgroup_class.coe_neg AddSubgroupClass.coe_neg -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_div (x y : H) : (x / y).1 = x.1 / y.1 := rfl #align subgroup_class.coe_div SubgroupClass.coe_div @@ -273,7 +273,7 @@ def subtype : H →* G := #align subgroup_class.subtype SubgroupClass.subtype #align add_subgroup_class.subtype AddSubgroupClass.subtype -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coeSubtype : (subtype H : H → G) = coe := rfl #align subgroup_class.coe_subtype SubgroupClass.coeSubtype @@ -281,13 +281,13 @@ theorem coeSubtype : (subtype H : H → G) = coe := variable {H} -@[simp, norm_cast, to_additive coe_smul] +@[to_additive (attr := simp, norm_cast) coe_smul] theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := (subtype H : H →* G).map_pow _ _ #align subgroup_class.coe_pow SubgroupClass.coe_pow #align add_subgroup_class.coe_smul AddSubgroupClass.coe_smul -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := (subtype H : H →* G).map_zpow _ _ #align subgroup_class.coe_zpow SubgroupClass.coe_zpow @@ -300,14 +300,14 @@ def inclusion {H K : S} (h : H ≤ K) : H →* K := #align subgroup_class.inclusion SubgroupClass.inclusion #align add_subgroup_class.inclusion AddSubgroupClass.inclusion -@[simp, to_additive] +@[to_additive (attr := simp)] theorem inclusion_self (x : H) : inclusion le_rfl x = x := by cases x rfl #align subgroup_class.inclusion_self SubgroupClass.inclusion_self #align add_subgroup_class.inclusion_self AddSubgroupClass.inclusion_self -@[simp, to_additive] +@[to_additive (attr := simp)] theorem inclusion_mk {h : H ≤ K} (x : G) (hx : x ∈ H) : inclusion h ⟨x, hx⟩ = ⟨x, h hx⟩ := rfl #align subgroup_class.inclusion_mk SubgroupClass.inclusion_mk @@ -327,14 +327,14 @@ theorem inclusion_inclusion {L : S} (hHK : H ≤ K) (hKL : K ≤ L) (x : H) : rfl #align subgroup_class.inclusion_inclusion SubgroupClass.inclusion_inclusion -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := by cases a simp only [inclusion, MonoidHom.mk'_apply] #align subgroup_class.coe_inclusion SubgroupClass.coe_inclusion #align add_subgroup_class.coe_inclusion AddSubgroupClass.coe_inclusion -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subtype_comp_inclusion {H K : S} (hH : H ≤ K) : (subtype K).comp (inclusion hH) = subtype H := by ext @@ -381,25 +381,25 @@ instance : SubgroupClass (Subgroup G) G where one_mem := Subgroup.one_mem' inv_mem := Subgroup.inv_mem' _ -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl #align subgroup.mem_carrier Subgroup.mem_carrier #align add_subgroup.mem_carrier AddSubgroup.mem_carrier -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_mk {s : Set G} {x : G} (h_one) (h_mul) (h_inv) : x ∈ mk s h_one h_mul h_inv ↔ x ∈ s := Iff.rfl #align subgroup.mem_mk Subgroup.mem_mk #align add_subgroup.mem_mk AddSubgroup.mem_mk -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_set_mk {s : Set G} (h_one) (h_mul) (h_inv) : (mk s h_one h_mul h_inv : Set G) = s := rfl #align subgroup.coe_set_mk Subgroup.coe_set_mk #align add_subgroup.coe_set_mk AddSubgroup.coe_set_mk -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mk_le_mk {s t : Set G} (h_one) (h_mul) (h_inv) (h_one') (h_mul') (h_inv') : mk s h_one h_mul h_inv ≤ mk t h_one' h_mul' h_inv' ↔ s ⊆ t := Iff.rfl @@ -417,13 +417,13 @@ initialize_simps_projections Subgroup (carrier → coe) initialize_simps_projections AddSubgroup (carrier → coe) -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_toSubmonoid (K : Subgroup G) : (K.toSubmonoid : Set G) = K := rfl #align subgroup.coe_to_submonoid Subgroup.coe_toSubmonoid #align add_subgroup.coe_to_add_submonoid AddSubgroup.coe_to_add_submonoid -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_toSubmonoid (K : Subgroup G) (x : G) : x ∈ K.toSubmonoid ↔ x ∈ K := Iff.rfl #align subgroup.mem_to_submonoid Subgroup.mem_toSubmonoid @@ -435,7 +435,7 @@ theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subgroup G → #align subgroup.to_submonoid_injective Subgroup.toSubmonoid_injective #align add_subgroup.to_add_submonoid_injective AddSubgroup.to_add_submonoid_injective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem toSubmonoid_eq {p q : Subgroup G} : p.toSubmonoid = q.toSubmonoid ↔ p = q := toSubmonoid_injective.eq_iff #align subgroup.to_submonoid_eq Subgroup.toSubmonoid_eq @@ -457,7 +457,7 @@ theorem toSubmonoid_mono : Monotone (toSubmonoid : Subgroup G → Submonoid G) : attribute [mono] AddSubgroup.to_add_submonoid_mono -@[simp, to_additive] +@[to_additive (attr := simp)] theorem toSubmonoid_le {p q : Subgroup G} : p.toSubmonoid ≤ q.toSubmonoid ↔ p ≤ q := Iff.rfl #align subgroup.to_submonoid_le Subgroup.toSubmonoid_le @@ -525,7 +525,7 @@ protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G #align subgroup.copy Subgroup.copy #align add_subgroup.copy AddSubgroup.copy -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_copy (K : Subgroup G) (s : Set G) (hs : s = ↑K) : (K.copy s hs : Set G) = s := rfl #align subgroup.coe_copy Subgroup.coe_copy @@ -683,49 +683,49 @@ instance hasZpow : Pow H ℤ := #align subgroup.has_zpow Subgroup.hasZpow #align add_subgroup.has_zsmul AddSubgroup.hasZsmul -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_mul (x y : H) : (↑(x * y) : G) = ↑x * ↑y := rfl #align subgroup.coe_mul Subgroup.coe_mul #align add_subgroup.coe_add AddSubgroup.coe_add -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_one : ((1 : H) : G) = 1 := rfl #align subgroup.coe_one Subgroup.coe_one #align add_subgroup.coe_zero AddSubgroup.coe_zero -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_inv (x : H) : ↑(x⁻¹ : H) = (x⁻¹ : G) := rfl #align subgroup.coe_inv Subgroup.coe_inv #align add_subgroup.coe_neg AddSubgroup.coe_neg -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_div (x y : H) : (↑(x / y) : G) = ↑x / ↑y := rfl #align subgroup.coe_div Subgroup.coe_div #align add_subgroup.coe_sub AddSubgroup.coe_sub -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_mk (x : G) (hx : x ∈ H) : ((⟨x, hx⟩ : H) : G) = x := rfl #align subgroup.coe_mk Subgroup.coe_mk #align add_subgroup.coe_mk AddSubgroup.coe_mk -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := rfl #align subgroup.coe_pow Subgroup.coe_pow #align add_subgroup.coe_nsmul AddSubgroup.coe_nsmul -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := rfl #align subgroup.coe_zpow Subgroup.coe_zpow #align add_subgroup.coe_zsmul AddSubgroup.coe_zsmul -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mk_eq_one_iff {g : G} {h} : (⟨g, h⟩ : H) = 1 ↔ g = 1 := show (⟨g, h⟩ : H) = (⟨1, H.one_mem⟩ : H) ↔ g = 1 by simp #align subgroup.mk_eq_one_iff Subgroup.mk_eq_one_iff @@ -773,7 +773,7 @@ def subtype : H →* G := #align subgroup.subtype Subgroup.subtype #align add_subgroup.subtype AddSubgroup.subtype -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coeSubtype : ⇑H.Subtype = coe := rfl #align subgroup.coe_subtype Subgroup.coeSubtype @@ -792,7 +792,7 @@ def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := #align subgroup.inclusion Subgroup.inclusion #align add_subgroup.inclusion AddSubgroup.inclusion -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_inclusion {H K : Subgroup G} {h : H ≤ K} (a : H) : (inclusion h a : G) = a := by cases a simp only [inclusion, coe_mk, MonoidHom.mk'_apply] @@ -805,7 +805,7 @@ theorem inclusion_injective {H K : Subgroup G} (h : H ≤ K) : Function.Injectiv #align subgroup.inclusion_injective Subgroup.inclusion_injective #align add_subgroup.inclusion_injective AddSubgroup.inclusion_injective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : K.Subtype.comp (inclusion hH) = H.Subtype := rfl @@ -837,25 +837,25 @@ instance : Bot (Subgroup G) := instance : Inhabited (Subgroup G) := ⟨⊥⟩ -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_bot {x : G} : x ∈ (⊥ : Subgroup G) ↔ x = 1 := Iff.rfl #align subgroup.mem_bot Subgroup.mem_bot #align add_subgroup.mem_bot AddSubgroup.mem_bot -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_top (x : G) : x ∈ (⊤ : Subgroup G) := Set.mem_univ x #align subgroup.mem_top Subgroup.mem_top #align add_subgroup.mem_top AddSubgroup.mem_top -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_top : ((⊤ : Subgroup G) : Set G) = Set.univ := rfl #align subgroup.coe_top Subgroup.coe_top #align add_subgroup.coe_top AddSubgroup.coe_top -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_bot : ((⊥ : Subgroup G) : Set G) = {1} := rfl #align subgroup.coe_bot Subgroup.coe_bot @@ -865,13 +865,13 @@ theorem coe_bot : ((⊥ : Subgroup G) : Set G) = {1} := instance : Unique (⊥ : Subgroup G) := ⟨⟨1⟩, fun g => Subtype.ext g.2⟩ -@[simp, to_additive] +@[to_additive (attr := simp)] theorem top_toSubmonoid : (⊤ : Subgroup G).toSubmonoid = ⊤ := rfl #align subgroup.top_to_submonoid Subgroup.top_toSubmonoid #align add_subgroup.top_to_add_submonoid AddSubgroup.top_to_add_submonoid -@[simp, to_additive] +@[to_additive (attr := simp)] theorem bot_toSubmonoid : (⊥ : Subgroup G).toSubmonoid = ⊥ := rfl #align subgroup.bot_to_submonoid Subgroup.bot_toSubmonoid @@ -942,13 +942,13 @@ instance : HasInf (Subgroup G) := { H₁.toSubmonoid ⊓ H₂.toSubmonoid with inv_mem' := fun _ ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = p ∩ p' := rfl #align subgroup.coe_inf Subgroup.coe_inf #align add_subgroup.coe_inf AddSubgroup.coe_inf -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_inf {p p' : Subgroup G} {x : G} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x ∈ p' := Iff.rfl #align subgroup.mem_inf Subgroup.mem_inf @@ -961,13 +961,13 @@ instance : InfSet (Subgroup G) := inv_mem' := fun x hx => Set.mem_binterᵢ fun i h => i.inv_mem (by apply Set.mem_interᵢ₂.1 hx i h) }⟩ -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_infₛ (H : Set (Subgroup G)) : ((infₛ H : Subgroup G) : Set G) = ⋂ s ∈ H, ↑s := rfl #align subgroup.coe_Inf Subgroup.coe_infₛ #align add_subgroup.coe_Inf AddSubgroup.coe_Inf -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_infₛ {S : Set (Subgroup G)} {x : G} : x ∈ infₛ S ↔ ∀ p ∈ S, x ∈ p := Set.mem_interᵢ₂ #align subgroup.mem_Inf Subgroup.mem_infₛ @@ -979,7 +979,7 @@ theorem mem_infᵢ {ι : Sort _} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, #align subgroup.mem_infi Subgroup.mem_infᵢ #align add_subgroup.mem_infi AddSubgroup.mem_infi -@[simp, norm_cast, to_additive] +@[to_additive (attr := simp, norm_cast)] theorem coe_infᵢ {ι : Sort _} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by simp only [infᵢ, coe_Inf, Set.binterᵢ_range] #align subgroup.coe_infi Subgroup.coe_infᵢ @@ -1033,7 +1033,7 @@ theorem mem_supₛ_of_mem {S : Set (Subgroup G)} {s : Subgroup G} (hs : s ∈ S) #align subgroup.mem_Sup_of_mem Subgroup.mem_supₛ_of_mem #align add_subgroup.mem_Sup_of_mem AddSubgroup.mem_supₛ_of_mem -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subsingleton_iff : Subsingleton (Subgroup G) ↔ Subsingleton G := ⟨fun h => ⟨fun x y => @@ -1044,7 +1044,7 @@ theorem subsingleton_iff : Subsingleton (Subgroup G) ↔ Subsingleton G := #align subgroup.subsingleton_iff Subgroup.subsingleton_iff #align add_subgroup.subsingleton_iff AddSubgroup.subsingleton_iff -@[simp, to_additive] +@[to_additive (attr := simp)] theorem nontrivial_iff : Nontrivial (Subgroup G) ↔ Nontrivial G := not_iff_not.mp ((not_nontrivial_iff_subsingleton.trans subsingleton_iff).trans @@ -1148,7 +1148,7 @@ theorem closure_induction₂ {p : G → G → Prop} {x} {y : G} (hx : x ∈ clos #align subgroup.closure_induction₂ Subgroup.closure_induction₂ #align add_subgroup.closure_induction₂ AddSubgroup.closure_induction₂ -@[simp, to_additive] +@[to_additive (attr := simp)] theorem closure_closure_coe_preimage {k : Set G} : closure ((coe : closure k → G) ⁻¹' k) = ⊤ := eq_top_iff.2 fun x => Subtype.recOn x fun x hx _ => @@ -1213,13 +1213,13 @@ theorem closure_eq : closure (K : Set G) = K := #align subgroup.closure_eq Subgroup.closure_eq #align add_subgroup.closure_eq AddSubgroup.closure_eq -@[simp, to_additive] +@[to_additive (attr := simp)] theorem closure_empty : closure (∅ : Set G) = ⊥ := (Subgroup.gi G).gc.l_bot #align subgroup.closure_empty Subgroup.closure_empty #align add_subgroup.closure_empty AddSubgroup.closure_empty -@[simp, to_additive] +@[to_additive (attr := simp)] theorem closure_univ : closure (univ : Set G) = ⊤ := @coe_top G _ ▸ closure_eq ⊤ #align subgroup.closure_univ Subgroup.closure_univ @@ -1331,13 +1331,13 @@ def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := #align subgroup.comap Subgroup.comap #align add_subgroup.comap AddSubgroup.comap -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_comap (K : Subgroup N) (f : G →* N) : (K.comap f : Set G) = f ⁻¹' K := rfl #align subgroup.coe_comap Subgroup.coe_comap #align add_subgroup.coe_comap AddSubgroup.coe_comap -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_comap {K : Subgroup N} {f : G →* N} {x : G} : x ∈ K.comap f ↔ f x ∈ K := Iff.rfl #align subgroup.mem_comap Subgroup.mem_comap @@ -1356,7 +1356,7 @@ theorem comap_comap (K : Subgroup P) (g : N →* P) (f : G →* N) : #align subgroup.comap_comap Subgroup.comap_comap #align add_subgroup.comap_comap AddSubgroup.comap_comap -@[simp, to_additive] +@[to_additive (attr := simp)] theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by ext rfl @@ -1375,13 +1375,13 @@ def map (f : G →* N) (H : Subgroup G) : Subgroup N := #align subgroup.map Subgroup.map #align add_subgroup.map AddSubgroup.map -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_map (f : G →* N) (K : Subgroup G) : (K.map f : Set N) = f '' K := rfl #align subgroup.coe_map Subgroup.coe_map #align add_subgroup.coe_map AddSubgroup.coe_map -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := mem_image_iff_bex #align subgroup.mem_map Subgroup.mem_map @@ -1405,7 +1405,7 @@ theorem map_mono {f : G →* N} {K K' : Subgroup G} : K ≤ K' → map f K ≤ m #align subgroup.map_mono Subgroup.map_mono #align add_subgroup.map_mono AddSubgroup.map_mono -@[simp, to_additive] +@[to_additive (attr := simp)] theorem map_id : K.map (MonoidHom.id G) = K := SetLike.coe_injective <| image_id _ #align subgroup.map_id Subgroup.map_id @@ -1417,7 +1417,7 @@ theorem map_map (g : N →* P) (f : G →* N) : (K.map f).map g = K.map (g.comp #align subgroup.map_map Subgroup.map_map #align add_subgroup.map_map AddSubgroup.map_map -@[simp, to_additive] +@[to_additive (attr := simp)] theorem map_one_eq_bot : K.map (1 : G →* N) = ⊥ := eq_bot_iff.mpr <| by rintro x ⟨y, _, rfl⟩ @@ -1533,13 +1533,13 @@ theorem map_inf_eq (H K : Subgroup G) (f : G →* N) (hf : Function.Injective f) #align subgroup.map_inf_eq Subgroup.map_inf_eq #align add_subgroup.map_inf_eq AddSubgroup.map_inf_eq -@[simp, to_additive] +@[to_additive (attr := simp)] theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := (gc_map_comap f).l_bot #align subgroup.map_bot Subgroup.map_bot #align add_subgroup.map_bot AddSubgroup.map_bot -@[simp, to_additive] +@[to_additive (attr := simp)] theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := by rw [eq_top_iff] intro x hx @@ -1548,7 +1548,7 @@ theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgr #align subgroup.map_top_of_surjective Subgroup.map_top_of_surjective #align add_subgroup.map_top_of_surjective AddSubgroup.map_top_of_surjective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem comap_top (f : G →* N) : (⊤ : Subgroup N).comap f = ⊤ := (gc_map_comap f).u_top #align subgroup.comap_top Subgroup.comap_top @@ -1573,13 +1573,13 @@ def subgroupOfEquivOfLe {G : Type _} [Group G] {H K : Subgroup G} (h : H ≤ K) #align subgroup.subgroup_of_equiv_of_le Subgroup.subgroupOfEquivOfLe #align add_subgroup.add_subgroup_of_equiv_of_le AddSubgroup.addSubgroupOfEquivOfLe -@[simp, to_additive] +@[to_additive (attr := simp)] theorem comap_subtype (H K : Subgroup G) : H.comap K.Subtype = H.subgroupOf K := rfl #align subgroup.comap_subtype Subgroup.comap_subtype #align add_subgroup.comap_subtype AddSubgroup.comap_subtype -@[simp, to_additive] +@[to_additive (attr := simp)] theorem comap_inclusion_subgroupOf {K₁ K₂ : Subgroup G} (h : K₁ ≤ K₂) (H : Subgroup G) : (H.subgroupOf K₂).comap (inclusion h) = H.subgroupOf K₁ := rfl @@ -1598,19 +1598,19 @@ theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h #align subgroup.mem_subgroup_of Subgroup.mem_subgroupOf #align add_subgroup.mem_add_subgroup_of AddSubgroup.mem_add_subgroupOf -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.Subtype = H ⊓ K := SetLike.ext' <| Subtype.image_preimage_coe _ _ #align subgroup.subgroup_of_map_subtype Subgroup.subgroupOf_map_subtype #align add_subgroup.add_subgroup_of_map_subtype AddSubgroup.add_subgroupOf_map_subtype -@[simp, to_additive] +@[to_additive (attr := simp)] theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := Eq.symm (Subgroup.ext fun g => Subtype.ext_iff) #align subgroup.bot_subgroup_of Subgroup.bot_subgroupOf #align add_subgroup.bot_add_subgroup_of AddSubgroup.bot_add_subgroupOf -@[simp, to_additive] +@[to_additive (attr := simp)] theorem top_subgroupOf : (⊤ : Subgroup G).subgroupOf H = ⊤ := rfl #align subgroup.top_subgroup_of Subgroup.top_subgroupOf @@ -1628,38 +1628,38 @@ theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := #align subgroup.subgroup_of_bot_eq_top Subgroup.subgroupOf_bot_eq_top #align add_subgroup.add_subgroup_of_bot_eq_top AddSubgroup.add_subgroupOf_bot_eq_top -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupOf_self : H.subgroupOf H = ⊤ := top_unique fun g hg => g.2 #align subgroup.subgroup_of_self Subgroup.subgroupOf_self #align add_subgroup.add_subgroup_of_self AddSubgroup.add_subgroupOf_self -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupOf_inj {H₁ H₂ K : Subgroup G} : H₁.subgroupOf K = H₂.subgroupOf K ↔ H₁ ⊓ K = H₂ ⊓ K := by simpa only [SetLike.ext_iff, mem_inf, mem_subgroup_of, and_congr_left_iff] using Subtype.forall #align subgroup.subgroup_of_inj Subgroup.subgroupOf_inj #align add_subgroup.add_subgroup_of_inj AddSubgroup.add_subgroupOf_inj -@[simp, to_additive] +@[to_additive (attr := simp)] theorem inf_subgroupOf_right (H K : Subgroup G) : (H ⊓ K).subgroupOf K = H.subgroupOf K := subgroupOf_inj.2 inf_right_idem #align subgroup.inf_subgroup_of_right Subgroup.inf_subgroupOf_right #align add_subgroup.inf_add_subgroup_of_right AddSubgroup.inf_add_subgroupOf_right -@[simp, to_additive] +@[to_additive (attr := simp)] theorem inf_subgroupOf_left (H K : Subgroup G) : (K ⊓ H).subgroupOf K = H.subgroupOf K := by rw [inf_comm, inf_subgroup_of_right] #align subgroup.inf_subgroup_of_left Subgroup.inf_subgroupOf_left #align add_subgroup.inf_add_subgroup_of_left AddSubgroup.inf_add_subgroupOf_left -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupOf_eq_bot {H K : Subgroup G} : H.subgroupOf K = ⊥ ↔ Disjoint H K := by rw [disjoint_iff, ← bot_subgroup_of, subgroup_of_inj, bot_inf_eq] #align subgroup.subgroup_of_eq_bot Subgroup.subgroupOf_eq_bot #align add_subgroup.add_subgroup_of_eq_bot AddSubgroup.add_subgroup_of_eq_bot -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H := by rw [← top_subgroup_of, subgroup_of_inj, top_inf_eq, inf_eq_right] #align subgroup.subgroup_of_eq_top Subgroup.subgroupOf_eq_top @@ -1836,7 +1836,7 @@ theorem le_pi_iff {I : Set η} {H : ∀ i, Subgroup (f i)} {J : Subgroup (∀ i, #align subgroup.le_pi_iff Subgroup.le_pi_iff #align add_subgroup.le_pi_iff AddSubgroup.le_pi_iff -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mulSingle_mem_pi [DecidableEq η] {I : Set η} {H : ∀ i, Subgroup (f i)} (i : η) (x : f i) : Pi.mulSingle i x ∈ pi I H ↔ i ∈ I → x ∈ H i := by constructor @@ -2030,7 +2030,7 @@ theorem coe_center : ↑(center G) = Set.center G := #align subgroup.coe_center Subgroup.coe_center #align add_subgroup.coe_center AddSubgroup.coe_center -@[simp, to_additive] +@[to_additive (attr := simp)] theorem center_toSubmonoid : (center G).toSubmonoid = Submonoid.center G := rfl #align subgroup.center_to_submonoid Subgroup.center_toSubmonoid @@ -2538,13 +2538,13 @@ def range (f : G →* N) : Subgroup N := #align monoid_hom.range MonoidHom.range #align add_monoid_hom.range AddMonoidHom.range -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_range (f : G →* N) : (f.range : Set N) = Set.range f := rfl #align monoid_hom.coe_range MonoidHom.coe_range #align add_monoid_hom.coe_range AddMonoidHom.coe_range -@[simp, to_additive] +@[to_additive (attr := simp)] theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := Iff.rfl #align monoid_hom.mem_range MonoidHom.mem_range @@ -2555,7 +2555,7 @@ theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by e #align monoid_hom.range_eq_map MonoidHom.range_eq_map #align add_monoid_hom.range_eq_map AddMonoidHom.range_eq_map -@[simp, to_additive] +@[to_additive (attr := simp)] theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, Subtype.coe_mk, iff_self_iff, forall_const] @@ -2571,7 +2571,7 @@ def rangeRestrict (f : G →* N) : G →* f.range := #align monoid_hom.range_restrict MonoidHom.rangeRestrict #align add_monoid_hom.range_restrict AddMonoidHom.rangeRestrict -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.range_restrict g : N) = f g := rfl #align monoid_hom.coe_range_restrict MonoidHom.coe_rangeRestrict @@ -2617,13 +2617,13 @@ theorem range_top_of_surjective {N} [Group N] (f : G →* N) (hf : Function.Surj #align monoid_hom.range_top_of_surjective MonoidHom.range_top_of_surjective #align add_monoid_hom.range_top_of_surjective AddMonoidHom.range_top_of_surjective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem range_one : (1 : G →* N).range = ⊥ := SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x #align monoid_hom.range_one MonoidHom.range_one #align add_monoid_hom.range_zero AddMonoidHom.range_zero -@[simp, to_additive] +@[to_additive (attr := simp)] theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := by rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] ext @@ -2631,7 +2631,7 @@ theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := by #align subgroup.subtype_range Subgroup.subtype_range #align add_subgroup.subtype_range AddSubgroup.subtype_range -@[simp, to_additive] +@[to_additive (attr := simp)] theorem Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : (inclusion h_le).range = H.subgroupOf K := Subgroup.ext fun g => Set.ext_iff.mp (Set.range_inclusion h_le) g @@ -2662,14 +2662,14 @@ def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : #align monoid_hom.of_left_inverse MonoidHom.ofLeftInverse #align add_monoid_hom.of_left_inverse AddMonoidHom.ofLeftInverse -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ofLeftInverse_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) (x : G) : ↑(ofLeftInverse h x) = f x := rfl #align monoid_hom.of_left_inverse_apply MonoidHom.ofLeftInverse_apply #align add_monoid_hom.of_left_inverse_apply AddMonoidHom.ofLeftInverse_apply -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ofLeftInverse_symm_apply {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) (x : f.range) : (ofLeftInverse h).symm x = g x := rfl @@ -2724,7 +2724,7 @@ theorem coe_ker (f : G →* M) : (f.ker : Set G) = (f : G → M) ⁻¹' {1} := #align monoid_hom.coe_ker MonoidHom.coe_ker #align add_monoid_hom.coe_ker AddMonoidHom.coe_ker -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker := by ext x simp [mem_ker, Units.ext_iff] @@ -2751,38 +2751,38 @@ theorem comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker #align monoid_hom.comap_ker MonoidHom.comap_ker #align add_monoid_hom.comap_ker AddMonoidHom.comap_ker -@[simp, to_additive] +@[to_additive (attr := simp)] theorem comap_bot (f : G →* N) : (⊥ : Subgroup N).comap f = f.ker := rfl #align monoid_hom.comap_bot MonoidHom.comap_bot #align add_monoid_hom.comap_bot AddMonoidHom.comap_bot -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_restrict (f : G →* N) : (f.restrict K).ker = f.ker.subgroupOf K := rfl #align monoid_hom.ker_restrict MonoidHom.ker_restrict #align add_monoid_hom.ker_restrict AddMonoidHom.ker_restrict -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s : S) (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := SetLike.ext fun x => Subtype.ext_iff #align monoid_hom.ker_cod_restrict MonoidHom.ker_codRestrict #align add_monoid_hom.ker_cod_restrict AddMonoidHom.ker_cod_restrict -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := ker_codRestrict _ _ _ #align monoid_hom.ker_range_restrict MonoidHom.ker_rangeRestrict #align add_monoid_hom.ker_range_restrict AddMonoidHom.ker_rangeRestrict -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_one : (1 : G →* M).ker = ⊤ := SetLike.ext fun x => eq_self_iff_true _ #align monoid_hom.ker_one MonoidHom.ker_one #align add_monoid_hom.ker_zero AddMonoidHom.ker_zero -@[simp, to_additive] +@[to_additive (attr := simp)] theorem ker_id : (MonoidHom.id G).ker = ⊥ := rfl #align monoid_hom.ker_id MonoidHom.ker_id @@ -2795,13 +2795,13 @@ theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := #align monoid_hom.ker_eq_bot_iff MonoidHom.ker_eq_bot_iff #align add_monoid_hom.ker_eq_bot_iff AddMonoidHom.ker_eq_bot_iff -@[simp, to_additive] +@[to_additive (attr := simp)] theorem Subgroup.ker_subtype (H : Subgroup G) : H.Subtype.ker = ⊥ := H.Subtype.ker_eq_bot_iff.mpr Subtype.coe_injective #align subgroup.ker_subtype Subgroup.ker_subtype #align add_subgroup.ker_subtype AddSubgroup.ker_subtype -@[simp, to_additive] +@[to_additive (attr := simp)] theorem Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := (inclusion h).ker_eq_bot_iff.mpr (Set.inclusion_injective h) #align subgroup.ker_inclusion Subgroup.ker_inclusion @@ -2842,7 +2842,7 @@ def eqLocus (f g : G →* M) : Subgroup G := #align monoid_hom.eq_locus MonoidHom.eqLocus #align add_monoid_hom.eq_locus AddMonoidHom.eqLocus -@[simp, to_additive] +@[to_additive (attr := simp)] theorem eqLocus_same (f : G →* N) : f.eqLocus f = ⊤ := SetLike.ext fun _ => eq_self_iff_true _ #align monoid_hom.eq_locus_same MonoidHom.eqLocus_same @@ -3051,7 +3051,7 @@ theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) { #align subgroup.map_le_map_iff_of_injective Subgroup.map_le_map_iff_of_injective #align add_subgroup.map_le_map_iff_of_injective AddSubgroup.map_le_map_iff_of_injective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : H.map G'.Subtype ≤ K.map G'.Subtype ↔ H ≤ K := map_le_map_iff_of_injective Subtype.coe_injective @@ -3133,7 +3133,7 @@ noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Func #align subgroup.equiv_map_of_injective Subgroup.equivMapOfInjective #align add_subgroup.equiv_map_of_injective AddSubgroup.equivMapOfInjective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_equivMapOfInjective_apply (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) (h : H) : (equivMapOfInjective H f hf h : N) = f h := rfl @@ -3224,7 +3224,7 @@ def liftOfRightInverseAux (hf : Function.RightInverse f_inv f) (g : G₁ →* G #align monoid_hom.lift_of_right_inverse_aux MonoidHom.liftOfRightInverseAux #align add_monoid_hom.lift_of_right_inverse_aux AddMonoidHom.liftOfRightInverseAux -@[simp, to_additive] +@[to_additive (attr := simp)] theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) (x : G₁) : (f.liftOfRightInverseAux f_inv hf g hg) (f x) = g x := by dsimp [lift_of_right_inverse_aux] @@ -3280,7 +3280,7 @@ noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : #align monoid_hom.lift_of_surjective MonoidHom.liftOfSurjective #align add_monoid_hom.lift_of_surjective AddMonoidHom.liftOfSurjective -@[simp, to_additive] +@[to_additive (attr := simp)] theorem liftOfRightInverse_comp_apply (hf : Function.RightInverse f_inv f) (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) (x : G₁) : (f.liftOfRightInverse f_inv hf g) (f x) = g x := @@ -3288,7 +3288,7 @@ theorem liftOfRightInverse_comp_apply (hf : Function.RightInverse f_inv f) #align monoid_hom.lift_of_right_inverse_comp_apply MonoidHom.liftOfRightInverse_comp_apply #align add_monoid_hom.lift_of_right_inverse_comp_apply AddMonoidHom.lift_of_right_inverse_comp_apply -@[simp, to_additive] +@[to_additive (attr := simp)] theorem liftOfRightInverse_comp (hf : Function.RightInverse f_inv f) (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) : (f.liftOfRightInverse f_inv hf g).comp f = g := MonoidHom.ext <| f.lift_of_right_inverse_comp_apply f_inv hf g @@ -3384,14 +3384,14 @@ def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') #align mul_equiv.subgroup_map MulEquiv.subgroupMap #align add_equiv.add_subgroup_map AddEquiv.addSubgroupMap -@[simp, to_additive] +@[to_additive (attr := simp)] theorem coe_subgroupMap_apply (e : G ≃* G') (H : Subgroup G) (g : H) : ((subgroupMap e H g : H.map (e : G →* G')) : G') = e g := rfl #align mul_equiv.coe_subgroup_map_apply MulEquiv.coe_subgroupMap_apply #align add_equiv.coe_add_subgroup_map_apply AddEquiv.coe_add_subgroupMap_apply -@[simp, to_additive] +@[to_additive (attr := simp)] theorem subgroupMap_symm_apply (e : G ≃* G') (H : Subgroup G) (g : H.map (e : G →* G')) : (e.subgroupMap H).symm g = ⟨e.symm g, SetLike.mem_coe.1 <| Set.mem_image_equiv.1 g.2⟩ := rfl @@ -3402,7 +3402,7 @@ end MulEquiv namespace Subgroup -@[simp, to_additive] +@[to_additive (attr := simp)] theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : H.equivMapOfInjective (e : G →* G') (EquivLike.injective e) = e.subgroupMap H := by ext From 7738cf5ffdd96fa27ae6cb9087af7b1a53a45372 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Tue, 24 Jan 2023 23:38:42 +0100 Subject: [PATCH 07/34] small fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 54 +++++++++++++++---------- 1 file changed, 33 insertions(+), 21 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 93b3a6fac8187..89edd199596be 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -267,28 +267,28 @@ instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCo #align add_subgroup_class.to_linear_ordered_add_comm_group AddSubgroupClass.toLinearOrderedAddCommGroup /-- The natural group hom from a subgroup of group `G` to `G`. -/ -@[to_additive "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] +@[to_additive (attr := coe) "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] def subtype : H →* G := ⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩ #align subgroup_class.subtype SubgroupClass.subtype #align add_subgroup_class.subtype AddSubgroupClass.subtype @[to_additive (attr := simp)] -theorem coeSubtype : (subtype H : H → G) = coe := +theorem coeSubtype : (subtype H : H → G) = ((↑) : H → G) := by rfl #align subgroup_class.coe_subtype SubgroupClass.coeSubtype #align add_subgroup_class.coe_subtype AddSubgroupClass.coeSubtype variable {H} -@[to_additive (attr := simp, norm_cast) coe_smul] -theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := +@[to_additive (attr := simp, norm_cast)] +theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := (subtype H : H →* G).map_pow _ _ #align subgroup_class.coe_pow SubgroupClass.coe_pow #align add_subgroup_class.coe_smul AddSubgroupClass.coe_smul @[to_additive (attr := simp, norm_cast)] -theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := +theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := (subtype H : H →* G).map_zpow _ _ #align subgroup_class.coe_zpow SubgroupClass.coe_zpow #align add_subgroup_class.coe_zsmul AddSubgroupClass.coe_zsmul @@ -296,7 +296,7 @@ theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] def inclusion {H K : S} (h : H ≤ K) : H →* K := - MonoidHom.mk' (fun x => ⟨x, h x.prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl + MonoidHom.mk' (fun x => ⟨x, h x.prop⟩) fun _ _=> rfl #align subgroup_class.inclusion SubgroupClass.inclusion #align add_subgroup_class.inclusion AddSubgroupClass.inclusion @@ -360,7 +360,8 @@ structure AddSubgroup (G : Type _) [AddGroup G] extends AddSubmonoid G where attribute [to_additive] Subgroup -attribute [to_additive AddSubgroup.toAddSubmonoid] Subgroup.toSubmonoid +-- Porting note: Removed, translation already exists +-- attribute [to_additive AddSubgroup.toAddSubmonoid] Subgroup.toSubmonoid /-- Reinterpret a `subgroup` as a `submonoid`. -/ add_decl_doc Subgroup.toSubmonoid @@ -373,13 +374,17 @@ namespace Subgroup @[to_additive] instance : SetLike (Subgroup G) G where coe s := s.carrier - coe_injective' p q h := by cases p <;> cases q <;> congr + coe_injective' p q h := by + obtain ⟨⟨⟨hp,_⟩,_⟩,_⟩ := p + obtain ⟨⟨⟨hq,_⟩,_⟩,_⟩ := q + congr +-- Porting note: Below can probably be written more uniformly @[to_additive] instance : SubgroupClass (Subgroup G) G where - mul_mem := Subgroup.mul_mem' - one_mem := Subgroup.one_mem' inv_mem := Subgroup.inv_mem' _ + one_mem _ := (Subgroup.toSubmonoid _).one_mem' + mul_mem := (Subgroup.toSubmonoid _).mul_mem' @[to_additive (attr := simp)] theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := @@ -387,6 +392,7 @@ theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := #align subgroup.mem_carrier Subgroup.mem_carrier #align add_subgroup.mem_carrier AddSubgroup.mem_carrier +-- Porting note: Do we still want these _mk lemmas? I thought a lot of them are obsolete now @[to_additive (attr := simp)] theorem mem_mk {s : Set G} {x : G} (h_one) (h_mul) (h_inv) : x ∈ mk s h_one h_mul h_inv ↔ x ∈ s := Iff.rfl @@ -421,41 +427,47 @@ initialize_simps_projections AddSubgroup (carrier → coe) theorem coe_toSubmonoid (K : Subgroup G) : (K.toSubmonoid : Set G) = K := rfl #align subgroup.coe_to_submonoid Subgroup.coe_toSubmonoid -#align add_subgroup.coe_to_add_submonoid AddSubgroup.coe_to_add_submonoid +#align add_subgroup.coe_to_add_submonoid AddSubgroup.coe_toAddSubmonoid @[to_additive (attr := simp)] theorem mem_toSubmonoid (K : Subgroup G) (x : G) : x ∈ K.toSubmonoid ↔ x ∈ K := Iff.rfl #align subgroup.mem_to_submonoid Subgroup.mem_toSubmonoid -#align add_subgroup.mem_to_add_submonoid AddSubgroup.mem_to_add_submonoid +#align add_subgroup.mem_to_add_submonoid AddSubgroup.mem_toAddSubmonoid @[to_additive] theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subgroup G → Submonoid G) := fun p q h => SetLike.ext'_iff.2 (show _ from SetLike.ext'_iff.1 h) #align subgroup.to_submonoid_injective Subgroup.toSubmonoid_injective -#align add_subgroup.to_add_submonoid_injective AddSubgroup.to_add_submonoid_injective +#align add_subgroup.to_add_submonoid_injective AddSubgroup.toAddSubmonoid_injective @[to_additive (attr := simp)] theorem toSubmonoid_eq {p q : Subgroup G} : p.toSubmonoid = q.toSubmonoid ↔ p = q := toSubmonoid_injective.eq_iff #align subgroup.to_submonoid_eq Subgroup.toSubmonoid_eq -#align add_subgroup.to_add_submonoid_eq AddSubgroup.to_add_submonoid_eq +#align add_subgroup.to_add_submonoid_eq AddSubgroup.toAddSubmonoid_eq -@[to_additive, mono] +-- Porting note: Unknown attribute mono +--@[to_additive, mono] +@[to_additive] theorem toSubmonoid_strictMono : StrictMono (toSubmonoid : Subgroup G → Submonoid G) := fun _ _ => id #align subgroup.to_submonoid_strict_mono Subgroup.toSubmonoid_strictMono -#align add_subgroup.to_add_submonoid_strict_mono AddSubgroup.to_add_submonoid_strictMono +#align add_subgroup.to_add_submonoid_strict_mono AddSubgroup.toAddSubmonoid_strictMono -attribute [mono] AddSubgroup.to_add_submonoid_strictMono +-- Porting note: Unknown attribute mono +-- attribute [mono] AddSubgroup.to_add_submonoid_strictMono -@[to_additive, mono] +-- Porting note: Unknown attribute mono +--@[to_additive, mono] +@[to_additive] theorem toSubmonoid_mono : Monotone (toSubmonoid : Subgroup G → Submonoid G) := - toSubmonoid_strictMono.Monotone + toSubmonoid_strictMono.monotone #align subgroup.to_submonoid_mono Subgroup.toSubmonoid_mono -#align add_subgroup.to_add_submonoid_mono AddSubgroup.to_add_submonoid_mono +#align add_subgroup.to_add_submonoid_mono AddSubgroup.toAddSubmonoid_mono -attribute [mono] AddSubgroup.to_add_submonoid_mono +-- Porting note: Unknown attribute mono +-- attribute [mono] AddSubgroup.to_add_submonoid_mono @[to_additive (attr := simp)] theorem toSubmonoid_le {p q : Subgroup G} : p.toSubmonoid ≤ q.toSubmonoid ↔ p ≤ q := From 0b9029ce56ad071a5d81a4d2984ef36d5a21bcda Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Wed, 25 Jan 2023 05:46:22 +0000 Subject: [PATCH 08/34] Fixed comments and more. --- Mathlib/GroupTheory/Subgroup/Basic.lean | 390 +++++++++++++----------- 1 file changed, 205 insertions(+), 185 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 89edd199596be..74c9199debda7 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -50,25 +50,25 @@ Notation used here: Definitions in the file: -* `subgroup G` : the type of subgroups of a group `G` +* `Subgroup G` : the type of subgroups of a group `G` -* `add_subgroup A` : the type of subgroups of an additive group `A` +* `AddSubgroup A` : the type of subgroups of an additive group `A` -* `complete_lattice (subgroup G)` : the subgroups of `G` form a complete lattice +* `CompleteLattice (Subgroup G)` : the subgroups of `G` form a complete lattice -* `subgroup.closure k` : the minimal subgroup that includes the set `k` +* `Subgroup.closure k` : the minimal subgroup that includes the set `k` -* `subgroup.subtype` : the natural group homomorphism from a subgroup of group `G` to `G` +* `Subgroup.subtype` : the natural group homomorphism from a subgroup of group `G` to `G` -* `subgroup.gi` : `closure` forms a Galois insertion with the coercion to set +* `Subgroup.gi` : `closure` forms a Galois insertion with the coercion to set -* `subgroup.comap H f` : the preimage of a subgroup `H` along the group homomorphism `f` is also a +* `Subgroup.comap H f` : the preimage of a subgroup `H` along the group homomorphism `f` is also a subgroup -* `subgroup.map f H` : the image of a subgroup `H` along the group homomorphism `f` is also a +* `Subgroup.map f H` : the image of a subgroup `H` along the group homomorphism `f` is also a subgroup -* `subgroup.prod H K` : the product of subgroups `H`, `K` of groups `G`, `N` respectively, `H × K` +* `Subgroup.prod H K` : the product of subgroups `H`, `K` of groups `G`, `N` respectively, `H × K` is a subgroup of `G × N` * `monoid_hom.range f` : the range of the group homomorphism `f` is a subgroup @@ -98,26 +98,26 @@ variable {A : Type _} [AddGroup A] section SubgroupClass -/-- `inv_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ +/-- `InvMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ class InvMemClass (S G : Type _) [Inv G] [SetLike S G] : Prop where inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s #align inv_mem_class InvMemClass export InvMemClass (inv_mem) -/-- `neg_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ +/-- `NegMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ class NegMemClass (S G : Type _) [Neg G] [SetLike S G] : Prop where neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s #align neg_mem_class NegMemClass export NegMemClass (neg_mem) -/-- `subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ +/-- `SubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ class SubgroupClass (S G : Type _) [DivInvMonoid G] [SetLike S G] extends SubmonoidClass S G, InvMemClass S G : Prop #align subgroup_class SubgroupClass -/-- `add_subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are +/-- `AddSubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are additive subgroups of `G`. -/ class AddSubgroupClass (S G : Type _) [SubNegMonoid G] [SetLike S G] extends AddSubmonoidClass S G, NegMemClass S G : Prop @@ -226,7 +226,7 @@ theorem coe_div (x y : H) : (x / y).1 = x.1 / y.1 := variable (H) --- Prefer subclasses of `group` over subclasses of `subgroup_class`. +-- Prefer subclasses of `Group` over subclasses of `SubgroupClass`. /-- A subgroup of a group inherits a group structure. -/ @[to_additive "An additive subgroup of an `add_group` inherits an `add_group` structure."] instance (priority := 75) toGroup : Group H := @@ -363,10 +363,10 @@ attribute [to_additive] Subgroup -- Porting note: Removed, translation already exists -- attribute [to_additive AddSubgroup.toAddSubmonoid] Subgroup.toSubmonoid -/-- Reinterpret a `subgroup` as a `submonoid`. -/ +/-- Reinterpret a `Subgroup` as a `Submonoid`. -/ add_decl_doc Subgroup.toSubmonoid -/-- Reinterpret an `add_subgroup` as an `add_submonoid`. -/ +/-- Reinterpret an `AddSubgroup` as an `AddSubmonoid`. -/ add_decl_doc AddSubgroup.toAddSubmonoid namespace Subgroup @@ -495,12 +495,12 @@ def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) map_rel_iff' a b := Iff.rfl #align subgroup.to_add_subgroup Subgroup.toAddSubgroup -/-- Additive subgroup of an additive group `additive G` are isomorphic to subgroup of `G`. -/ +/-- Additive subgroup of an additive group `Additive G` are isomorphic to subgroup of `G`. -/ abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := Subgroup.toAddSubgroup.symm #align add_subgroup.to_subgroup' AddSubgroup.toSubgroup' -/-- Additive supgroups of an additive group `A` are isomorphic to subgroups of `multiplicative A`. +/-- Additive supgroups of an additive group `A` are isomorphic to subgroups of `Multiplicative A`. -/ @[simps] def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) @@ -512,7 +512,7 @@ def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) map_rel_iff' a b := Iff.rfl #align add_subgroup.to_subgroup AddSubgroup.toSubgroup -/-- Subgroups of an additive group `multiplicative A` are isomorphic to additive subgroups of `A`. +/-- Subgroups of an additive group `Multiplicative A` are isomorphic to additive subgroups of `A`. -/ abbrev Subgroup.toAddSubgroup' : Subgroup (Multiplicative A) ≃o AddSubgroup A := AddSubgroup.toSubgroup.symm @@ -550,14 +550,14 @@ theorem copy_eq (K : Subgroup G) (s : Set G) (hs : s = ↑K) : K.copy s hs = K : #align add_subgroup.copy_eq AddSubgroup.copy_eq /-- Two subgroups are equal if they have the same elements. -/ -@[ext, to_additive "Two `add_subgroup`s are equal if they have the same elements."] +@[ext, to_additive "Two `AddSubgroup`s are equal if they have the same elements."] theorem ext {H K : Subgroup G} (h : ∀ x, x ∈ H ↔ x ∈ K) : H = K := SetLike.ext h #align subgroup.ext Subgroup.ext #align add_subgroup.ext AddSubgroup.ext /-- A subgroup contains the group's 1. -/ -@[to_additive "An `add_subgroup` contains the group's 0."] +@[to_additive "An `AddSubgroup` contains the group's 0."] protected theorem one_mem : (1 : G) ∈ H := one_mem _ #align subgroup.one_mem Subgroup.one_mem @@ -627,7 +627,8 @@ protected theorem zpow_mem {x : G} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K := #align subgroup.zpow_mem Subgroup.zpow_mem #align add_subgroup.zsmul_mem AddSubgroup.zsmul_mem -/- ./././Mathport/Syntax/Translate/Basic.lean:632:2: warning: expanding binder collection (x y «expr ∈ » s) -/ +/- ./././Mathport/Syntax/Translate/Basic.lean:632:2: warning: + expanding binder collection (x y «expr ∈ » s) -/ /-- Construct a subgroup from a nonempty set that is closed under division. -/ @[to_additive "Construct a subgroup from a nonempty set that is closed under subtraction"] def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y ∈ s), x * y⁻¹ ∈ s) : @@ -671,7 +672,7 @@ instance hasDiv : Div H := #align subgroup.has_div Subgroup.hasDiv #align add_subgroup.has_sub AddSubgroup.hasSub -/-- An `add_subgroup` of an `add_group` inherits a natural scaling. -/ +/-- An `AddSubgroup` of an `AddGroup` inherits a natural scaling. -/ instance AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ #align add_subgroup.has_nsmul AddSubgroup.hasNsmul @@ -683,7 +684,7 @@ instance hasNpow : Pow H ℕ := #align subgroup.has_npow Subgroup.hasNpow #align add_subgroup.has_nsmul AddSubgroup.hasNsmul -/-- An `add_subgroup` of an `add_group` inherits an integer scaling. -/ +/-- An `AddSubgroup` of an `AddGroup` inherits an integer scaling. -/ instance AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := ⟨fun n a => ⟨n • a, H.zsmul_mem a.2 n⟩⟩ #align add_subgroup.has_zsmul AddSubgroup.hasZsmul @@ -751,7 +752,7 @@ instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := #align subgroup.to_group Subgroup.toGroup #align add_subgroup.to_add_group AddSubgroup.toAddGroup -/-- A subgroup of a `comm_group` is a `comm_group`. -/ +/-- A subgroup of a `CommGroup` is a `CommGroup`. -/ @[to_additive "An `add_subgroup` of an `add_comm_group` is an `add_comm_group`."] instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H := Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) @@ -759,8 +760,8 @@ instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H : #align subgroup.to_comm_group Subgroup.toCommGroup #align add_subgroup.to_add_comm_group AddSubgroup.toAddCommGroup -/-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/ -@[to_additive "An `add_subgroup` of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] +/-- A subgroup of an `OrderedCommGroup` is an `OrderedCommGroup`. -/ +@[to_additive "An `AddSubgroup` of an `AddOrderedCommGroup` is an `AddOrderedCommGroup`."] instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : OrderedCommGroup H := Subtype.coe_injective.OrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) @@ -768,9 +769,9 @@ instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : #align subgroup.to_ordered_comm_group Subgroup.toOrderedCommGroup #align add_subgroup.to_ordered_add_comm_group AddSubgroup.toOrderedAddCommGroup -/-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ +/-- A subgroup of a `LinearOrderedCommGroup` is a `LinearOrderedCommGroup`. -/ @[to_additive - "An `add_subgroup` of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] + "An `AddSubgroup` of a `LinearOrderedAddCommGroup` is a\n `LinearOrderedAddCommGroup`."] instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : Subgroup G) : LinearOrderedCommGroup H := Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) @@ -831,9 +832,9 @@ instance : Top (Subgroup G) := /-- The top subgroup is isomorphic to the group. -This is the group version of `submonoid.top_equiv`. -/ +This is the group version of `Submonoid.topEquiv`. -/ @[to_additive - "The top additive subgroup is isomorphic to the additive group.\n\nThis is the additive group version of `add_submonoid.top_equiv`.", + "The top additive subgroup is isomorphic to the additive group.\n\nThis is the additive group version of `AddSubmonoid.topEquiv`.", simps] def topEquiv : (⊤ : Subgroup G) ≃* G := Submonoid.topEquiv @@ -881,13 +882,13 @@ instance : Unique (⊥ : Subgroup G) := theorem top_toSubmonoid : (⊤ : Subgroup G).toSubmonoid = ⊤ := rfl #align subgroup.top_to_submonoid Subgroup.top_toSubmonoid -#align add_subgroup.top_to_add_submonoid AddSubgroup.top_to_add_submonoid +#align add_subgroup.top_to_add_submonoid AddSubgroup.top_toAddSubmonoid @[to_additive (attr := simp)] theorem bot_toSubmonoid : (⊥ : Subgroup G).toSubmonoid = ⊥ := rfl #align subgroup.bot_to_submonoid Subgroup.bot_toSubmonoid -#align add_subgroup.bot_to_add_submonoid AddSubgroup.bot_to_add_submonoid +#align add_subgroup.bot_to_add_submonoid AddSubgroup.bot_toAddSubmonoid @[to_additive] theorem eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := @@ -1078,7 +1079,7 @@ theorem eq_top_iff' : H = ⊤ ↔ ∀ x : G, x ∈ H := #align subgroup.eq_top_iff' Subgroup.eq_top_iff' #align add_subgroup.eq_top_iff' AddSubgroup.eq_top_iff' -/-- The `subgroup` generated by a set. -/ +/-- The `Subgroup` generated by a set. -/ @[to_additive "The `add_subgroup` generated by a set"] def closure (k : Set G) : Subgroup G := infₛ { K | k ⊆ K } @@ -1132,7 +1133,7 @@ theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x #align subgroup.closure_induction Subgroup.closure_induction #align add_subgroup.closure_induction AddSubgroup.closure_induction -/-- A dependent version of `subgroup.closure_induction`. -/ +/-- A dependent version of `Subgroup.closure_induction`. -/ @[elab_as_elim, to_additive "A dependent version of `add_subgroup.closure_induction`. "] theorem closure_induction' {p : ∀ x, x ∈ closure k → Prop} (Hs : ∀ (x) (h : x ∈ k), p x (subset_closure h)) (H1 : p 1 (one_mem _)) @@ -1265,7 +1266,7 @@ theorem supᵢ_eq_closure {ι : Sort _} (p : ι → Subgroup G) : /-- The subgroup generated by an element of a group equals the set of integer number powers of the element. -/ @[to_additive - "The `add_subgroup` generated by an element of an `add_group` equals the set of\nnatural number multiples of the element."] + "The `AddSubgroup` generated by an element of an `AddGroup` equals the set of\nnatural number multiples of the element."] theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := by refine' ⟨fun hy => closure_induction hy _ _ _ _, fun ⟨n, hn⟩ => @@ -1335,7 +1336,7 @@ variable {N : Type _} [Group N] {P : Type _} [Group P] /-- The preimage of a subgroup along a monoid homomorphism is a subgroup. -/ @[to_additive - "The preimage of an `add_subgroup` along an `add_monoid` homomorphism\nis an `add_subgroup`."] + "The preimage of an `AddSubgroup` along an `AddMonoid` homomorphism\nis an `AddSubgroup`."] def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := { H.toSubmonoid.comap f with carrier := f ⁻¹' H @@ -1377,7 +1378,7 @@ theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by /-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ @[to_additive - "The image of an `add_subgroup` along an `add_monoid` homomorphism\nis an `add_subgroup`."] + "The image of an `AddSubgroup` along an `AddMonoid` homomorphism\nis an `AddSubgroup`."] def map (f : G →* N) (H : Subgroup G) : Subgroup N := { H.toSubmonoid.map f with carrier := f '' H @@ -1407,7 +1408,7 @@ theorem mem_map_of_mem (f : G →* N) {K : Subgroup G} {x : G} (hx : x ∈ K) : @[to_additive] theorem apply_coe_mem_map (f : G →* N) (K : Subgroup G) (x : K) : f x ∈ K.map f := - mem_map_of_mem f x.Prop + mem_map_of_mem f x.prop #align subgroup.apply_coe_mem_map Subgroup.apply_coe_mem_map #align add_subgroup.apply_coe_mem_map AddSubgroup.apply_coe_mem_map @@ -1586,7 +1587,7 @@ def subgroupOfEquivOfLe {G : Type _} [Group G] {H K : Subgroup G} (h : H ≤ K) #align add_subgroup.add_subgroup_of_equiv_of_le AddSubgroup.addSubgroupOfEquivOfLe @[to_additive (attr := simp)] -theorem comap_subtype (H K : Subgroup G) : H.comap K.Subtype = H.subgroupOf K := +theorem comap_subtype (H K : Subgroup G) : H.comap K.subtype = H.subgroupOf K := rfl #align subgroup.comap_subtype Subgroup.comap_subtype #align add_subgroup.comap_subtype AddSubgroup.comap_subtype @@ -1596,90 +1597,90 @@ theorem comap_inclusion_subgroupOf {K₁ K₂ : Subgroup G} (h : K₁ ≤ K₂) (H.subgroupOf K₂).comap (inclusion h) = H.subgroupOf K₁ := rfl #align subgroup.comap_inclusion_subgroup_of Subgroup.comap_inclusion_subgroupOf -#align add_subgroup.comap_inclusion_add_subgroup_of AddSubgroup.comap_inclusion_add_subgroupOf +#align add_subgroup.comap_inclusion_add_subgroup_of AddSubgroup.comap_inclusion_addSubgroupOf @[to_additive] -theorem coe_subgroupOf (H K : Subgroup G) : (H.subgroupOf K : Set K) = K.Subtype ⁻¹' H := +theorem coe_subgroupOf (H K : Subgroup G) : (H.subgroupOf K : Set K) = K.subtype ⁻¹' H := rfl #align subgroup.coe_subgroup_of Subgroup.coe_subgroupOf -#align add_subgroup.coe_add_subgroup_of AddSubgroup.coe_add_subgroupOf +#align add_subgroup.coe_add_subgroup_of AddSubgroup.coe_addSubgroupOf @[to_additive] theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h : G) ∈ H := Iff.rfl #align subgroup.mem_subgroup_of Subgroup.mem_subgroupOf -#align add_subgroup.mem_add_subgroup_of AddSubgroup.mem_add_subgroupOf +#align add_subgroup.mem_add_subgroup_of AddSubgroup.mem_addSubgroupOf @[to_additive (attr := simp)] -theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.Subtype = H ⊓ K := +theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subtype = H ⊓ K := SetLike.ext' <| Subtype.image_preimage_coe _ _ #align subgroup.subgroup_of_map_subtype Subgroup.subgroupOf_map_subtype -#align add_subgroup.add_subgroup_of_map_subtype AddSubgroup.add_subgroupOf_map_subtype +#align add_subgroup.add_subgroup_of_map_subtype AddSubgroup.addSubgroupOf_map_subtype @[to_additive (attr := simp)] theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := Eq.symm (Subgroup.ext fun g => Subtype.ext_iff) #align subgroup.bot_subgroup_of Subgroup.bot_subgroupOf -#align add_subgroup.bot_add_subgroup_of AddSubgroup.bot_add_subgroupOf +#align add_subgroup.bot_add_subgroup_of AddSubgroup.bot_addSubgroupOf @[to_additive (attr := simp)] theorem top_subgroupOf : (⊤ : Subgroup G).subgroupOf H = ⊤ := rfl #align subgroup.top_subgroup_of Subgroup.top_subgroupOf -#align add_subgroup.top_add_subgroup_of AddSubgroup.top_add_subgroupOf +#align add_subgroup.top_add_subgroup_of AddSubgroup.top_addSubgroupOf @[to_additive] theorem subgroupOf_bot_eq_bot : H.subgroupOf ⊥ = ⊥ := Subsingleton.elim _ _ #align subgroup.subgroup_of_bot_eq_bot Subgroup.subgroupOf_bot_eq_bot -#align add_subgroup.add_subgroup_of_bot_eq_bot AddSubgroup.add_subgroupOf_bot_eq_bot +#align add_subgroup.add_subgroup_of_bot_eq_bot AddSubgroup.addSubgroupOf_bot_eq_bot @[to_additive] theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := Subsingleton.elim _ _ #align subgroup.subgroup_of_bot_eq_top Subgroup.subgroupOf_bot_eq_top -#align add_subgroup.add_subgroup_of_bot_eq_top AddSubgroup.add_subgroupOf_bot_eq_top +#align add_subgroup.add_subgroup_of_bot_eq_top AddSubgroup.addSubgroupOf_bot_eq_top @[to_additive (attr := simp)] theorem subgroupOf_self : H.subgroupOf H = ⊤ := top_unique fun g hg => g.2 #align subgroup.subgroup_of_self Subgroup.subgroupOf_self -#align add_subgroup.add_subgroup_of_self AddSubgroup.add_subgroupOf_self +#align add_subgroup.add_subgroup_of_self AddSubgroup.addSubgroupOf_self @[to_additive (attr := simp)] theorem subgroupOf_inj {H₁ H₂ K : Subgroup G} : H₁.subgroupOf K = H₂.subgroupOf K ↔ H₁ ⊓ K = H₂ ⊓ K := by - simpa only [SetLike.ext_iff, mem_inf, mem_subgroup_of, and_congr_left_iff] using Subtype.forall + simpa only [SetLike.ext_iff, mem_inf, mem_subgroupOf, and_congr_left_iff] using Subtype.forall #align subgroup.subgroup_of_inj Subgroup.subgroupOf_inj -#align add_subgroup.add_subgroup_of_inj AddSubgroup.add_subgroupOf_inj +#align add_subgroup.add_subgroup_of_inj AddSubgroup.addSubgroupOf_inj @[to_additive (attr := simp)] theorem inf_subgroupOf_right (H K : Subgroup G) : (H ⊓ K).subgroupOf K = H.subgroupOf K := subgroupOf_inj.2 inf_right_idem #align subgroup.inf_subgroup_of_right Subgroup.inf_subgroupOf_right -#align add_subgroup.inf_add_subgroup_of_right AddSubgroup.inf_add_subgroupOf_right +#align add_subgroup.inf_add_subgroup_of_right AddSubgroup.inf_addSubgroupOf_right @[to_additive (attr := simp)] theorem inf_subgroupOf_left (H K : Subgroup G) : (K ⊓ H).subgroupOf K = H.subgroupOf K := by - rw [inf_comm, inf_subgroup_of_right] + rw [inf_comm, inf_subgroupOf_right] #align subgroup.inf_subgroup_of_left Subgroup.inf_subgroupOf_left -#align add_subgroup.inf_add_subgroup_of_left AddSubgroup.inf_add_subgroupOf_left +#align add_subgroup.inf_add_subgroup_of_left AddSubgroup.inf_addSubgroupOf_left @[to_additive (attr := simp)] theorem subgroupOf_eq_bot {H K : Subgroup G} : H.subgroupOf K = ⊥ ↔ Disjoint H K := by - rw [disjoint_iff, ← bot_subgroup_of, subgroup_of_inj, bot_inf_eq] + rw [disjoint_iff, ← bot_subgroupOf, subgroupOf_inj, bot_inf_eq] #align subgroup.subgroup_of_eq_bot Subgroup.subgroupOf_eq_bot -#align add_subgroup.add_subgroup_of_eq_bot AddSubgroup.add_subgroup_of_eq_bot +#align add_subgroup.add_subgroup_of_eq_bot AddSubgroup.addSubgroupOf_eq_bot @[to_additive (attr := simp)] theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H := by - rw [← top_subgroup_of, subgroup_of_inj, top_inf_eq, inf_eq_right] + rw [← top_subgroupOf, subgroupOf_inj, top_inf_eq, inf_eq_right] #align subgroup.subgroup_of_eq_top Subgroup.subgroupOf_eq_top -#align add_subgroup.add_subgroup_of_eq_top AddSubgroup.add_subgroupOf_eq_top +#align add_subgroup.add_subgroup_of_eq_top AddSubgroup.addSubgroupOf_eq_top -/-- Given `subgroup`s `H`, `K` of groups `G`, `N` respectively, `H × K` as a subgroup of `G × N`. -/ -@[to_additive Prod - "Given `add_subgroup`s `H`, `K` of `add_group`s `A`, `B` respectively, `H × K`\nas an `add_subgroup` of `A × B`."] +/-- Given `Subgroup`s `H`, `K` of groups `G`, `N` respectively, `H × K` as a subgroup of `G × N`. -/ +@[to_additive prod + "Given `AddSubgroup`s `H`, `K` of `AddGroup`s `A`, `B` respectively, `H × K`\nas an `AddSubgroup` of `A × B`."] def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := { Submonoid.prod H.toSubmonoid K.toSubmonoid with inv_mem' := fun _ hx => ⟨H.inv_mem' hx.1, K.inv_mem' hx.2⟩ } @@ -1688,13 +1689,13 @@ def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[to_additive coe_prod] -theorem coe_prod (H : Subgroup G) (K : Subgroup N) : (H.Prod K : Set (G × N)) = H ×ˢ K := +theorem coe_prod (H : Subgroup G) (K : Subgroup N) : (H.prod K : Set (G × N)) = H ×ˢ K := rfl #align subgroup.coe_prod Subgroup.coe_prod #align add_subgroup.coe_prod AddSubgroup.coe_prod @[to_additive mem_prod] -theorem mem_prod {H : Subgroup G} {K : Subgroup N} {p : G × N} : p ∈ H.Prod K ↔ p.1 ∈ H ∧ p.2 ∈ K := +theorem mem_prod {H : Subgroup G} {K : Subgroup N} {p : G × N} : p ∈ H.prod K ↔ p.1 ∈ H ∧ p.2 ∈ K := Iff.rfl #align subgroup.mem_prod Subgroup.mem_prod #align add_subgroup.mem_prod AddSubgroup.mem_prod @@ -1706,65 +1707,65 @@ theorem prod_mono : ((· ≤ ·) ⇒ (· ≤ ·) ⇒ (· ≤ ·)) (@prod G _ N _ #align add_subgroup.prod_mono AddSubgroup.prod_mono @[to_additive prod_mono_right] -theorem prod_mono_right (K : Subgroup G) : Monotone fun t : Subgroup N => K.Prod t := +theorem prod_mono_right (K : Subgroup G) : Monotone fun t : Subgroup N => K.prod t := prod_mono (le_refl K) #align subgroup.prod_mono_right Subgroup.prod_mono_right #align add_subgroup.prod_mono_right AddSubgroup.prod_mono_right @[to_additive prod_mono_left] -theorem prod_mono_left (H : Subgroup N) : Monotone fun K : Subgroup G => K.Prod H := fun s₁ s₂ hs => +theorem prod_mono_left (H : Subgroup N) : Monotone fun K : Subgroup G => K.prod H := fun _ _ hs => prod_mono hs (le_refl H) #align subgroup.prod_mono_left Subgroup.prod_mono_left #align add_subgroup.prod_mono_left AddSubgroup.prod_mono_left @[to_additive prod_top] -theorem prod_top (K : Subgroup G) : K.Prod (⊤ : Subgroup N) = K.comap (MonoidHom.fst G N) := +theorem prod_top (K : Subgroup G) : K.prod (⊤ : Subgroup N) = K.comap (MonoidHom.fst G N) := ext fun x => by simp [mem_prod, MonoidHom.coe_fst] #align subgroup.prod_top Subgroup.prod_top #align add_subgroup.prod_top AddSubgroup.prod_top @[to_additive top_prod] -theorem top_prod (H : Subgroup N) : (⊤ : Subgroup G).Prod H = H.comap (MonoidHom.snd G N) := +theorem top_prod (H : Subgroup N) : (⊤ : Subgroup G).prod H = H.comap (MonoidHom.snd G N) := ext fun x => by simp [mem_prod, MonoidHom.coe_snd] #align subgroup.top_prod Subgroup.top_prod #align add_subgroup.top_prod AddSubgroup.top_prod -@[simp, to_additive top_prod_top] -theorem top_prod_top : (⊤ : Subgroup G).Prod (⊤ : Subgroup N) = ⊤ := +@[to_additive (attr := simp) top_prod_top] +theorem top_prod_top : (⊤ : Subgroup G).prod (⊤ : Subgroup N) = ⊤ := (top_prod _).trans <| comap_top _ #align subgroup.top_prod_top Subgroup.top_prod_top #align add_subgroup.top_prod_top AddSubgroup.top_prod_top @[to_additive] -theorem bot_prod_bot : (⊥ : Subgroup G).Prod (⊥ : Subgroup N) = ⊥ := +theorem bot_prod_bot : (⊥ : Subgroup G).prod (⊥ : Subgroup N) = ⊥ := SetLike.coe_injective <| by simp [coe_prod, Prod.one_eq_mk] #align subgroup.bot_prod_bot Subgroup.bot_prod_bot #align add_subgroup.bot_sum_bot AddSubgroup.bot_sum_bot @[to_additive le_prod_iff] theorem le_prod_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : - J ≤ H.Prod K ↔ map (MonoidHom.fst G N) J ≤ H ∧ map (MonoidHom.snd G N) J ≤ K := by + J ≤ H.prod K ↔ map (MonoidHom.fst G N) J ≤ H ∧ map (MonoidHom.snd G N) J ≤ K := by simpa only [← Subgroup.toSubmonoid_le] using Submonoid.le_prod_iff #align subgroup.le_prod_iff Subgroup.le_prod_iff #align add_subgroup.le_prod_iff AddSubgroup.le_prod_iff @[to_additive prod_le_iff] theorem prod_le_iff {H : Subgroup G} {K : Subgroup N} {J : Subgroup (G × N)} : - H.Prod K ≤ J ↔ map (MonoidHom.inl G N) H ≤ J ∧ map (MonoidHom.inr G N) K ≤ J := by + H.prod K ≤ J ↔ map (MonoidHom.inl G N) H ≤ J ∧ map (MonoidHom.inr G N) K ≤ J := by simpa only [← Subgroup.toSubmonoid_le] using Submonoid.prod_le_iff #align subgroup.prod_le_iff Subgroup.prod_le_iff #align add_subgroup.prod_le_iff AddSubgroup.prod_le_iff -@[simp, to_additive prod_eq_bot_iff] -theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.Prod K = ⊥ ↔ H = ⊥ ∧ K = ⊥ := by +@[to_additive (attr := simp) prod_eq_bot_iff] +theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.prod K = ⊥ ↔ H = ⊥ ∧ K = ⊥ := by simpa only [← Subgroup.toSubmonoid_eq] using Submonoid.prod_eq_bot_iff #align subgroup.prod_eq_bot_iff Subgroup.prod_eq_bot_iff #align add_subgroup.prod_eq_bot_iff AddSubgroup.prod_eq_bot_iff /-- Product of subgroups is isomorphic to their product as groups. -/ -@[to_additive prod_equiv +@[to_additive prodEquiv "Product of additive subgroups is isomorphic to their product\nas additive groups"] -def prodEquiv (H : Subgroup G) (K : Subgroup N) : H.Prod K ≃* H × K := +def prodEquiv (H : Subgroup G) (K : Subgroup N) : H.prod K ≃* H × K := { Equiv.Set.prod ↑H ↑K with map_mul' := fun x y => rfl } #align subgroup.prod_equiv Subgroup.prodEquiv #align add_subgroup.prod_equiv AddSubgroup.prodEquiv @@ -1774,11 +1775,11 @@ section Pi variable {η : Type _} {f : η → Type _} -- defined here and not in group_theory.submonoid.operations to have access to algebra.group.pi -/-- A version of `set.pi` for submonoids. Given an index set `I` and a family of submodules -`s : Π i, submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that -`f i` belongs to `pi I s` whenever `i ∈ I`. -/ +/-- A version of `Set.pi` for submonoids. Given an index set `I` and a family of submodules +`s : Π i, Submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that +`f i` belongs to `Pi I s` whenever `i ∈ I`. -/ @[to_additive - " A version of `set.pi` for `add_submonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, add_submonoid f i`, `pi I s` is the `add_submonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] + " A version of `Set.pi` for `add_submonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : Submonoid (∀ i, f i) where carrier := I.pi fun i => (s i).carrier @@ -1789,11 +1790,11 @@ def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid ( variable [∀ i, Group (f i)] -/-- A version of `set.pi` for subgroups. Given an index set `I` and a family of submodules -`s : Π i, subgroup f i`, `pi I s` is the subgroup of dependent functions `f : Π i, f i` such that +/-- A version of `Set.pi` for subgroups. Given an index set `I` and a family of submodules +`s : Π i, Subgroup f i`, `pi I s` is the subgroup of dependent functions `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ @[to_additive - " A version of `set.pi` for `add_subgroup`s. Given an index set `I` and a family\nof submodules `s : Π i, add_subgroup f i`, `pi I s` is the `add_subgroup` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] + " A version of `Set.pi` for `AddSubgroup`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubgroup f i`, `pi I s` is the `AddSubgroup` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def pi (I : Set η) (H : ∀ i, Subgroup (f i)) : Subgroup (∀ i, f i) := { Submonoid.pi I fun i => (H i).toSubmonoid with inv_mem' := fun p hp i hI => (H i).inv_mem (hp i hI) } @@ -1882,7 +1883,7 @@ structure Normal : Prop where conj_mem : ∀ n, n ∈ H → ∀ g : G, g * n * g⁻¹ ∈ H #align subgroup.normal Subgroup.Normal -attribute [class] normal +attribute [class] Normal end Subgroup @@ -1894,9 +1895,9 @@ structure Normal (H : AddSubgroup A) : Prop where conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H #align add_subgroup.normal AddSubgroup.Normal -attribute [to_additive AddSubgroup.Normal] Subgroup.Normal +attribute [to_additive] Subgroup.Normal -attribute [class] normal +attribute [class] Normal end AddSubgroup @@ -1932,12 +1933,12 @@ end Normal variable (H) /-- A subgroup is characteristic if it is fixed by all automorphisms. - Several equivalent conditions are provided by lemmas of the form `characteristic.iff...` -/ + Several equivalent conditions are provided by lemmas of the form `Characteristic.iff...` -/ structure Characteristic : Prop where fixed : ∀ ϕ : G ≃* G, H.comap ϕ.toMonoidHom = H #align subgroup.characteristic Subgroup.Characteristic -attribute [class] characteristic +attribute [class] Characteristic instance (priority := 100) normal_of_characteristic [h : H.Characteristic] : H.Normal := ⟨fun a ha b => (SetLike.ext_iff.mp (h.fixed (MulAut.conj b)) a).mpr ha⟩ @@ -1950,14 +1951,14 @@ namespace AddSubgroup variable (H : AddSubgroup A) /-- A add_subgroup is characteristic if it is fixed by all automorphisms. - Several equivalent conditions are provided by lemmas of the form `characteristic.iff...` -/ + Several equivalent conditions are provided by lemmas of the form `Characteristic.iff...` -/ structure Characteristic : Prop where fixed : ∀ ϕ : A ≃+ A, H.comap ϕ.toAddMonoidHom = H #align add_subgroup.characteristic AddSubgroup.Characteristic attribute [to_additive AddSubgroup.Characteristic] Subgroup.Characteristic -attribute [class] characteristic +attribute [class] Characteristic instance (priority := 100) normal_of_characteristic [h : H.Characteristic] : H.Normal := ⟨fun a ha b => (SetLike.ext_iff.mp (h.fixed (AddAut.conj b)) a).mpr ha⟩ @@ -2016,13 +2017,13 @@ theorem characteristic_iff_le_map : H.Characteristic ↔ ∀ ϕ : G ≃* G, H instance botCharacteristic : Characteristic (⊥ : Subgroup G) := characteristic_iff_le_map.mpr fun ϕ => bot_le #align subgroup.bot_characteristic Subgroup.botCharacteristic -#align add_subgroup.bot_characteristic AddSubgroup.bot_characteristic +#align add_subgroup.bot_characteristic AddSubgroup.botCharacteristic @[to_additive] instance topCharacteristic : Characteristic (⊤ : Subgroup G) := characteristic_iff_map_le.mpr fun ϕ => le_top #align subgroup.top_characteristic Subgroup.topCharacteristic -#align add_subgroup.top_characteristic AddSubgroup.top_characteristic +#align add_subgroup.top_characteristic AddSubgroup.topCharacteristic variable (G) @@ -2046,7 +2047,7 @@ theorem coe_center : ↑(center G) = Set.center G := theorem center_toSubmonoid : (center G).toSubmonoid = Submonoid.center G := rfl #align subgroup.center_to_submonoid Subgroup.center_toSubmonoid -#align add_subgroup.center_to_add_submonoid AddSubgroup.center_to_add_submonoid +#align add_subgroup.center_to_add_submonoid AddSubgroup.center_toAddSubmonoid variable {G} @@ -2066,7 +2067,7 @@ instance centerCharacteristic : (center G).Characteristic := by rw [← ϕ.injective.eq_iff, ϕ.map_mul, ϕ.map_mul] exact hg (ϕ h) #align subgroup.center_characteristic Subgroup.centerCharacteristic -#align add_subgroup.center_characteristic AddSubgroup.center_characteristic +#align add_subgroup.center_characteristic AddSubgroup.centerCharacteristic theorem CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := by rw [eq_top_iff'] @@ -2106,9 +2107,9 @@ def normalizer : Subgroup G -- variant for sets. -- TODO should this replace `normalizer`? -/-- The `set_normalizer` of `S` is the subgroup of `G` whose elements satisfy `g*S*g⁻¹=S` -/ +/-- The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy `g*S*g⁻¹=S` -/ @[to_additive - "The `set_normalizer` of `S` is the subgroup of `G` whose elements satisfy\n`g+S-g=S`."] + "The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy\n`g+S-g=S`."] def setNormalizer (S : Set G) : Subgroup G where carrier := { g : G | ∀ n, n ∈ S ↔ g * n * g⁻¹ ∈ S } @@ -2263,7 +2264,9 @@ theorem mem_centralizer_iff_commutator_eq_one {g : G} : g ∈ H.centralizer ↔ ∀ h ∈ H, h * g * h⁻¹ * g⁻¹ = 1 := by simp only [mem_centralizer_iff, mul_inv_eq_iff_eq_mul, one_mul] #align subgroup.mem_centralizer_iff_commutator_eq_one Subgroup.mem_centralizer_iff_commutator_eq_one -#align add_subgroup.mem_centralizer_iff_commutator_eq_zero AddSubgroup.mem_centralizer_iff_commutator_eq_zero +#align + add_subgroup.mem_centralizer_iff_commutator_eq_zero + AddSubgroup.mem_centralizer_iff_commutator_eq_zero @[to_additive] theorem centralizer_top : centralizer ⊤ = center G := @@ -2306,7 +2309,7 @@ structure AddSubgroup.IsCommutative (H : AddSubgroup A) : Prop where is_comm : IsCommutative H (· + ·) #align add_subgroup.is_commutative AddSubgroup.IsCommutative -attribute [to_additive AddSubgroup.IsCommutative] Subgroup.IsCommutative +attribute [to_additive] Subgroup.IsCommutative attribute [class] AddSubgroup.IsCommutative @@ -2367,7 +2370,7 @@ namespace Group variable {s : Set G} -/-- Given a set `s`, `conjugates_of_set s` is the set of all conjugates of +/-- Given a set `s`, `conjugatesOfSet s` is the set of all conjugates of the elements of `s`. -/ def conjugatesOfSet (s : Set G) : Set G := ⋃ a ∈ s, conjugatesOf a @@ -2420,7 +2423,9 @@ def normalClosure (s : Set G) : Subgroup G := theorem conjugatesOfSet_subset_normalClosure : conjugatesOfSet s ⊆ normalClosure s := subset_closure -#align subgroup.conjugates_of_set_subset_normal_closure Subgroup.conjugatesOfSet_subset_normalClosure +#align + subgroup.conjugates_of_set_subset_normal_closure + Subgroup.conjugatesOfSet_subset_normalClosure theorem subset_normalClosure : s ⊆ normalClosure s := Set.Subset.trans subset_conjugatesOfSet conjugatesOfSet_subset_normalClosure @@ -2435,8 +2440,8 @@ instance normalClosure_normal : (normalClosure s).Normal := ⟨fun n h g => by refine' Subgroup.closure_induction h (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ - · exact conjugates_of_set_subset_normal_closure (conj_mem_conjugates_of_set hx) - · simpa using (normal_closure s).one_mem + · exact conjugatesOfSet_subset_normalClosure (conj_mem_conjugates_of_set hx) + · simpa using (normalClosure s).one_mem · rw [← conj_mul] exact mul_mem ihx ihy · rw [← conj_inv] @@ -2447,7 +2452,7 @@ instance normalClosure_normal : (normalClosure s).Normal := theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : normalClosure s ≤ N := by intro a w refine' closure_induction w (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ - · exact conjugates_of_set_subset h hx + · exact conjugatesOfSet_subset h hx · exact one_mem _ · exact mul_mem ihx ihy · exact inv_mem ihx @@ -2463,14 +2468,14 @@ theorem normalClosure_mono {s t : Set G} (h : s ⊆ t) : normalClosure s ≤ nor theorem normalClosure_eq_infᵢ : normalClosure s = ⨅ (N : Subgroup G) (_ : Normal N) (hs : s ⊆ N), N := - le_antisymm (le_infᵢ fun N => le_infᵢ fun hN => le_infᵢ normal_closure_le_normal) + le_antisymm (le_infᵢ fun N => le_infᵢ fun hN => le_infᵢ normalClosure_le_normal) (infᵢ_le_of_le (normalClosure s) (infᵢ_le_of_le (by infer_instance) (infᵢ_le_of_le subset_normalClosure le_rfl))) #align subgroup.normal_closure_eq_infi Subgroup.normalClosure_eq_infᵢ @[simp] theorem normalClosure_eq_self (H : Subgroup G) [H.Normal] : normalClosure ↑H = H := - le_antisymm (normalClosure_le_normal rfl.Subset) le_normalClosure + le_antisymm (normalClosure_le_normal rfl.subset) le_normalClosure #align subgroup.normal_closure_eq_self Subgroup.normalClosure_eq_self @[simp] @@ -2479,7 +2484,7 @@ theorem normalClosure_idempotent : normalClosure ↑(normalClosure s) = normalCl #align subgroup.normal_closure_idempotent Subgroup.normalClosure_idempotent theorem closure_le_normalClosure {s : Set G} : closure s ≤ normalClosure s := by - simp only [subset_normal_closure, closure_le] + simp only [subset_normalClosure, closure_le] #align subgroup.closure_le_normal_closure Subgroup.closure_le_normalClosure @[simp] @@ -2489,7 +2494,7 @@ theorem normalClosure_closure_eq_normalClosure {s : Set G} : #align subgroup.normal_closure_closure_eq_normal_closure Subgroup.normalClosure_closure_eq_normalClosure /-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, -as shown by `subgroup.normal_core_eq_supr`. -/ +as shown by `subgroup.normalCore_eq_supr`. -/ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } @@ -2521,18 +2526,18 @@ theorem normalCore_eq_supᵢ (H : Subgroup G) : H.normalCore = ⨆ (N : Subgroup G) (_ : Normal N) (hs : N ≤ H), N := le_antisymm (le_supᵢ_of_le H.normalCore - (le_supᵢ_of_le H.normal_core_normal (le_supᵢ_of_le H.normal_core_le le_rfl))) - (supᵢ_le fun N => supᵢ_le fun hN => supᵢ_le normal_le_normal_core.mpr) + (le_supᵢ_of_le H.normalCore_normal (le_supᵢ_of_le H.normalCore_le le_rfl))) + (supᵢ_le fun N => supᵢ_le fun hN => supᵢ_le normal_le_normalCore.mpr) #align subgroup.normal_core_eq_supr Subgroup.normalCore_eq_supᵢ @[simp] theorem normalCore_eq_self (H : Subgroup G) [H.Normal] : H.normalCore = H := - le_antisymm H.normal_core_le (normal_le_normalCore.mpr le_rfl) + le_antisymm H.normalCore_le (normal_le_normalCore.mpr le_rfl) #align subgroup.normal_core_eq_self Subgroup.normalCore_eq_self @[simp] theorem normalCore_idempotent (H : Subgroup G) : H.normalCore.normalCore = H.normalCore := - H.normalCore.normal_core_eq_self + H.normalCore.normalCore_eq_self #align subgroup.normal_core_idempotent Subgroup.normalCore_idempotent end Subgroup @@ -2577,33 +2582,33 @@ theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by /-- The canonical surjective group homomorphism `G →* f(G)` induced by a group homomorphism `G →* N`. -/ @[to_additive - "The canonical surjective `add_group` homomorphism `G →+ f(G)` induced by a group\nhomomorphism `G →+ N`."] + "The canonical surjective `AddGroup` homomorphism `G →+ f(G)` induced by a group\nhomomorphism `G →+ N`."] def rangeRestrict (f : G →* N) : G →* f.range := codRestrict f _ fun x => ⟨x, rfl⟩ #align monoid_hom.range_restrict MonoidHom.rangeRestrict #align add_monoid_hom.range_restrict AddMonoidHom.rangeRestrict @[to_additive (attr := simp)] -theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.range_restrict g : N) = f g := +theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.rangeRestrict g : N) = f g := rfl #align monoid_hom.coe_range_restrict MonoidHom.coe_rangeRestrict #align add_monoid_hom.coe_range_restrict AddMonoidHom.coe_rangeRestrict @[to_additive] theorem coe_comp_rangeRestrict (f : G →* N) : - (coe : f.range → N) ∘ (⇑f.range_restrict : G → f.range) = f := + (coe : f.range → N) ∘ (⇑f.rangeRestrict : G → f.range) = f := rfl #align monoid_hom.coe_comp_range_restrict MonoidHom.coe_comp_rangeRestrict #align add_monoid_hom.coe_comp_range_restrict AddMonoidHom.coe_comp_rangeRestrict @[to_additive] -theorem subtype_comp_rangeRestrict (f : G →* N) : f.range.Subtype.comp f.range_restrict = f := - ext <| f.coe_range_restrict +theorem subtype_comp_rangeRestrict (f : G →* N) : f.range.subtype.comp f.rangeRestrict = f := + ext <| f.coe_rangeRestrict #align monoid_hom.subtype_comp_range_restrict MonoidHom.subtype_comp_rangeRestrict #align add_monoid_hom.subtype_comp_range_restrict AddMonoidHom.subtype_comp_rangeRestrict @[to_additive] -theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.range_restrict := +theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.rangeRestrict := fun ⟨_, g, rfl⟩ => ⟨g, rfl⟩ #align monoid_hom.range_restrict_surjective MonoidHom.rangeRestrict_surjective #align add_monoid_hom.range_restrict_surjective AddMonoidHom.rangeRestrict_surjective @@ -2636,7 +2641,7 @@ theorem range_one : (1 : G →* N).range = ⊥ := #align add_monoid_hom.range_zero AddMonoidHom.range_zero @[to_additive (attr := simp)] -theorem Subgroup.subtype_range (H : Subgroup G) : H.Subtype.range = H := by +theorem Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] ext simp @@ -2658,10 +2663,10 @@ theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type _} [Group G₁] [Group G₂] refine' exists_congr _ simp [Subtype.ext_iff] #align monoid_hom.subgroup_of_range_eq_of_le MonoidHom.subgroupOf_range_eq_of_le -#align add_monoid_hom.add_subgroup_of_range_eq_of_le AddMonoidHom.add_subgroup_of_range_eq_of_le +#align add_monoid_hom.add_subgroup_of_range_eq_of_le AddMonoidHom.addSubgroupOf_range_eq_of_le -/-- Computable alternative to `monoid_hom.of_injective`. -/ -@[to_additive "Computable alternative to `add_monoid_hom.of_injective`."] +/-- Computable alternative to `MonoidHom.ofInjective`. -/ +@[to_additive "Computable alternative to `AddMonoidHom.ofInjective`."] def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : G ≃* f.range := { f.range_restrict with toFun := f.range_restrict @@ -2713,7 +2718,7 @@ variable {M : Type _} [MulOneClass M] /-- The multiplicative kernel of a monoid homomorphism is the subgroup of elements `x : G` such that `f x = 1` -/ @[to_additive - "The additive kernel of an `add_monoid` homomorphism is the `add_subgroup` of elements\nsuch that `f x = 0`"] + "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements\nsuch that `f x = 0`"] def ker (f : G →* M) : Subgroup G := { f.mker with inv_mem' := fun x (hx : f x = 1) => @@ -2741,7 +2746,7 @@ theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker ext x simp [mem_ker, Units.ext_iff] #align monoid_hom.ker_to_hom_units MonoidHom.ker_toHomUnits -#align add_monoid_hom.ker_to_hom_add_units AddMonoidHom.ker_to_hom_add_units +#align add_monoid_hom.ker_to_hom_add_units AddMonoidHom.ker_toHomAddUnits @[to_additive] theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := by @@ -2808,7 +2813,7 @@ theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := #align add_monoid_hom.ker_eq_bot_iff AddMonoidHom.ker_eq_bot_iff @[to_additive (attr := simp)] -theorem Subgroup.ker_subtype (H : Subgroup G) : H.Subtype.ker = ⊥ := +theorem Subgroup.ker_subtype (H : Subgroup G) : H.subtype.ker = ⊥ := H.Subtype.ker_eq_bot_iff.mpr Subtype.coe_injective #align subgroup.ker_subtype Subgroup.ker_subtype #align add_subgroup.ker_subtype AddSubgroup.ker_subtype @@ -2822,17 +2827,17 @@ theorem Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h). @[to_additive] theorem prodMap_comap_prod {G' : Type _} {N' : Type _} [Group G'] [Group N'] (f : G →* N) (g : G' →* N') (S : Subgroup N) (S' : Subgroup N') : - (S.Prod S').comap (prodMap f g) = (S.comap f).Prod (S'.comap g) := + (S.prod S').comap (prodMap f g) = (S.comap f).prod (S'.comap g) := SetLike.coe_injective <| Set.preimage_prod_map_prod f g _ _ #align monoid_hom.prod_map_comap_prod MonoidHom.prodMap_comap_prod -#align add_monoid_hom.sum_map_comap_sum AddMonoidHom.sum_map_comap_sum +#align add_monoid_hom.sum_map_comap_sum AddMonoidHom.sumMap_comap_sum @[to_additive] theorem ker_prodMap {G' : Type _} {N' : Type _} [Group G'] [Group N'] (f : G →* N) (g : G' →* N') : - (prodMap f g).ker = f.ker.Prod g.ker := by - rw [← comap_bot, ← comap_bot, ← comap_bot, ← prod_map_comap_prod, bot_prod_bot] + (prodMap f g).ker = f.ker.prod g.ker := by + rw [← comap_bot, ← comap_bot, ← comap_bot, ← prodMap_comap_prod, bot_prod_bot] #align monoid_hom.ker_prod_map MonoidHom.ker_prodMap -#align add_monoid_hom.ker_sum_map AddMonoidHom.ker_sum_map +#align add_monoid_hom.ker_sum_map AddMonoidHom.ker_sumMap @[to_additive] instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := @@ -2876,7 +2881,7 @@ theorem eq_of_eqOn_top {f g : G →* M} (h : Set.EqOn f g (⊤ : Subgroup G)) : @[to_additive] theorem eq_of_eqOn_dense {s : Set G} (hs : closure s = ⊤) {f g : G →* M} (h : s.EqOn f g) : f = g := - eq_of_eq_on_top <| hs ▸ eqOn_closure h + eq_of_eqOn_top <| hs ▸ eqOn_closure h #align monoid_hom.eq_of_eq_on_dense MonoidHom.eq_of_eqOn_dense #align add_monoid_hom.eq_of_eq_on_dense AddMonoidHom.eq_of_eqOn_dense @@ -2891,7 +2896,7 @@ theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) /-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup generated by the image of the set. -/ @[to_additive - "The image under an `add_monoid` hom of the `add_subgroup` generated by a set equals\nthe `add_subgroup` generated by the image of the set."] + "The image under an `add_monoid` hom of the `AddSubgroup` generated by a set equals\nthe `AddSubgroup` generated by the image of the set."] theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := Set.image_preimage.l_comm_of_u_comm (Subgroup.gc_map_comap f) (Subgroup.gi N).gc (Subgroup.gi G).gc fun t => rfl @@ -2940,8 +2945,8 @@ theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := #align add_subgroup.map_le_range AddSubgroup.map_le_range @[to_additive] -theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.Subtype ≤ H := - (K.map_le_range H.Subtype).trans (le_of_eq H.subtype_range) +theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.subtype ≤ H := + (K.map_le_range H.subtype).trans (le_of_eq H.subtype_range) #align subgroup.map_subtype_le Subgroup.map_subtype_le #align add_subgroup.map_subtype_le AddSubgroup.map_subtype_le @@ -3093,7 +3098,7 @@ theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ke #align add_subgroup.map_injective_of_ker_le AddSubgroup.map_injective_of_ker_le @[to_additive] -theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).Subtype ⁻¹' s) = ⊤ := by +theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by apply map_injective (closure s).subtype_injective rwa [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, Set.image_preimage_eq_of_subset] @@ -3136,9 +3141,9 @@ theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : #align add_subgroup.codisjoint_add_subgroup_of_sup AddSubgroup.codisjoint_add_subgroup_of_sup /-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, -use `mul_equiv.subgroup_map` for better definitional equalities. -/ +use `MulEquiv.subgroupMap` for better definitional equalities. -/ @[to_additive - "An additive subgroup is isomorphic to its image under an injective function. If you\nhave an isomorphism, use `add_equiv.add_subgroup_map` for better definitional equalities."] + "An additive subgroup is isomorphic to its image under an injective function. If you\nhave an isomorphism, use `AddEquiv.addSubgroupMap` for better definitional equalities."] noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) : H ≃* H.map f := { Equiv.Set.image f H hf with map_mul' := fun _ _ => Subtype.ext (f.map_mul _ _) } @@ -3179,8 +3184,12 @@ theorem comap_normalizer_eq_of_injective_of_le_range {N : Type _} [Group N] (H : rw [map_comap_eq_self h] · refine' le_trans (le_normalizer_map f) (le_of_eq _) rw [map_comap_eq_self (le_trans le_normalizer h)] -#align subgroup.comap_normalizer_eq_of_injective_of_le_range Subgroup.comap_normalizer_eq_of_injective_of_le_range -#align add_subgroup.comap_normalizer_eq_of_injective_of_le_range AddSubgroup.comap_normalizer_eq_of_injective_of_le_range +#align + subgroup.comap_normalizer_eq_of_injective_of_le_range + Subgroup.comap_normalizer_eq_of_injective_of_le_range +#align + add_subgroup.comap_normalizer_eq_of_injective_of_le_range + AddSubgroup.comap_normalizer_eq_of_injective_of_le_range @[to_additive] theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : @@ -3189,7 +3198,7 @@ theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : exact Subtype.coe_injective simpa #align subgroup.subgroup_of_normalizer_eq Subgroup.subgroupOf_normalizer_eq -#align add_subgroup.add_subgroup_of_normalizer_eq AddSubgroup.add_subgroupOf_normalizer_eq +#align add_subgroup.add_subgroup_of_normalizer_eq AddSubgroup.addSubgroupOf_normalizer_eq /-- The image of the normalizer is equal to the normalizer of the image of an isomorphism. -/ @[to_additive @@ -3221,8 +3230,8 @@ variable {G₁ G₂ G₃ : Type _} [Group G₁] [Group G₂] [Group G₃] variable (f : G₁ →* G₂) (f_inv : G₂ → G₁) -/-- Auxiliary definition used to define `lift_of_right_inverse` -/ -@[to_additive "Auxiliary definition used to define `lift_of_right_inverse`"] +/-- Auxiliary definition used to define `liftOfRightInverse` -/ +@[to_additive "Auxiliary definition used to define `liftOfRightInverse`"] def liftOfRightInverseAux (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) : G₂ →* G₃ where toFun b := g (f_inv b) @@ -3244,16 +3253,20 @@ theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g apply hg rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one] simp only [hf _] -#align monoid_hom.lift_of_right_inverse_aux_comp_apply MonoidHom.liftOfRightInverseAux_comp_apply -#align add_monoid_hom.lift_of_right_inverse_aux_comp_apply AddMonoidHom.liftOfRightInverseAux_comp_apply +#align + monoid_hom.lift_of_right_inverse_aux_comp_apply + MonoidHom.liftOfRightInverseAux_comp_apply +#align + add_monoid_hom.lift_of_right_inverse_aux_comp_apply + AddMonoidHom.liftOfRightInverseAux_comp_apply -/-- `lift_of_right_inverse f hf g hg` is the unique group homomorphism `φ` +/-- `liftOfRightInverse f hf g hg` is the unique group homomorphism `φ` -* such that `φ.comp f = g` (`monoid_hom.lift_of_right_inverse_comp`), -* where `f : G₁ →+* G₂` has a right_inverse `f_inv` (`hf`), +* such that `φ.comp f = g` (`monoid_hom.liftOfRightInverse_comp`), +* where `f : G₁ →+* G₂` has a RightInverse `f_inv` (`hf`), * and `g : G₂ →+* G₃` satisfies `hg : f.ker ≤ g.ker`. -See `monoid_hom.eq_lift_of_right_inverse` for the uniqueness lemma. +See `MonoidHom.eq_liftOfRightInverse` for the uniqueness lemma. ``` G₁. @@ -3266,7 +3279,7 @@ See `monoid_hom.eq_lift_of_right_inverse` for the uniqueness lemma. ``` -/ @[to_additive - "`lift_of_right_inverse f f_inv hf g hg` is the unique additive group homomorphism `φ`\n\n* such that `φ.comp f = g` (`add_monoid_hom.lift_of_right_inverse_comp`),\n* where `f : G₁ →+ G₂` has a right_inverse `f_inv` (`hf`),\n* and `g : G₂ →+ G₃` satisfies `hg : f.ker ≤ g.ker`.\n\nSee `add_monoid_hom.eq_lift_of_right_inverse` for the uniqueness lemma.\n\n```\n G₁.\n | \\\n f | \\ g\n | \\\n v \\⌟\n G₂----> G₃\n ∃!φ\n```"] + "`liftOfRightInverse f f_inv hf g hg` is the unique additive group homomorphism `φ`\n\n* such that `φ.comp f = g` (`AddMonoidHom.liftOfRightInverse_comp`),\n* where `f : G₁ →+ G₂` has a RightInverse `f_inv` (`hf`),\n* and `g : G₂ →+ G₃` satisfies `hg : f.ker ≤ g.ker`.\n\nSee `AddMonoidHom.eq_liftOfRightInverse` for the uniqueness lemma.\n\n```\n G₁.\n | \\\n f | \\ g\n | \\\n v \\⌟\n G₂----> G₃\n ∃!φ\n```"] def liftOfRightInverse (hf : Function.RightInverse f_inv f) : { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) where @@ -3274,18 +3287,18 @@ def liftOfRightInverse (hf : Function.RightInverse f_inv f) : invFun φ := ⟨φ.comp f, fun x hx => (mem_ker _).mpr <| by simp [(mem_ker _).mp hx]⟩ left_inv g := by ext - simp only [comp_apply, lift_of_right_inverse_aux_comp_apply, Subtype.coe_mk, Subtype.val_eq_coe] + simp only [comp_apply, liftOfRightInverseAux_comp_apply, Subtype.coe_mk, Subtype.val_eq_coe] right_inv φ := by ext b - simp [lift_of_right_inverse_aux, hf b] + simp [liftOfRightInverseAux, hf b] #align monoid_hom.lift_of_right_inverse MonoidHom.liftOfRightInverse #align add_monoid_hom.lift_of_right_inverse AddMonoidHom.liftOfRightInverse -/-- A non-computable version of `monoid_hom.lift_of_right_inverse` for when no computable right -inverse is available, that uses `function.surj_inv`. -/ +/-- A non-computable version of `MonoidHom.liftOfRightInverse` for when no computable right +inverse is available, that uses `Function.surjInv`. -/ @[simp, to_additive - "A non-computable version of `add_monoid_hom.lift_of_right_inverse` for when no\ncomputable right inverse is available."] + "A non-computable version of `AddMonoidHom.liftOfRightInverse` for when no\ncomputable right inverse is available."] noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) := f.liftOfRightInverse (Function.surjInv hf) (Function.rightInverse_surjInv hf) @@ -3296,31 +3309,31 @@ noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : theorem liftOfRightInverse_comp_apply (hf : Function.RightInverse f_inv f) (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) (x : G₁) : (f.liftOfRightInverse f_inv hf g) (f x) = g x := - f.lift_of_right_inverse_aux_comp_apply f_inv hf g.1 g.2 x + f.liftOfRightInverseAux_comp_apply f_inv hf g.1 g.2 x #align monoid_hom.lift_of_right_inverse_comp_apply MonoidHom.liftOfRightInverse_comp_apply -#align add_monoid_hom.lift_of_right_inverse_comp_apply AddMonoidHom.lift_of_right_inverse_comp_apply +#align add_monoid_hom.lift_of_right_inverse_comp_apply AddMonoidHom.liftOfRightInverse_comp_apply @[to_additive (attr := simp)] theorem liftOfRightInverse_comp (hf : Function.RightInverse f_inv f) (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) : (f.liftOfRightInverse f_inv hf g).comp f = g := - MonoidHom.ext <| f.lift_of_right_inverse_comp_apply f_inv hf g + MonoidHom.ext <| f.liftOfRightInverse_comp_apply f_inv hf g #align monoid_hom.lift_of_right_inverse_comp MonoidHom.liftOfRightInverse_comp -#align add_monoid_hom.lift_of_right_inverse_comp AddMonoidHom.lift_of_right_inverse_comp +#align add_monoid_hom.lift_of_right_inverse_comp AddMonoidHom.liftOfRightInverse_comp @[to_additive] theorem eq_liftOfRightInverse (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) (h : G₂ →* G₃) (hh : h.comp f = g) : h = f.liftOfRightInverse f_inv hf ⟨g, hg⟩ := by simp_rw [← hh] - exact ((f.lift_of_right_inverse f_inv hf).apply_symm_apply _).symm + exact ((f.liftOfRightInverse f_inv hf).apply_symm_apply _).symm #align monoid_hom.eq_lift_of_right_inverse MonoidHom.eq_liftOfRightInverse -#align add_monoid_hom.eq_lift_of_right_inverse AddMonoidHom.eq_lift_of_right_inverse +#align add_monoid_hom.eq_lift_of_right_inverse AddMonoidHom.eq_liftOfRightInverse end MonoidHom variable {N : Type _} [Group N] --- Here `H.normal` is an explicit argument so we can use dot notation with `comap`. +-- Here `H.Normal` is an explicit argument so we can use dot notation with `comap`. @[to_additive] theorem Subgroup.Normal.comap {H : Subgroup N} (hH : H.Normal) (f : G →* N) : (H.comap f).Normal := ⟨fun _ => by simp (config := { contextual := true }) [Subgroup.mem_comap, hH.conj_mem]⟩ @@ -3334,31 +3347,31 @@ instance (priority := 100) Subgroup.normal_comap {H : Subgroup N} [nH : H.Normal #align subgroup.normal_comap Subgroup.normal_comap #align add_subgroup.normal_comap AddSubgroup.normal_comap --- Here `H.normal` is an explicit argument so we can use dot notation with `subgroup_of`. +-- Here `H.Normal` is an explicit argument so we can use dot notation with `subgroupOf`. @[to_additive] theorem Subgroup.Normal.subgroupOf {H : Subgroup G} (hH : H.Normal) (K : Subgroup G) : (H.subgroupOf K).Normal := hH.comap _ #align subgroup.normal.subgroup_of Subgroup.Normal.subgroupOf -#align add_subgroup.normal.add_subgroup_of AddSubgroup.Normal.add_subgroupOf +#align add_subgroup.normal.add_subgroup_of AddSubgroup.Normal.addSubgroupOf @[to_additive] instance (priority := 100) Subgroup.normal_subgroupOf {H N : Subgroup G} [N.Normal] : (N.subgroupOf H).Normal := Subgroup.normal_comap _ #align subgroup.normal_subgroup_of Subgroup.normal_subgroupOf -#align add_subgroup.normal_add_subgroup_of AddSubgroup.normal_add_subgroupOf +#align add_subgroup.normal_add_subgroup_of AddSubgroup.normal_addSubgroupOf namespace MonoidHom -/-- The `monoid_hom` from the preimage of a subgroup to itself. -/ -@[to_additive "the `add_monoid_hom` from the preimage of an additive subgroup to itself.", simps] +/-- The `MonoidHom` from the preimage of a subgroup to itself. -/ +@[to_additive "the `AddMonoidHom` from the preimage of an additive subgroup to itself.", simps] def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := f.submonoidComap H'.toSubmonoid #align monoid_hom.subgroup_comap MonoidHom.subgroupComap #align add_monoid_hom.add_subgroup_comap AddMonoidHom.addSubgroupComap -/-- The `monoid_hom` from a subgroup to its image. -/ +/-- The `MonoidHom` from a subgroup to its image. -/ @[to_additive "the `add_monoid_hom` from an additive subgroup to its image", simps] def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := f.submonoidMap H.toSubmonoid @@ -3368,9 +3381,9 @@ def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := @[to_additive] theorem subgroupMap_surjective (f : G →* G') (H : Subgroup G) : Function.Surjective (f.subgroupMap H) := - f.submonoid_map_surjective H.toSubmonoid + f.submonoidMap_surjective H.toSubmonoid #align monoid_hom.subgroup_map_surjective MonoidHom.subgroupMap_surjective -#align add_monoid_hom.add_subgroup_map_surjective AddMonoidHom.add_subgroupMap_surjective +#align add_monoid_hom.add_subgroup_map_surjective AddMonoidHom.addSubgroupMap_surjective end MonoidHom @@ -3388,9 +3401,9 @@ def subgroupCongr (h : H = K) : H ≃* K := #align add_equiv.add_subgroup_congr AddEquiv.addSubgroupCongr /-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, -use `subgroup.equiv_map_of_injective`. -/ +use `Subgroup.equiv_map_of_injective`. -/ @[to_additive - "An additive subgroup is isomorphic to its image under an an isomorphism. If you only\nhave an injective map, use `add_subgroup.equiv_map_of_injective`."] + "An additive subgroup is isomorphic to its image under an an isomorphism. If you only\nhave an injective map, use `AddSubgroup.equiv_map_of_injective`."] def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') := MulEquiv.submonoidMap (e : G ≃* G') H.toSubmonoid #align mul_equiv.subgroup_map MulEquiv.subgroupMap @@ -3401,14 +3414,14 @@ theorem coe_subgroupMap_apply (e : G ≃* G') (H : Subgroup G) (g : H) : ((subgroupMap e H g : H.map (e : G →* G')) : G') = e g := rfl #align mul_equiv.coe_subgroup_map_apply MulEquiv.coe_subgroupMap_apply -#align add_equiv.coe_add_subgroup_map_apply AddEquiv.coe_add_subgroupMap_apply +#align add_equiv.coe_add_subgroup_map_apply AddEquiv.coe_addSubgroupMap_apply @[to_additive (attr := simp)] theorem subgroupMap_symm_apply (e : G ≃* G') (H : Subgroup G) (g : H.map (e : G →* G')) : (e.subgroupMap H).symm g = ⟨e.symm g, SetLike.mem_coe.1 <| Set.mem_image_equiv.1 g.2⟩ := rfl #align mul_equiv.subgroup_map_symm_apply MulEquiv.subgroupMap_symm_apply -#align add_equiv.add_subgroup_map_symm_apply AddEquiv.add_subgroup_map_symm_apply +#align add_equiv.add_subgroup_map_symm_apply AddEquiv.addSubgroupMap_symm_apply end MulEquiv @@ -3420,7 +3433,9 @@ theorem equivMapOfInjective_coe_mulEquiv (H : Subgroup G) (e : G ≃* G') : ext rfl #align subgroup.equiv_map_of_injective_coe_mul_equiv Subgroup.equivMapOfInjective_coe_mulEquiv -#align add_subgroup.equiv_map_of_injective_coe_add_equiv AddSubgroup.equiv_map_of_injective_coe_add_equiv +#align + add_subgroup.equiv_map_of_injective_coe_add_equiv + AddSubgroup.equivMapOfInjective_coe_addEquiv variable {C : Type _} [CommGroup C] {s t : Subgroup C} {x : C} @@ -3475,7 +3490,7 @@ theorem normal_subgroupOf_iff {H K : Subgroup G} (hHK : H ≤ K) : ⟨fun hN h k hH hK => hN.conj_mem ⟨h, hHK hH⟩ hH ⟨k, hK⟩, fun hN => { conj_mem := fun h hm k => hN h.1 k.1 hm k.2 }⟩ #align subgroup.normal_subgroup_of_iff Subgroup.normal_subgroupOf_iff -#align add_subgroup.normal_add_subgroup_of_iff AddSubgroup.normal_add_subgroupOf_iff +#align add_subgroup.normal_add_subgroup_of_iff AddSubgroup.normal_addSubgroupOf_iff @[to_additive] instance prod_subgroupOf_prod_normal {H₁ K₁ : Subgroup G} {H₂ K₂ : Subgroup N} @@ -3487,11 +3502,11 @@ instance prod_subgroupOf_prod_normal {H₁ K₁ : Subgroup G} {H₂ K₂ : Subgr h₂.conj_mem ⟨(n : G × N).snd, (mem_prod.mp n.2).2⟩ hgHK.2 ⟨(g : G × N).snd, (mem_prod.mp g.2).2⟩⟩ #align subgroup.prod_subgroup_of_prod_normal Subgroup.prod_subgroupOf_prod_normal -#align add_subgroup.sum_add_subgroup_of_sum_normal AddSubgroup.sum_add_subgroupOf_sum_normal +#align add_subgroup.sum_add_subgroup_of_sum_normal AddSubgroup.sum_addSubgroupOf_sum_normal @[to_additive] instance prod_normal (H : Subgroup G) (K : Subgroup N) [hH : H.Normal] [hK : K.Normal] : - (H.Prod K).Normal + (H.prod K).Normal where conj_mem n hg g := ⟨hH.conj_mem n.fst (Subgroup.mem_prod.mp hg).1 g.fst, hK.conj_mem n.snd (Subgroup.mem_prod.mp hg).2 g.snd⟩ @@ -3506,7 +3521,7 @@ theorem inf_subgroupOf_inf_normal_of_right (A B' B : Subgroup G) (hB : B' ≤ B) ⟨mul_mem (mul_mem (mem_inf.1 g.2).1 (mem_inf.1 n.2).1) (inv_mem (mem_inf.1 g.2).1), (normal_subgroupOf_iff hB).mp hN n g hn.2 (mem_inf.mp g.2).2⟩ } #align subgroup.inf_subgroup_of_inf_normal_of_right Subgroup.inf_subgroupOf_inf_normal_of_right -#align add_subgroup.inf_add_subgroup_of_inf_normal_of_right AddSubgroup.inf_add_subgroupOf_inf_normal_of_right +#align add_subgroup.inf_add_subgroup_of_inf_normal_of_right AddSubgroup.inf_addSubgroupOf_inf_normal_of_right @[to_additive] theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) (hA : A' ≤ A) @@ -3516,7 +3531,9 @@ theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) ( ⟨(normal_subgroupOf_iff hA).mp hN n g hn.1 (mem_inf.mp g.2).1, mul_mem (mul_mem (mem_inf.1 g.2).2 (mem_inf.1 n.2).2) (inv_mem (mem_inf.1 g.2).2)⟩ } #align subgroup.inf_subgroup_of_inf_normal_of_left Subgroup.inf_subgroupOf_inf_normal_of_left -#align add_subgroup.inf_add_subgroup_of_inf_normal_of_left AddSubgroup.inf_add_subgroupOf_inf_normal_of_left +#align + add_subgroup.inf_add_subgroup_of_inf_normal_of_left + AddSubgroup.inf_addSubgroupOf_inf_normal_of_left @[to_additive] instance normal_inf_normal (H K : Subgroup G) [hH : H.Normal] [hK : K.Normal] : (H ⊓ K).Normal := @@ -3533,12 +3550,12 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : · simp only [subgroup_of, map_comap_eq, map_sup, subtype_range] rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] #align subgroup.subgroup_of_sup Subgroup.subgroupOf_sup -#align add_subgroup.add_subgroup_of_sup AddSubgroup.add_subgroupOf_sup +#align add_subgroup.add_subgroup_of_sup AddSubgroup.addSubgroupOf_sup @[to_additive] theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgroupOf K).Normal] {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := by - have := (normal_subgroup_of_iff hK).mp hN (a * b) b h hb + have := (normal_subgroupOf_iff hK).mp hN (a * b) b h hb rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this #align subgroup.subgroup_normal.mem_comm Subgroup.SubgroupNormal.mem_comm #align add_subgroup.subgroup_normal.mem_comm AddSubgroup.SubgroupNormal.mem_comm @@ -3546,11 +3563,14 @@ theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgr /-- Elements of disjoint, normal subgroups commute. -/ @[to_additive "Elements of disjoint, normal subgroups commute."] theorem 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 := by +-- Porting note: Goal was `Commute x y`. Removed ambiguity. + (hdis : Disjoint H₁ H₂) (x y : G) (hx : x ∈ H₁) (hy : y ∈ H₂) : _root_.Commute x y := by suffices x * y * x⁻¹ * y⁻¹ = 1 by show x * y = y * x · rw [mul_assoc, mul_eq_one_iff_eq_inv] at this - simpa + -- Porting note: Previous code was: + -- simpa + simp only [this, mul_inv_rev, inv_inv] apply hdis.le_bot constructor · suffices x * (y * x⁻¹ * y⁻¹) ∈ H₁ by simpa [mul_assoc] @@ -3592,7 +3612,7 @@ theorem mul_injective_of_disjoint {H₁ H₂ : Subgroup G} (h : Disjoint H₁ H Function.Injective (fun g => g.1 * g.2 : H₁ × H₂ → G) := by intro x y hxy rw [← inv_mul_eq_iff_eq_mul, ← mul_assoc, ← mul_inv_eq_one, mul_assoc] at hxy - replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).Prop (x.2 * y.2⁻¹).Prop hxy + replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).prop (x.2 * y.2⁻¹).prop hxy rwa [coe_mul, coe_mul, coe_inv, coe_inv, inv_mul_eq_one, mul_inv_eq_one, ← Subtype.ext_iff, ← Subtype.ext_iff, eq_comm, ← Prod.ext_iff] at hxy #align subgroup.mul_injective_of_disjoint Subgroup.mul_injective_of_disjoint From 7714bdfcf7787f88599510050ffb4aa6325ab972 Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Wed, 25 Jan 2023 06:23:38 +0000 Subject: [PATCH 09/34] More fixes. --- Mathlib/GroupTheory/Subgroup/Basic.lean | 108 ++++++++++++------------ 1 file changed, 56 insertions(+), 52 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 74c9199debda7..97bae39b5bc1f 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -473,7 +473,7 @@ theorem toSubmonoid_mono : Monotone (toSubmonoid : Subgroup G → Submonoid G) : theorem toSubmonoid_le {p q : Subgroup G} : p.toSubmonoid ≤ q.toSubmonoid ↔ p ≤ q := Iff.rfl #align subgroup.to_submonoid_le Subgroup.toSubmonoid_le -#align add_subgroup.to_add_submonoid_le AddSubgroup.to_add_submonoid_le +#align add_subgroup.to_add_submonoid_le AddSubgroup.toAddSubmonoid_le end Subgroup @@ -550,7 +550,7 @@ theorem copy_eq (K : Subgroup G) (s : Set G) (hs : s = ↑K) : K.copy s hs = K : #align add_subgroup.copy_eq AddSubgroup.copy_eq /-- Two subgroups are equal if they have the same elements. -/ -@[ext, to_additive "Two `AddSubgroup`s are equal if they have the same elements."] +@[to_additive (attr := ext) "Two `AddSubgroup`s are equal if they have the same elements."] theorem ext {H K : Subgroup G} (h : ∀ x, x ∈ H ↔ x ∈ K) : H = K := SetLike.ext h #align subgroup.ext Subgroup.ext @@ -564,21 +564,21 @@ protected theorem one_mem : (1 : G) ∈ H := #align add_subgroup.zero_mem AddSubgroup.zero_mem /-- A subgroup is closed under multiplication. -/ -@[to_additive "An `add_subgroup` is closed under addition."] +@[to_additive "An `AddSubgroup` is closed under addition."] protected theorem mul_mem {x y : G} : x ∈ H → y ∈ H → x * y ∈ H := mul_mem #align subgroup.mul_mem Subgroup.mul_mem #align add_subgroup.add_mem AddSubgroup.add_mem /-- A subgroup is closed under inverse. -/ -@[to_additive "An `add_subgroup` is closed under inverse."] +@[to_additive "An `AddSubgroup` is closed under inverse."] protected theorem inv_mem {x : G} : x ∈ H → x⁻¹ ∈ H := inv_mem #align subgroup.inv_mem Subgroup.inv_mem #align add_subgroup.neg_mem AddSubgroup.neg_mem /-- A subgroup is closed under division. -/ -@[to_additive "An `add_subgroup` is closed under subtraction."] +@[to_additive "An `AddSubgroup` is closed under subtraction."] protected theorem div_mem {x y : G} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := div_mem hx hy #align subgroup.div_mem Subgroup.div_mem @@ -645,28 +645,28 @@ def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y #align add_subgroup.of_sub AddSubgroup.ofSub /-- A subgroup of a group inherits a multiplication. -/ -@[to_additive "An `add_subgroup` of an `add_group` inherits an addition."] +@[to_additive "An `AddSubgroup` of an `AddGroup` inherits an addition."] instance hasMul : Mul H := - H.toSubmonoid.HasMul + H.toSubmonoid.mul #align subgroup.has_mul Subgroup.hasMul #align add_subgroup.has_add AddSubgroup.hasAdd /-- A subgroup of a group inherits a 1. -/ -@[to_additive "An `add_subgroup` of an `add_group` inherits a zero."] +@[to_additive "An `AddSubgroup` of an `AddGroup` inherits a zero."] instance hasOne : One H := - H.toSubmonoid.HasOne + H.toSubmonoid.one #align subgroup.has_one Subgroup.hasOne #align add_subgroup.has_zero AddSubgroup.hasZero /-- A subgroup of a group inherits an inverse. -/ -@[to_additive "A `add_subgroup` of a `add_group` inherits an inverse."] +@[to_additive "A `AddSubgroup` of a `AddGroup` inherits an inverse."] instance hasInv : Inv H := ⟨fun a => ⟨a⁻¹, H.inv_mem a.2⟩⟩ #align subgroup.has_inv Subgroup.hasInv #align add_subgroup.has_neg AddSubgroup.hasNeg /-- A subgroup of a group inherits a division -/ -@[to_additive "An `add_subgroup` of an `add_group` inherits a subtraction."] +@[to_additive "An `AddSubgroup` of an `AddGroup` inherits a subtraction."] instance hasDiv : Div H := ⟨fun a b => ⟨a / b, H.div_mem a.2 b.2⟩⟩ #align subgroup.has_div Subgroup.hasDiv @@ -745,7 +745,7 @@ theorem mk_eq_one_iff {g : G} {h} : (⟨g, h⟩ : H) = 1 ↔ g = 1 := #align add_subgroup.mk_eq_zero_iff AddSubgroup.mk_eq_zero_iff /-- A subgroup of a group inherits a group structure. -/ -@[to_additive "An `add_subgroup` of an `add_group` inherits an `add_group` structure."] +@[to_additive "An `AddSubgroup` of an `AddGroup` inherits an `AddGroup` structure."] instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := Subtype.coe_injective.Group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -753,7 +753,7 @@ instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := #align add_subgroup.to_add_group AddSubgroup.toAddGroup /-- A subgroup of a `CommGroup` is a `CommGroup`. -/ -@[to_additive "An `add_subgroup` of an `add_comm_group` is an `add_comm_group`."] +@[to_additive "An `AddSubgroup` of an `AddCommGroup` is an `AddCommGroup`."] instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H := Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -780,14 +780,14 @@ instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : S #align add_subgroup.to_linear_ordered_add_comm_group AddSubgroup.toLinearOrderedAddCommGroup /-- The natural group hom from a subgroup of group `G` to `G`. -/ -@[to_additive "The natural group hom from an `add_subgroup` of `add_group` `G` to `G`."] +@[to_additive "The natural group hom from an `AddSubgroup` of `AddGroup` `G` to `G`."] def subtype : H →* G := ⟨coe, rfl, fun _ _ => rfl⟩ #align subgroup.subtype Subgroup.subtype #align add_subgroup.subtype AddSubgroup.subtype @[to_additive (attr := simp)] -theorem coeSubtype : ⇑H.Subtype = coe := +theorem coeSubtype : ⇑H.subtype = coe := rfl #align subgroup.coe_subtype Subgroup.coeSubtype #align add_subgroup.coe_subtype AddSubgroup.coeSubtype @@ -826,7 +826,7 @@ theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : #align add_subgroup.subtype_comp_inclusion AddSubgroup.subtype_comp_inclusion /-- The subgroup `G` of the group `G`. -/ -@[to_additive "The `add_subgroup G` of the `add_group G`."] +@[to_additive "The `AddSubgroup G` of the `AddGroup G`."] instance : Top (Subgroup G) := ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ _ => Set.mem_univ _ }⟩ @@ -842,7 +842,7 @@ def topEquiv : (⊤ : Subgroup G) ≃* G := #align add_subgroup.top_equiv AddSubgroup.topEquiv /-- The trivial subgroup `{1}` of an group `G`. -/ -@[to_additive "The trivial `add_subgroup` `{0}` of an `add_group` `G`."] +@[to_additive "The trivial `AddSubgroup` `{0}` of an `AddGroup` `G`."] instance : Bot (Subgroup G) := ⟨{ (⊥ : Submonoid G) with inv_mem' := fun _ => by simp [*] }⟩ @@ -978,28 +978,28 @@ instance : InfSet (Subgroup G) := theorem coe_infₛ (H : Set (Subgroup G)) : ((infₛ H : Subgroup G) : Set G) = ⋂ s ∈ H, ↑s := rfl #align subgroup.coe_Inf Subgroup.coe_infₛ -#align add_subgroup.coe_Inf AddSubgroup.coe_Inf +#align add_subgroup.coe_Inf AddSubgroup.coe_infₛ @[to_additive (attr := simp)] theorem mem_infₛ {S : Set (Subgroup G)} {x : G} : x ∈ infₛ S ↔ ∀ p ∈ S, x ∈ p := Set.mem_interᵢ₂ #align subgroup.mem_Inf Subgroup.mem_infₛ -#align add_subgroup.mem_Inf AddSubgroup.mem_Inf +#align add_subgroup.mem_Inf AddSubgroup.mem_infₛ @[to_additive] theorem mem_infᵢ {ι : Sort _} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by - simp only [infᵢ, mem_Inf, Set.forall_range_iff] + simp only [infᵢ, mem_inf, Set.forall_range_iff] #align subgroup.mem_infi Subgroup.mem_infᵢ -#align add_subgroup.mem_infi AddSubgroup.mem_infi +#align add_subgroup.mem_infi AddSubgroup.mem_infᵢ @[to_additive (attr := simp, norm_cast)] theorem coe_infᵢ {ι : Sort _} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by - simp only [infᵢ, coe_Inf, Set.binterᵢ_range] + simp only [infᵢ, coe_inf, Set.binterᵢ_range] #align subgroup.coe_infi Subgroup.coe_infᵢ -#align add_subgroup.coe_infi AddSubgroup.coe_infi +#align add_subgroup.coe_infi AddSubgroup.coe_infᵢ /-- Subgroups of a group form a complete lattice. -/ -@[to_additive "The `add_subgroup`s of an `add_group` form a complete lattice."] +@[to_additive "The `AddSubgroup`s of an `AddGroup` form a complete lattice."] instance : CompleteLattice (Subgroup G) := { completeLatticeOfInf (Subgroup G) fun s => @@ -1080,7 +1080,7 @@ theorem eq_top_iff' : H = ⊤ ↔ ∀ x : G, x ∈ H := #align add_subgroup.eq_top_iff' AddSubgroup.eq_top_iff' /-- The `Subgroup` generated by a set. -/ -@[to_additive "The `add_subgroup` generated by a set"] +@[to_additive "The `AddSubgroup` generated by a set"] def closure (k : Set G) : Subgroup G := infₛ { K | k ⊆ K } #align subgroup.closure Subgroup.closure @@ -1095,8 +1095,8 @@ theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K #align add_subgroup.mem_closure AddSubgroup.mem_closure /-- The subgroup generated by a set includes the set. -/ -@[simp, to_additive "The `add_subgroup` generated by a set includes the set."] -theorem subset_closure : k ⊆ closure k := fun x hx => mem_closure.2 fun K hK => hK hx +@[to_additive (attr := simp) "The `AddSubgroup` generated by a set includes the set."] +theorem subset_closure : k ⊆ closure k := fun x hx => mem_closure.2 fun _ hK => hK hx #align subgroup.subset_closure Subgroup.subset_closure #align add_subgroup.subset_closure AddSubgroup.subset_closure @@ -1109,7 +1109,7 @@ theorem not_mem_of_not_mem_closure {P : G} (hP : P ∉ closure k) : P ∉ k := f open Set /-- A subgroup `K` includes `closure k` if and only if it includes `k`. -/ -@[simp, to_additive "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] +@[to_additive (attr := simp) "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] theorem closure_le : closure k ≤ K ↔ k ⊆ K := ⟨Subset.trans subset_closure, fun h => infₛ_le h⟩ #align subgroup.closure_le Subgroup.closure_le @@ -1134,7 +1134,7 @@ theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x #align add_subgroup.closure_induction AddSubgroup.closure_induction /-- A dependent version of `Subgroup.closure_induction`. -/ -@[elab_as_elim, to_additive "A dependent version of `add_subgroup.closure_induction`. "] +@[elab_as_elim, to_additive "A dependent version of `AddSubgroup.closure_induction`. "] theorem closure_induction' {p : ∀ x, x ∈ closure k → Prop} (Hs : ∀ (x) (h : x ∈ k), p x (subset_closure h)) (H1 : p 1 (one_mem _)) (Hmul : ∀ x hx y hy, p x hx → p y hy → p (x * y) (mul_mem hx hy)) @@ -1220,7 +1220,7 @@ theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k #align add_subgroup.closure_mono AddSubgroup.closure_mono /-- Closure of a subgroup `K` equals `K`. -/ -@[simp, to_additive "Additive closure of an additive subgroup `K` equals `K`"] +@[to_additive (attr := simp) "Additive closure of an additive subgroup `K` equals `K`"] theorem closure_eq : closure (K : Set G) = K := (Subgroup.gi G).l_u_eq K #align subgroup.closure_eq Subgroup.closure_eq @@ -1259,7 +1259,7 @@ theorem closure_eq_bot_iff (G : Type _) [Group G] (S : Set G) : closure S = ⊥ @[to_additive] theorem supᵢ_eq_closure {ι : Sort _} (p : ι → Subgroup G) : - (⨆ i, p i) = closure (⋃ i, (p i : Set G)) := by simp_rw [closure_Union, closure_eq] + (⨆ i, p i) = closure (⋃ i, (p i : Set G)) := by simp_rw [closure_unionᵢ, closure_eq] #align subgroup.supr_eq_closure Subgroup.supᵢ_eq_closure #align add_subgroup.supr_eq_closure AddSubgroup.supᵢ_eq_closure @@ -1306,8 +1306,8 @@ theorem mem_supᵢ_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} {x : G} : x ∈ (supᵢ K : Subgroup G) ↔ ∃ i, x ∈ K i := by refine' ⟨_, fun ⟨i, hi⟩ => (SetLike.le_def.1 <| le_supᵢ K i) hi⟩ suffices x ∈ closure (⋃ i, (K i : Set G)) → ∃ i, x ∈ K i by - simpa only [closure_Union, closure_eq (K _)] using this - refine' fun hx => closure_induction hx (fun _ => mem_Union.1) _ _ _ + simpa only [closure_unionᵢ, closure_eq (K _)] using this + refine' fun hx => closure_induction hx (fun _ => mem_unionᵢ.1) _ _ _ · exact hι.elim fun i => ⟨i, (K i).one_mem⟩ · rintro x y ⟨i, hi⟩ ⟨j, hj⟩ rcases hK i j with ⟨k, hki, hkj⟩ @@ -1320,7 +1320,7 @@ theorem mem_supᵢ_of_directed {ι} [hι : Nonempty ι] {K : ι → Subgroup G} @[to_additive] theorem coe_supᵢ_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : Directed (· ≤ ·) S) : ((⨆ i, S i : Subgroup G) : Set G) = ⋃ i, ↑(S i) := - Set.ext fun x => by simp [mem_supr_of_directed hS] + Set.ext fun x => by simp [mem_supᵢ_of_directed hS] #align subgroup.coe_supr_of_directed Subgroup.coe_supᵢ_of_directed #align add_subgroup.coe_supr_of_directed AddSubgroup.coe_supᵢ_of_directed @@ -1328,7 +1328,7 @@ theorem coe_supᵢ_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : theorem mem_supₛ_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) {x : G} : x ∈ supₛ K ↔ ∃ s ∈ K, x ∈ s := by haveI : Nonempty K := Kne.to_subtype - simp only [supₛ_eq_supᵢ', mem_supr_of_directed hK.directed_coe, SetCoe.exists, Subtype.coe_mk] + simp only [supₛ_eq_supᵢ', mem_supᵢ_of_directed hK.directed_coe, SetCoe.exists, Subtype.coe_mk] #align subgroup.mem_Sup_of_directed_on Subgroup.mem_supₛ_of_directedOn #align add_subgroup.mem_Sup_of_directed_on AddSubgroup.mem_supₛ_of_directedOn @@ -1501,7 +1501,7 @@ theorem map_sup (H K : Subgroup G) (f : G →* N) : (H ⊔ K).map f = H.map f @[to_additive] theorem map_supᵢ {ι : Sort _} (f : G →* N) (s : ι → Subgroup G) : (supᵢ s).map f = ⨆ i, (s i).map f := - (gc_map_comap f).l_supr + (gc_map_comap f).l_supᵢ #align subgroup.map_supr Subgroup.map_supᵢ #align add_subgroup.map_supr AddSubgroup.map_supᵢ @@ -1528,9 +1528,9 @@ theorem comap_inf (H K : Subgroup N) (f : G →* N) : (H ⊓ K).comap f = H.coma @[to_additive] theorem comap_infᵢ {ι : Sort _} (f : G →* N) (s : ι → Subgroup N) : (infᵢ s).comap f = ⨅ i, (s i).comap f := - (gc_map_comap f).u_infi + (gc_map_comap f).u_infᵢ #align subgroup.comap_infi Subgroup.comap_infᵢ -#align add_subgroup.comap_infi AddSubgroup.comap_infi +#align add_subgroup.comap_infi AddSubgroup.comap_infᵢ @[to_additive] theorem map_inf_le (H K : Subgroup G) (f : G →* N) : map f (H ⊓ K) ≤ map f H ⊓ map f K := @@ -1570,7 +1570,7 @@ theorem comap_top (f : G →* N) : (⊤ : Subgroup N).comap f = ⊤ := /-- For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`. -/ @[to_additive "For any subgroups `H` and `K`, view `H ⊓ K` as a subgroup of `K`."] def subgroupOf (H K : Subgroup G) : Subgroup K := - H.comap K.Subtype + H.comap K.subtype #align subgroup.subgroup_of Subgroup.subgroupOf #align add_subgroup.add_subgroup_of AddSubgroup.addSubgroupOf @@ -1774,12 +1774,12 @@ section Pi variable {η : Type _} {f : η → Type _} --- defined here and not in group_theory.submonoid.operations to have access to algebra.group.pi +-- defined here and not in GroupTheory.Submonoid.Operations to have access to Algebra.Group.Pi /-- A version of `Set.pi` for submonoids. Given an index set `I` and a family of submodules `s : Π i, Submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that `f i` belongs to `Pi I s` whenever `i ∈ I`. -/ @[to_additive - " A version of `Set.pi` for `add_submonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] + " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : Submonoid (∀ i, f i) where carrier := I.pi fun i => (s i).carrier @@ -1918,7 +1918,9 @@ variable (nH : H.Normal) @[to_additive] theorem mem_comm {a b : G} (h : a * b ∈ H) : b * a ∈ H := by have : a⁻¹ * (a * b) * a⁻¹⁻¹ ∈ H := nH.conj_mem (a * b) h a⁻¹ - simpa + -- Porting note: Previous code was: + -- simpa + simp_all only [inv_mul_cancel_left, inv_inv] #align subgroup.normal.mem_comm Subgroup.Normal.mem_comm #align add_subgroup.normal.mem_comm AddSubgroup.Normal.mem_comm @@ -1956,7 +1958,7 @@ structure Characteristic : Prop where fixed : ∀ ϕ : A ≃+ A, H.comap ϕ.toAddMonoidHom = H #align add_subgroup.characteristic AddSubgroup.Characteristic -attribute [to_additive AddSubgroup.Characteristic] Subgroup.Characteristic +attribute [to_additive] Subgroup.Characteristic attribute [class] Characteristic @@ -2227,7 +2229,9 @@ theorem normalizerCondition_iff_only_full_group_self_normalizing : apply forall_congr'; intro H simp only [lt_iff_le_and_ne, le_normalizer, true_and_iff, le_top, Ne.def] tauto -#align normalizer_condition_iff_only_full_group_self_normalizing normalizerCondition_iff_only_full_group_self_normalizing +#align + normalizer_condition_iff_only_full_group_self_normalizing + normalizerCondition_iff_only_full_group_self_normalizing variable (H) @@ -2494,7 +2498,7 @@ theorem normalClosure_closure_eq_normalClosure {s : Set G} : #align subgroup.normal_closure_closure_eq_normal_closure Subgroup.normalClosure_closure_eq_normalClosure /-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, -as shown by `subgroup.normalCore_eq_supr`. -/ +as shown by `subgroup.normalCore_eq_supᵢ`. -/ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } @@ -2549,7 +2553,7 @@ variable {N : Type _} {P : Type _} [Group N] [Group P] (K : Subgroup G) open Subgroup /-- The range of a monoid homomorphism from a group is a subgroup. -/ -@[to_additive "The range of an `add_monoid_hom` from an `add_group` is an `add_subgroup`."] +@[to_additive "The range of an `AddMonoidHom` from an `AddGroup` is an `AddSubgroup`."] def range (f : G →* N) : Subgroup N := Subgroup.copy ((⊤ : Subgroup G).map f) (Set.range f) (by simp [Set.ext_iff]) #align monoid_hom.range MonoidHom.range @@ -2627,7 +2631,7 @@ theorem range_top_iff_surjective {N} [Group N] {f : G →* N} : #align add_monoid_hom.range_top_iff_surjective AddMonoidHom.range_top_iff_surjective /-- The range of a surjective monoid homomorphism is the whole of the codomain. -/ -@[to_additive "The range of a surjective `add_monoid` homomorphism is the whole of the codomain."] +@[to_additive "The range of a surjective `AddMonoid` homomorphism is the whole of the codomain."] theorem range_top_of_surjective {N} [Group N] (f : G →* N) (hf : Function.Surjective f) : f.range = (⊤ : Subgroup N) := range_top_iff_surjective.2 hf @@ -2896,7 +2900,7 @@ theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) /-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup generated by the image of the set. -/ @[to_additive - "The image under an `add_monoid` hom of the `AddSubgroup` generated by a set equals\nthe `AddSubgroup` generated by the image of the set."] + "The image under an `AddMonoid` hom of the `AddSubgroup` generated by a set equals\nthe `AddSubgroup` generated by the image of the set."] theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := Set.image_preimage.l_comm_of_u_comm (Subgroup.gc_map_comap f) (Subgroup.gi N).gc (Subgroup.gi G).gc fun t => rfl @@ -3070,7 +3074,7 @@ theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) { @[to_additive (attr := simp)] theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : - H.map G'.Subtype ≤ K.map G'.Subtype ↔ H ≤ K := + H.map G'.subtype ≤ K.map G'.subtype ↔ H ≤ K := map_le_map_iff_of_injective Subtype.coe_injective #align subgroup.map_subtype_le_map_subtype Subgroup.map_subtype_le_map_subtype #align add_subgroup.map_subtype_le_map_subtype AddSubgroup.map_subtype_le_map_subtype @@ -3130,15 +3134,15 @@ theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := comap_sup_eq_of_le_range L.Subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) #align subgroup.sup_subgroup_of_eq Subgroup.sup_subgroupOf_eq -#align add_subgroup.sup_add_subgroup_of_eq AddSubgroup.sup_add_subgroupOf_eq +#align add_subgroup.sup_add_subgroup_of_eq AddSubgroup.sup_addSubgroupOf_eq @[to_additive] theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : Codisjoint (H.subgroupOf (H ⊔ K)) (K.subgroupOf (H ⊔ K)) := by - rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self] + rw [codisjoint_iff, sup_subgroupOf_eq, subgroupOf_self] exacts[le_sup_left, le_sup_right] #align subgroup.codisjoint_subgroup_of_sup Subgroup.codisjoint_subgroupOf_sup -#align add_subgroup.codisjoint_add_subgroup_of_sup AddSubgroup.codisjoint_add_subgroup_of_sup +#align add_subgroup.codisjoint_add_subgroup_of_sup AddSubgroup.codisjoint_addSubgroupOf_sup /-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, use `MulEquiv.subgroupMap` for better definitional equalities. -/ From 4ba794d4ab44d14dd28f99696a573e5e93565d30 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Wed, 25 Jan 2023 09:24:54 +0100 Subject: [PATCH 10/34] further changes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 43 ++++++++++++++----------- 1 file changed, 25 insertions(+), 18 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 97bae39b5bc1f..921bbb372f01f 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -392,28 +392,30 @@ theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := #align subgroup.mem_carrier Subgroup.mem_carrier #align add_subgroup.mem_carrier AddSubgroup.mem_carrier --- Porting note: Do we still want these _mk lemmas? I thought a lot of them are obsolete now @[to_additive (attr := simp)] -theorem mem_mk {s : Set G} {x : G} (h_one) (h_mul) (h_inv) : x ∈ mk s h_one h_mul h_inv ↔ x ∈ s := +theorem mem_mk {s : Set G} {x : G} (h_one) (h_mul) (h_inv) : + x ∈ mk ⟨⟨s, h_one⟩, h_mul⟩ h_inv ↔ x ∈ s := Iff.rfl #align subgroup.mem_mk Subgroup.mem_mk #align add_subgroup.mem_mk AddSubgroup.mem_mk @[to_additive (attr := simp)] -theorem coe_set_mk {s : Set G} (h_one) (h_mul) (h_inv) : (mk s h_one h_mul h_inv : Set G) = s := +theorem coe_set_mk {s : Set G} (h_one) (h_mul) (h_inv) : + (mk ⟨⟨s, h_one⟩, h_mul⟩ h_inv : Set G) = s := rfl #align subgroup.coe_set_mk Subgroup.coe_set_mk #align add_subgroup.coe_set_mk AddSubgroup.coe_set_mk @[to_additive (attr := simp)] theorem mk_le_mk {s t : Set G} (h_one) (h_mul) (h_inv) (h_one') (h_mul') (h_inv') : - mk s h_one h_mul h_inv ≤ mk t h_one' h_mul' h_inv' ↔ s ⊆ t := + mk ⟨⟨s, h_one⟩, h_mul⟩ h_inv ≤ mk ⟨⟨t, h_one'⟩, h_mul'⟩ h_inv' ↔ s ⊆ t := Iff.rfl #align subgroup.mk_le_mk Subgroup.mk_le_mk #align add_subgroup.mk_le_mk AddSubgroup.mk_le_mk /-- See Note [custom simps projection] -/ --@[to_additive "See Note [custom simps projection]"] +@[to_additive] def Simps.coe (S : Subgroup G) : Set G := S #align subgroup.simps.coe Subgroup.Simps.coe @@ -437,7 +439,11 @@ theorem mem_toSubmonoid (K : Subgroup G) (x : G) : x ∈ K.toSubmonoid ↔ x ∈ @[to_additive] theorem toSubmonoid_injective : Function.Injective (toSubmonoid : Subgroup G → Submonoid G) := - fun p q h => SetLike.ext'_iff.2 (show _ from SetLike.ext'_iff.1 h) + -- fun p q h => SetLike.ext'_iff.2 (show _ from SetLike.ext'_iff.1 h) + fun p q h => by + have := SetLike.ext'_iff.1 h + rw [coe_toSubmonoid, coe_toSubmonoid] at this + exact SetLike.ext'_iff.2 this #align subgroup.to_submonoid_injective Subgroup.toSubmonoid_injective #align add_subgroup.to_add_submonoid_injective AddSubgroup.toAddSubmonoid_injective @@ -488,11 +494,11 @@ section mul_add @[simps] def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where - toFun S := { S.toSubmonoid.toAddSubmonoid with neg_mem' := fun _ => S.inv_mem' } - invFun S := { S.toAddSubmonoid.toSubmonoid' with inv_mem' := fun _ => S.neg_mem' } - left_inv x := by cases x <;> rfl - right_inv x := by cases x <;> rfl - map_rel_iff' a b := Iff.rfl + toFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } + invFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } + left_inv x := by cases x; rfl + right_inv x := by cases x; rfl + map_rel_iff' := Iff.rfl #align subgroup.to_add_subgroup Subgroup.toAddSubgroup /-- Additive subgroup of an additive group `Additive G` are isomorphic to subgroup of `G`. -/ @@ -505,11 +511,11 @@ abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := @[simps] def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) where - toFun S := { S.toAddSubmonoid.toSubmonoid with inv_mem' := fun _ => S.neg_mem' } - invFun S := { S.toSubmonoid.toAddSubmonoid' with neg_mem' := fun _ => S.inv_mem' } - left_inv x := by cases x <;> rfl - right_inv x := by cases x <;> rfl - map_rel_iff' a b := Iff.rfl + toFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } + invFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } + left_inv x := by cases x; rfl + right_inv x := by cases x; rfl + map_rel_iff' := Iff.rfl #align add_subgroup.to_subgroup AddSubgroup.toSubgroup /-- Subgroups of an additive group `Multiplicative A` are isomorphic to additive subgroups of `A`. @@ -527,13 +533,14 @@ variable (H K : Subgroup G) /-- Copy of a subgroup with a new `carrier` equal to the old one. Useful to fix definitional equalities.-/ @[to_additive - "Copy of an additive subgroup with a new `carrier` equal to the old one.\nUseful to fix definitional equalities"] + "Copy of an additive subgroup with a new `carrier` equal to the old one. + Useful to fix definitional equalities"] protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G where carrier := s one_mem' := hs.symm ▸ K.one_mem' - mul_mem' _ _ := hs.symm ▸ K.mul_mem' - inv_mem' _ := hs.symm ▸ K.inv_mem' + mul_mem' := hs.symm ▸ K.mul_mem' + inv_mem' := hs.symm ▸ K.inv_mem' #align subgroup.copy Subgroup.copy #align add_subgroup.copy AddSubgroup.copy From 4d9390cf81c5ab11b9e76d2f53131878c676d334 Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Wed, 25 Jan 2023 10:04:55 +0000 Subject: [PATCH 11/34] Fix `simps` projection --- Mathlib/GroupTheory/Subgroup/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 921bbb372f01f..ecd9d1f8b38f1 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -421,9 +421,9 @@ def Simps.coe (S : Subgroup G) : Set G := #align subgroup.simps.coe Subgroup.Simps.coe #align add_subgroup.simps.coe AddSubgroup.Simps.coe -initialize_simps_projections Subgroup (carrier → coe) +initialize_simps_projections Subgroup (toSubmonoid_toSubsemigroup_carrier → coe) -initialize_simps_projections AddSubgroup (carrier → coe) +initialize_simps_projections AddSubgroup (toAddSubmonoid_toAddSubsemigroup_carrier → coe) @[to_additive (attr := simp)] theorem coe_toSubmonoid (K : Subgroup G) : (K.toSubmonoid : Set G) = K := From 1008741e6cf48ab5d21b65781e6718496d679fe9 Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Wed, 25 Jan 2023 10:12:16 +0000 Subject: [PATCH 12/34] Use `simp` because the triangle is being obstinate again --- Mathlib/GroupTheory/Subgroup/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index ecd9d1f8b38f1..4362a3c9ac8c0 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -540,7 +540,7 @@ protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G carrier := s one_mem' := hs.symm ▸ K.one_mem' mul_mem' := hs.symm ▸ K.mul_mem' - inv_mem' := hs.symm ▸ K.inv_mem' + inv_mem' {x} hx := by simpa [hs] using hx -- porting note: `▸` didn't work here #align subgroup.copy Subgroup.copy #align add_subgroup.copy AddSubgroup.copy From b8ca528c63575ec3f39b591ee7778facc7c90c83 Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Wed, 25 Jan 2023 10:22:36 +0000 Subject: [PATCH 13/34] Fix many of the typeclass issues: unbundled subclasses of `outParam` classes should not repeat the parents' `outParam` --- Mathlib/GroupTheory/Subgroup/Basic.lean | 2 +- Mathlib/GroupTheory/Submonoid/Basic.lean | 8 ++++---- Mathlib/GroupTheory/Subsemigroup/Basic.lean | 4 ++-- 3 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 4362a3c9ac8c0..6f634ea5870a8 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -540,7 +540,7 @@ protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G carrier := s one_mem' := hs.symm ▸ K.one_mem' mul_mem' := hs.symm ▸ K.mul_mem' - inv_mem' {x} hx := by simpa [hs] using hx -- porting note: `▸` didn't work here + inv_mem' hx := by simpa [hs] using hx -- porting note: `▸` didn't work here #align subgroup.copy Subgroup.copy #align add_subgroup.copy AddSubgroup.copy diff --git a/Mathlib/GroupTheory/Submonoid/Basic.lean b/Mathlib/GroupTheory/Submonoid/Basic.lean index 0c16570084036..10675e27d1dcd 100644 --- a/Mathlib/GroupTheory/Submonoid/Basic.lean +++ b/Mathlib/GroupTheory/Submonoid/Basic.lean @@ -72,7 +72,7 @@ variable [MulOneClass M] {s : Set M} variable [AddZeroClass A] {t : Set A} /-- `OneMemClass S M` says `S` is a type of subsets `s ≤ M`, such that `1 ∈ s` for all `s`. -/ -class OneMemClass (S : Type _) (M : outParam <| Type _) [One M] [SetLike S M] : Prop where +class OneMemClass (S : Type _) (M : Type _) [One M] [SetLike S M] : Prop where /-- By definition, if we have `OneMemClass S M`, we have `1 ∈ s` for all `s : S`. -/ one_mem : ∀ s : S, (1 : M) ∈ s #align one_mem_class OneMemClass @@ -80,7 +80,7 @@ class OneMemClass (S : Type _) (M : outParam <| Type _) [One M] [SetLike S M] : export OneMemClass (one_mem) /-- `ZeroMemClass S M` says `S` is a type of subsets `s ≤ M`, such that `0 ∈ s` for all `s`. -/ -class ZeroMemClass (S : Type _) (M : outParam <| Type _) [Zero M] [SetLike S M] : Prop where +class ZeroMemClass (S : Type _) (M : Type _) [Zero M] [SetLike S M] : Prop where /-- By definition, if we have `ZeroMemClass S M`, we have `0 ∈ s` for all `s : S`. -/ zero_mem : ∀ s : S, (0 : M) ∈ s #align zero_mem_class ZeroMemClass @@ -104,7 +104,7 @@ add_decl_doc Submonoid.toSubsemigroup /-- `SubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `1` and are closed under `(*)` -/ -class SubmonoidClass (S : Type _) (M : outParam <| Type _) [MulOneClass M] [SetLike S M] extends +class SubmonoidClass (S : Type _) (M : Type _) [MulOneClass M] [SetLike S M] extends MulMemClass S M, OneMemClass S M : Prop #align submonoid_class SubmonoidClass @@ -125,7 +125,7 @@ add_decl_doc AddSubmonoid.toAddSubsemigroup /-- `AddSubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `0` and are closed under `(+)` -/ -class AddSubmonoidClass (S : Type _) (M : outParam <| Type _) [AddZeroClass M] [SetLike S M] extends +class AddSubmonoidClass (S : Type _) (M : Type _) [AddZeroClass M] [SetLike S M] extends AddMemClass S M, ZeroMemClass S M : Prop #align add_submonoid_class AddSubmonoidClass diff --git a/Mathlib/GroupTheory/Subsemigroup/Basic.lean b/Mathlib/GroupTheory/Subsemigroup/Basic.lean index 6de911354498d..3b793511657cb 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Basic.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Basic.lean @@ -62,7 +62,7 @@ variable [Mul M] {s : Set M} variable [Add A] {t : Set A} /-- `MulMemClass S M` says `S` is a type of sets `s : Set M` that are closed under `(*)` -/ -class MulMemClass (S : Type _) (M : outParam <| Type _) [Mul M] [SetLike S M] : Prop where +class MulMemClass (S : Type _) (M : Type _) [Mul M] [SetLike S M] : Prop where /-- A substructure satisfying `MulMemClass` is closed under multiplication. -/ mul_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a * b ∈ s #align mul_mem_class MulMemClass @@ -70,7 +70,7 @@ class MulMemClass (S : Type _) (M : outParam <| Type _) [Mul M] [SetLike S M] : export MulMemClass (mul_mem) /-- `AddMemClass S M` says `S` is a type of sets `s : Set M` that are closed under `(+)` -/ -class AddMemClass (S : Type _) (M : outParam <| Type _) [Add M] [SetLike S M] : Prop where +class AddMemClass (S : Type _) (M : Type _) [Add M] [SetLike S M] : Prop where /-- A substructure satisfying `AddMemClass` is closed under addition. -/ add_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a + b ∈ s #align add_mem_class AddMemClass From faf49b193ef68a6ea6289b97021d6cbce650b12f Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Wed, 25 Jan 2023 11:33:13 +0100 Subject: [PATCH 14/34] docstrings & minor fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 48 ++++++++++++------------- 1 file changed, 23 insertions(+), 25 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 4362a3c9ac8c0..b273a47cc4224 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -36,11 +36,11 @@ Special thanks goes to Amelia Livingston and Yury Kudryashov for their help and Notation used here: -- `G N` are `group`s +- `G N` are `Group`s -- `A` is an `add_group` +- `A` is an `AddGroup` -- `H K` are `subgroup`s of `G` or `add_subgroup`s of `A` +- `H K` are `Subgroup`s of `G` or `AddSubgroup`s of `A` - `x` is an element of type `G` or type `A` @@ -71,12 +71,12 @@ Definitions in the file: * `Subgroup.prod H K` : the product of subgroups `H`, `K` of groups `G`, `N` respectively, `H × K` is a subgroup of `G × N` -* `monoid_hom.range f` : the range of the group homomorphism `f` is a subgroup +* `MonoidHom.range f` : the range of the group homomorphism `f` is a subgroup -* `monoid_hom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` +* `MonoidHom.ker f` : the kernel of a group homomorphism `f` is the subgroup of elements `x : G` such that `f x = 1` -* `monoid_hom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that +* `MonoidHom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that `f x = g x` form a subgroup of `G` ## Implementation notes @@ -235,7 +235,7 @@ instance (priority := 75) toGroup : Group H := #align subgroup_class.to_group SubgroupClass.toGroup #align add_subgroup_class.to_add_group AddSubgroupClass.toAddGroup --- Prefer subclasses of `comm_group` over subclasses of `subgroup_class`. +-- Prefer subclasses of `comm_group` over subclasses of `Subgroup_class`. /-- A subgroup of a `comm_group` is a `comm_group`. -/ @[to_additive "An additive subgroup of an `add_comm_group` is an `add_comm_group`."] instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [SubgroupClass S G] : @@ -245,7 +245,7 @@ instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [ #align subgroup_class.to_comm_group SubgroupClass.toCommGroup #align add_subgroup_class.to_add_comm_group AddSubgroupClass.toAddCommGroup --- Prefer subclasses of `group` over subclasses of `subgroup_class`. +-- Prefer subclasses of `group` over subclasses of `Subgroup_class`. /-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/ @[to_additive "An additive subgroup of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [SetLike S G] @@ -255,7 +255,7 @@ instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [ #align subgroup_class.to_ordered_comm_group SubgroupClass.toOrderedCommGroup #align add_subgroup_class.to_ordered_add_comm_group AddSubgroupClass.toOrderedAddCommGroup --- Prefer subclasses of `group` over subclasses of `subgroup_class`. +-- Prefer subclasses of `group` over subclasses of `Subgroup_class`. /-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ @[to_additive "An additive subgroup of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] @@ -490,7 +490,7 @@ end Subgroup section mul_add -/-- Supgroups of a group `G` are isomorphic to additive subgroups of `additive G`. -/ +/-- Subgroups of a group `G` are isomorphic to additive subgroups of `Additive G`. -/ @[simps] def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where @@ -646,8 +646,8 @@ def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y have inv_mem : ∀ x, x ∈ s → x⁻¹ ∈ s := fun x hx => by simpa using hs 1 one_mem x hx { carrier := s one_mem' := one_mem - inv_mem' := inv_mem - mul_mem' := fun x y hx hy => by simpa using hs x hx y⁻¹ (inv_mem y hy) } + inv_mem' := inv_mem _ + mul_mem' := fun hx hy => by simpa using hs _ hx _ (inv_mem _ hy) } #align subgroup.of_div Subgroup.ofDiv #align add_subgroup.of_sub AddSubgroup.ofSub @@ -680,7 +680,7 @@ instance hasDiv : Div H := #align add_subgroup.has_sub AddSubgroup.hasSub /-- An `AddSubgroup` of an `AddGroup` inherits a natural scaling. -/ -instance AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := +instance _root_.AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ #align add_subgroup.has_nsmul AddSubgroup.hasNsmul @@ -689,10 +689,9 @@ instance AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H instance hasNpow : Pow H ℕ := ⟨fun a n => ⟨a ^ n, H.pow_mem a.2 n⟩⟩ #align subgroup.has_npow Subgroup.hasNpow -#align add_subgroup.has_nsmul AddSubgroup.hasNsmul /-- An `AddSubgroup` of an `AddGroup` inherits an integer scaling. -/ -instance AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := +instance _root_.AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := ⟨fun n a => ⟨n • a, H.zsmul_mem a.2 n⟩⟩ #align add_subgroup.has_zsmul AddSubgroup.hasZsmul @@ -701,7 +700,6 @@ instance AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H instance hasZpow : Pow H ℤ := ⟨fun a n => ⟨a ^ n, H.zpow_mem a.2 n⟩⟩ #align subgroup.has_zpow Subgroup.hasZpow -#align add_subgroup.has_zsmul AddSubgroup.hasZsmul @[to_additive (attr := simp, norm_cast)] theorem coe_mul (x y : H) : (↑(x * y) : G) = ↑x * ↑y := @@ -734,13 +732,13 @@ theorem coe_mk (x : G) (hx : x ∈ H) : ((⟨x, hx⟩ : H) : G) = x := #align add_subgroup.coe_mk AddSubgroup.coe_mk @[to_additive (attr := simp, norm_cast)] -theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := +theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := rfl #align subgroup.coe_pow Subgroup.coe_pow -#align add_subgroup.coe_nsmul AddSubgroup.coe_nsmul +#align add_subgroup.coe_nsmul AddSubgroup.coe_smul @[to_additive (attr := simp, norm_cast)] -theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := +theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := rfl #align subgroup.coe_zpow Subgroup.coe_zpow #align add_subgroup.coe_zsmul AddSubgroup.coe_zsmul @@ -789,12 +787,12 @@ instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : S /-- The natural group hom from a subgroup of group `G` to `G`. -/ @[to_additive "The natural group hom from an `AddSubgroup` of `AddGroup` `G` to `G`."] def subtype : H →* G := - ⟨coe, rfl, fun _ _ => rfl⟩ + ⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩ #align subgroup.subtype Subgroup.subtype #align add_subgroup.subtype AddSubgroup.subtype @[to_additive (attr := simp)] -theorem coeSubtype : ⇑H.subtype = coe := +theorem coeSubtype : ⇑ H.subtype = ((↑) : H → G) := rfl #align subgroup.coe_subtype Subgroup.coeSubtype #align add_subgroup.coe_subtype AddSubgroup.coeSubtype @@ -807,8 +805,8 @@ theorem subtype_injective : Function.Injective (subtype H) := /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] -def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := - MonoidHom.mk' (fun x => ⟨x, h x.Prop⟩) fun ⟨a, ha⟩ ⟨b, hb⟩ => rfl +def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := by + #align subgroup.inclusion Subgroup.inclusion #align add_subgroup.inclusion AddSubgroup.inclusion @@ -2505,7 +2503,7 @@ theorem normalClosure_closure_eq_normalClosure {s : Set G} : #align subgroup.normal_closure_closure_eq_normal_closure Subgroup.normalClosure_closure_eq_normalClosure /-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, -as shown by `subgroup.normalCore_eq_supᵢ`. -/ +as shown by `Subgroup.normalCore_eq_supᵢ`. -/ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } @@ -3273,7 +3271,7 @@ theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g /-- `liftOfRightInverse f hf g hg` is the unique group homomorphism `φ` -* such that `φ.comp f = g` (`monoid_hom.liftOfRightInverse_comp`), +* such that `φ.comp f = g` (`MonoidHom.liftOfRightInverse_comp`), * where `f : G₁ →+* G₂` has a RightInverse `f_inv` (`hf`), * and `g : G₂ →+* G₃` satisfies `hg : f.ker ≤ g.ker`. From 04f7e2b15b4fed666c4c0f0b87aa08f7f2caf88c Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Wed, 25 Jan 2023 11:40:07 +0100 Subject: [PATCH 15/34] long lines --- Mathlib/GroupTheory/Subgroup/Basic.lean | 148 +++++++++++++++++------- 1 file changed, 106 insertions(+), 42 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 019768340d941..75a336205f97f 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -160,7 +160,8 @@ theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := by #align sub_mem_comm_iff sub_mem_comm_iff @[to_additive (attr := simp)] -theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ P x⁻¹) ↔ ∃ x ∈ H, P x := by +theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : + (∃ x : G, x ∈ H ∧ P x⁻¹) ↔ ∃ x ∈ H, P x := by constructor <;> · rintro ⟨x, x_in, hx⟩ exact ⟨x⁻¹, inv_mem x_in, by simp [hx]⟩ @@ -258,16 +259,22 @@ instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [ -- Prefer subclasses of `group` over subclasses of `Subgroup_class`. /-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ @[to_additive - "An additive subgroup of a `linear_ordered_add_comm_group` is a\n `linear_ordered_add_comm_group`."] + "An additive subgroup of a `linear_ordered_add_comm_group` is a + `linear_ordered_add_comm_group`."] instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : LinearOrderedCommGroup H := Subtype.coe_injective.linearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -#align subgroup_class.to_linear_ordered_comm_group SubgroupClass.toLinearOrderedCommGroup -#align add_subgroup_class.to_linear_ordered_add_comm_group AddSubgroupClass.toLinearOrderedAddCommGroup +#align + subgroup_class.to_linear_ordered_comm_group + SubgroupClass.toLinearOrderedCommGroup +#align + add_subgroup_class.to_linear_ordered_add_comm_group + AddSubgroupClass.toLinearOrderedAddCommGroup /-- The natural group hom from a subgroup of group `G` to `G`. -/ -@[to_additive (attr := coe) "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] +@[to_additive (attr := coe) + "The natural group hom from an additive subgroup of `AddGroup` `G` to `G`."] def subtype : H →* G := ⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩ #align subgroup_class.subtype SubgroupClass.subtype @@ -776,7 +783,8 @@ instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : /-- A subgroup of a `LinearOrderedCommGroup` is a `LinearOrderedCommGroup`. -/ @[to_additive - "An `AddSubgroup` of a `LinearOrderedAddCommGroup` is a\n `LinearOrderedAddCommGroup`."] + "An `AddSubgroup` of a `LinearOrderedAddCommGroup` is a + `LinearOrderedAddCommGroup`."] instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : Subgroup G) : LinearOrderedCommGroup H := Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) @@ -805,7 +813,8 @@ theorem subtype_injective : Function.Injective (subtype H) := /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] -def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := by +def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := + #align subgroup.inclusion Subgroup.inclusion #align add_subgroup.inclusion AddSubgroup.inclusion @@ -839,7 +848,9 @@ instance : Top (Subgroup G) := This is the group version of `Submonoid.topEquiv`. -/ @[to_additive - "The top additive subgroup is isomorphic to the additive group.\n\nThis is the additive group version of `AddSubmonoid.topEquiv`.", + "The top additive subgroup is isomorphic to the additive group. + + This is the additive group version of `AddSubmonoid.topEquiv`.", simps] def topEquiv : (⊤ : Subgroup G) ≃* G := Submonoid.topEquiv @@ -1114,7 +1125,8 @@ theorem not_mem_of_not_mem_closure {P : G} (hP : P ∉ closure k) : P ∉ k := f open Set /-- A subgroup `K` includes `closure k` if and only if it includes `k`. -/ -@[to_additive (attr := simp) "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] +@[to_additive (attr := simp) + "An additive subgroup `K` includes `closure k` if and only if it includes `k`"] theorem closure_le : closure k ≤ K ↔ k ⊆ K := ⟨Subset.trans subset_closure, fun h => infₛ_le h⟩ #align subgroup.closure_le Subgroup.closure_le @@ -1131,7 +1143,9 @@ is preserved under multiplication and inverse, then `p` holds for all elements o of `k`. -/ @[elab_as_elim, to_additive - "An induction principle for additive closure membership. If `p`\nholds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` holds\nfor all elements of the additive closure of `k`."] + "An induction principle for additive closure membership. If `p` + holds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` + holds for all elements of the additive closure of `k`."] theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x ∈ k, p x) (H1 : p 1) (Hmul : ∀ x y, p x → p y → p (x * y)) (Hinv : ∀ x, p x → p x⁻¹) : p x := (@closure_le _ _ ⟨p, Hmul, H1, Hinv⟩ _).2 Hk h @@ -1154,7 +1168,8 @@ theorem closure_induction' {p : ∀ x, x ∈ closure k → Prop} /-- An induction principle for closure membership for predicates with two arguments. -/ @[elab_as_elim, to_additive - "An induction principle for additive closure membership, for\npredicates with two arguments."] + "An induction principle for additive closure membership, for + predicates with two arguments."] theorem closure_induction₂ {p : G → G → Prop} {x} {y : G} (hx : x ∈ closure k) (hy : y ∈ closure k) (Hk : ∀ x ∈ k, ∀ y ∈ k, p x y) (H1_left : ∀ x, p 1 x) (H1_right : ∀ x, p x 1) (Hmul_left : ∀ x₁ x₂ y, p x₁ y → p x₂ y → p (x₁ * x₂) y) @@ -1181,7 +1196,8 @@ theorem closure_closure_coe_preimage {k : Set G} : closure ((coe : closure k → /-- If all the elements of a set `s` commute, then `closure s` is a commutative group. -/ @[to_additive - "If all the elements of a set `s` commute, then `closure s` is an additive\ncommutative group."] + "If all the elements of a set `s` commute, then `closure s` is an additive + commutative group."] def closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * y = y * x) : CommGroup (closure k) := { (closure k).toGroup with @@ -1218,7 +1234,8 @@ variable {G} /-- Subgroup closure of a set is monotone in its argument: if `h ⊆ k`, then `closure h ≤ closure k`. -/ @[to_additive - "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`,\nthen `closure h ≤ closure k`"] + "Additive subgroup closure of a set is monotone in its argument: if `h ⊆ k`, + then `closure h ≤ closure k`"] theorem closure_mono ⦃h k : Set G⦄ (h' : h ⊆ k) : closure h ≤ closure k := (Subgroup.gi G).gc.monotone_l h' #align subgroup.closure_mono Subgroup.closure_mono @@ -1271,7 +1288,8 @@ theorem supᵢ_eq_closure {ι : Sort _} (p : ι → Subgroup G) : /-- The subgroup generated by an element of a group equals the set of integer number powers of the element. -/ @[to_additive - "The `AddSubgroup` generated by an element of an `AddGroup` equals the set of\nnatural number multiples of the element."] + "The `AddSubgroup` generated by an element of an `AddGroup` equals the set of + natural number multiples of the element."] theorem mem_closure_singleton {x y : G} : y ∈ closure ({x} : Set G) ↔ ∃ n : ℤ, x ^ n = y := by refine' ⟨fun hy => closure_induction hy _ _ _ _, fun ⟨n, hn⟩ => @@ -1341,7 +1359,8 @@ variable {N : Type _} [Group N] {P : Type _} [Group P] /-- The preimage of a subgroup along a monoid homomorphism is a subgroup. -/ @[to_additive - "The preimage of an `AddSubgroup` along an `AddMonoid` homomorphism\nis an `AddSubgroup`."] + "The preimage of an `AddSubgroup` along an `AddMonoid` homomorphism + is an `AddSubgroup`."] def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := { H.toSubmonoid.comap f with carrier := f ⁻¹' H @@ -1383,7 +1402,8 @@ theorem comap_id (K : Subgroup N) : K.comap (MonoidHom.id _) = K := by /-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ @[to_additive - "The image of an `AddSubgroup` along an `AddMonoid` homomorphism\nis an `AddSubgroup`."] + "The image of an `AddSubgroup` along an `AddMonoid` homomorphism + is an `AddSubgroup`."] def map (f : G →* N) (H : Subgroup G) : Subgroup N := { H.toSubmonoid.map f with carrier := f '' H @@ -1685,7 +1705,8 @@ theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H /-- Given `Subgroup`s `H`, `K` of groups `G`, `N` respectively, `H × K` as a subgroup of `G × N`. -/ @[to_additive prod - "Given `AddSubgroup`s `H`, `K` of `AddGroup`s `A`, `B` respectively, `H × K`\nas an `AddSubgroup` of `A × B`."] + "Given `AddSubgroup`s `H`, `K` of `AddGroup`s `A`, `B` respectively, `H × K` + as an `AddSubgroup` of `A × B`."] def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := { Submonoid.prod H.toSubmonoid K.toSubmonoid with inv_mem' := fun _ hx => ⟨H.inv_mem' hx.1, K.inv_mem' hx.2⟩ } @@ -1769,7 +1790,8 @@ theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.prod K = ⊥ ↔ H /-- Product of subgroups is isomorphic to their product as groups. -/ @[to_additive prodEquiv - "Product of additive subgroups is isomorphic to their product\nas additive groups"] + "Product of additive subgroups is isomorphic to their product + as additive groups"] def prodEquiv (H : Subgroup G) (K : Subgroup N) : H.prod K ≃* H × K := { Equiv.Set.prod ↑H ↑K with map_mul' := fun x y => rfl } #align subgroup.prod_equiv Subgroup.prodEquiv @@ -1784,7 +1806,9 @@ variable {η : Type _} {f : η → Type _} `s : Π i, Submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that `f i` belongs to `Pi I s` whenever `i ∈ I`. -/ @[to_additive - " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] + " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family + of submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions + `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : Submonoid (∀ i, f i) where carrier := I.pi fun i => (s i).carrier @@ -1799,7 +1823,9 @@ variable [∀ i, Group (f i)] `s : Π i, Subgroup f i`, `pi I s` is the subgroup of dependent functions `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ @[to_additive - " A version of `Set.pi` for `AddSubgroup`s. Given an index set `I` and a family\nof submodules `s : Π i, AddSubgroup f i`, `pi I s` is the `AddSubgroup` of dependent functions\n`f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] + " A version of `Set.pi` for `AddSubgroup`s. Given an index set `I` and a family + of submodules `s : Π i, AddSubgroup f i`, `pi I s` is the `AddSubgroup` of dependent functions + `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def pi (I : Set η) (H : ∀ i, Subgroup (f i)) : Subgroup (∀ i, f i) := { Submonoid.pi I fun i => (H i).toSubmonoid with inv_mem' := fun p hp i hI => (H i).inv_mem (hp i hI) } @@ -1895,7 +1921,7 @@ end Subgroup namespace AddSubgroup /- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: #[`conj_mem] [] -/ -/-- An add_subgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ +/-- An AddSubgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ structure Normal (H : AddSubgroup A) : Prop where conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H #align add_subgroup.normal AddSubgroup.Normal @@ -2036,7 +2062,8 @@ variable (G) /-- The center of a group `G` is the set of elements that commute with everything in `G` -/ @[to_additive - "The center of an additive group `G` is the set of elements that commute with\neverything in `G`"] + "The center of an additive group `G` is the set of elements that commute with + everything in `G`"] def center : Subgroup G := { Submonoid.center G with carrier := Set.center G @@ -2116,7 +2143,8 @@ def normalizer : Subgroup G -- TODO should this replace `normalizer`? /-- The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy `g*S*g⁻¹=S` -/ @[to_additive - "The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy\n`g+S-g=S`."] + "The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy + `g+S-g=S`."] def setNormalizer (S : Set G) : Subgroup G where carrier := { g : G | ∀ n, n ∈ S ↔ g * n * g⁻¹ ∈ S } @@ -2252,7 +2280,8 @@ section Centralizer /-- The `centralizer` of `H` is the subgroup of `g : G` commuting with every `h : H`. -/ @[to_additive - "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with\nevery `h : H`."] + "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with + every `h : H`."] def centralizer : Subgroup G := { Submonoid.centralizer ↑H with carrier := Set.centralizer H @@ -2301,8 +2330,12 @@ instance Subgroup.Centralizer.characteristic [hH : H.Characteristic] : refine' subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.Injective _ rw [map_mul, map_mul] exact hg (ϕ h) (subgroup.characteristic_iff_le_comap.mp hH ϕ hh) -#align subgroup.subgroup.centralizer.characteristic Subgroup.Subgroup.Centralizer.characteristic -#align add_subgroup.subgroup.centralizer.characteristic AddSubgroup.Subgroup.Centralizer.characteristic +#align + subgroup.subgroup.centralizer.characteristic + Subgroup.Subgroup.Centralizer.characteristic +#align + add_subgroup.subgroup.centralizer.characteristic + AddSubgroup.Subgroup.Centralizer.characteristic end Centralizer @@ -2500,7 +2533,9 @@ theorem closure_le_normalClosure {s : Set G} : closure s ≤ normalClosure s := theorem normalClosure_closure_eq_normalClosure {s : Set G} : normalClosure ↑(closure s) = normalClosure s := le_antisymm (normalClosure_le_normal closure_le_normalClosure) (normalClosure_mono subset_closure) -#align subgroup.normal_closure_closure_eq_normal_closure Subgroup.normalClosure_closure_eq_normalClosure +#align + subgroup.normal_closure_closure_eq_normal_closure + Subgroup.normalClosure_closure_eq_normalClosure /-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, as shown by `Subgroup.normalCore_eq_supᵢ`. -/ @@ -2591,7 +2626,8 @@ theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by /-- The canonical surjective group homomorphism `G →* f(G)` induced by a group homomorphism `G →* N`. -/ @[to_additive - "The canonical surjective `AddGroup` homomorphism `G →+ f(G)` induced by a group\nhomomorphism `G →+ N`."] + "The canonical surjective `AddGroup` homomorphism `G →+ f(G)` induced by a group + homomorphism `G →+ N`."] def rangeRestrict (f : G →* N) : G →* f.range := codRestrict f _ fun x => ⟨x, rfl⟩ #align monoid_hom.range_restrict MonoidHom.rangeRestrict @@ -2703,7 +2739,8 @@ theorem ofLeftInverse_symm_apply {f : G →* N} {g : N →* G} (h : Function.Lef #align add_monoid_hom.of_left_inverse_symm_apply AddMonoidHom.ofLeftInverse_symm_apply /-- The range of an injective group homomorphism is isomorphic to its domain. -/ -@[to_additive "The range of an injective additive group homomorphism is isomorphic to its\ndomain."] +@[to_additive "The range of an injective additive group homomorphism is isomorphic to its +domain."] noncomputable def ofInjective {f : G →* N} (hf : Function.Injective f) : G ≃* f.range := MulEquiv.ofBijective (f.codRestrict f.range fun x => ⟨x, rfl⟩) ⟨fun x y h => hf (Subtype.ext_iff.mp h), @@ -2727,7 +2764,8 @@ variable {M : Type _} [MulOneClass M] /-- The multiplicative kernel of a monoid homomorphism is the subgroup of elements `x : G` such that `f x = 1` -/ @[to_additive - "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements\nsuch that `f x = 0`"] + "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements + such that `f x = 0`"] def ker (f : G →* M) : Subgroup G := { f.mker with inv_mem' := fun x (hx : f x = 1) => @@ -2876,7 +2914,8 @@ theorem eqLocus_same (f : G →* N) : f.eqLocus f = ⊤ := /-- If two monoid homomorphisms are equal on a set, then they are equal on its subgroup closure. -/ @[to_additive - "If two monoid homomorphisms are equal on a set, then they are equal on its subgroup\nclosure."] + "If two monoid homomorphisms are equal on a set, then they are equal on its subgroup + closure."] theorem eqOn_closure {f g : G →* M} {s : Set G} (h : Set.EqOn f g s) : Set.EqOn f g (closure s) := show closure s ≤ f.eqLocus g from (closure_le _).2 h #align monoid_hom.eq_on_closure MonoidHom.eqOn_closure @@ -2905,7 +2944,8 @@ theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) /-- The image under a monoid homomorphism of the subgroup generated by a set equals the subgroup generated by the image of the set. -/ @[to_additive - "The image under an `AddMonoid` hom of the `AddSubgroup` generated by a set equals\nthe `AddSubgroup` generated by the image of the set."] + "The image under an `AddMonoid` hom of the `AddSubgroup` generated by a set equals + the `AddSubgroup` generated by the image of the set."] theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := Set.image_preimage.l_comm_of_u_comm (Subgroup.gc_map_comap f) (Subgroup.gi N).gc (Subgroup.gi G).gc fun t => rfl @@ -3152,7 +3192,8 @@ theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : /-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, use `MulEquiv.subgroupMap` for better definitional equalities. -/ @[to_additive - "An additive subgroup is isomorphic to its image under an injective function. If you\nhave an isomorphism, use `AddEquiv.addSubgroupMap` for better definitional equalities."] + "An additive subgroup is isomorphic to its image under an injective function. If you + have an isomorphism, use `AddEquiv.addSubgroupMap` for better definitional equalities."] noncomputable def equivMapOfInjective (H : Subgroup G) (f : G →* N) (hf : Function.Injective f) : H ≃* H.map f := { Equiv.Set.image f H hf with map_mul' := fun _ _ => Subtype.ext (f.map_mul _ _) } @@ -3169,7 +3210,8 @@ theorem coe_equivMapOfInjective_apply (H : Subgroup G) (f : G →* N) (hf : Func /-- The preimage of the normalizer is equal to the normalizer of the preimage of a surjective function. -/ @[to_additive - "The preimage of the normalizer is equal to the normalizer of the preimage of\na surjective function."] + "The preimage of the normalizer is equal to the normalizer of the preimage of + a surjective function."] theorem comap_normalizer_eq_of_surjective (H : Subgroup G) {f : N →* G} (hf : Function.Surjective f) : H.normalizer.comap f = (H.comap f).normalizer := le_antisymm (le_normalizer_comap f) @@ -3211,7 +3253,8 @@ theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : /-- The image of the normalizer is equal to the normalizer of the image of an isomorphism. -/ @[to_additive - "The image of the normalizer is equal to the normalizer of the image of an\nisomorphism."] + "The image of the normalizer is equal to the normalizer of the image of an + isomorphism."] theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : H.normalizer.map f.toMonoidHom = (H.map f.toMonoidHom).normalizer := by ext x @@ -3224,7 +3267,8 @@ theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : /-- The image of the normalizer is equal to the normalizer of the image of a bijective function. -/ @[to_additive - "The image of the normalizer is equal to the normalizer of the image of a bijective\n function."] + "The image of the normalizer is equal to the normalizer of the image of a bijective + function."] theorem map_normalizer_eq_of_bijective (H : Subgroup G) {f : G →* N} (hf : Function.Bijective f) : H.normalizer.map f = (H.map f).normalizer := map_equiv_normalizer_eq H (MulEquiv.ofBijective f hf) @@ -3288,7 +3332,20 @@ See `MonoidHom.eq_liftOfRightInverse` for the uniqueness lemma. ``` -/ @[to_additive - "`liftOfRightInverse f f_inv hf g hg` is the unique additive group homomorphism `φ`\n\n* such that `φ.comp f = g` (`AddMonoidHom.liftOfRightInverse_comp`),\n* where `f : G₁ →+ G₂` has a RightInverse `f_inv` (`hf`),\n* and `g : G₂ →+ G₃` satisfies `hg : f.ker ≤ g.ker`.\n\nSee `AddMonoidHom.eq_liftOfRightInverse` for the uniqueness lemma.\n\n```\n G₁.\n | \\\n f | \\ g\n | \\\n v \\⌟\n G₂----> G₃\n ∃!φ\n```"] + "`liftOfRightInverse f f_inv hf g hg` is the unique additive group homomorphism `φ` + * such that `φ.comp f = g` (`AddMonoidHom.liftOfRightInverse_comp`), + * where `f : G₁ →+ G₂` has a RightInverse `f_inv` (`hf`), + * and `g : G₂ →+ G₃` satisfies `hg : f.ker ≤ g.ker`. + See `AddMonoidHom.eq_liftOfRightInverse` for the uniqueness lemma. + ``` + G₁. + | \\ + f | \\ g + | \\ + v \\⌟ + G₂----> G₃ + ∃!φ + ```"] def liftOfRightInverse (hf : Function.RightInverse f_inv f) : { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) where @@ -3307,7 +3364,8 @@ def liftOfRightInverse (hf : Function.RightInverse f_inv f) : inverse is available, that uses `Function.surjInv`. -/ @[simp, to_additive - "A non-computable version of `AddMonoidHom.liftOfRightInverse` for when no\ncomputable right inverse is available."] + "A non-computable version of `AddMonoidHom.liftOfRightInverse` for when no + computable right inverse is available."] noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : { g : G₁ →* G₃ // f.ker ≤ g.ker } ≃ (G₂ →* G₃) := f.liftOfRightInverse (Function.surjInv hf) (Function.rightInverse_surjInv hf) @@ -3403,7 +3461,8 @@ variable {H K : Subgroup G} /-- Makes the identity isomorphism from a proof two subgroups of a multiplicative group are equal. -/ @[to_additive - "Makes the identity additive isomorphism from a proof\ntwo subgroups of an additive group are equal."] + "Makes the identity additive isomorphism from a proof + two subgroups of an additive group are equal."] def subgroupCongr (h : H = K) : H ≃* K := { Equiv.setCongr <| congr_arg _ h with map_mul' := fun _ _ => rfl } #align mul_equiv.subgroup_congr MulEquiv.subgroupCongr @@ -3412,7 +3471,8 @@ def subgroupCongr (h : H = K) : H ≃* K := /-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, use `Subgroup.equiv_map_of_injective`. -/ @[to_additive - "An additive subgroup is isomorphic to its image under an an isomorphism. If you only\nhave an injective map, use `AddSubgroup.equiv_map_of_injective`."] + "An additive subgroup is isomorphic to its image under an an isomorphism. If you only + have an injective map, use `AddSubgroup.equiv_map_of_injective`."] def subgroupMap (e : G ≃* G') (H : Subgroup G) : H ≃* H.map (e : G →* G') := MulEquiv.submonoidMap (e : G ≃* G') H.toSubmonoid #align mul_equiv.subgroup_map MulEquiv.subgroupMap @@ -3529,8 +3589,12 @@ theorem inf_subgroupOf_inf_normal_of_right (A B' B : Subgroup G) (hB : B' ≤ B) conj_mem := fun n hn g => ⟨mul_mem (mul_mem (mem_inf.1 g.2).1 (mem_inf.1 n.2).1) (inv_mem (mem_inf.1 g.2).1), (normal_subgroupOf_iff hB).mp hN n g hn.2 (mem_inf.mp g.2).2⟩ } -#align subgroup.inf_subgroup_of_inf_normal_of_right Subgroup.inf_subgroupOf_inf_normal_of_right -#align add_subgroup.inf_add_subgroup_of_inf_normal_of_right AddSubgroup.inf_addSubgroupOf_inf_normal_of_right +#align + subgroup.inf_subgroup_of_inf_normal_of_right + Subgroup.inf_subgroupOf_inf_normal_of_right +#align + add_subgroup.inf_add_subgroup_of_inf_normal_of_right + AddSubgroup.inf_addSubgroupOf_inf_normal_of_right @[to_additive] theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) (hA : A' ≤ A) From e3db17bf7cecf7921e7a6b3e23630ca607a8fc2c Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Wed, 25 Jan 2023 11:41:21 +0100 Subject: [PATCH 16/34] restore Subgroup.inclusion --- Mathlib/GroupTheory/Subgroup/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 75a336205f97f..ea4bb5cd560f0 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -814,7 +814,7 @@ theorem subtype_injective : Function.Injective (subtype H) := /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] def inclusion {H K : Subgroup G} (h : H ≤ K) : H →* K := - + MonoidHom.mk' (fun x => ⟨x, h x.2⟩) fun _ _ => rfl #align subgroup.inclusion Subgroup.inclusion #align add_subgroup.inclusion AddSubgroup.inclusion From 8923ba1d8a5a9bf68c176b100588c36df4f0fa36 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Wed, 25 Jan 2023 11:49:31 +0100 Subject: [PATCH 17/34] some redundant arguments --- Mathlib/GroupTheory/Subgroup/Basic.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index ea4bb5cd560f0..94ee1752e9027 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -21,7 +21,7 @@ import Mathlib.Tactic.ApplyFun # Subgroups This file defines multiplicative and additive subgroups as an extension of submonoids, in a bundled -form (unbundled subgroups are in `deprecated/subgroups.lean`). +form (unbundled subgroups are in `Deprecated/Subgroups.lean`). We prove subgroups of a group form a complete lattice, and results about images and preimages of subgroups under group homomorphisms. The bundled subgroups use bundled monoid homomorphisms. @@ -834,7 +834,7 @@ theorem inclusion_injective {H K : Subgroup G} (h : H ≤ K) : Function.Injectiv @[to_additive (attr := simp)] theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : - K.Subtype.comp (inclusion hH) = H.Subtype := + K.subtype.comp (inclusion hH) = H.subtype := rfl #align subgroup.subtype_comp_inclusion Subgroup.subtype_comp_inclusion #align add_subgroup.subtype_comp_inclusion AddSubgroup.subtype_comp_inclusion @@ -842,7 +842,7 @@ theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : /-- The subgroup `G` of the group `G`. -/ @[to_additive "The `AddSubgroup G` of the `AddGroup G`."] instance : Top (Subgroup G) := - ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ _ => Set.mem_univ _ }⟩ + ⟨{ (⊤ : Submonoid G) with inv_mem' := fun _ => Set.mem_univ _ }⟩ /-- The top subgroup is isomorphic to the group. @@ -860,7 +860,7 @@ def topEquiv : (⊤ : Subgroup G) ≃* G := /-- The trivial subgroup `{1}` of an group `G`. -/ @[to_additive "The trivial `AddSubgroup` `{0}` of an `AddGroup` `G`."] instance : Bot (Subgroup G) := - ⟨{ (⊥ : Submonoid G) with inv_mem' := fun _ => by simp [*] }⟩ + ⟨{ (⊥ : Submonoid G) with inv_mem' := by simp}⟩ @[to_additive] instance : Inhabited (Subgroup G) := @@ -969,7 +969,7 @@ theorem bot_or_exists_ne_one (H : Subgroup G) : H = ⊥ ∨ ∃ x ∈ H, x ≠ ( instance : HasInf (Subgroup G) := ⟨fun H₁ H₂ => { H₁.toSubmonoid ⊓ H₂.toSubmonoid with - inv_mem' := fun _ ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ + inv_mem' := fun ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ @[to_additive (attr := simp)] theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = p ∩ p' := From e604b3b993d494821682ff2338d4f4ba1522e7f8 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 13:56:18 +0000 Subject: [PATCH 18/34] minifix --- Mathlib/GroupTheory/Subgroup/Basic.lean | 17 +++++++++-------- 1 file changed, 9 insertions(+), 8 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 94ee1752e9027..efbad83e3eabd 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -292,7 +292,7 @@ variable {H} theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := (subtype H : H →* G).map_pow _ _ #align subgroup_class.coe_pow SubgroupClass.coe_pow -#align add_subgroup_class.coe_smul AddSubgroupClass.coe_smul +#align add_subgroup_class.coe_smul AddSubgroupClass.coe_nsmul @[to_additive (attr := simp, norm_cast)] theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := @@ -759,7 +759,7 @@ theorem mk_eq_one_iff {g : G} {h} : (⟨g, h⟩ : H) = 1 ↔ g = 1 := /-- A subgroup of a group inherits a group structure. -/ @[to_additive "An `AddSubgroup` of an `AddGroup` inherits an `AddGroup` structure."] instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := - Subtype.coe_injective.Group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup.to_group Subgroup.toGroup #align add_subgroup.to_add_group AddSubgroup.toAddGroup @@ -767,7 +767,7 @@ instance toGroup {G : Type _} [Group G] (H : Subgroup G) : Group H := /-- A subgroup of a `CommGroup` is a `CommGroup`. -/ @[to_additive "An `AddSubgroup` of an `AddCommGroup` is an `AddCommGroup`."] instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H := - Subtype.coe_injective.CommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.commGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup.to_comm_group Subgroup.toCommGroup #align add_subgroup.to_add_comm_group AddSubgroup.toAddCommGroup @@ -776,7 +776,7 @@ instance toCommGroup {G : Type _} [CommGroup G] (H : Subgroup G) : CommGroup H : @[to_additive "An `AddSubgroup` of an `AddOrderedCommGroup` is an `AddOrderedCommGroup`."] instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : OrderedCommGroup H := - Subtype.coe_injective.OrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + Subtype.coe_injective.orderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup.to_ordered_comm_group Subgroup.toOrderedCommGroup #align add_subgroup.to_ordered_add_comm_group AddSubgroup.toOrderedAddCommGroup @@ -787,7 +787,7 @@ instance toOrderedCommGroup {G : Type _} [OrderedCommGroup G] (H : Subgroup G) : `LinearOrderedAddCommGroup`."] instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : Subgroup G) : LinearOrderedCommGroup H := - Subtype.coe_injective.LinearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) + Subtype.coe_injective.linearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup.to_linear_ordered_comm_group Subgroup.toLinearOrderedCommGroup #align add_subgroup.to_linear_ordered_add_comm_group AddSubgroup.toLinearOrderedAddCommGroup @@ -938,8 +938,9 @@ theorem coe_eq_singleton {H : Subgroup G} : (∃ g : G, (H : Set G) = {g}) ↔ H #align add_subgroup.coe_eq_singleton AddSubgroup.coe_eq_singleton @[to_additive] -theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x ∈ H, x ≠ (1 : G) := - Subtype.nontrivial_iff_exists_ne (fun x => x ∈ H) (1 : H) +theorem nontrivial_iff_exists_ne_one (H : Subgroup G) : Nontrivial H ↔ ∃ x ∈ H, x ≠ (1 : G) := by + rw [Subtype.nontrivial_iff_exists_ne (fun x => x ∈ H) (1 : H)] + simp #align subgroup.nontrivial_iff_exists_ne_one Subgroup.nontrivial_iff_exists_ne_one #align add_subgroup.nontrivial_iff_exists_ne_zero AddSubgroup.nontrivial_iff_exists_ne_zero @@ -952,7 +953,7 @@ theorem bot_or_nontrivial (H : Subgroup G) : H = ⊥ ∨ Nontrivial H := by exact H.eq_bot_iff_forall.mpr h · right simp only [not_forall] at h - simpa only [nontrivial_iff_exists_ne_one] + simpa [nontrivial_iff_exists_ne_one] using h #align subgroup.bot_or_nontrivial Subgroup.bot_or_nontrivial #align add_subgroup.bot_or_nontrivial AddSubgroup.bot_or_nontrivial From af42022067ad91a0f0af73b5f82261f0f1708034 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 15:06:34 +0100 Subject: [PATCH 19/34] fix some errors --- Mathlib/GroupTheory/Subgroup/Basic.lean | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index efbad83e3eabd..30cecfda5418f 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -629,7 +629,7 @@ protected theorem mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ #align subgroup.mul_mem_cancel_left Subgroup.mul_mem_cancel_left #align add_subgroup.add_mem_cancel_left AddSubgroup.add_mem_cancel_left -@[to_additive AddSubgroup.nsmul_mem] +@[to_additive] protected theorem pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := pow_mem hx #align subgroup.pow_mem Subgroup.pow_mem @@ -687,9 +687,9 @@ instance hasDiv : Div H := #align add_subgroup.has_sub AddSubgroup.hasSub /-- An `AddSubgroup` of an `AddGroup` inherits a natural scaling. -/ -instance _root_.AddSubgroup.hasNsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := +instance _root_.AddSubgroup.hasNSMul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ -#align add_subgroup.has_nsmul AddSubgroup.hasNsmul +#align add_subgroup.has_nsmul AddSubgroup.hasNSMul /-- A subgroup of a group inherits a natural power -/ @[to_additive] @@ -742,7 +742,7 @@ theorem coe_mk (x : G) (hx : x ∈ H) : ((⟨x, hx⟩ : H) : G) = x := theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := rfl #align subgroup.coe_pow Subgroup.coe_pow -#align add_subgroup.coe_nsmul AddSubgroup.coe_smul +#align add_subgroup.coe_nsmul AddSubgroup.coe_nsmul @[to_additive (attr := simp, norm_cast)] theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := @@ -973,7 +973,7 @@ instance : HasInf (Subgroup G) := inv_mem' := fun ⟨hx, hx'⟩ => ⟨H₁.inv_mem hx, H₂.inv_mem hx'⟩ }⟩ @[to_additive (attr := simp)] -theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = p ∩ p' := +theorem coe_inf (p p' : Subgroup G) : ((p ⊓ p' : Subgroup G) : Set G) = (p : Set G) ∩ p' := rfl #align subgroup.coe_inf Subgroup.coe_inf #align add_subgroup.coe_inf AddSubgroup.coe_inf @@ -1113,7 +1113,7 @@ theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K /-- The subgroup generated by a set includes the set. -/ @[to_additive (attr := simp) "The `AddSubgroup` generated by a set includes the set."] -theorem subset_closure : k ⊆ closure k := fun x hx => mem_closure.2 fun _ hK => hK hx +theorem subset_closure : k ⊆ closure k := fun _ hx => mem_closure.2 fun _ hK => hK hx #align subgroup.subset_closure Subgroup.subset_closure #align add_subgroup.subset_closure AddSubgroup.subset_closure @@ -1183,7 +1183,7 @@ theorem closure_induction₂ {p : G → G → Prop} {x} {y : G} (hx : x ∈ clos #align add_subgroup.closure_induction₂ AddSubgroup.closure_induction₂ @[to_additive (attr := simp)] -theorem closure_closure_coe_preimage {k : Set G} : closure ((coe : closure k → G) ⁻¹' k) = ⊤ := +theorem closure_closure_coe_preimage {k : Set G} : closure (((↑) : closure k → G) ⁻¹' k) = ⊤ := eq_top_iff.2 fun x => Subtype.recOn x fun x hx _ => by @@ -1321,7 +1321,7 @@ theorem le_closure_toSubmonoid (S : Set G) : Submonoid.closure S ≤ (closure S) @[to_additive] theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = ⊤) : closure S = ⊤ := - (eq_top_iff' _).2 fun x => le_closure_toSubmonoid _ <| h.symm ▸ trivial + (eq_top_iff' _).2 fun _ => le_closure_toSubmonoid _ <| h.symm ▸ trivial #align subgroup.closure_eq_top_of_mclosure_eq_top Subgroup.closure_eq_top_of_mclosure_eq_top #align add_subgroup.closure_eq_top_of_mclosure_eq_top AddSubgroup.closure_eq_top_of_mclosure_eq_top @@ -1581,7 +1581,7 @@ theorem map_bot (f : G →* N) : (⊥ : Subgroup G).map f = ⊥ := @[to_additive (attr := simp)] theorem map_top_of_surjective (f : G →* N) (h : Function.Surjective f) : Subgroup.map f ⊤ = ⊤ := by rw [eq_top_iff] - intro x hx + intro x _ obtain ⟨y, hy⟩ := h x exact ⟨y, trivial, hy⟩ #align subgroup.map_top_of_surjective Subgroup.map_top_of_surjective @@ -2049,13 +2049,13 @@ theorem characteristic_iff_le_map : H.Characteristic ↔ ∀ ϕ : G ≃* G, H @[to_additive] instance botCharacteristic : Characteristic (⊥ : Subgroup G) := - characteristic_iff_le_map.mpr fun ϕ => bot_le + characteristic_iff_le_map.mpr fun _ϕ => bot_le #align subgroup.bot_characteristic Subgroup.botCharacteristic #align add_subgroup.bot_characteristic AddSubgroup.botCharacteristic @[to_additive] instance topCharacteristic : Characteristic (⊤ : Subgroup G) := - characteristic_iff_map_le.mpr fun ϕ => le_top + characteristic_iff_map_le.mpr fun _ϕ => le_top #align subgroup.top_characteristic Subgroup.topCharacteristic #align add_subgroup.top_characteristic AddSubgroup.topCharacteristic From 7f152a1e66fbfce4fd83cb112e3fcb82c1e99185 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 15:07:54 +0100 Subject: [PATCH 20/34] change instance names --- Mathlib/GroupTheory/Subgroup/Basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 30cecfda5418f..88d12331869b4 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -687,15 +687,15 @@ instance hasDiv : Div H := #align add_subgroup.has_sub AddSubgroup.hasSub /-- An `AddSubgroup` of an `AddGroup` inherits a natural scaling. -/ -instance _root_.AddSubgroup.hasNSMul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := +instance _root_.AddSubgroup.instNSMul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ -#align add_subgroup.has_nsmul AddSubgroup.hasNSMul +#align add_subgroup.has_nsmul AddSubgroup.instNSMul /-- A subgroup of a group inherits a natural power -/ @[to_additive] -instance hasNpow : Pow H ℕ := +protected instance instNPow : Pow H ℕ := ⟨fun a n => ⟨a ^ n, H.pow_mem a.2 n⟩⟩ -#align subgroup.has_npow Subgroup.hasNpow +#align subgroup.has_npow Subgroup.instNPow /-- An `AddSubgroup` of an `AddGroup` inherits an integer scaling. -/ instance _root_.AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := From 600b4fa789d80bbe7dc78ec0e859efa88c3556c6 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 14:57:09 +0000 Subject: [PATCH 21/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 104 +++++++++++------------- 1 file changed, 49 insertions(+), 55 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 88d12331869b4..00c7d7aff3def 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -1221,7 +1221,7 @@ variable (G) /-- `closure` forms a Galois insertion with the coercion to set. -/ @[to_additive "`closure` forms a Galois insertion with the coercion to set."] -protected def gi : GaloisInsertion (@closure G _) coe +protected def gi : GaloisInsertion (@closure G _) (↑) where choice s _ := closure s gc s t := @closure_le _ _ t s @@ -1269,7 +1269,7 @@ theorem closure_union (s t : Set G) : closure (s ∪ t) = closure s ⊔ closure @[to_additive] theorem closure_unionᵢ {ι} (s : ι → Set G) : closure (⋃ i, s i) = ⨆ i, closure (s i) := - (Subgroup.gi G).gc.l_supr + (Subgroup.gi G).gc.l_supᵢ #align subgroup.closure_Union Subgroup.closure_unionᵢ #align add_subgroup.closure_Union AddSubgroup.closure_unionᵢ @@ -1316,7 +1316,7 @@ theorem closure_singleton_one : closure ({1} : Set G) = ⊥ := by theorem le_closure_toSubmonoid (S : Set G) : Submonoid.closure S ≤ (closure S).toSubmonoid := Submonoid.closure_le.2 subset_closure #align subgroup.le_closure_to_submonoid Subgroup.le_closure_toSubmonoid -#align add_subgroup.le_closure_to_add_submonoid AddSubgroup.le_closure_to_add_submonoid +#align add_subgroup.le_closure_to_add_submonoid AddSubgroup.le_closure_toAddSubmonoid @[to_additive] theorem closure_eq_top_of_mclosure_eq_top {S : Set G} (h : Submonoid.closure S = ⊤) : @@ -1365,7 +1365,7 @@ variable {N : Type _} [Group N] {P : Type _} [Group P] def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := { H.toSubmonoid.comap f with carrier := f ⁻¹' H - inv_mem' := fun a ha => show f a⁻¹ ∈ H by rw [f.map_inv] <;> exact H.inv_mem ha } + inv_mem' := fun {a} ha => show f a⁻¹ ∈ H by rw [f.map_inv] <;> exact H.inv_mem ha } #align subgroup.comap Subgroup.comap #align add_subgroup.comap AddSubgroup.comap @@ -1710,7 +1710,7 @@ theorem subgroupOf_eq_top {H K : Subgroup G} : H.subgroupOf K = ⊤ ↔ K ≤ H as an `AddSubgroup` of `A × B`."] def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := { Submonoid.prod H.toSubmonoid K.toSubmonoid with - inv_mem' := fun _ hx => ⟨H.inv_mem' hx.1, K.inv_mem' hx.2⟩ } + inv_mem' := fun hx => ⟨H.inv_mem' hx.1, K.inv_mem' hx.2⟩ } #align subgroup.prod Subgroup.prod #align add_subgroup.prod AddSubgroup.prod @@ -1794,7 +1794,7 @@ theorem prod_eq_bot_iff {H : Subgroup G} {K : Subgroup N} : H.prod K = ⊥ ↔ H "Product of additive subgroups is isomorphic to their product as additive groups"] def prodEquiv (H : Subgroup G) (K : Subgroup N) : H.prod K ≃* H × K := - { Equiv.Set.prod ↑H ↑K with map_mul' := fun x y => rfl } + { Equiv.Set.prod (H : Set G) (K : Set N) with map_mul' := fun _ _ => rfl } #align subgroup.prod_equiv Subgroup.prodEquiv #align add_subgroup.prod_equiv AddSubgroup.prodEquiv @@ -1810,11 +1810,11 @@ variable {η : Type _} {f : η → Type _} " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family of submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] -def Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : +def _root_.Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : Submonoid (∀ i, f i) where carrier := I.pi fun i => (s i).carrier one_mem' i _ := (s i).one_mem - mul_mem' p q hp hq i hI := (s i).mul_mem (hp i hI) (hq i hI) + mul_mem' hp hq i hI := (s i).mul_mem (hp i hI) (hq i hI) #align submonoid.pi Submonoid.pi #align add_submonoid.pi AddSubmonoid.pi @@ -1829,7 +1829,7 @@ variable [∀ i, Group (f i)] `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def pi (I : Set η) (H : ∀ i, Subgroup (f i)) : Subgroup (∀ i, f i) := { Submonoid.pi I fun i => (H i).toSubmonoid with - inv_mem' := fun p hp i hI => (H i).inv_mem (hp i hI) } + inv_mem' := fun hp i hI => (H i).inv_mem (hp i hI) } #align subgroup.pi Subgroup.pi #align add_subgroup.pi AddSubgroup.pi @@ -1889,9 +1889,9 @@ theorem mulSingle_mem_pi [DecidableEq η] {I : Set η} {H : ∀ i, Subgroup (f i simpa using h i hi · intro h j hj by_cases heq : j = i - · subst HEq + · subst heq simpa using h hj - · simp [HEq, one_mem] + · simp [heq, one_mem] #align subgroup.mul_single_mem_pi Subgroup.mulSingle_mem_pi #align add_subgroup.single_mem_pi AddSubgroup.single_mem_pi @@ -1902,7 +1902,7 @@ theorem pi_eq_bot_iff (H : ∀ i, Subgroup (f i)) : pi Set.univ H = ⊥ ↔ ∀ constructor · intro h i x hx have : MonoidHom.single f i x = 1 := - h (MonoidHom.single f i x) ((mul_single_mem_pi i x).mpr fun _ => hx) + h (MonoidHom.single f i x) ((mulSingle_mem_pi i x).mpr fun _ => hx) simpa using congr_fun this i · exact fun h x hx => funext fun i => h _ _ (hx i trivial) #align subgroup.pi_eq_bot_iff Subgroup.pi_eq_bot_iff @@ -2068,7 +2068,7 @@ variable (G) def center : Subgroup G := { Submonoid.center G with carrier := Set.center G - inv_mem' := fun a => Set.inv_mem_center } + inv_mem' := Set.inv_mem_center } #align subgroup.center Subgroup.center #align add_subgroup.center AddSubgroup.center @@ -2104,14 +2104,14 @@ instance centerCharacteristic : (center G).Characteristic := by #align subgroup.center_characteristic Subgroup.centerCharacteristic #align add_subgroup.center_characteristic AddSubgroup.centerCharacteristic -theorem CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := by +theorem _root_.CommGroup.center_eq_top {G : Type _} [CommGroup G] : center G = ⊤ := by rw [eq_top_iff'] intro x y exact mul_comm y x #align comm_group.center_eq_top CommGroup.center_eq_top /-- A group is commutative if the center is the whole group -/ -def Group.commGroupOfCenterEqTop (h : center G = ⊤) : CommGroup G := +def _root_.Group.commGroupOfCenterEqTop (h : center G = ⊤) : CommGroup G := { (_ : Group G) with mul_comm := by rw [eq_top_iff'] at h @@ -2119,24 +2119,21 @@ def Group.commGroupOfCenterEqTop (h : center G = ⊤) : CommGroup G := exact h y x } #align group.comm_group_of_center_eq_top Group.commGroupOfCenterEqTop -variable {G} (H) +variable (H) section Normalizer /-- 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."] -def normalizer : Subgroup G - where +def normalizer : Subgroup G where carrier := { g : G | ∀ n, n ∈ H ↔ g * n * g⁻¹ ∈ H } one_mem' := by simp - mul_mem' a b (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) (hb : ∀ n, n ∈ H ↔ b * n * b⁻¹ ∈ H) n := - by + mul_mem' {a b} (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) (hb : ∀ n, n ∈ H ↔ b * n * b⁻¹ ∈ H) n := by rw [hb, ha] - simp [mul_assoc] - inv_mem' a (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) n := - by + simp only [mul_assoc, mul_inv_rev] + inv_mem' {a} (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) n := by rw [ha (a⁻¹ * n * a⁻¹⁻¹)] - simp [mul_assoc] + simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_right_inv, mul_one] #align subgroup.normalizer Subgroup.normalizer #align add_subgroup.normalizer AddSubgroup.normalizer @@ -2146,18 +2143,15 @@ def normalizer : Subgroup G @[to_additive "The `setNormalizer` of `S` is the subgroup of `G` whose elements satisfy `g+S-g=S`."] -def setNormalizer (S : Set G) : Subgroup G - where +def setNormalizer (S : Set G) : Subgroup G where carrier := { g : G | ∀ n, n ∈ S ↔ g * n * g⁻¹ ∈ S } one_mem' := by simp - mul_mem' a b (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) (hb : ∀ n, n ∈ S ↔ b * n * b⁻¹ ∈ S) n := - by + mul_mem' {a b} (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) (hb : ∀ n, n ∈ S ↔ b * n * b⁻¹ ∈ S) n := by rw [hb, ha] - simp [mul_assoc] - inv_mem' a (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n := - by + simp only [mul_assoc, mul_inv_rev] + inv_mem' {a} (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n := by rw [ha (a⁻¹ * n * a⁻¹⁻¹)] - simp [mul_assoc] + simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_right_inv, mul_one] #align subgroup.set_normalizer Subgroup.setNormalizer #align add_subgroup.set_normalizer AddSubgroup.setNormalizer @@ -2190,7 +2184,7 @@ theorem le_normalizer : H ≤ normalizer H := fun x xH n => by @[to_additive] instance (priority := 100) normal_in_normalizer : (H.subgroupOf H.normalizer).Normal := - ⟨fun x xH g => by simpa using (g.2 x).1 xH⟩ + ⟨fun x xH g => by simpa only [mem_subgroupOf] using (g.2 x.1).1 xH⟩ #align subgroup.normal_in_normalizer Subgroup.normal_in_normalizer #align add_subgroup.normal_in_normalizer AddSubgroup.normal_in_normalizer @@ -2214,7 +2208,7 @@ open Classical theorem le_normalizer_of_normal [hK : (H.subgroupOf K).Normal] (HK : H ≤ K) : K ≤ H.normalizer := fun x hx y => ⟨fun yH => hK.conj_mem ⟨y, HK yH⟩ yH ⟨x, hx⟩, fun yH => by - simpa [mem_subgroup_of, mul_assoc] using + simpa [mem_subgroupOf, mul_assoc] using hK.conj_mem ⟨x * y * x⁻¹, HK yH⟩ yH ⟨x⁻¹, K.inv_mem hx⟩⟩ #align subgroup.le_normalizer_of_normal Subgroup.le_normalizer_of_normal #align add_subgroup.le_normalizer_of_normal AddSubgroup.le_normalizer_of_normal @@ -2250,7 +2244,7 @@ theorem le_normalizer_map (f : G →* N) : H.normalizer.map f ≤ (H.map f).norm variable (G) /-- Every proper subgroup `H` of `G` is a proper normal subgroup of the normalizer of `H` in `G`. -/ -def NormalizerCondition := +def _root_.NormalizerCondition := ∀ H : Subgroup G, H < ⊤ → H < normalizer H #align normalizer_condition NormalizerCondition @@ -2258,7 +2252,7 @@ variable {G} /-- Alternative phrasing of the normalizer condition: Only the full group is self-normalizing. This may be easier to work with, as it avoids inequalities and negations. -/ -theorem normalizerCondition_iff_only_full_group_self_normalizing : +theorem _root_.normalizerCondition_iff_only_full_group_self_normalizing : NormalizerCondition G ↔ ∀ H : Subgroup G, H.normalizer = H → H = ⊤ := by apply forall_congr'; intro H simp only [lt_iff_le_and_ne, le_normalizer, true_and_iff, le_top, Ne.def] @@ -2284,9 +2278,9 @@ section Centralizer "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with every `h : H`."] def centralizer : Subgroup G := - { Submonoid.centralizer ↑H with + { Submonoid.centralizer (H : Set G) with carrier := Set.centralizer H - inv_mem' := fun g => Set.inv_mem_centralizer } + inv_mem' := Set.inv_mem_centralizer } #align subgroup.centralizer Subgroup.centralizer #align add_subgroup.centralizer AddSubgroup.centralizer @@ -2328,9 +2322,9 @@ theorem centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := @[to_additive] instance Subgroup.Centralizer.characteristic [hH : H.Characteristic] : H.centralizer.Characteristic := by - refine' subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.Injective _ + refine' Subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.injective _ rw [map_mul, map_mul] - exact hg (ϕ h) (subgroup.characteristic_iff_le_comap.mp hH ϕ hh) + exact hg (ϕ h) (Subgroup.characteristic_iff_le_comap.mp hH ϕ hh) #align subgroup.subgroup.centralizer.characteristic Subgroup.Subgroup.Centralizer.characteristic @@ -2348,8 +2342,8 @@ structure IsCommutative : Prop where attribute [class] IsCommutative /-- Commutivity of an additive subgroup -/ -structure AddSubgroup.IsCommutative (H : AddSubgroup A) : Prop where - is_comm : IsCommutative H (· + ·) +structure _root_.AddSubgroup.IsCommutative (H : AddSubgroup A) : Prop where + is_comm : _root_.IsCommutative H (· + ·) #align add_subgroup.is_commutative AddSubgroup.IsCommutative attribute [to_additive] Subgroup.IsCommutative @@ -2372,7 +2366,7 @@ instance map_isCommutative (f : G →* G') [H.IsCommutative] : (H.map f).IsCommu ⟨⟨by rintro ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩ rw [Subtype.ext_iff, coe_mul, coe_mul, Subtype.coe_mk, Subtype.coe_mk, ← map_mul, ← map_mul] - exact congr_arg f (subtype.ext_iff.mp (mul_comm ⟨a, ha⟩ ⟨b, hb⟩))⟩⟩ + exact congr_arg f (Subtype.ext_iff.mp (mul_comm (⟨a, ha⟩ : H) ⟨b, hb⟩))⟩⟩ #align subgroup.map_is_commutative Subgroup.map_isCommutative #align add_subgroup.map_is_commutative AddSubgroup.map_isCommutative @@ -2390,14 +2384,14 @@ theorem comap_injective_isCommutative {f : G' →* G} (hf : Injective f) [H.IsCo @[to_additive] instance subgroupOf_isCommutative [H.IsCommutative] : (H.subgroupOf K).IsCommutative := - H.comap_injective_is_commutative Subtype.coe_injective + H.comap_injective_isCommutative Subtype.coe_injective #align subgroup.subgroup_of_is_commutative Subgroup.subgroupOf_isCommutative -#align add_subgroup.add_subgroup_of_is_commutative AddSubgroup.add_subgroupOf_isCommutative +#align add_subgroup.add_subgroup_of_is_commutative AddSubgroup.addSubgroupOf_isCommutative @[to_additive] theorem le_centralizer_iff_isCommutative : K ≤ K.centralizer ↔ K.IsCommutative := - ⟨fun h => ⟨⟨fun x y => Subtype.ext (h y.2 x x.2)⟩⟩, fun h x hx y hy => - congr_arg coe (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ + ⟨fun h => ⟨⟨fun x y => Subtype.ext (h y.2 x x.2)⟩⟩, + fun h x hx y hy => congr_arg Subtype.val (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ #align subgroup.le_centralizer_iff_is_commutative Subgroup.le_centralizer_iff_isCommutative #align add_subgroup.le_centralizer_iff_is_commutative AddSubgroup.le_centralizer_iff_isCommutative @@ -2446,8 +2440,8 @@ theorem conjugatesOfSet_subset {s : Set G} {N : Subgroup G} [N.Normal] (h : s /-- The set of conjugates of `s` is closed under conjugation. -/ theorem conj_mem_conjugatesOfSet {x c : G} : x ∈ conjugatesOfSet s → c * x * c⁻¹ ∈ conjugatesOfSet s := fun H => by - rcases mem_conjugates_of_set_iff.1 H with ⟨a, h₁, h₂⟩ - exact mem_conjugates_of_set_iff.2 ⟨a, h₁, h₂.trans (isConj_iff.2 ⟨c, rfl⟩)⟩ + rcases mem_conjugatesOfSet_iff.1 H with ⟨a, h₁, h₂⟩ + exact mem_conjugatesOfSet_iff.2 ⟨a, h₁, h₂.trans (isConj_iff.2 ⟨c, rfl⟩)⟩ #align group.conj_mem_conjugates_of_set Group.conj_mem_conjugatesOfSet end Group @@ -2483,7 +2477,7 @@ instance normalClosure_normal : (normalClosure s).Normal := ⟨fun n h g => by refine' Subgroup.closure_induction h (fun x hx => _) _ (fun x y ihx ihy => _) fun x ihx => _ - · exact conjugatesOfSet_subset_normalClosure (conj_mem_conjugates_of_set hx) + · exact conjugatesOfSet_subset_normalClosure (conj_mem_conjugatesOfSet hx) · simpa using (normalClosure s).one_mem · rw [← conj_mul] exact mul_mem ihx ihy @@ -2544,8 +2538,8 @@ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } one_mem' a := by rw [mul_one, mul_inv_self] <;> exact H.one_mem - inv_mem' a h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) - mul_mem' a b ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) + inv_mem' {a} h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) + mul_mem' {a b} ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) #align subgroup.normal_core Subgroup.normalCore theorem normalCore_le (H : Subgroup G) : H.normalCore ≤ H := fun a h => by @@ -2560,11 +2554,11 @@ instance normalCore_normal (H : Subgroup G) : H.normalCore.Normal := theorem normal_le_normalCore {H : Subgroup G} {N : Subgroup G} [hN : N.Normal] : N ≤ H.normalCore ↔ N ≤ H := - ⟨ge_trans H.normal_core_le, fun h_le n hn g => h_le (hN.conj_mem n hn g)⟩ + ⟨ge_trans H.normalCore_le, fun h_le n hn g => h_le (hN.conj_mem n hn g)⟩ #align subgroup.normal_le_normal_core Subgroup.normal_le_normalCore theorem normalCore_mono {H K : Subgroup G} (h : H ≤ K) : H.normalCore ≤ K.normalCore := - normal_le_normalCore.mpr (H.normal_core_le.trans h) + normal_le_normalCore.mpr (H.normalCore_le.trans h) #align subgroup.normal_core_mono Subgroup.normalCore_mono theorem normalCore_eq_supᵢ (H : Subgroup G) : @@ -2620,7 +2614,7 @@ theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by e @[to_additive (attr := simp)] theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, Subtype.coe_mk, - iff_self_iff, forall_const] + iff_self_iff, forall_const, exists_prop, forall_const] #align monoid_hom.restrict_range MonoidHom.restrict_range #align add_monoid_hom.restrict_range AddMonoidHom.restrict_range From 73e8a0f5d01385dead024effb93d4825766855f7 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:18:01 +0000 Subject: [PATCH 22/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 77 +++++++++++++------------ 1 file changed, 39 insertions(+), 38 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 00c7d7aff3def..dc2404073a413 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -2636,7 +2636,7 @@ theorem coe_rangeRestrict (f : G →* N) (g : G) : (f.rangeRestrict g : N) = f g @[to_additive] theorem coe_comp_rangeRestrict (f : G →* N) : - (coe : f.range → N) ∘ (⇑f.rangeRestrict : G → f.range) = f := + ((↑) : f.range → N) ∘ (⇑f.rangeRestrict : G → f.range) = f := rfl #align monoid_hom.coe_comp_range_restrict MonoidHom.coe_comp_rangeRestrict #align add_monoid_hom.coe_comp_range_restrict AddMonoidHom.coe_comp_rangeRestrict @@ -2681,7 +2681,7 @@ theorem range_one : (1 : G →* N).range = ⊥ := #align add_monoid_hom.range_zero AddMonoidHom.range_zero @[to_additive (attr := simp)] -theorem Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by +theorem _root_.Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] ext simp @@ -2689,7 +2689,7 @@ theorem Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by #align add_subgroup.subtype_range AddSubgroup.subtype_range @[to_additive (attr := simp)] -theorem Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : +theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : (inclusion h_le).range = H.subgroupOf K := Subgroup.ext fun g => Set.ext_iff.mp (Set.range_inclusion h_le) g #align subgroup.inclusion_range Subgroup.inclusion_range @@ -2708,14 +2708,14 @@ theorem subgroupOf_range_eq_of_le {G₁ G₂ : Type _} [Group G₁] [Group G₂] /-- Computable alternative to `MonoidHom.ofInjective`. -/ @[to_additive "Computable alternative to `AddMonoidHom.ofInjective`."] def ofLeftInverse {f : G →* N} {g : N →* G} (h : Function.LeftInverse g f) : G ≃* f.range := - { f.range_restrict with - toFun := f.range_restrict - invFun := g ∘ f.range.Subtype + { f.rangeRestrict with + toFun := f.rangeRestrict + invFun := g ∘ f.range.subtype left_inv := h right_inv := by rintro ⟨x, y, rfl⟩ apply Subtype.ext - rw [coe_range_restrict, Function.comp_apply, Subgroup.coeSubtype, Subtype.coe_mk, h] } + rw [coe_rangeRestrict, Function.comp_apply, Subgroup.coeSubtype, Subtype.coe_mk, h] } #align monoid_hom.of_left_inverse MonoidHom.ofLeftInverse #align add_monoid_hom.of_left_inverse AddMonoidHom.ofLeftInverse @@ -2762,12 +2762,11 @@ variable {M : Type _} [MulOneClass M] "The additive kernel of an `AddMonoid` homomorphism is the `AddSubgroup` of elements such that `f x = 0`"] def ker (f : G →* M) : Subgroup G := - { f.mker with - inv_mem' := fun x (hx : f x = 1) => + { MonoidHom.mker f with + inv_mem' := fun {x} (hx : f x = 1) => calc f x⁻¹ = f x * f x⁻¹ := by rw [hx, one_mul] - _ = 1 := by rw [← map_mul, mul_inv_self, map_one] - } + _ = 1 := by rw [← map_mul, mul_inv_self, map_one] } #align monoid_hom.ker MonoidHom.ker #align add_monoid_hom.ker AddMonoidHom.ker @@ -2827,7 +2826,7 @@ theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := SetLike.ext fun x => Subtype.ext_iff #align monoid_hom.ker_cod_restrict MonoidHom.ker_codRestrict -#align add_monoid_hom.ker_cod_restrict AddMonoidHom.ker_cod_restrict +#align add_monoid_hom.ker_cod_restrict AddMonoidHom.ker_codRestrict @[to_additive (attr := simp)] theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := @@ -2855,13 +2854,13 @@ theorem ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ Function.Injective f := #align add_monoid_hom.ker_eq_bot_iff AddMonoidHom.ker_eq_bot_iff @[to_additive (attr := simp)] -theorem Subgroup.ker_subtype (H : Subgroup G) : H.subtype.ker = ⊥ := - H.Subtype.ker_eq_bot_iff.mpr Subtype.coe_injective +theorem _root_.Subgroup.ker_subtype (H : Subgroup G) : H.subtype.ker = ⊥ := + H.subtype.ker_eq_bot_iff.mpr Subtype.coe_injective #align subgroup.ker_subtype Subgroup.ker_subtype #align add_subgroup.ker_subtype AddSubgroup.ker_subtype @[to_additive (attr := simp)] -theorem Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := +theorem _root_.Subgroup.ker_inclusion {H K : Subgroup G} (h : H ≤ K) : (inclusion h).ker = ⊥ := (inclusion h).ker_eq_bot_iff.mpr (Set.inclusion_injective h) #align subgroup.ker_inclusion Subgroup.ker_inclusion #align add_subgroup.ker_inclusion AddSubgroup.ker_inclusion @@ -2897,7 +2896,7 @@ variable {M : Type _} [Monoid M] /-- The subgroup of elements `x : G` such that `f x = g x` -/ @[to_additive "The additive subgroup of elements `x : G` such that `f x = g x`"] def eqLocus (f g : G →* M) : Subgroup G := - { eqLocusM f g with inv_mem' := fun x => eq_on_inv f g } + { eqLocusM f g with inv_mem' := eq_on_inv f g } #align monoid_hom.eq_locus MonoidHom.eqLocus #align add_monoid_hom.eq_locus AddMonoidHom.eqLocus @@ -3115,7 +3114,7 @@ theorem map_le_map_iff_of_injective {f : G →* N} (hf : Function.Injective f) { @[to_additive (attr := simp)] theorem map_subtype_le_map_subtype {G' : Subgroup G} {H K : Subgroup G'} : H.map G'.subtype ≤ K.map G'.subtype ↔ H ≤ K := - map_le_map_iff_of_injective Subtype.coe_injective + map_le_map_iff_of_injective $ by apply Subtype.coe_injective #align subgroup.map_subtype_le_map_subtype Subgroup.map_subtype_le_map_subtype #align add_subgroup.map_subtype_le_map_subtype AddSubgroup.map_subtype_le_map_subtype @@ -3144,7 +3143,7 @@ theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ke @[to_additive] theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by apply map_injective (closure s).subtype_injective - rwa [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, + rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, Set.image_preimage_eq_of_subset] rw [coeSubtype, Subtype.range_coe_subtype] exact subset_closure @@ -3172,7 +3171,7 @@ theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : @[to_additive] theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := - comap_sup_eq_of_le_range L.Subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) + comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) #align subgroup.sup_subgroup_of_eq Subgroup.sup_subgroupOf_eq #align add_subgroup.sup_add_subgroup_of_eq AddSubgroup.sup_addSubgroupOf_eq @@ -3254,8 +3253,11 @@ theorem map_equiv_normalizer_eq (H : Subgroup G) (f : G ≃* N) : H.normalizer.map f.toMonoidHom = (H.map f.toMonoidHom).normalizer := by ext x simp only [mem_normalizer_iff, mem_map_equiv] - rw [f.to_equiv.forall_congr] - simp + rw [f.toEquiv.forall_congr] + intro + erw [f.toEquiv.symm_apply_apply] + simp only [map_mul, map_inv] + erw [f.toEquiv.symm_apply_apply] #align subgroup.map_equiv_normalizer_eq Subgroup.map_equiv_normalizer_eq #align add_subgroup.map_equiv_normalizer_eq AddSubgroup.map_equiv_normalizer_eq @@ -3296,7 +3298,7 @@ def liftOfRightInverseAux (hf : Function.RightInverse f_inv f) (g : G₁ →* G @[to_additive (attr := simp)] theorem liftOfRightInverseAux_comp_apply (hf : Function.RightInverse f_inv f) (g : G₁ →* G₃) (hg : f.ker ≤ g.ker) (x : G₁) : (f.liftOfRightInverseAux f_inv hf g hg) (f x) = g x := by - dsimp [lift_of_right_inverse_aux] + dsimp [liftOfRightInverseAux] rw [← mul_inv_eq_one, ← g.map_inv, ← g.map_mul, ← g.mem_ker] apply hg rw [f.mem_ker, f.map_mul, f.map_inv, mul_inv_eq_one] @@ -3348,7 +3350,7 @@ def liftOfRightInverse (hf : Function.RightInverse f_inv f) : invFun φ := ⟨φ.comp f, fun x hx => (mem_ker _).mpr <| by simp [(mem_ker _).mp hx]⟩ left_inv g := by ext - simp only [comp_apply, liftOfRightInverseAux_comp_apply, Subtype.coe_mk, Subtype.val_eq_coe] + simp only [comp_apply, liftOfRightInverseAux_comp_apply, Subtype.coe_mk] right_inv φ := by ext b simp [liftOfRightInverseAux, hf b] @@ -3370,7 +3372,7 @@ noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : @[to_additive (attr := simp)] theorem liftOfRightInverse_comp_apply (hf : Function.RightInverse f_inv f) (g : { g : G₁ →* G₃ // f.ker ≤ g.ker }) (x : G₁) : - (f.liftOfRightInverse f_inv hf g) (f x) = g x := + (f.liftOfRightInverse f_inv hf g) (f x) = g.1 x := f.liftOfRightInverseAux_comp_apply f_inv hf g.1 g.2 x #align monoid_hom.lift_of_right_inverse_comp_apply MonoidHom.liftOfRightInverse_comp_apply #align add_monoid_hom.lift_of_right_inverse_comp_apply AddMonoidHom.liftOfRightInverse_comp_apply @@ -3507,13 +3509,13 @@ variable {C : Type _} [CommGroup C] {s t : Subgroup C} {x : C} theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := ⟨fun h => by rw [← closure_eq s, ← closure_eq t, ← closure_union] at h - apply closure_induction h + refine Subgroup.closure_induction h ?_ ?_ ?_ ?_ · rintro y (h | h) · exact ⟨y, h, 1, t.one_mem, by simp⟩ · exact ⟨1, s.one_mem, y, h, by simp⟩ · exact ⟨1, s.one_mem, 1, ⟨t.one_mem, mul_one 1⟩⟩ · rintro _ _ ⟨y₁, hy₁, z₁, hz₁, rfl⟩ ⟨y₂, hy₂, z₂, hz₂, rfl⟩ - exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc] <;> cc⟩ + exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc, mul_left_comm]⟩ · rintro _ ⟨y, hy, z, hz, rfl⟩ exact ⟨_, inv_mem hy, _, inv_mem hz, mul_comm z y ▸ (mul_inv_rev z y).symm⟩, by rintro ⟨y, hy, z, hz, rfl⟩ <;> exact mul_mem_sup hy hz⟩ @@ -3522,7 +3524,7 @@ theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := @[to_additive] theorem mem_sup' : x ∈ s ⊔ t ↔ ∃ (y : s)(z : t), (y : C) * z = x := - mem_sup.trans <| by simp only [SetLike.exists, coe_mk] + mem_sup.trans <| by simp only [SetLike.exists, coe_mk, exists_prop] #align subgroup.mem_sup' Subgroup.mem_sup' #align add_subgroup.mem_sup' AddSubgroup.mem_sup' @@ -3580,9 +3582,9 @@ instance prod_normal (H : Subgroup G) (K : Subgroup N) [hH : H.Normal] [hK : K.N @[to_additive] theorem inf_subgroupOf_inf_normal_of_right (A B' B : Subgroup G) (hB : B' ≤ B) [hN : (B'.subgroupOf B).Normal] : ((A ⊓ B').subgroupOf (A ⊓ B)).Normal := - { - conj_mem := fun n hn g => - ⟨mul_mem (mul_mem (mem_inf.1 g.2).1 (mem_inf.1 n.2).1) (inv_mem (mem_inf.1 g.2).1), + { conj_mem := fun {n} hn g => + ⟨mul_mem (mul_mem (mem_inf.1 g.2).1 (mem_inf.1 n.2).1) $ + show ↑g⁻¹ ∈ A from (inv_mem (mem_inf.1 g.2).1), (normal_subgroupOf_iff hB).mp hN n g hn.2 (mem_inf.mp g.2).2⟩ } #align subgroup.inf_subgroup_of_inf_normal_of_right @@ -3594,10 +3596,10 @@ theorem inf_subgroupOf_inf_normal_of_right (A B' B : Subgroup G) (hB : B' ≤ B) @[to_additive] theorem inf_subgroupOf_inf_normal_of_left {A' A : Subgroup G} (B : Subgroup G) (hA : A' ≤ A) [hN : (A'.subgroupOf A).Normal] : ((A' ⊓ B).subgroupOf (A ⊓ B)).Normal := - { - conj_mem := fun n hn g => + { conj_mem := fun n hn g => ⟨(normal_subgroupOf_iff hA).mp hN n g hn.1 (mem_inf.mp g.2).1, - mul_mem (mul_mem (mem_inf.1 g.2).2 (mem_inf.1 n.2).2) (inv_mem (mem_inf.1 g.2).2)⟩ } + mul_mem (mul_mem (mem_inf.1 g.2).2 (mem_inf.1 n.2).2) $ + show ↑g⁻¹ ∈ B from (inv_mem (mem_inf.1 g.2).2)⟩ } #align subgroup.inf_subgroup_of_inf_normal_of_left Subgroup.inf_subgroupOf_inf_normal_of_left #align add_subgroup.inf_add_subgroup_of_inf_normal_of_left @@ -3615,7 +3617,7 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : refine' map_injective_of_ker_le B.subtype (ker_le_comap _ _) (le_trans (ker_le_comap B.subtype _) le_sup_left) _ - · simp only [subgroup_of, map_comap_eq, map_sup, subtype_range] + · simp only [subgroupOf, map_comap_eq, map_sup, subtype_range] rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] #align subgroup.subgroup_of_sup Subgroup.subgroupOf_sup #align add_subgroup.add_subgroup_of_sup AddSubgroup.addSubgroupOf_sup @@ -3709,13 +3711,12 @@ theorem normalClosure_eq_top_of {N : Subgroup G} [hn : N.Normal] {g g' : G} {hg simp only [MonoidHom.codRestrict_apply, MulEquiv.coe_toMonoidHom, MulAut.conj_apply, coe_mk, MonoidHom.restrict_apply, Subtype.mk_eq_mk, ← mul_assoc, mul_inv_self, one_mul] rw [mul_assoc, mul_inv_self, mul_one] - have ht' := map_mono (eq_top_iff.1 ht) - rw [← MonoidHom.range_eq_map, MonoidHom.range_top_of_surjective _ hs] at ht' - refine' eq_top_iff.2 (le_trans ht' (map_le_iff_le_comap.2 (normal_closure_le_normal _))) + rw [eq_top_iff, ← MonoidHom.range_top_of_surjective _ hs, MonoidHom.range_eq_map] + refine' le_trans (map_mono (eq_top_iff.1 ht)) (map_le_iff_le_comap.2 (normalClosure_le_normal _)) rw [Set.singleton_subset_iff, SetLike.mem_coe] simp only [MonoidHom.codRestrict_apply, MulEquiv.coe_toMonoidHom, MulAut.conj_apply, coe_mk, MonoidHom.restrict_apply, mem_comap] - exact subset_normal_closure (Set.mem_singleton _) + exact subset_normalClosure (Set.mem_singleton _) #align is_conj.normal_closure_eq_top_of IsConj.normalClosure_eq_top_of end IsConj From 7e73377d8269dcafddfb846bbc7c42026ae076f8 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:36:58 +0000 Subject: [PATCH 23/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 16 +++++++--------- 1 file changed, 7 insertions(+), 9 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index dc2404073a413..64d22687dfe09 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -988,7 +988,7 @@ theorem mem_inf {p p' : Subgroup G} {x : G} : x ∈ p ⊓ p' ↔ x ∈ p ∧ x instance : InfSet (Subgroup G) := ⟨fun s => { (⨅ S ∈ s, Subgroup.toSubmonoid S).copy (⋂ S ∈ s, ↑S) (by simp) with - inv_mem' := fun x hx => + inv_mem' := fun {x} hx => Set.mem_binterᵢ fun i h => i.inv_mem (by apply Set.mem_interᵢ₂.1 hx i h) }⟩ @[to_additive (attr := simp, norm_cast)] @@ -1005,23 +1005,21 @@ theorem mem_infₛ {S : Set (Subgroup G)} {x : G} : x ∈ infₛ S ↔ ∀ p ∈ @[to_additive] theorem mem_infᵢ {ι : Sort _} {S : ι → Subgroup G} {x : G} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := by - simp only [infᵢ, mem_inf, Set.forall_range_iff] + simp only [infᵢ, mem_infₛ, Set.forall_range_iff] #align subgroup.mem_infi Subgroup.mem_infᵢ #align add_subgroup.mem_infi AddSubgroup.mem_infᵢ @[to_additive (attr := simp, norm_cast)] theorem coe_infᵢ {ι : Sort _} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Set G) = ⋂ i, S i := by - simp only [infᵢ, coe_inf, Set.binterᵢ_range] + simp only [infᵢ, coe_infₛ, Set.binterᵢ_range] #align subgroup.coe_infi Subgroup.coe_infᵢ #align add_subgroup.coe_infi AddSubgroup.coe_infᵢ /-- Subgroups of a group form a complete lattice. -/ @[to_additive "The `AddSubgroup`s of an `AddGroup` form a complete lattice."] instance : CompleteLattice (Subgroup G) := - { - completeLatticeOfInf (Subgroup G) fun s => - IsGLB.of_image (fun H K => show (H : Set G) ≤ K ↔ H ≤ K from SetLike.coe_subset_coe) - isGLB_binfᵢ with + { completeLatticeOfInf (Subgroup G) fun s => + IsGLB.of_image SetLike.coe_subset_coe isGLB_binfᵢ with bot := ⊥ bot_le := fun S x hx => (mem_bot.1 hx).symm ▸ S.one_mem top := ⊤ @@ -1107,7 +1105,7 @@ variable {k : Set G} @[to_additive] theorem mem_closure {x : G} : x ∈ closure k ↔ ∀ K : Subgroup G, k ⊆ K → x ∈ K := - mem_Inf + mem_infₛ #align subgroup.mem_closure Subgroup.mem_closure #align add_subgroup.mem_closure AddSubgroup.mem_closure @@ -1149,7 +1147,7 @@ of `k`. -/ holds for all elements of the additive closure of `k`."] theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (Hk : ∀ x ∈ k, p x) (H1 : p 1) (Hmul : ∀ x y, p x → p y → p (x * y)) (Hinv : ∀ x, p x → p x⁻¹) : p x := - (@closure_le _ _ ⟨p, Hmul, H1, Hinv⟩ _).2 Hk h + (@closure_le _ _ ⟨⟨⟨setOf p, fun {x y} ↦ Hmul x y⟩, H1⟩, fun {x} ↦ Hinv x⟩ k).2 Hk h #align subgroup.closure_induction Subgroup.closure_induction #align add_subgroup.closure_induction AddSubgroup.closure_induction From 5f859301a88b793ca1708a3a1cad12dd1cea2496 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:44:49 +0000 Subject: [PATCH 24/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 23 ++++++++++++----------- 1 file changed, 12 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 64d22687dfe09..c3619e369593f 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -1031,13 +1031,13 @@ instance : CompleteLattice (Subgroup G) := @[to_additive] theorem mem_sup_left {S T : Subgroup G} : ∀ {x : G}, x ∈ S → x ∈ S ⊔ T := - show S ≤ S ⊔ T from le_sup_left + have : S ≤ S ⊔ T := le_sup_left; fun h ↦ this h #align subgroup.mem_sup_left Subgroup.mem_sup_left #align add_subgroup.mem_sup_left AddSubgroup.mem_sup_left @[to_additive] theorem mem_sup_right {S T : Subgroup G} : ∀ {x : G}, x ∈ T → x ∈ S ⊔ T := - show T ≤ S ⊔ T from le_sup_right + have : T ≤ S ⊔ T := le_sup_right; fun h ↦ this h #align subgroup.mem_sup_right Subgroup.mem_sup_right #align add_subgroup.mem_sup_right AddSubgroup.mem_sup_right @@ -1050,14 +1050,14 @@ theorem mul_mem_sup {S T : Subgroup G} {x y : G} (hx : x ∈ S) (hy : y ∈ T) : @[to_additive] theorem mem_supᵢ_of_mem {ι : Sort _} {S : ι → Subgroup G} (i : ι) : ∀ {x : G}, x ∈ S i → x ∈ supᵢ S := - show S i ≤ supᵢ S from le_supᵢ _ _ + have : S i ≤ supᵢ S := le_supᵢ _ _; fun h ↦ this h #align subgroup.mem_supr_of_mem Subgroup.mem_supᵢ_of_mem #align add_subgroup.mem_supr_of_mem AddSubgroup.mem_supᵢ_of_mem @[to_additive] theorem mem_supₛ_of_mem {S : Set (Subgroup G)} {s : Subgroup G} (hs : s ∈ S) : ∀ {x : G}, x ∈ s → x ∈ supₛ S := - show s ≤ supₛ S from le_supₛ hs + have : s ≤ supₛ S := le_supₛ hs; fun h ↦ this h #align subgroup.mem_Sup_of_mem Subgroup.mem_supₛ_of_mem #align add_subgroup.mem_Sup_of_mem AddSubgroup.mem_supₛ_of_mem @@ -1350,7 +1350,8 @@ theorem coe_supᵢ_of_directed {ι} [Nonempty ι] {S : ι → Subgroup G} (hS : theorem mem_supₛ_of_directedOn {K : Set (Subgroup G)} (Kne : K.Nonempty) (hK : DirectedOn (· ≤ ·) K) {x : G} : x ∈ supₛ K ↔ ∃ s ∈ K, x ∈ s := by haveI : Nonempty K := Kne.to_subtype - simp only [supₛ_eq_supᵢ', mem_supᵢ_of_directed hK.directed_coe, SetCoe.exists, Subtype.coe_mk] + simp only [supₛ_eq_supᵢ', mem_supᵢ_of_directed hK.directed_val, SetCoe.exists, Subtype.coe_mk, + exists_prop] #align subgroup.mem_Sup_of_directed_on Subgroup.mem_supₛ_of_directedOn #align add_subgroup.mem_Sup_of_directed_on AddSubgroup.mem_supₛ_of_directedOn @@ -1419,8 +1420,8 @@ theorem coe_map (f : G →* N) (K : Subgroup G) : (K.map f : Set N) = f '' K := #align add_subgroup.coe_map AddSubgroup.coe_map @[to_additive (attr := simp)] -theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := - mem_image_iff_bex +theorem mem_map {f : G →* N} {K : Subgroup G} {y : N} : y ∈ K.map f ↔ ∃ x ∈ K, f x = y := by + erw [mem_image_iff_bex]; simp #align subgroup.mem_map Subgroup.mem_map #align add_subgroup.mem_map AddSubgroup.mem_map @@ -1464,8 +1465,8 @@ theorem map_one_eq_bot : K.map (1 : G →* N) = ⊥ := @[to_additive] theorem mem_map_equiv {f : G ≃* N} {K : Subgroup G} {x : N} : - x ∈ K.map f.toMonoidHom ↔ f.symm x ∈ K := - @Set.mem_image_equiv _ _ (↑K) f.toEquiv x + x ∈ K.map f.toMonoidHom ↔ f.symm x ∈ K := by + erw [@Set.mem_image_equiv _ _ (↑K) f.toEquiv x]; rfl #align subgroup.mem_map_equiv Subgroup.mem_map_equiv #align add_subgroup.mem_map_equiv AddSubgroup.mem_map_equiv @@ -2411,8 +2412,8 @@ def conjugatesOfSet (s : Set G) : Set G := ⋃ a ∈ s, conjugatesOf a #align group.conjugates_of_set Group.conjugatesOfSet -theorem mem_conjugatesOfSet_iff {x : G} : x ∈ conjugatesOfSet s ↔ ∃ a ∈ s, IsConj a x := - Set.mem_unionᵢ₂ +theorem mem_conjugatesOfSet_iff {x : G} : x ∈ conjugatesOfSet s ↔ ∃ a ∈ s, IsConj a x := by + erw [Set.mem_unionᵢ₂]; simp only [conjugatesOf, isConj_iff, Set.mem_setOf_eq, exists_prop] #align group.mem_conjugates_of_set_iff Group.mem_conjugatesOfSet_iff theorem subset_conjugatesOfSet : s ⊆ conjugatesOfSet s := fun (x : G) (h : x ∈ s) => From faba79d6632eea5acff36105b7ca2145dd12fb94 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:47:03 +0000 Subject: [PATCH 25/34] fixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index c3619e369593f..71a1e76f45d99 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -3537,7 +3537,7 @@ theorem mem_closure_pair {x y z : C} : @[to_additive] instance : IsModularLattice (Subgroup C) := - ⟨fun x y z xz a ha => by + ⟨fun {x} y z xz a ha => by rw [mem_inf, mem_sup] at ha rcases ha with ⟨⟨b, hb, c, hc, rfl⟩, haz⟩ rw [mem_sup] @@ -3560,8 +3560,8 @@ theorem normal_subgroupOf_iff {H K : Subgroup G} (hHK : H ≤ K) : @[to_additive] instance prod_subgroupOf_prod_normal {H₁ K₁ : Subgroup G} {H₂ K₂ : Subgroup N} [h₁ : (H₁.subgroupOf K₁).Normal] [h₂ : (H₂.subgroupOf K₂).Normal] : - ((H₁.Prod H₂).subgroupOf (K₁.Prod K₂)).Normal - where conj_mem n hgHK g := + ((H₁.prod H₂).subgroupOf (K₁.prod K₂)).Normal where + conj_mem n hgHK g := ⟨h₁.conj_mem ⟨(n : G × N).fst, (mem_prod.mp n.2).1⟩ hgHK.1 ⟨(g : G × N).fst, (mem_prod.mp g.2).1⟩, h₂.conj_mem ⟨(n : G × N).snd, (mem_prod.mp n.2).2⟩ hgHK.2 From 81ffc4e520f1daf04af8b6f72910918e60c6b6c8 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:49:17 +0000 Subject: [PATCH 26/34] auto: fix lints --- Mathlib/GroupTheory/Subgroup/Basic.lean | 58 ++++++++++++------------- 1 file changed, 29 insertions(+), 29 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 71a1e76f45d99..c52d2804ee96b 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -1018,16 +1018,16 @@ theorem coe_infᵢ {ι : Sort _} {S : ι → Subgroup G} : (↑(⨅ i, S i) : Se /-- Subgroups of a group form a complete lattice. -/ @[to_additive "The `AddSubgroup`s of an `AddGroup` form a complete lattice."] instance : CompleteLattice (Subgroup G) := - { completeLatticeOfInf (Subgroup G) fun s => + { completeLatticeOfInf (Subgroup G) fun _s => IsGLB.of_image SetLike.coe_subset_coe isGLB_binfᵢ with bot := ⊥ - bot_le := fun S x hx => (mem_bot.1 hx).symm ▸ S.one_mem + bot_le := fun S _x hx => (mem_bot.1 hx).symm ▸ S.one_mem top := ⊤ - le_top := fun S x hx => mem_top x + le_top := fun _S x _hx => mem_top x inf := (· ⊓ ·) - le_inf := fun a b c ha hb x hx => ⟨ha hx, hb hx⟩ - inf_le_left := fun a b x => And.left - inf_le_right := fun a b x => And.right } + le_inf := fun _a _b _c ha hb _x hx => ⟨ha hx, hb hx⟩ + inf_le_left := fun _a _b _x => And.left + inf_le_right := fun _a _b _x => And.right } @[to_additive] theorem mem_sup_left {S T : Subgroup G} : ∀ {x : G}, x ∈ S → x ∈ S ⊔ T := @@ -1223,8 +1223,8 @@ protected def gi : GaloisInsertion (@closure G _) (↑) where choice s _ := closure s gc s t := @closure_le _ _ t s - le_l_u s := subset_closure - choice_eq s h := rfl + le_l_u _s := subset_closure + choice_eq _s _h := rfl #align subgroup.gi Subgroup.gi #align add_subgroup.gi AddSubgroup.gi @@ -1364,7 +1364,7 @@ variable {N : Type _} [Group N] {P : Type _} [Group P] def comap {N : Type _} [Group N] (f : G →* N) (H : Subgroup N) : Subgroup G := { H.toSubmonoid.comap f with carrier := f ⁻¹' H - inv_mem' := fun {a} ha => show f a⁻¹ ∈ H by rw [f.map_inv] <;> exact H.inv_mem ha } + inv_mem' := fun {a} ha => show f a⁻¹ ∈ H by rw [f.map_inv]; exact H.inv_mem ha } #align subgroup.comap Subgroup.comap #align add_subgroup.comap AddSubgroup.comap @@ -1605,9 +1605,9 @@ def subgroupOfEquivOfLe {G : Type _} [Group G] {H K : Subgroup G} (h : H ≤ K) where toFun g := ⟨g.1, g.2⟩ invFun g := ⟨⟨g.1, h g.2⟩, g.2⟩ - left_inv g := Subtype.ext (Subtype.ext rfl) - right_inv g := Subtype.ext rfl - map_mul' g h := rfl + left_inv _g := Subtype.ext (Subtype.ext rfl) + right_inv _g := Subtype.ext rfl + map_mul' _g _h := rfl #align subgroup.subgroup_of_equiv_of_le Subgroup.subgroupOfEquivOfLe #align add_subgroup.add_subgroup_of_equiv_of_le AddSubgroup.addSubgroupOfEquivOfLe @@ -1644,7 +1644,7 @@ theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subty @[to_additive (attr := simp)] theorem bot_subgroupOf : (⊥ : Subgroup G).subgroupOf H = ⊥ := - Eq.symm (Subgroup.ext fun g => Subtype.ext_iff) + Eq.symm (Subgroup.ext fun _g => Subtype.ext_iff) #align subgroup.bot_subgroup_of Subgroup.bot_subgroupOf #align add_subgroup.bot_add_subgroup_of AddSubgroup.bot_addSubgroupOf @@ -1668,7 +1668,7 @@ theorem subgroupOf_bot_eq_top : H.subgroupOf ⊥ = ⊤ := @[to_additive (attr := simp)] theorem subgroupOf_self : H.subgroupOf H = ⊤ := - top_unique fun g hg => g.2 + top_unique fun g _hg => g.2 #align subgroup.subgroup_of_self Subgroup.subgroupOf_self #align add_subgroup.add_subgroup_of_self AddSubgroup.addSubgroupOf_self @@ -1728,7 +1728,7 @@ theorem mem_prod {H : Subgroup G} {K : Subgroup N} {p : G × N} : p ∈ H.prod K @[to_additive prod_mono] theorem prod_mono : ((· ≤ ·) ⇒ (· ≤ ·) ⇒ (· ≤ ·)) (@prod G _ N _) (@prod G _ N _) := - fun s s' hs t t' ht => Set.prod_mono hs ht + fun _s _s' hs _t _t' ht => Set.prod_mono hs ht #align subgroup.prod_mono Subgroup.prod_mono #align add_subgroup.prod_mono AddSubgroup.prod_mono @@ -2190,7 +2190,7 @@ instance (priority := 100) normal_in_normalizer : (H.subgroupOf H.normalizer).No @[to_additive] theorem normalizer_eq_top : H.normalizer = ⊤ ↔ H.Normal := eq_top_iff.trans - ⟨fun h => ⟨fun a ha b => (h (mem_top b) a).mp ha⟩, fun h a ha b => + ⟨fun h => ⟨fun a ha b => (h (mem_top b) a).mp ha⟩, fun h a _ha b => ⟨fun hb => h.conj_mem b hb a, fun hb => by rwa [h.mem_comm_iff, inv_mul_cancel_left] at hb⟩⟩ #align subgroup.normalizer_eq_top Subgroup.normalizer_eq_top #align add_subgroup.normalizer_eq_top AddSubgroup.normalizer_eq_top @@ -2308,7 +2308,7 @@ theorem centralizer_top : centralizer ⊤ = center G := @[to_additive] theorem le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := - ⟨fun h x hx y hy => (h hy x hx).symm, fun h x hx y hy => (h hy x hx).symm⟩ + ⟨fun h x hx _y hy => (h hy x hx).symm, fun h x hx _y hy => (h hy x hx).symm⟩ #align subgroup.le_centralizer_iff Subgroup.le_centralizer_iff #align add_subgroup.le_centralizer_iff AddSubgroup.le_centralizer_iff @@ -2433,7 +2433,7 @@ theorem conjugates_subset_normal {N : Subgroup G} [tn : N.Normal] {a : G} (h : a theorem conjugatesOfSet_subset {s : Set G} {N : Subgroup G} [N.Normal] (h : s ⊆ N) : conjugatesOfSet s ⊆ N := - Set.unionᵢ₂_subset fun x H => conjugates_subset_normal (h H) + Set.unionᵢ₂_subset fun _x H => conjugates_subset_normal (h H) #align group.conjugates_of_set_subset Group.conjugatesOfSet_subset /-- The set of conjugates of `s` is closed under conjugation. -/ @@ -2536,7 +2536,7 @@ as shown by `Subgroup.normalCore_eq_supᵢ`. -/ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } - one_mem' a := by rw [mul_one, mul_inv_self] <;> exact H.one_mem + one_mem' a := by rw [mul_one, mul_inv_self]; exact H.one_mem inv_mem' {a} h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) mul_mem' {a b} ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) #align subgroup.normal_core Subgroup.normalCore @@ -2548,7 +2548,7 @@ theorem normalCore_le (H : Subgroup G) : H.normalCore ≤ H := fun a h => by instance normalCore_normal (H : Subgroup G) : H.normalCore.Normal := ⟨fun a h b c => by - rw [mul_assoc, mul_assoc, ← mul_inv_rev, ← mul_assoc, ← mul_assoc] <;> exact h (c * b)⟩ + rw [mul_assoc, mul_assoc, ← mul_inv_rev, ← mul_assoc, ← mul_assoc]; exact h (c * b)⟩ #align subgroup.normal_core_normal Subgroup.normalCore_normal theorem normal_le_normalCore {H : Subgroup G} {N : Subgroup G} [hN : N.Normal] : @@ -2606,7 +2606,7 @@ theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := #align add_monoid_hom.mem_range AddMonoidHom.mem_range @[to_additive] -theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext <;> simp +theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext; simp #align monoid_hom.range_eq_map MonoidHom.range_eq_map #align add_monoid_hom.range_eq_map AddMonoidHom.range_eq_map @@ -2654,7 +2654,7 @@ theorem rangeRestrict_surjective (f : G →* N) : Function.Surjective f.rangeRes @[to_additive] theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by - rw [range_eq_map, range_eq_map] <;> exact (⊤ : Subgroup G).map_map g f + rw [range_eq_map, range_eq_map]; exact (⊤ : Subgroup G).map_map g f #align monoid_hom.map_range MonoidHom.map_range #align add_monoid_hom.map_range AddMonoidHom.map_range @@ -2823,7 +2823,7 @@ theorem ker_restrict (f : G →* N) : (f.restrict K).ker = f.ker.subgroupOf K := @[to_additive (attr := simp)] theorem ker_codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : G →* N) (s : S) (h : ∀ x, f x ∈ s) : (f.codRestrict s h).ker = f.ker := - SetLike.ext fun x => Subtype.ext_iff + SetLike.ext fun _x => Subtype.ext_iff #align monoid_hom.ker_cod_restrict MonoidHom.ker_codRestrict #align add_monoid_hom.ker_cod_restrict AddMonoidHom.ker_codRestrict @@ -2835,7 +2835,7 @@ theorem ker_rangeRestrict (f : G →* N) : ker (rangeRestrict f) = ker f := @[to_additive (attr := simp)] theorem ker_one : (1 : G →* M).ker = ⊤ := - SetLike.ext fun x => eq_self_iff_true _ + SetLike.ext fun _x => eq_self_iff_true _ #align monoid_hom.ker_one MonoidHom.ker_one #align add_monoid_hom.ker_zero AddMonoidHom.ker_zero @@ -2916,7 +2916,7 @@ theorem eqOn_closure {f g : G →* M} {s : Set G} (h : Set.EqOn f g s) : Set.EqO @[to_additive] theorem eq_of_eqOn_top {f g : G →* M} (h : Set.EqOn f g (⊤ : Subgroup G)) : f = g := - ext fun x => h trivial + ext fun _x => h trivial #align monoid_hom.eq_of_eq_on_top MonoidHom.eq_of_eqOn_top #align add_monoid_hom.eq_of_eq_on_top AddMonoidHom.eq_of_eqOn_top @@ -2930,7 +2930,7 @@ end EqLocus @[to_additive] theorem closure_preimage_le (f : G →* N) (s : Set N) : closure (f ⁻¹' s) ≤ (closure s).comap f := - (closure_le _).2 fun x hx => by rw [SetLike.mem_coe, mem_comap] <;> exact subset_closure hx + (closure_le _).2 fun x hx => by rw [SetLike.mem_coe, mem_comap]; exact subset_closure hx #align monoid_hom.closure_preimage_le MonoidHom.closure_preimage_le #align add_monoid_hom.closure_preimage_le AddMonoidHom.closure_preimage_le @@ -2941,7 +2941,7 @@ generated by the image of the set. -/ the `AddSubgroup` generated by the image of the set."] theorem map_closure (f : G →* N) (s : Set G) : (closure s).map f = closure (f '' s) := Set.image_preimage.l_comm_of_u_comm (Subgroup.gc_map_comap f) (Subgroup.gi N).gc - (Subgroup.gi G).gc fun t => rfl + (Subgroup.gi G).gc fun _t => rfl #align monoid_hom.map_closure MonoidHom.map_closure #align add_monoid_hom.map_closure AddMonoidHom.map_closure @@ -3517,7 +3517,7 @@ theorem mem_sup : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := exact ⟨_, mul_mem hy₁ hy₂, _, mul_mem hz₁ hz₂, by simp [mul_assoc, mul_left_comm]⟩ · rintro _ ⟨y, hy, z, hz, rfl⟩ exact ⟨_, inv_mem hy, _, inv_mem hz, mul_comm z y ▸ (mul_inv_rev z y).symm⟩, by - rintro ⟨y, hy, z, hz, rfl⟩ <;> exact mul_mem_sup hy hz⟩ + rintro ⟨y, hy, z, hz, rfl⟩; exact mul_mem_sup hy hz⟩ #align subgroup.mem_sup Subgroup.mem_sup #align add_subgroup.mem_sup AddSubgroup.mem_sup @@ -3661,7 +3661,7 @@ theorem disjoint_def {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x : @[to_additive] theorem disjoint_def' {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := - disjoint_def.trans ⟨fun h x y hx hy hxy => h hx <| hxy.symm ▸ hy, fun h x hx hx' => h hx hx' rfl⟩ + disjoint_def.trans ⟨fun h _x _y hx hy hxy => h hx <| hxy.symm ▸ hy, fun h _x hx hx' => h hx hx' rfl⟩ #align subgroup.disjoint_def' Subgroup.disjoint_def' #align add_subgroup.disjoint_def' AddSubgroup.disjoint_def' From 718cd395de7bdf04667f8f37bffcaa808f65c287 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:52:22 +0000 Subject: [PATCH 27/34] auto: naming --- Mathlib/GroupTheory/Subgroup/Basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index c52d2804ee96b..0c481a92a00d7 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -247,7 +247,7 @@ instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [ #align add_subgroup_class.to_add_comm_group AddSubgroupClass.toAddCommGroup -- Prefer subclasses of `group` over subclasses of `Subgroup_class`. -/-- A subgroup of an `ordered_comm_group` is an `ordered_comm_group`. -/ +/-- A subgroup of an `OrderedCommGroup` is an `OrderedCommGroup`. -/ @[to_additive "An additive subgroup of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : OrderedCommGroup H := @@ -257,7 +257,7 @@ instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [ #align add_subgroup_class.to_ordered_add_comm_group AddSubgroupClass.toOrderedAddCommGroup -- Prefer subclasses of `group` over subclasses of `Subgroup_class`. -/-- A subgroup of a `linear_ordered_comm_group` is a `linear_ordered_comm_group`. -/ +/-- A subgroup of a `LinearOrderedCommGroup` is a `LinearOrderedCommGroup`. -/ @[to_additive "An additive subgroup of a `linear_ordered_add_comm_group` is a `linear_ordered_add_comm_group`."] @@ -491,7 +491,7 @@ theorem toSubmonoid_le {p q : Subgroup G} : p.toSubmonoid ≤ q.toSubmonoid ↔ end Subgroup /-! -### Conversion to/from `additive`/`multiplicative` +### Conversion to/from `Additive`/`Multiplicative` -/ From ec16bef7e02ae77db20177ff36ef108a00f29b5a Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:53:58 +0000 Subject: [PATCH 28/34] fix long lines --- Mathlib/GroupTheory/Subgroup/Basic.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 0c481a92a00d7..5cbc294e06937 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -1805,10 +1805,9 @@ variable {η : Type _} {f : η → Type _} /-- A version of `Set.pi` for submonoids. Given an index set `I` and a family of submodules `s : Π i, Submonoid f i`, `pi I s` is the submonoid of dependent functions `f : Π i, f i` such that `f i` belongs to `Pi I s` whenever `i ∈ I`. -/ -@[to_additive - " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family - of submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions - `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] +@[to_additive " A version of `Set.pi` for `AddSubmonoid`s. Given an index set `I` and a family + of submodules `s : Π i, AddSubmonoid f i`, `pi I s` is the `AddSubmonoid` of dependent functions + `f : Π i, f i` such that `f i` belongs to `pi I s` whenever `i ∈ I`. -/ "] def _root_.Submonoid.pi [∀ i, MulOneClass (f i)] (I : Set η) (s : ∀ i, Submonoid (f i)) : Submonoid (∀ i, f i) where carrier := I.pi fun i => (s i).carrier @@ -1920,7 +1919,8 @@ end Subgroup namespace AddSubgroup -/- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: #[`conj_mem] [] -/ +/- ./././Mathport/Syntax/Translate/Command.lean:388:30: + infer kinds are unsupported in Lean 4: #[`conj_mem] [] -/ /-- An AddSubgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ structure Normal (H : AddSubgroup A) : Prop where conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H @@ -3661,7 +3661,7 @@ theorem disjoint_def {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x : @[to_additive] theorem disjoint_def' {H₁ H₂ : Subgroup G} : Disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := - disjoint_def.trans ⟨fun h _x _y hx hy hxy => h hx <| hxy.symm ▸ hy, fun h _x hx hx' => h hx hx' rfl⟩ + disjoint_def.trans ⟨fun h _x _y hx hy hxy ↦ h hx <| hxy.symm ▸ hy, fun h _x hx hx' ↦ h hx hx' rfl⟩ #align subgroup.disjoint_def' Subgroup.disjoint_def' #align add_subgroup.disjoint_def' AddSubgroup.disjoint_def' From db40de47724530e3d9a704f5bd39b507517125fc Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 25 Jan 2023 15:56:26 +0000 Subject: [PATCH 29/34] fix a lint --- Mathlib/GroupTheory/Subgroup/Basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 5cbc294e06937..30e38191bb1ae 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -2319,17 +2319,17 @@ theorem centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := #align add_subgroup.centralizer_le AddSubgroup.centralizer_le @[to_additive] -instance Subgroup.Centralizer.characteristic [hH : H.Characteristic] : +instance Centralizer.characteristic [hH : H.Characteristic] : H.centralizer.Characteristic := by refine' Subgroup.characteristic_iff_comap_le.mpr fun ϕ g hg h hh => ϕ.injective _ rw [map_mul, map_mul] exact hg (ϕ h) (Subgroup.characteristic_iff_le_comap.mp hH ϕ hh) #align subgroup.subgroup.centralizer.characteristic - Subgroup.Subgroup.Centralizer.characteristic + Subgroup.Centralizer.characteristic #align add_subgroup.subgroup.centralizer.characteristic - AddSubgroup.Subgroup.Centralizer.characteristic + AddSubgroup.Centralizer.characteristic end Centralizer From 0953d4ed76503a751d33bceb934e4e3f1f08090a Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 17:02:53 +0100 Subject: [PATCH 30/34] fix remaining errors, add porting note --- Mathlib/GroupTheory/Subgroup/Basic.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 5cbc294e06937..eb964a8ec2329 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -2502,8 +2502,9 @@ theorem normalClosure_mono {s t : Set G} (h : s ⊆ t) : normalClosure s ≤ nor normalClosure_le_normal (Set.Subset.trans h subset_normalClosure) #align subgroup.normal_closure_mono Subgroup.normalClosure_mono +-- Porting note: the elaborator trips up on using underscores for names in `⨅` theorem normalClosure_eq_infᵢ : - normalClosure s = ⨅ (N : Subgroup G) (_ : Normal N) (hs : s ⊆ N), N := + normalClosure s = ⨅ (N : Subgroup G) (_hN : Normal N) (_hs : s ⊆ N), N := le_antisymm (le_infᵢ fun N => le_infᵢ fun hN => le_infᵢ normalClosure_le_normal) (infᵢ_le_of_le (normalClosure s) (infᵢ_le_of_le (by infer_instance) (infᵢ_le_of_le subset_normalClosure le_rfl))) @@ -2561,11 +2562,11 @@ theorem normalCore_mono {H K : Subgroup G} (h : H ≤ K) : H.normalCore ≤ K.no #align subgroup.normal_core_mono Subgroup.normalCore_mono theorem normalCore_eq_supᵢ (H : Subgroup G) : - H.normalCore = ⨆ (N : Subgroup G) (_ : Normal N) (hs : N ≤ H), N := + H.normalCore = ⨆ (N : Subgroup G) (_hN : Normal N) (_hs : N ≤ H), N := le_antisymm (le_supᵢ_of_le H.normalCore (le_supᵢ_of_le H.normalCore_normal (le_supᵢ_of_le H.normalCore_le le_rfl))) - (supᵢ_le fun N => supᵢ_le fun hN => supᵢ_le normal_le_normalCore.mpr) + (supᵢ_le fun _ => supᵢ_le fun _ => supᵢ_le normal_le_normalCore.mpr) #align subgroup.normal_core_eq_supr Subgroup.normalCore_eq_supᵢ @[simp] From 74c8a009e47143ffbf77113efef36b6d1ecb2f26 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 17:30:57 +0100 Subject: [PATCH 31/34] lintfixes --- Mathlib/GroupTheory/Subgroup/Basic.lean | 34 +++++++++++++++++-------- 1 file changed, 23 insertions(+), 11 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 5ddb9143e5a55..d4bde097f5b85 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -100,6 +100,7 @@ section SubgroupClass /-- `InvMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ class InvMemClass (S G : Type _) [Inv G] [SetLike S G] : Prop where + /-- `s` is closed under inverses -/ inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s #align inv_mem_class InvMemClass @@ -107,6 +108,7 @@ export InvMemClass (inv_mem) /-- `NegMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ class NegMemClass (S G : Type _) [Neg G] [SetLike S G] : Prop where +/-- `s` is closed under negation -/ neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s #align neg_mem_class NegMemClass @@ -159,7 +161,7 @@ theorem div_mem_comm_iff {a b : G} : a / b ∈ H ↔ b / a ∈ H := by #align div_mem_comm_iff div_mem_comm_iff #align sub_mem_comm_iff sub_mem_comm_iff -@[to_additive (attr := simp)] +@[to_additive /-(attr := simp)-/] -- porting note: `simp` cannot simplify LHS theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : (∃ x : G, x ∈ H ∧ P x⁻¹) ↔ ∃ x ∈ H, P x := by constructor <;> @@ -168,6 +170,8 @@ theorem exists_inv_mem_iff_exists_mem {P : G → Prop} : #align exists_inv_mem_iff_exists_mem exists_inv_mem_iff_exists_mem #align exists_neg_mem_iff_exists_mem exists_neg_mem_iff_exists_mem + + @[to_additive] theorem mul_mem_cancel_right {x y : G} (h : x ∈ H) : y * x ∈ H ↔ y ∈ H := -- Porting note: whut? why do we need this? @@ -290,13 +294,13 @@ variable {H} @[to_additive (attr := simp, norm_cast)] theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := - (subtype H : H →* G).map_pow _ _ + rfl #align subgroup_class.coe_pow SubgroupClass.coe_pow #align add_subgroup_class.coe_smul AddSubgroupClass.coe_nsmul @[to_additive (attr := simp, norm_cast)] theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := - (subtype H : H →* G).map_zpow _ _ + rfl #align subgroup_class.coe_zpow SubgroupClass.coe_zpow #align add_subgroup_class.coe_zsmul AddSubgroupClass.coe_zsmul @@ -356,12 +360,14 @@ end SubgroupClass /-- A subgroup of a group `G` is a subset containing 1, closed under multiplication and closed under multiplicative inverse. -/ structure Subgroup (G : Type _) [Group G] extends Submonoid G where + /-- `G` is closed under inverses -/ inv_mem' {x} : x ∈ carrier → x⁻¹ ∈ carrier #align subgroup Subgroup /-- An additive subgroup of an additive group `G` is a subset containing 0, closed under addition and additive inverse. -/ structure AddSubgroup (G : Type _) [AddGroup G] extends AddSubmonoid G where +/-- `G` is closed under negation -/ neg_mem' {x} : x ∈ carrier → -x ∈ carrier #align add_subgroup AddSubgroup @@ -393,7 +399,7 @@ instance : SubgroupClass (Subgroup G) G where one_mem _ := (Subgroup.toSubmonoid _).one_mem' mul_mem := (Subgroup.toSubmonoid _).mul_mem' -@[to_additive (attr := simp)] +@[to_additive (attr := simp, nolint simpNF)] -- Porting note: dsimp can not prove this theorem mem_carrier {s : Subgroup G} {x : G} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl #align subgroup.mem_carrier Subgroup.mem_carrier @@ -422,7 +428,8 @@ theorem mk_le_mk {s t : Set G} (h_one) (h_mul) (h_inv) (h_one') (h_mul') (h_inv' /-- See Note [custom simps projection] -/ --@[to_additive "See Note [custom simps projection]"] -@[to_additive] +-- Porting note: temporarily removed brackets to not confuse syntax highlighting +@[to_additive "See Note custom simps projection "] def Simps.coe (S : Subgroup G) : Set G := S #align subgroup.simps.coe Subgroup.Simps.coe @@ -732,7 +739,8 @@ theorem coe_div (x y : H) : (↑(x / y) : G) = ↑x / ↑y := #align subgroup.coe_div Subgroup.coe_div #align add_subgroup.coe_sub AddSubgroup.coe_sub -@[to_additive (attr := simp, norm_cast)] +-- Porting note: removed simp, theorem has variable as head symbol +@[to_additive (attr := norm_cast)] theorem coe_mk (x : G) (hx : x ∈ H) : ((⟨x, hx⟩ : H) : G) = x := rfl #align subgroup.coe_mk Subgroup.coe_mk @@ -744,7 +752,7 @@ theorem coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = (x : G) ^ n := #align subgroup.coe_pow Subgroup.coe_pow #align add_subgroup.coe_nsmul AddSubgroup.coe_nsmul -@[to_additive (attr := simp, norm_cast)] +@[to_additive (attr := norm_cast)] -- Porting note: dsimp can provce this theorem coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = (x : G) ^ n := rfl #align subgroup.coe_zpow Subgroup.coe_zpow @@ -1910,6 +1918,7 @@ end Pi /-- A subgroup is normal if whenever `n ∈ H`, then `g * n * g⁻¹ ∈ H` for every `g : G` -/ structure Normal : Prop where + /- `N` is closed under conjugation -/ conj_mem : ∀ n, n ∈ H → ∀ g : G, g * n * g⁻¹ ∈ H #align subgroup.normal Subgroup.Normal @@ -1919,10 +1928,9 @@ end Subgroup namespace AddSubgroup -/- ./././Mathport/Syntax/Translate/Command.lean:388:30: - infer kinds are unsupported in Lean 4: #[`conj_mem] [] -/ /-- An AddSubgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ structure Normal (H : AddSubgroup A) : Prop where + /- `N` is closed under additive conjugation -/ conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H #align add_subgroup.normal AddSubgroup.Normal @@ -1968,6 +1976,7 @@ variable (H) /-- A subgroup is characteristic if it is fixed by all automorphisms. Several equivalent conditions are provided by lemmas of the form `Characteristic.iff...` -/ structure Characteristic : Prop where + /-- `H` is fixed by all automorphisms -/ fixed : ∀ ϕ : G ≃* G, H.comap ϕ.toMonoidHom = H #align subgroup.characteristic Subgroup.Characteristic @@ -1986,6 +1995,7 @@ variable (H : AddSubgroup A) /-- A add_subgroup is characteristic if it is fixed by all automorphisms. Several equivalent conditions are provided by lemmas of the form `Characteristic.iff...` -/ structure Characteristic : Prop where + /-- `H` is fixed by all automorphisms -/ fixed : ∀ ϕ : A ≃+ A, H.comap ϕ.toAddMonoidHom = H #align add_subgroup.characteristic AddSubgroup.Characteristic @@ -2335,6 +2345,7 @@ end Centralizer /-- Commutivity of a subgroup -/ structure IsCommutative : Prop where + /-- `*` is commutative on `H` -/ is_comm : IsCommutative H (· * ·) #align subgroup.is_commutative Subgroup.IsCommutative @@ -2342,6 +2353,7 @@ attribute [class] IsCommutative /-- Commutivity of an additive subgroup -/ structure _root_.AddSubgroup.IsCommutative (H : AddSubgroup A) : Prop where + /-- `+` is commutative on `H` -/ is_comm : _root_.IsCommutative H (· + ·) #align add_subgroup.is_commutative AddSubgroup.IsCommutative @@ -2515,7 +2527,7 @@ theorem normalClosure_eq_self (H : Subgroup G) [H.Normal] : normalClosure ↑H = le_antisymm (normalClosure_le_normal rfl.subset) le_normalClosure #align subgroup.normal_closure_eq_self Subgroup.normalClosure_eq_self -@[simp] +-- @[simp] -- Porting note: simp can prove this theorem normalClosure_idempotent : normalClosure ↑(normalClosure s) = normalClosure s := normalClosure_eq_self _ #align subgroup.normal_closure_idempotent Subgroup.normalClosure_idempotent @@ -2574,7 +2586,7 @@ theorem normalCore_eq_self (H : Subgroup G) [H.Normal] : H.normalCore = H := le_antisymm H.normalCore_le (normal_le_normalCore.mpr le_rfl) #align subgroup.normal_core_eq_self Subgroup.normalCore_eq_self -@[simp] +-- @[simp] -- Porting note: simp can prove this theorem normalCore_idempotent (H : Subgroup G) : H.normalCore.normalCore = H.normalCore := H.normalCore.normalCore_eq_self #align subgroup.normal_core_idempotent Subgroup.normalCore_idempotent From a59af1b14ce47552a09bc72bead0ab2fb19d24ea Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 17:33:12 +0100 Subject: [PATCH 32/34] fix --- Mathlib/GroupTheory/Subgroup/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index d4bde097f5b85..f56ce98613889 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -1918,7 +1918,7 @@ end Pi /-- A subgroup is normal if whenever `n ∈ H`, then `g * n * g⁻¹ ∈ H` for every `g : G` -/ structure Normal : Prop where - /- `N` is closed under conjugation -/ + /-- `N` is closed under conjugation -/ conj_mem : ∀ n, n ∈ H → ∀ g : G, g * n * g⁻¹ ∈ H #align subgroup.normal Subgroup.Normal @@ -1930,7 +1930,7 @@ namespace AddSubgroup /-- An AddSubgroup is normal if whenever `n ∈ H`, then `g + n - g ∈ H` for every `g : G` -/ structure Normal (H : AddSubgroup A) : Prop where - /- `N` is closed under additive conjugation -/ + /-- `N` is closed under additive conjugation -/ conj_mem : ∀ n, n ∈ H → ∀ g : A, g + n + -g ∈ H #align add_subgroup.normal AddSubgroup.Normal From c12a65399ef635f176f8521e3465b4eaeb8984e0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 26 Jan 2023 09:56:48 +0100 Subject: [PATCH 33/34] Apply suggestions from code review --- Mathlib/GroupTheory/Subgroup/Basic.lean | 99 +++++++++++-------------- 1 file changed, 45 insertions(+), 54 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index f56ce98613889..713bb252a6958 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -192,29 +192,29 @@ namespace SubgroupClass /-- A subgroup of a group inherits an inverse. -/ @[to_additive "An additive subgroup of a `add_group` inherits an inverse."] -instance hasInv : Inv H := +instance inv : Inv H := ⟨fun a => ⟨a⁻¹, inv_mem a.2⟩⟩ -#align subgroup_class.has_inv SubgroupClass.hasInv -#align add_subgroup_class.has_neg AddSubgroupClass.hasNeg +#align subgroup_class.has_inv SubgroupClass.inv +#align add_subgroup_class.has_neg AddSubgroupClass.neg /-- A subgroup of a group inherits a division -/ @[to_additive "An additive subgroup of an `add_group` inherits a subtraction."] -instance hasDiv : Div H := +instance div : Div H := ⟨fun a b => ⟨a / b, div_mem a.2 b.2⟩⟩ -#align subgroup_class.has_div SubgroupClass.hasDiv -#align add_subgroup_class.has_sub AddSubgroupClass.hasSub +#align subgroup_class.has_div SubgroupClass.div +#align add_subgroup_class.has_sub AddSubgroupClass.sub /-- An additive subgroup of an `add_group` inherits an integer scaling. -/ -instance _root_.AddSubgroupClass.hasZsmul {M S} [SubNegMonoid M] [SetLike S M] +instance _root_.AddSubgroupClass.zsmul {M S} [SubNegMonoid M] [SetLike S M] [AddSubgroupClass S M] {H : S} : SMul ℤ H := ⟨fun n a => ⟨n • a.1, zsmul_mem a.2 n⟩⟩ -#align add_subgroup_class.has_zsmul AddSubgroupClass.hasZsmul +#align add_subgroup_class.has_zsmul AddSubgroupClass.zsmul /-- A subgroup of a group inherits an integer power. -/ @[to_additive] -instance hasZpow {M S} [DivInvMonoid M] [SetLike S M] [SubgroupClass S M] {H : S} : Pow H ℤ := +instance zpow {M S} [DivInvMonoid M] [SetLike S M] [SubgroupClass S M] {H : S} : Pow H ℤ := ⟨fun a n => ⟨a.1 ^ n, zpow_mem a.2 n⟩⟩ -#align subgroup_class.has_zpow SubgroupClass.hasZpow +#align subgroup_class.has_zpow SubgroupClass.zpow -- Porting note: additive align statement is given above @[to_additive (attr := simp, norm_cast)] @@ -233,16 +233,16 @@ variable (H) -- Prefer subclasses of `Group` over subclasses of `SubgroupClass`. /-- A subgroup of a group inherits a group structure. -/ -@[to_additive "An additive subgroup of an `add_group` inherits an `add_group` structure."] +@[to_additive "An additive subgroup of an `AddGroup` inherits an `AddGroup` structure."] instance (priority := 75) toGroup : Group H := Subtype.coe_injective.group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl #align subgroup_class.to_group SubgroupClass.toGroup #align add_subgroup_class.to_add_group AddSubgroupClass.toAddGroup --- Prefer subclasses of `comm_group` over subclasses of `Subgroup_class`. -/-- A subgroup of a `comm_group` is a `comm_group`. -/ -@[to_additive "An additive subgroup of an `add_comm_group` is an `add_comm_group`."] +-- Prefer subclasses of `CommGroup` over subclasses of `SubgroupClass`. +/-- A subgroup of a `CommGroup` is a `CommGroup`. -/ +@[to_additive "An additive subgroup of an `AddCommGroup` is an `AddCommGroup`."] instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [SubgroupClass S G] : CommGroup H := Subtype.coe_injective.commGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) @@ -250,9 +250,9 @@ instance (priority := 75) toCommGroup {G : Type _} [CommGroup G] [SetLike S G] [ #align subgroup_class.to_comm_group SubgroupClass.toCommGroup #align add_subgroup_class.to_add_comm_group AddSubgroupClass.toAddCommGroup --- Prefer subclasses of `group` over subclasses of `Subgroup_class`. +-- Prefer subclasses of `Group` over subclasses of `SubgroupClass`. /-- A subgroup of an `OrderedCommGroup` is an `OrderedCommGroup`. -/ -@[to_additive "An additive subgroup of an `add_ordered_comm_group` is an `add_ordered_comm_group`."] +@[to_additive "An additive subgroup of an `AddOrderedCommGroup` is an `AddOrderedCommGroup`."] instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : OrderedCommGroup H := Subtype.coe_injective.orderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) @@ -260,11 +260,11 @@ instance (priority := 75) toOrderedCommGroup {G : Type _} [OrderedCommGroup G] [ #align subgroup_class.to_ordered_comm_group SubgroupClass.toOrderedCommGroup #align add_subgroup_class.to_ordered_add_comm_group AddSubgroupClass.toOrderedAddCommGroup --- Prefer subclasses of `group` over subclasses of `Subgroup_class`. +-- Prefer subclasses of `Group` over subclasses of `SubgroupClass`. /-- A subgroup of a `LinearOrderedCommGroup` is a `LinearOrderedCommGroup`. -/ @[to_additive - "An additive subgroup of a `linear_ordered_add_comm_group` is a - `linear_ordered_add_comm_group`."] + "An additive subgroup of a `LinearOrderedAddCommGroup` is a + `LinearOrderedAddCommGroup`."] instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] [SetLike S G] [SubgroupClass S G] : LinearOrderedCommGroup H := Subtype.coe_injective.linearOrderedCommGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) @@ -506,8 +506,7 @@ section mul_add /-- Subgroups of a group `G` are isomorphic to additive subgroups of `Additive G`. -/ @[simps] -def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) - where +def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where toFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } invFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } left_inv x := by cases x; rfl @@ -523,8 +522,7 @@ abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := /-- Additive supgroups of an additive group `A` are isomorphic to subgroups of `Multiplicative A`. -/ @[simps] -def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) - where +def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) where toFun S := { AddSubmonoid.toSubmonoid S.toAddSubmonoid with inv_mem' := S.neg_mem' } invFun S := { Submonoid.toAddSubmonoid S.toSubmonoid with neg_mem' := S.inv_mem' } left_inv x := by cases x; rfl @@ -549,8 +547,7 @@ equalities.-/ @[to_additive "Copy of an additive subgroup with a new `carrier` equal to the old one. Useful to fix definitional equalities"] -protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G - where +protected def copy (K : Subgroup G) (s : Set G) (hs : s = K) : Subgroup G where carrier := s one_mem' := hs.symm ▸ K.one_mem' mul_mem' := hs.symm ▸ K.mul_mem' @@ -648,8 +645,6 @@ protected theorem zpow_mem {x : G} (hx : x ∈ K) : ∀ n : ℤ, x ^ n ∈ K := #align subgroup.zpow_mem Subgroup.zpow_mem #align add_subgroup.zsmul_mem AddSubgroup.zsmul_mem -/- ./././Mathport/Syntax/Translate/Basic.lean:632:2: warning: - expanding binder collection (x y «expr ∈ » s) -/ /-- Construct a subgroup from a nonempty set that is closed under division. -/ @[to_additive "Construct a subgroup from a nonempty set that is closed under subtraction"] def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y ∈ s), x * y⁻¹ ∈ s) : @@ -667,53 +662,53 @@ def ofDiv (s : Set G) (hsn : s.Nonempty) (hs : ∀ (x) (_ : x ∈ s) (y) (_ : y /-- A subgroup of a group inherits a multiplication. -/ @[to_additive "An `AddSubgroup` of an `AddGroup` inherits an addition."] -instance hasMul : Mul H := +instance mul : Mul H := H.toSubmonoid.mul -#align subgroup.has_mul Subgroup.hasMul -#align add_subgroup.has_add AddSubgroup.hasAdd +#align subgroup.has_mul Subgroup.mul +#align add_subgroup.has_add AddSubgroup.add /-- A subgroup of a group inherits a 1. -/ @[to_additive "An `AddSubgroup` of an `AddGroup` inherits a zero."] -instance hasOne : One H := +instance one : One H := H.toSubmonoid.one -#align subgroup.has_one Subgroup.hasOne -#align add_subgroup.has_zero AddSubgroup.hasZero +#align subgroup.has_one Subgroup.one +#align add_subgroup.has_zero AddSubgroup.zero /-- A subgroup of a group inherits an inverse. -/ @[to_additive "A `AddSubgroup` of a `AddGroup` inherits an inverse."] -instance hasInv : Inv H := +instance inv : Inv H := ⟨fun a => ⟨a⁻¹, H.inv_mem a.2⟩⟩ -#align subgroup.has_inv Subgroup.hasInv -#align add_subgroup.has_neg AddSubgroup.hasNeg +#align subgroup.has_inv Subgroup.inv +#align add_subgroup.has_neg AddSubgroup.neg /-- A subgroup of a group inherits a division -/ @[to_additive "An `AddSubgroup` of an `AddGroup` inherits a subtraction."] -instance hasDiv : Div H := +instance div : Div H := ⟨fun a b => ⟨a / b, H.div_mem a.2 b.2⟩⟩ -#align subgroup.has_div Subgroup.hasDiv -#align add_subgroup.has_sub AddSubgroup.hasSub +#align subgroup.has_div Subgroup.div +#align add_subgroup.has_sub AddSubgroup.sub /-- An `AddSubgroup` of an `AddGroup` inherits a natural scaling. -/ -instance _root_.AddSubgroup.instNSMul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := +instance _root_.AddSubgroup.nsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℕ H := ⟨fun n a => ⟨n • a, H.nsmul_mem a.2 n⟩⟩ -#align add_subgroup.has_nsmul AddSubgroup.instNSMul +#align add_subgroup.has_nsmul AddSubgroup.nsmul /-- A subgroup of a group inherits a natural power -/ @[to_additive] -protected instance instNPow : Pow H ℕ := +protected instance npow : Pow H ℕ := ⟨fun a n => ⟨a ^ n, H.pow_mem a.2 n⟩⟩ -#align subgroup.has_npow Subgroup.instNPow +#align subgroup.has_npow Subgroup.npow /-- An `AddSubgroup` of an `AddGroup` inherits an integer scaling. -/ -instance _root_.AddSubgroup.hasZsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := +instance _root_.AddSubgroup.zsmul {G} [AddGroup G] {H : AddSubgroup G} : SMul ℤ H := ⟨fun n a => ⟨n • a, H.zsmul_mem a.2 n⟩⟩ -#align add_subgroup.has_zsmul AddSubgroup.hasZsmul +#align add_subgroup.has_zsmul AddSubgroup.zsmul /-- A subgroup of a group inherits an integer power -/ @[to_additive] -instance hasZpow : Pow H ℤ := +instance zpow : Pow H ℤ := ⟨fun a n => ⟨a ^ n, H.zpow_mem a.2 n⟩⟩ -#align subgroup.has_zpow Subgroup.hasZpow +#align subgroup.has_zpow Subgroup.zpow @[to_additive (attr := simp, norm_cast)] theorem coe_mul (x y : H) : (↑(x * y) : G) = ↑x * ↑y := @@ -1503,11 +1498,9 @@ theorem comap_equiv_eq_map_symm (f : N ≃* G) (K : Subgroup G) : theorem map_symm_eq_iff_map_eq {H : Subgroup N} {e : G ≃* N} : H.map ↑e.symm = K ↔ K.map ↑e = H := by constructor <;> rintro rfl - · - rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.symm_trans_self, + · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.symm_trans_self, MulEquiv.coe_monoidHom_refl, map_id] - · - rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.self_trans_symm, + · rw [map_map, ← MulEquiv.coe_monoidHom_trans, MulEquiv.self_trans_symm, MulEquiv.coe_monoidHom_refl, map_id] #align subgroup.map_symm_eq_iff_map_eq Subgroup.map_symm_eq_iff_map_eq #align add_subgroup.map_symm_eq_iff_map_eq AddSubgroup.map_symm_eq_iff_map_eq @@ -1721,7 +1714,6 @@ def prod (H : Subgroup G) (K : Subgroup N) : Subgroup (G × N) := #align subgroup.prod Subgroup.prod #align add_subgroup.prod AddSubgroup.prod -/- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @[to_additive coe_prod] theorem coe_prod (H : Subgroup G) (K : Subgroup N) : (H.prod K : Set (G × N)) = H ×ˢ K := rfl @@ -2546,8 +2538,7 @@ theorem normalClosure_closure_eq_normalClosure {s : Set G} : /-- The normal core of a subgroup `H` is the largest normal subgroup of `G` contained in `H`, as shown by `Subgroup.normalCore_eq_supᵢ`. -/ -def normalCore (H : Subgroup G) : Subgroup G - where +def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } one_mem' a := by rw [mul_one, mul_inv_self]; exact H.one_mem inv_mem' {a} h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) From 7d4c9e79dcb4bd9546a33d7c1ca3137f0a017fdb Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Thu, 26 Jan 2023 10:25:39 +0000 Subject: [PATCH 34/34] to_additive some simps --- Mathlib/GroupTheory/Subgroup/Basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 713bb252a6958..42a20970545dd 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -3362,8 +3362,7 @@ def liftOfRightInverse (hf : Function.RightInverse f_inv f) : /-- A non-computable version of `MonoidHom.liftOfRightInverse` for when no computable right inverse is available, that uses `Function.surjInv`. -/ -@[simp, - to_additive +@[to_additive (attr := simp) "A non-computable version of `AddMonoidHom.liftOfRightInverse` for when no computable right inverse is available."] noncomputable abbrev liftOfSurjective (hf : Function.Surjective f) : @@ -3432,14 +3431,15 @@ instance (priority := 100) Subgroup.normal_subgroupOf {H N : Subgroup G} [N.Norm namespace MonoidHom /-- The `MonoidHom` from the preimage of a subgroup to itself. -/ -@[to_additive "the `AddMonoidHom` from the preimage of an additive subgroup to itself.", simps] +@[to_additive (attr := simps) "the `AddMonoidHom` from the preimage of an +additive subgroup to itself."] def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := f.submonoidComap H'.toSubmonoid #align monoid_hom.subgroup_comap MonoidHom.subgroupComap #align add_monoid_hom.add_subgroup_comap AddMonoidHom.addSubgroupComap /-- The `MonoidHom` from a subgroup to its image. -/ -@[to_additive "the `add_monoid_hom` from an additive subgroup to its image", simps] +@[to_additive (attr := simps) "the `add_monoid_hom` from an additive subgroup to its image"] def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := f.submonoidMap H.toSubmonoid #align monoid_hom.subgroup_map MonoidHom.subgroupMap