Documentation

Mathlib.GroupTheory.Coset.Basic

Cosets #

This file develops the basic theory of left and right cosets.

When G is a group and a : G, s : Set G, with open scoped Pointwise we can write:

If instead G is an additive group, we can write (with open scoped Pointwise still)

Main definitions #

Notation #

TODO #

Properly merge with pointwise actions on sets, by renaming and deduplicating lemmas as appropriate.

theorem mem_leftAddCoset {α : Type u_1} [Add α] {s : Set α} {x : α} (a : α) (hxS : x s) :
a + x a +ᵥ s
theorem mem_leftCoset {α : Type u_1} [Mul α] {s : Set α} {x : α} (a : α) (hxS : x s) :
a * x a s
theorem mem_rightAddCoset {α : Type u_1} [Add α] {s : Set α} {x : α} (a : α) (hxS : x s) :
theorem mem_rightCoset {α : Type u_1} [Mul α] {s : Set α} {x : α} (a : α) (hxS : x s) :
def LeftAddCosetEquivalence {α : Type u_1} [Add α] (s : Set α) (a : α) (b : α) :

Equality of two left cosets a + s and b + s.

Equations
Instances For
    def LeftCosetEquivalence {α : Type u_1} [Mul α] (s : Set α) (a : α) (b : α) :

    Equality of two left cosets a * s and b * s.

    Equations
    Instances For
      def RightAddCosetEquivalence {α : Type u_1} [Add α] (s : Set α) (a : α) (b : α) :

      Equality of two right cosets s + a and s + b.

      Equations
      Instances For
        def RightCosetEquivalence {α : Type u_1} [Mul α] (s : Set α) (a : α) (b : α) :

        Equality of two right cosets s * a and s * b.

        Equations
        Instances For
          theorem leftAddCoset_assoc {α : Type u_1} [AddSemigroup α] (s : Set α) (a : α) (b : α) :
          a +ᵥ (b +ᵥ s) = a + b +ᵥ s
          theorem leftCoset_assoc {α : Type u_1} [Semigroup α] (s : Set α) (a : α) (b : α) :
          a b s = (a * b) s
          theorem rightAddCoset_assoc {α : Type u_1} [AddSemigroup α] (s : Set α) (a : α) (b : α) :
          theorem rightCoset_assoc {α : Type u_1} [Semigroup α] (s : Set α) (a : α) (b : α) :
          theorem leftAddCoset_rightAddCoset {α : Type u_1} [AddSemigroup α] (s : Set α) (a : α) (b : α) :
          theorem leftCoset_rightCoset {α : Type u_1} [Semigroup α] (s : Set α) (a : α) (b : α) :
          theorem zero_leftAddCoset {α : Type u_1} [AddMonoid α] (s : Set α) :
          0 +ᵥ s = s
          theorem one_leftCoset {α : Type u_1} [Monoid α] (s : Set α) :
          1 s = s
          theorem rightAddCoset_zero {α : Type u_1} [AddMonoid α] (s : Set α) :
          theorem rightCoset_one {α : Type u_1} [Monoid α] (s : Set α) :
          theorem mem_own_leftAddCoset {α : Type u_1} [AddMonoid α] (s : AddSubmonoid α) (a : α) :
          a a +ᵥ s
          theorem mem_own_leftCoset {α : Type u_1} [Monoid α] (s : Submonoid α) (a : α) :
          a a s
          theorem mem_own_rightAddCoset {α : Type u_1} [AddMonoid α] (s : AddSubmonoid α) (a : α) :
          theorem mem_own_rightCoset {α : Type u_1} [Monoid α] (s : Submonoid α) (a : α) :
          theorem mem_leftAddCoset_leftAddCoset {α : Type u_1} [AddMonoid α] (s : AddSubmonoid α) {a : α} (ha : a +ᵥ s = s) :
          a s
          theorem mem_leftCoset_leftCoset {α : Type u_1} [Monoid α] (s : Submonoid α) {a : α} (ha : a s = s) :
          a s
          theorem mem_rightAddCoset_rightAddCoset {α : Type u_1} [AddMonoid α] (s : AddSubmonoid α) {a : α} (ha : AddOpposite.op a +ᵥ s = s) :
          a s
          theorem mem_rightCoset_rightCoset {α : Type u_1} [Monoid α] (s : Submonoid α) {a : α} (ha : MulOpposite.op a s = s) :
          a s
          theorem mem_leftAddCoset_iff {α : Type u_1} [AddGroup α] {s : Set α} {x : α} (a : α) :
          x a +ᵥ s -a + x s
          theorem mem_leftCoset_iff {α : Type u_1} [Group α] {s : Set α} {x : α} (a : α) :
          x a s a⁻¹ * x s
          theorem mem_rightAddCoset_iff {α : Type u_1} [AddGroup α] {s : Set α} {x : α} (a : α) :
          theorem mem_rightCoset_iff {α : Type u_1} [Group α] {s : Set α} {x : α} (a : α) :
          theorem leftAddCoset_mem_leftAddCoset {α : Type u_1} [AddGroup α] (s : AddSubgroup α) {a : α} (ha : a s) :
          a +ᵥ s = s
          theorem leftCoset_mem_leftCoset {α : Type u_1} [Group α] (s : Subgroup α) {a : α} (ha : a s) :
          a s = s
          theorem rightAddCoset_mem_rightAddCoset {α : Type u_1} [AddGroup α] (s : AddSubgroup α) {a : α} (ha : a s) :
          AddOpposite.op a +ᵥ s = s
          theorem rightCoset_mem_rightCoset {α : Type u_1} [Group α] (s : Subgroup α) {a : α} (ha : a s) :
          MulOpposite.op a s = s
          theorem orbit_addSubgroup_eq_rightCoset {α : Type u_1} [AddGroup α] (s : AddSubgroup α) (a : α) :
          theorem orbit_subgroup_eq_rightCoset {α : Type u_1} [Group α] (s : Subgroup α) (a : α) :
          theorem orbit_addSubgroup_eq_self_of_mem {α : Type u_1} [AddGroup α] (s : AddSubgroup α) {a : α} (ha : a s) :
          AddAction.orbit (↥s) a = s
          theorem orbit_subgroup_eq_self_of_mem {α : Type u_1} [Group α] (s : Subgroup α) {a : α} (ha : a s) :
          MulAction.orbit (↥s) a = s
          theorem orbit_addSubgroup_zero_eq_self {α : Type u_1} [AddGroup α] (s : AddSubgroup α) :
          AddAction.orbit (↥s) 0 = s
          theorem orbit_subgroup_one_eq_self {α : Type u_1} [Group α] (s : Subgroup α) :
          MulAction.orbit (↥s) 1 = s
          theorem eq_addCosets_of_normal {α : Type u_1} [AddGroup α] (s : AddSubgroup α) (N : s.Normal) (g : α) :
          g +ᵥ s = AddOpposite.op g +ᵥ s
          theorem eq_cosets_of_normal {α : Type u_1} [Group α] (s : Subgroup α) (N : s.Normal) (g : α) :
          g s = MulOpposite.op g s
          theorem normal_of_eq_addCosets {α : Type u_1} [AddGroup α] (s : AddSubgroup α) (h : ∀ (g : α), g +ᵥ s = AddOpposite.op g +ᵥ s) :
          s.Normal
          theorem normal_of_eq_cosets {α : Type u_1} [Group α] (s : Subgroup α) (h : ∀ (g : α), g s = MulOpposite.op g s) :
          s.Normal
          theorem normal_iff_eq_addCosets {α : Type u_1} [AddGroup α] (s : AddSubgroup α) :
          s.Normal ∀ (g : α), g +ᵥ s = AddOpposite.op g +ᵥ s
          theorem normal_iff_eq_cosets {α : Type u_1} [Group α] (s : Subgroup α) :
          s.Normal ∀ (g : α), g s = MulOpposite.op g s
          theorem leftAddCoset_eq_iff {α : Type u_1} [AddGroup α] (s : AddSubgroup α) {x : α} {y : α} :
          x +ᵥ s = y +ᵥ s -x + y s
          theorem leftCoset_eq_iff {α : Type u_1} [Group α] (s : Subgroup α) {x : α} {y : α} :
          x s = y s x⁻¹ * y s
          theorem rightAddCoset_eq_iff {α : Type u_1} [AddGroup α] (s : AddSubgroup α) {x : α} {y : α} :
          theorem rightCoset_eq_iff {α : Type u_1} [Group α] (s : Subgroup α) {x : α} {y : α} :
          theorem QuotientGroup.leftRel_prod {α : Type u_1} [Group α] (s : Subgroup α) {β : Type u_2} [Group β] (s' : Subgroup β) :
          theorem QuotientAddGroup.leftRel_pi {ι : Type u_2} {β : ιType u_3} [(i : ι) → AddGroup (β i)] (s' : (i : ι) → AddSubgroup (β i)) :
          theorem QuotientGroup.leftRel_pi {ι : Type u_2} {β : ιType u_3} [(i : ι) → Group (β i)] (s' : (i : ι) → Subgroup (β i)) :
          QuotientGroup.leftRel (Subgroup.pi Set.univ s') = piSetoid
          theorem QuotientGroup.rightRel_prod {α : Type u_1} [Group α] (s : Subgroup α) {β : Type u_2} [Group β] (s' : Subgroup β) :
          theorem QuotientAddGroup.rightRel_pi {ι : Type u_2} {β : ιType u_3} [(i : ι) → AddGroup (β i)] (s' : (i : ι) → AddSubgroup (β i)) :
          theorem QuotientGroup.rightRel_pi {ι : Type u_2} {β : ιType u_3} [(i : ι) → Group (β i)] (s' : (i : ι) → Subgroup (β i)) :
          QuotientGroup.rightRel (Subgroup.pi Set.univ s') = piSetoid
          theorem QuotientAddGroup.strictMono_comap_prod_image {α : Type u_1} [AddGroup α] (s : AddSubgroup α) :
          StrictMono fun (t : AddSubgroup α) => (AddSubgroup.comap s.subtype t, QuotientAddGroup.mk '' t)

          Given an additive subgroup s, the function that sends an additive subgroup t to the pair consisting of its intersection with s and its image in the quotient α ⧸ s is strictly monotone, even though it is not injective in general.

          theorem QuotientGroup.strictMono_comap_prod_image {α : Type u_1} [Group α] (s : Subgroup α) :
          StrictMono fun (t : Subgroup α) => (Subgroup.comap s.subtype t, QuotientGroup.mk '' t)

          Given a subgroup s, the function that sends a subgroup t to the pair consisting of its intersection with s and its image in the quotient α ⧸ s is strictly monotone, even though it is not injective in general.

          theorem QuotientAddGroup.eq_class_eq_leftCoset {α : Type u_1} [AddGroup α] (s : AddSubgroup α) (g : α) :
          {x : α | x = g} = g +ᵥ s
          theorem QuotientGroup.eq_class_eq_leftCoset {α : Type u_1} [Group α] (s : Subgroup α) (g : α) :
          {x : α | x = g} = g s
          theorem QuotientGroup.orbit_mk_eq_smul {α : Type u_1} [Group α] {s : Subgroup α} (x : α) :
          def AddSubgroup.leftCosetEquivAddSubgroup {α : Type u_1} [AddGroup α] {s : AddSubgroup α} (g : α) :
          (g +ᵥ s) s

          The natural bijection between the cosets g + s and s.

          Equations
          Instances For
            def Subgroup.leftCosetEquivSubgroup {α : Type u_1} [Group α] {s : Subgroup α} (g : α) :
            (g s) s

            The natural bijection between a left coset g * s and s.

            Equations
            Instances For
              def AddSubgroup.rightCosetEquivAddSubgroup {α : Type u_1} [AddGroup α] {s : AddSubgroup α} (g : α) :
              (AddOpposite.op g +ᵥ s) s

              The natural bijection between the cosets s + g and s.

              Equations
              Instances For
                def Subgroup.rightCosetEquivSubgroup {α : Type u_1} [Group α] {s : Subgroup α} (g : α) :
                (MulOpposite.op g s) s

                The natural bijection between a right coset s * g and s.

                Equations
                Instances For
                  noncomputable def AddSubgroup.addGroupEquivQuotientProdAddSubgroup {α : Type u_1} [AddGroup α] {s : AddSubgroup α} :
                  α (α s) × s

                  A (non-canonical) bijection between an add_group α and the product (α/s) × s

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    noncomputable def Subgroup.groupEquivQuotientProdSubgroup {α : Type u_1} [Group α] {s : Subgroup α} :
                    α (α s) × s

                    A (non-canonical) bijection between a group α and the product (α/s) × s

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def AddSubgroup.quotientEquivSumOfLE' {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientAddGroup.mk) :
                      α s (α t) × t s.addSubgroupOf t

                      If H ≤ K, then G/H ≃ G/K × K/H constructively, using the provided right inverse of the quotient map G → G/K. The classical version is AddSubgroup.quotientEquivSumOfLE.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem Subgroup.quotientEquivProdOfLE'_symm_apply {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientGroup.mk) (a : (α t) × t s.subgroupOf t) :
                        (Subgroup.quotientEquivProdOfLE' h_le f hf).symm a = Quotient.map' (fun (b : t) => f a.1 * b) a.2
                        @[simp]
                        theorem AddSubgroup.quotientEquivSumOfLE'_symm_apply {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientAddGroup.mk) (a : (α t) × t s.addSubgroupOf t) :
                        (AddSubgroup.quotientEquivSumOfLE' h_le f hf).symm a = Quotient.map' (fun (b : t) => f a.1 + b) a.2
                        @[simp]
                        theorem AddSubgroup.quotientEquivSumOfLE'_apply {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientAddGroup.mk) (a : α s) :
                        (AddSubgroup.quotientEquivSumOfLE' h_le f hf) a = (Quotient.map' id a, Quotient.map' (fun (g : α) => -f (Quotient.mk'' g) + g, ) a)
                        @[simp]
                        theorem Subgroup.quotientEquivProdOfLE'_apply {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientGroup.mk) (a : α s) :
                        (Subgroup.quotientEquivProdOfLE' h_le f hf) a = (Quotient.map' id a, Quotient.map' (fun (g : α) => (f (Quotient.mk'' g))⁻¹ * g, ) a)
                        def Subgroup.quotientEquivProdOfLE' {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) (f : α tα) (hf : Function.RightInverse f QuotientGroup.mk) :
                        α s (α t) × t s.subgroupOf t

                        If H ≤ K, then G/H ≃ G/K × K/H constructively, using the provided right inverse of the quotient map G → G/K. The classical version is Subgroup.quotientEquivProdOfLE.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          noncomputable def AddSubgroup.quotientEquivSumOfLE {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) :
                          α s (α t) × t s.addSubgroupOf t

                          If H ≤ K, then G/H ≃ G/K × K/H nonconstructively. The constructive version is quotientEquivProdOfLE'.

                          Equations
                          Instances For
                            @[simp]
                            theorem AddSubgroup.quotientEquivSumOfLE_symm_apply {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) (a : (α t) × t s.addSubgroupOf t) :
                            (AddSubgroup.quotientEquivSumOfLE h_le).symm a = Quotient.map' (fun (b : t) => Quotient.out' a.1 + b) a.2
                            @[simp]
                            theorem AddSubgroup.quotientEquivSumOfLE_apply {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h_le : s t) (a : α s) :
                            (AddSubgroup.quotientEquivSumOfLE h_le) a = (Quotient.map' id a, Quotient.map' (fun (g : α) => -(Quotient.mk'' g).out' + g, ) a)
                            @[simp]
                            theorem Subgroup.quotientEquivProdOfLE_apply {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) (a : α s) :
                            (Subgroup.quotientEquivProdOfLE h_le) a = (Quotient.map' id a, Quotient.map' (fun (g : α) => (Quotient.mk'' g).out'⁻¹ * g, ) a)
                            @[simp]
                            theorem Subgroup.quotientEquivProdOfLE_symm_apply {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) (a : (α t) × t s.subgroupOf t) :
                            (Subgroup.quotientEquivProdOfLE h_le).symm a = Quotient.map' (fun (b : t) => Quotient.out' a.1 * b) a.2
                            noncomputable def Subgroup.quotientEquivProdOfLE {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h_le : s t) :
                            α s (α t) × t s.subgroupOf t

                            If H ≤ K, then G/H ≃ G/K × K/H nonconstructively. The constructive version is quotientEquivProdOfLE'.

                            Equations
                            Instances For
                              def AddSubgroup.quotientAddSubgroupOfEmbeddingOfLE {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (H : AddSubgroup α) (h : s t) :
                              s H.addSubgroupOf s t H.addSubgroupOf t

                              If s ≤ t, then there is an embedding s ⧸ H.addSubgroupOf s ↪ t ⧸ H.addSubgroupOf t.

                              Equations
                              Instances For
                                def Subgroup.quotientSubgroupOfEmbeddingOfLE {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (H : Subgroup α) (h : s t) :
                                s H.subgroupOf s t H.subgroupOf t

                                If s ≤ t, then there is an embedding s ⧸ H.subgroupOf s ↪ t ⧸ H.subgroupOf t.

                                Equations
                                Instances For
                                  @[simp]
                                  theorem Subgroup.quotientSubgroupOfEmbeddingOfLE_apply_mk {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (H : Subgroup α) (h : s t) (g : s) :
                                  def AddSubgroup.quotientAddSubgroupOfMapOfLE {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (H : AddSubgroup α) (h : s t) :
                                  H s.addSubgroupOf HH t.addSubgroupOf H

                                  If s ≤ t, then there is a map H ⧸ s.addSubgroupOf H → H ⧸ t.addSubgroupOf H.

                                  Equations
                                  Instances For
                                    def Subgroup.quotientSubgroupOfMapOfLE {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (H : Subgroup α) (h : s t) :
                                    H s.subgroupOf HH t.subgroupOf H

                                    If s ≤ t, then there is a map H ⧸ s.subgroupOf H → H ⧸ t.subgroupOf H.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem AddSubgroup.quotientAddSubgroupOfMapOfLE_apply_mk {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (H : AddSubgroup α) (h : s t) (g : H) :
                                      @[simp]
                                      theorem Subgroup.quotientSubgroupOfMapOfLE_apply_mk {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (H : Subgroup α) (h : s t) (g : H) :
                                      def AddSubgroup.quotientMapOfLE {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h : s t) :
                                      α sα t

                                      If s ≤ t, then there is a map α ⧸ s → α ⧸ t.

                                      Equations
                                      Instances For
                                        def Subgroup.quotientMapOfLE {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h : s t) :
                                        α sα t

                                        If s ≤ t, then there is a map α ⧸ s → α ⧸ t.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem AddSubgroup.quotientMapOfLE_apply_mk {α : Type u_1} [AddGroup α] {s : AddSubgroup α} {t : AddSubgroup α} (h : s t) (g : α) :
                                          @[simp]
                                          theorem Subgroup.quotientMapOfLE_apply_mk {α : Type u_1} [Group α] {s : Subgroup α} {t : Subgroup α} (h : s t) (g : α) :
                                          def AddSubgroup.quotientiInfAddSubgroupOfEmbedding {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) (H : AddSubgroup α) :
                                          H (⨅ (i : ι), f i).addSubgroupOf H (i : ι) → H (f i).addSubgroupOf H

                                          The natural embedding H ⧸ (⨅ i, f i).addSubgroupOf H) ↪ Π i, H ⧸ (f i).addSubgroupOf H.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem AddSubgroup.quotientiInfAddSubgroupOfEmbedding_apply {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) (H : AddSubgroup α) (q : H (⨅ (i : ι), f i).addSubgroupOf H) (i : ι) :
                                            @[simp]
                                            theorem Subgroup.quotientiInfSubgroupOfEmbedding_apply {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) (H : Subgroup α) (q : H (⨅ (i : ι), f i).subgroupOf H) (i : ι) :
                                            def Subgroup.quotientiInfSubgroupOfEmbedding {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) (H : Subgroup α) :
                                            H (⨅ (i : ι), f i).subgroupOf H (i : ι) → H (f i).subgroupOf H

                                            The natural embedding H ⧸ (⨅ i, f i).subgroupOf H ↪ Π i, H ⧸ (f i).subgroupOf H.

                                            Equations
                                            Instances For
                                              @[simp]
                                              theorem AddSubgroup.quotientiInfAddSubgroupOfEmbedding_apply_mk {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) (H : AddSubgroup α) (g : H) (i : ι) :
                                              @[simp]
                                              theorem Subgroup.quotientiInfSubgroupOfEmbedding_apply_mk {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) (H : Subgroup α) (g : H) (i : ι) :
                                              def AddSubgroup.quotientiInfEmbedding {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) :
                                              α ⨅ (i : ι), f i (i : ι) → α f i

                                              The natural embedding α ⧸ (⨅ i, f i) ↪ Π i, α ⧸ f i.

                                              Equations
                                              Instances For
                                                @[simp]
                                                theorem AddSubgroup.quotientiInfEmbedding_apply {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) (q : α ⨅ (i : ι), f i) (i : ι) :
                                                @[simp]
                                                theorem Subgroup.quotientiInfEmbedding_apply {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) (q : α ⨅ (i : ι), f i) (i : ι) :
                                                def Subgroup.quotientiInfEmbedding {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) :
                                                α ⨅ (i : ι), f i (i : ι) → α f i

                                                The natural embedding α ⧸ (⨅ i, f i) ↪ Π i, α ⧸ f i.

                                                Equations
                                                Instances For
                                                  @[simp]
                                                  theorem AddSubgroup.quotientiInfEmbedding_apply_mk {α : Type u_1} [AddGroup α] {ι : Type u_2} (f : ιAddSubgroup α) (g : α) (i : ι) :
                                                  @[simp]
                                                  theorem Subgroup.quotientiInfEmbedding_apply_mk {α : Type u_1} [Group α] {ι : Type u_2} (f : ιSubgroup α) (g : α) (i : ι) :
                                                  def AddMonoidHom.fiberEquivKer {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) :
                                                  (f ⁻¹' {f a}) f.ker

                                                  An equivalence between any non-empty fiber of an AddMonoidHom and its kernel.

                                                  Equations
                                                  Instances For
                                                    def MonoidHom.fiberEquivKer {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) :
                                                    (f ⁻¹' {f a}) f.ker

                                                    An equivalence between any non-empty fiber of a MonoidHom and its kernel.

                                                    Equations
                                                    Instances For
                                                      @[simp]
                                                      theorem AddMonoidHom.fiberEquivKer_apply {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) (g : (f ⁻¹' {f a})) :
                                                      ((f.fiberEquivKer a) g) = -a + g
                                                      @[simp]
                                                      theorem MonoidHom.fiberEquivKer_apply {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) (g : (f ⁻¹' {f a})) :
                                                      ((f.fiberEquivKer a) g) = a⁻¹ * g
                                                      @[simp]
                                                      theorem AddMonoidHom.fiberEquivKer_symm_apply {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) (g : f.ker) :
                                                      ((f.fiberEquivKer a).symm g) = a + g
                                                      @[simp]
                                                      theorem MonoidHom.fiberEquivKer_symm_apply {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) (g : f.ker) :
                                                      ((f.fiberEquivKer a).symm g) = a * g
                                                      noncomputable def AddMonoidHom.fiberEquivKerOfSurjective {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] {f : α →+ H} (hf : Function.Surjective f) (h : H) :
                                                      (f ⁻¹' {h}) f.ker

                                                      An equivalence between any fiber of a surjective AddMonoidHom and its kernel.

                                                      Equations
                                                      Instances For
                                                        noncomputable def MonoidHom.fiberEquivKerOfSurjective {α : Type u_1} [Group α] {H : Type u_2} [Group H] {f : α →* H} (hf : Function.Surjective f) (h : H) :
                                                        (f ⁻¹' {h}) f.ker

                                                        An equivalence between any fiber of a surjective MonoidHom and its kernel.

                                                        Equations
                                                        Instances For
                                                          def AddMonoidHom.fiberEquiv {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) (b : α) :
                                                          (f ⁻¹' {f a}) (f ⁻¹' {f b})

                                                          An equivalence between any two non-empty fibers of an AddMonoidHom.

                                                          Equations
                                                          • f.fiberEquiv a b = (f.fiberEquivKer a).trans (f.fiberEquivKer b).symm
                                                          Instances For
                                                            def MonoidHom.fiberEquiv {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) (b : α) :
                                                            (f ⁻¹' {f a}) (f ⁻¹' {f b})

                                                            An equivalence between any two non-empty fibers of a MonoidHom.

                                                            Equations
                                                            • f.fiberEquiv a b = (f.fiberEquivKer a).trans (f.fiberEquivKer b).symm
                                                            Instances For
                                                              @[simp]
                                                              theorem AddMonoidHom.fiberEquiv_apply {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) (b : α) (g : (f ⁻¹' {f a})) :
                                                              ((f.fiberEquiv a b) g) = b + (-a + g)
                                                              @[simp]
                                                              theorem MonoidHom.fiberEquiv_apply {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) (b : α) (g : (f ⁻¹' {f a})) :
                                                              ((f.fiberEquiv a b) g) = b * (a⁻¹ * g)
                                                              @[simp]
                                                              theorem AddMonoidHom.fiberEquiv_symm_apply {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] (f : α →+ H) (a : α) (b : α) (g : (f ⁻¹' {f b})) :
                                                              ((f.fiberEquiv a b).symm g) = a + (-b + g)
                                                              @[simp]
                                                              theorem MonoidHom.fiberEquiv_symm_apply {α : Type u_1} [Group α] {H : Type u_2} [Group H] (f : α →* H) (a : α) (b : α) (g : (f ⁻¹' {f b})) :
                                                              ((f.fiberEquiv a b).symm g) = a * (b⁻¹ * g)
                                                              noncomputable def AddMonoidHom.fiberEquivOfSurjective {α : Type u_1} [AddGroup α] {H : Type u_2} [AddGroup H] {f : α →+ H} (hf : Function.Surjective f) (h : H) (h' : H) :
                                                              (f ⁻¹' {h}) (f ⁻¹' {h'})

                                                              An equivalence between any two fibers of a surjective AddMonoidHom.

                                                              Equations
                                                              Instances For
                                                                noncomputable def MonoidHom.fiberEquivOfSurjective {α : Type u_1} [Group α] {H : Type u_2} [Group H] {f : α →* H} (hf : Function.Surjective f) (h : H) (h' : H) :
                                                                (f ⁻¹' {h}) (f ⁻¹' {h'})

                                                                An equivalence between any two fibers of a surjective MonoidHom.

                                                                Equations
                                                                Instances For
                                                                  noncomputable def QuotientAddGroup.preimageMkEquivAddSubgroupProdSet {α : Type u_1} [AddGroup α] (s : AddSubgroup α) (t : Set (α s)) :
                                                                  (QuotientAddGroup.mk ⁻¹' t) s × t

                                                                  If s is a subgroup of the additive group α, and t is a subset of α ⧸ s, then there is a (typically non-canonical) bijection between the preimage of t in α and the product s × t.

                                                                  Equations
                                                                  • One or more equations did not get rendered due to their size.
                                                                  Instances For
                                                                    noncomputable def QuotientGroup.preimageMkEquivSubgroupProdSet {α : Type u_1} [Group α] (s : Subgroup α) (t : Set (α s)) :
                                                                    (QuotientGroup.mk ⁻¹' t) s × t

                                                                    If s is a subgroup of the group α, and t is a subset of α ⧸ s, then there is a (typically non-canonical) bijection between the preimage of t in α and the product s × t.

                                                                    Equations
                                                                    • One or more equations did not get rendered due to their size.
                                                                    Instances For
                                                                      theorem QuotientAddGroup.univ_eq_iUnion_vadd {α : Type u_1} [AddGroup α] (H : AddSubgroup α) :
                                                                      Set.univ = ⋃ (x : α H), Quotient.out' x +ᵥ H

                                                                      An additive group is made up of a disjoint union of cosets of an additive subgroup.

                                                                      theorem QuotientGroup.univ_eq_iUnion_smul {α : Type u_1} [Group α] (H : Subgroup α) :
                                                                      Set.univ = ⋃ (x : α H), Quotient.out' x H

                                                                      A group is made up of a disjoint union of cosets of a subgroup.