Documentation

Mathlib.Topology.Algebra.Group.Basic

Topological groups #

This file defines the following typeclasses:

There is an instance deducing ContinuousSub from IsTopologicalGroup but we use a separate typeclass because, e.g., and ℝ≥0 have continuous subtraction but are not additive groups.

We also define Homeomorph versions of several Equivs: Homeomorph.mulLeft, Homeomorph.mulRight, Homeomorph.inv, and prove a few facts about neighbourhood filters in groups.

Tags #

topological space, group, topological group

Groups with continuous multiplication #

In this section we prove a few statements about groups with continuous (*).

def Homeomorph.mulLeft {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
G ≃ₜ G

Multiplication from the left in a topological group as a homeomorphism.

Equations
Instances For

    Addition from the left in a topological additive group as a homeomorphism.

    Equations
    Instances For
      @[simp]
      theorem Homeomorph.coe_mulLeft {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
      (Homeomorph.mulLeft a) = fun (x : G) => a * x
      @[simp]
      theorem Homeomorph.coe_addLeft {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
      (Homeomorph.addLeft a) = fun (x : G) => a + x
      theorem isOpenMap_mul_left {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
      IsOpenMap fun (x : G) => a * x
      theorem isOpenMap_add_left {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
      IsOpenMap fun (x : G) => a + x
      theorem IsOpen.leftCoset {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] {U : Set G} (h : IsOpen U) (x : G) :
      IsOpen (x U)
      theorem IsOpen.left_addCoset {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] {U : Set G} (h : IsOpen U) (x : G) :
      IsOpen (x +ᵥ U)
      theorem isClosedMap_mul_left {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
      IsClosedMap fun (x : G) => a * x
      theorem isClosedMap_add_left {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
      IsClosedMap fun (x : G) => a + x
      theorem IsClosed.leftCoset {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] {U : Set G} (h : IsClosed U) (x : G) :
      theorem IsClosed.left_addCoset {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] {U : Set G} (h : IsClosed U) (x : G) :
      def Homeomorph.mulRight {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
      G ≃ₜ G

      Multiplication from the right in a topological group as a homeomorphism.

      Equations
      Instances For

        Addition from the right in a topological additive group as a homeomorphism.

        Equations
        Instances For
          @[simp]
          theorem Homeomorph.coe_mulRight {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
          (Homeomorph.mulRight a) = fun (x : G) => x * a
          @[simp]
          theorem Homeomorph.coe_addRight {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
          (Homeomorph.addRight a) = fun (x : G) => x + a
          theorem isOpenMap_mul_right {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
          IsOpenMap fun (x : G) => x * a
          theorem isOpenMap_add_right {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
          IsOpenMap fun (x : G) => x + a
          theorem IsOpen.rightCoset {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] {U : Set G} (h : IsOpen U) (x : G) :
          theorem IsOpen.right_addCoset {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] {U : Set G} (h : IsOpen U) (x : G) :
          theorem isClosedMap_mul_right {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] (a : G) :
          IsClosedMap fun (x : G) => x * a
          theorem isClosedMap_add_right {G : Type w} [TopologicalSpace G] [AddGroup G] [ContinuousAdd G] (a : G) :
          IsClosedMap fun (x : G) => x + a
          theorem IsClosed.rightCoset {G : Type w} [TopologicalSpace G] [Group G] [ContinuousMul G] {U : Set G} (h : IsClosed U) (x : G) :

          ContinuousInv and ContinuousNeg #

          theorem ContinuousInv.induced {α : Type u_1} {β : Type u_2} {F : Type u_3} [FunLike F α β] [Group α] [Group β] [MonoidHomClass F α β] [ : TopologicalSpace β] [ContinuousInv β] (f : F) :
          theorem ContinuousNeg.induced {α : Type u_1} {β : Type u_2} {F : Type u_3} [FunLike F α β] [AddGroup α] [AddGroup β] [AddMonoidHomClass F α β] [ : TopologicalSpace β] [ContinuousNeg β] (f : F) :
          theorem Specializes.inv {G : Type w} [TopologicalSpace G] [Inv G] [ContinuousInv G] {x y : G} (h : x y) :
          theorem Specializes.neg {G : Type w} [TopologicalSpace G] [Neg G] [ContinuousNeg G] {x y : G} (h : x y) :
          (-x) (-y)
          theorem Inseparable.neg {G : Type w} [TopologicalSpace G] [Neg G] [ContinuousNeg G] {x y : G} (h : Inseparable x y) :
          theorem Specializes.zpow {G : Type u_1} [DivInvMonoid G] [TopologicalSpace G] [ContinuousMul G] [ContinuousInv G] {x y : G} (h : x y) (m : ) :
          (x ^ m) (y ^ m)
          theorem Specializes.zsmul {G : Type u_1} [SubNegMonoid G] [TopologicalSpace G] [ContinuousAdd G] [ContinuousNeg G] {x y : G} (h : x y) (m : ) :
          (m x) (m y)
          theorem Inseparable.zpow {G : Type u_1} [DivInvMonoid G] [TopologicalSpace G] [ContinuousMul G] [ContinuousInv G] {x y : G} (h : Inseparable x y) (m : ) :
          Inseparable (x ^ m) (y ^ m)
          theorem Inseparable.zsmul {G : Type u_1} [SubNegMonoid G] [TopologicalSpace G] [ContinuousAdd G] [ContinuousNeg G] {x y : G} (h : Inseparable x y) (m : ) :
          Inseparable (m x) (m y)
          instance Pi.continuousInv {ι : Type u_1} {C : ιType u_2} [(i : ι) → TopologicalSpace (C i)] [(i : ι) → Inv (C i)] [∀ (i : ι), ContinuousInv (C i)] :
          ContinuousInv ((i : ι) → C i)
          instance Pi.continuousNeg {ι : Type u_1} {C : ιType u_2} [(i : ι) → TopologicalSpace (C i)] [(i : ι) → Neg (C i)] [∀ (i : ι), ContinuousNeg (C i)] :
          ContinuousNeg ((i : ι) → C i)
          instance Pi.has_continuous_inv' {G : Type w} [TopologicalSpace G] [Inv G] [ContinuousInv G] {ι : Type u_1} :
          ContinuousInv (ιG)

          A version of Pi.continuousInv for non-dependent functions. It is needed because sometimes Lean fails to use Pi.continuousInv for non-dependent functions.

          instance Pi.has_continuous_neg' {G : Type w} [TopologicalSpace G] [Neg G] [ContinuousNeg G] {ι : Type u_1} :
          ContinuousNeg (ιG)

          A version of Pi.continuousNeg for non-dependent functions. It is needed because sometimes Lean fails to use Pi.continuousNeg for non-dependent functions.

          theorem isClosed_setOf_map_inv (G₁ : Type u_2) (G₂ : Type u_3) [TopologicalSpace G₂] [T2Space G₂] [Inv G₁] [Inv G₂] [ContinuousInv G₂] :
          IsClosed {f : G₁G₂ | ∀ (x : G₁), f x⁻¹ = (f x)⁻¹}
          theorem isClosed_setOf_map_neg (G₁ : Type u_2) (G₂ : Type u_3) [TopologicalSpace G₂] [T2Space G₂] [Neg G₁] [Neg G₂] [ContinuousNeg G₂] :
          IsClosed {f : G₁G₂ | ∀ (x : G₁), f (-x) = -f x}

          Inversion in a topological group as a homeomorphism.

          Equations
          Instances For

            Negation in a topological group as a homeomorphism.

            Equations
            Instances For
              theorem nhds_neg (G : Type w) [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] (a : G) :
              nhds (-a) = -nhds a
              theorem IsOpen.inv {G : Type w} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] {s : Set G} (hs : IsOpen s) :
              theorem IsOpen.neg {G : Type w} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] {s : Set G} (hs : IsOpen s) :
              theorem IsClosed.neg {G : Type w} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] {s : Set G} (hs : IsClosed s) :
              @[simp]
              @[simp]
              theorem continuous_neg_iff {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} :
              @[simp]
              theorem continuousAt_inv_iff {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] [TopologicalSpace α] {f : αG} {x : α} :
              @[simp]
              theorem continuousAt_neg_iff {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} {x : α} :
              @[simp]
              theorem continuousOn_inv_iff {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] [TopologicalSpace α] {f : αG} {s : Set α} :
              @[simp]
              theorem continuousOn_neg_iff {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} {s : Set α} :
              theorem Continuous.of_inv {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] [TopologicalSpace α] {f : αG} :

              Alias of the forward direction of continuous_inv_iff.

              theorem Continuous.of_neg {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} :
              theorem ContinuousAt.of_inv {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] [TopologicalSpace α] {f : αG} {x : α} :

              Alias of the forward direction of continuousAt_inv_iff.

              theorem ContinuousAt.of_neg {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} {x : α} :
              theorem ContinuousOn.of_inv {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] [TopologicalSpace α] {f : αG} {s : Set α} :

              Alias of the forward direction of continuousOn_inv_iff.

              theorem ContinuousOn.of_neg {G : Type w} {α : Type u} [TopologicalSpace G] [InvolutiveNeg G] [ContinuousNeg G] [TopologicalSpace α] {f : αG} {s : Set α} :
              theorem continuousInv_sInf {G : Type w} [Inv G] {ts : Set (TopologicalSpace G)} (h : tts, ContinuousInv G) :
              theorem continuousNeg_sInf {G : Type w} [Neg G] {ts : Set (TopologicalSpace G)} (h : tts, ContinuousNeg G) :
              theorem continuousInv_iInf {G : Type w} {ι' : Sort u_1} [Inv G] {ts' : ι'TopologicalSpace G} (h' : ∀ (i : ι'), ContinuousInv G) :
              theorem continuousNeg_iInf {G : Type w} {ι' : Sort u_1} [Neg G] {ts' : ι'TopologicalSpace G} (h' : ∀ (i : ι'), ContinuousNeg G) :
              theorem continuousInv_inf {G : Type w} [Inv G] {t₁ t₂ : TopologicalSpace G} (h₁ : ContinuousInv G) (h₂ : ContinuousInv G) :
              theorem continuousNeg_inf {G : Type w} [Neg G] {t₁ t₂ : TopologicalSpace G} (h₁ : ContinuousNeg G) (h₂ : ContinuousNeg G) :
              theorem Topology.IsInducing.continuousInv {G : Type u_1} {H : Type u_2} [Inv G] [Inv H] [TopologicalSpace G] [TopologicalSpace H] [ContinuousInv H] {f : GH} (hf : IsInducing f) (hf_inv : ∀ (x : G), f x⁻¹ = (f x)⁻¹) :
              theorem Topology.IsInducing.continuousNeg {G : Type u_1} {H : Type u_2} [Neg G] [Neg H] [TopologicalSpace G] [TopologicalSpace H] [ContinuousNeg H] {f : GH} (hf : IsInducing f) (hf_inv : ∀ (x : G), f (-x) = -f x) :
              @[deprecated Topology.IsInducing.continuousInv (since := "2024-10-28")]
              theorem Inducing.continuousInv {G : Type u_1} {H : Type u_2} [Inv G] [Inv H] [TopologicalSpace G] [TopologicalSpace H] [ContinuousInv H] {f : GH} (hf : Topology.IsInducing f) (hf_inv : ∀ (x : G), f x⁻¹ = (f x)⁻¹) :

              Alias of Topology.IsInducing.continuousInv.

              Topological groups #

              A topological group is a group in which the multiplication and inversion operations are continuous. Topological additive groups are defined in the same way. Equivalently, we can require that the division operation x y ↦ x * y⁻¹ (resp., subtraction) is continuous.

              theorem IsTopologicalGroup.continuous_conj_prod {G : Type w} [TopologicalSpace G] [Inv G] [Mul G] [ContinuousMul G] [ContinuousInv G] :
              Continuous fun (g : G × G) => g.1 * g.2 * g.1⁻¹

              Conjugation is jointly continuous on G × G when both mul and inv are continuous.

              theorem IsTopologicalAddGroup.continuous_addConj_prod {G : Type w} [TopologicalSpace G] [Neg G] [Add G] [ContinuousAdd G] [ContinuousNeg G] :
              Continuous fun (g : G × G) => g.1 + g.2 + -g.1

              Conjugation is jointly continuous on G × G when both add and neg are continuous.

              @[deprecated IsTopologicalAddGroup.continuous_addConj_prod (since := "2025-03-11")]
              theorem IsTopologicalAddGroup.continuous_conj_sum {G : Type w} [TopologicalSpace G] [Neg G] [Add G] [ContinuousAdd G] [ContinuousNeg G] :
              Continuous fun (g : G × G) => g.1 + g.2 + -g.1

              Alias of IsTopologicalAddGroup.continuous_addConj_prod.


              Conjugation is jointly continuous on G × G when both add and neg are continuous.

              theorem IsTopologicalGroup.continuous_conj {G : Type w} [TopologicalSpace G] [Inv G] [Mul G] [ContinuousMul G] (g : G) :
              Continuous fun (h : G) => g * h * g⁻¹

              Conjugation by a fixed element is continuous when mul is continuous.

              theorem IsTopologicalAddGroup.continuous_conj {G : Type w} [TopologicalSpace G] [Neg G] [Add G] [ContinuousAdd G] (g : G) :
              Continuous fun (h : G) => g + h + -g

              Conjugation by a fixed element is continuous when add is continuous.

              theorem IsTopologicalGroup.continuous_conj' {G : Type w} [TopologicalSpace G] [Inv G] [Mul G] [ContinuousMul G] [ContinuousInv G] (h : G) :
              Continuous fun (g : G) => g * h * g⁻¹

              Conjugation acting on fixed element of the group is continuous when both mul and inv are continuous.

              theorem IsTopologicalAddGroup.continuous_conj' {G : Type w} [TopologicalSpace G] [Neg G] [Add G] [ContinuousAdd G] [ContinuousNeg G] (h : G) :
              Continuous fun (g : G) => g + h + -g

              Conjugation acting on fixed element of the additive group is continuous when both add and neg are continuous.

              theorem continuous_zpow {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (z : ) :
              Continuous fun (a : G) => a ^ z
              theorem continuous_zsmul {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (z : ) :
              Continuous fun (a : G) => z a
              theorem Continuous.zpow {G : Type w} {α : Type u} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [TopologicalSpace α] {f : αG} (h : Continuous f) (z : ) :
              Continuous fun (b : α) => f b ^ z
              theorem Continuous.zsmul {G : Type w} {α : Type u} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] [TopologicalSpace α] {f : αG} (h : Continuous f) (z : ) :
              Continuous fun (b : α) => z f b
              theorem continuousOn_zpow {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {s : Set G} (z : ) :
              ContinuousOn (fun (x : G) => x ^ z) s
              theorem continuousOn_zsmul {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {s : Set G} (z : ) :
              ContinuousOn (fun (x : G) => z x) s
              theorem continuousAt_zpow {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x : G) (z : ) :
              ContinuousAt (fun (x : G) => x ^ z) x
              theorem continuousAt_zsmul {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x : G) (z : ) :
              ContinuousAt (fun (x : G) => z x) x
              theorem Filter.Tendsto.zpow {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {α : Type u_1} {l : Filter α} {f : αG} {x : G} (hf : Tendsto f l (nhds x)) (z : ) :
              Tendsto (fun (x : α) => f x ^ z) l (nhds (x ^ z))
              theorem Filter.Tendsto.zsmul {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {α : Type u_1} {l : Filter α} {f : αG} {x : G} (hf : Tendsto f l (nhds x)) (z : ) :
              Tendsto (fun (x : α) => z f x) l (nhds (z x))
              theorem ContinuousWithinAt.zpow {G : Type w} {α : Type u} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [TopologicalSpace α] {f : αG} {x : α} {s : Set α} (hf : ContinuousWithinAt f s x) (z : ) :
              ContinuousWithinAt (fun (x : α) => f x ^ z) s x
              theorem ContinuousWithinAt.zsmul {G : Type w} {α : Type u} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] [TopologicalSpace α] {f : αG} {x : α} {s : Set α} (hf : ContinuousWithinAt f s x) (z : ) :
              ContinuousWithinAt (fun (x : α) => z f x) s x
              theorem ContinuousAt.zpow {G : Type w} {α : Type u} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [TopologicalSpace α] {f : αG} {x : α} (hf : ContinuousAt f x) (z : ) :
              ContinuousAt (fun (x : α) => f x ^ z) x
              theorem ContinuousAt.zsmul {G : Type w} {α : Type u} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] [TopologicalSpace α] {f : αG} {x : α} (hf : ContinuousAt f x) (z : ) :
              ContinuousAt (fun (x : α) => z f x) x
              theorem ContinuousOn.zpow {G : Type w} {α : Type u} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [TopologicalSpace α] {f : αG} {s : Set α} (hf : ContinuousOn f s) (z : ) :
              ContinuousOn (fun (x : α) => f x ^ z) s
              theorem ContinuousOn.zsmul {G : Type w} {α : Type u} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] [TopologicalSpace α] {f : αG} {s : Set α} (hf : ContinuousOn f s) (z : ) :
              ContinuousOn (fun (x : α) => z f x) s
              @[deprecated tendsto_neg_nhdsGT (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsGT.

              @[deprecated tendsto_inv_nhdsGT (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsGT.

              @[deprecated tendsto_neg_nhdsLT (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsLT.

              @[deprecated tendsto_inv_nhdsLT (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsLT.

              @[deprecated tendsto_neg_nhdsGT_neg (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsGT_neg.

              @[deprecated tendsto_inv_nhdsGT_inv (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsGT_inv.

              @[deprecated tendsto_neg_nhdsLT_neg (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsLT_neg.

              @[deprecated tendsto_inv_nhdsLT_inv (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsLT_inv.

              @[deprecated tendsto_neg_nhdsGE (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsGE.

              @[deprecated tendsto_inv_nhdsGE (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsGE.

              @[deprecated tendsto_neg_nhdsLE (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsLE.

              @[deprecated tendsto_inv_nhdsLE (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsLE.

              @[deprecated tendsto_neg_nhdsGE_neg (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsGE_neg.

              @[deprecated tendsto_inv_nhdsGE_inv (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsGE_inv.

              @[deprecated tendsto_neg_nhdsLE_neg (since := "2024-12-22")]

              Alias of tendsto_neg_nhdsLE_neg.

              @[deprecated tendsto_inv_nhdsLE_inv (since := "2024-12-22")]

              Alias of tendsto_inv_nhdsLE_inv.

              instance Pi.topologicalGroup {β : Type v} {C : βType u_1} [(b : β) → TopologicalSpace (C b)] [(b : β) → Group (C b)] [∀ (b : β), IsTopologicalGroup (C b)] :
              IsTopologicalGroup ((b : β) → C b)
              instance Pi.topologicalAddGroup {β : Type v} {C : βType u_1} [(b : β) → TopologicalSpace (C b)] [(b : β) → AddGroup (C b)] [∀ (b : β), IsTopologicalAddGroup (C b)] :
              IsTopologicalAddGroup ((b : β) → C b)

              If multiplication is continuous in α, then it also is in αᵐᵒᵖ.

              If addition is continuous in α, then it also is in αᵃᵒᵖ.

              theorem inv_mem_nhds_one (G : Type w) [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {S : Set G} (hS : S nhds 1) :
              theorem neg_mem_nhds_zero (G : Type w) [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {S : Set G} (hS : S nhds 0) :

              The map (x, y) ↦ (x, x * y) as a homeomorphism. This is a shear mapping.

              Equations
              Instances For

                The map (x, y) ↦ (x, x + y) as a homeomorphism. This is a shear mapping.

                Equations
                Instances For
                  @[simp]
                  theorem Homeomorph.shearMulRight_coe (G : Type w) [TopologicalSpace G] [Group G] [IsTopologicalGroup G] :
                  (Homeomorph.shearMulRight G) = fun (z : G × G) => (z.1, z.1 * z.2)
                  @[simp]
                  theorem Homeomorph.shearAddRight_coe (G : Type w) [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] :
                  (Homeomorph.shearAddRight G) = fun (z : G × G) => (z.1, z.1 + z.2)
                  @[simp]
                  @[deprecated Topology.IsInducing.topologicalGroup (since := "2024-10-28")]
                  theorem Inducing.topologicalGroup {G : Type w} {H : Type x} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {F : Type u_1} [Group H] [TopologicalSpace H] [FunLike F H G] [MonoidHomClass F H G] (f : F) (hf : Topology.IsInducing f) :

                  Alias of Topology.IsInducing.topologicalGroup.

                  theorem topologicalGroup_induced {G : Type w} {H : Type x} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {F : Type u_1} [Group H] [FunLike F H G] [MonoidHomClass F H G] (f : F) :

                  The (topological-space) closure of a subgroup of a topological group is itself a subgroup.

                  Equations
                  Instances For

                    The (topological-space) closure of an additive subgroup of an additive topological group is itself an additive subgroup.

                    Equations
                    Instances For

                      The topological closure of a normal subgroup is normal.

                      The topological closure of a normal additive subgroup is normal.

                      The connected component of 1 is a subgroup of G.

                      Equations
                      Instances For

                        The connected component of 0 is a subgroup of G.

                        Equations
                        Instances For
                          @[reducible, inline]
                          abbrev Subgroup.commGroupTopologicalClosure {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [T2Space G] (s : Subgroup G) (hs : ∀ (x y : s), x * y = y * x) :

                          If a subgroup of a topological group is commutative, then so is its topological closure.

                          See note [reducible non-instances].

                          Equations
                          Instances For
                            @[reducible, inline]

                            If a subgroup of an additive topological group is commutative, then so is its topological closure.

                            See note [reducible non-instances].

                            Equations
                            Instances For
                              theorem exists_nhds_split_inv {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {s : Set G} (hs : s nhds 1) :
                              Vnhds 1, vV, wV, v / w s
                              theorem exists_nhds_half_neg {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {s : Set G} (hs : s nhds 0) :
                              Vnhds 0, vV, wV, v - w s
                              theorem nhds_translation_mul_inv {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x : G) :
                              Filter.comap (fun (x_1 : G) => x_1 * x⁻¹) (nhds 1) = nhds x
                              theorem nhds_translation_add_neg {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x : G) :
                              Filter.comap (fun (x_1 : G) => x_1 + -x) (nhds 0) = nhds x
                              @[simp]
                              theorem map_mul_left_nhds {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x y : G) :
                              Filter.map (fun (x_1 : G) => x * x_1) (nhds y) = nhds (x * y)
                              @[simp]
                              theorem map_add_left_nhds {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x y : G) :
                              Filter.map (fun (x_1 : G) => x + x_1) (nhds y) = nhds (x + y)
                              theorem map_mul_left_nhds_one {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x : G) :
                              Filter.map (fun (x_1 : G) => x * x_1) (nhds 1) = nhds x
                              theorem map_add_left_nhds_zero {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x : G) :
                              Filter.map (fun (x_1 : G) => x + x_1) (nhds 0) = nhds x
                              @[simp]
                              theorem map_mul_right_nhds {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x y : G) :
                              Filter.map (fun (x_1 : G) => x_1 * x) (nhds y) = nhds (y * x)
                              @[simp]
                              theorem map_add_right_nhds {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x y : G) :
                              Filter.map (fun (x_1 : G) => x_1 + x) (nhds y) = nhds (y + x)
                              theorem map_mul_right_nhds_one {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x : G) :
                              Filter.map (fun (x_1 : G) => x_1 * x) (nhds 1) = nhds x
                              theorem map_add_right_nhds_zero {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x : G) :
                              Filter.map (fun (x_1 : G) => x_1 + x) (nhds 0) = nhds x
                              theorem Filter.HasBasis.nhds_of_one {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {ι : Sort u_1} {p : ιProp} {s : ιSet G} (hb : (nhds 1).HasBasis p s) (x : G) :
                              (nhds x).HasBasis p fun (i : ι) => {y : G | y / x s i}
                              theorem Filter.HasBasis.nhds_of_zero {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {ι : Sort u_1} {p : ιProp} {s : ιSet G} (hb : (nhds 0).HasBasis p s) (x : G) :
                              (nhds x).HasBasis p fun (i : ι) => {y : G | y - x s i}
                              theorem mem_closure_iff_nhds_one {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {x : G} {s : Set G} :
                              x closure s Unhds 1, ys, y / x U
                              theorem mem_closure_iff_nhds_zero {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {x : G} {s : Set G} :
                              x closure s Unhds 0, ys, y - x U
                              theorem continuous_of_continuousAt_one {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {M : Type u_1} {hom : Type u_2} [MulOneClass M] [TopologicalSpace M] [ContinuousMul M] [FunLike hom G M] [MonoidHomClass hom G M] (f : hom) (hf : ContinuousAt (⇑f) 1) :

                              A monoid homomorphism (a bundled morphism of a type that implements MonoidHomClass) from a topological group to a topological monoid is continuous provided that it is continuous at one. See also uniformContinuous_of_continuousAt_one.

                              theorem continuous_of_continuousAt_zero {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {M : Type u_1} {hom : Type u_2} [AddZeroClass M] [TopologicalSpace M] [ContinuousAdd M] [FunLike hom G M] [AddMonoidHomClass hom G M] (f : hom) (hf : ContinuousAt (⇑f) 0) :

                              An additive monoid homomorphism (a bundled morphism of a type that implements AddMonoidHomClass) from an additive topological group to an additive topological monoid is continuous provided that it is continuous at zero. See also uniformContinuous_of_continuousAt_zero.

                              theorem continuous_of_continuousAt_one₂ {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {H : Type u_1} {M : Type u_2} [CommMonoid M] [TopologicalSpace M] [ContinuousMul M] [Group H] [TopologicalSpace H] [IsTopologicalGroup H] (f : G →* H →* M) (hf : ContinuousAt (fun (x : G × H) => (f x.1) x.2) (1, 1)) (hl : ∀ (x : G), ContinuousAt (⇑(f x)) 1) (hr : ∀ (y : H), ContinuousAt (fun (x : G) => (f x) y) 1) :
                              Continuous fun (x : G × H) => (f x.1) x.2
                              theorem continuous_of_continuousAt_zero₂ {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {H : Type u_1} {M : Type u_2} [AddCommMonoid M] [TopologicalSpace M] [ContinuousAdd M] [AddGroup H] [TopologicalSpace H] [IsTopologicalAddGroup H] (f : G →+ H →+ M) (hf : ContinuousAt (fun (x : G × H) => (f x.1) x.2) (0, 0)) (hl : ∀ (x : G), ContinuousAt (⇑(f x)) 0) (hr : ∀ (y : H), ContinuousAt (fun (x : G) => (f x) y) 0) :
                              Continuous fun (x : G × H) => (f x.1) x.2

                              Let A and B be topological groups, and let φ : A → B be a continuous surjective group homomorphism. Assume furthermore that φ is a quotient map (i.e., V ⊆ B is open iff φ⁻¹ V is open). Then φ is an open quotient map, and in particular an open map.

                              Let A and B be topological additive groups, and let φ : A → B be a continuous surjective additive group homomorphism. Assume furthermore that φ is a quotient map (i.e., V ⊆ B is open iff φ⁻¹ V is open). Then φ is an open quotient map, and in particular an open map.

                              theorem IsTopologicalGroup.ext {G : Type u_1} [Group G] {t t' : TopologicalSpace G} (tg : IsTopologicalGroup G) (tg' : IsTopologicalGroup G) (h : nhds 1 = nhds 1) :
                              t = t'
                              theorem IsTopologicalAddGroup.ext {G : Type u_1} [AddGroup G] {t t' : TopologicalSpace G} (tg : IsTopologicalAddGroup G) (tg' : IsTopologicalAddGroup G) (h : nhds 0 = nhds 0) :
                              t = t'
                              theorem IsTopologicalGroup.ext_iff {G : Type u_1} [Group G] {t t' : TopologicalSpace G} (tg : IsTopologicalGroup G) (tg' : IsTopologicalGroup G) :
                              t = t' nhds 1 = nhds 1
                              theorem ContinuousInv.of_nhds_one {G : Type u_1} [Group G] [TopologicalSpace G] (hinv : Filter.Tendsto (fun (x : G) => x⁻¹) (nhds 1) (nhds 1)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ * x) (nhds 1)) (hconj : ∀ (x₀ : G), Filter.Tendsto (fun (x : G) => x₀ * x * x₀⁻¹) (nhds 1) (nhds 1)) :
                              theorem ContinuousNeg.of_nhds_zero {G : Type u_1} [AddGroup G] [TopologicalSpace G] (hinv : Filter.Tendsto (fun (x : G) => -x) (nhds 0) (nhds 0)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ + x) (nhds 0)) (hconj : ∀ (x₀ : G), Filter.Tendsto (fun (x : G) => x₀ + x + -x₀) (nhds 0) (nhds 0)) :
                              theorem IsTopologicalGroup.of_nhds_one' {G : Type u} [Group G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 * x2) (nhds 1 ×ˢ nhds 1) (nhds 1)) (hinv : Filter.Tendsto (fun (x : G) => x⁻¹) (nhds 1) (nhds 1)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ * x) (nhds 1)) (hright : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x * x₀) (nhds 1)) :
                              theorem IsTopologicalAddGroup.of_nhds_zero' {G : Type u} [AddGroup G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 + x2) (nhds 0 ×ˢ nhds 0) (nhds 0)) (hinv : Filter.Tendsto (fun (x : G) => -x) (nhds 0) (nhds 0)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ + x) (nhds 0)) (hright : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x + x₀) (nhds 0)) :
                              theorem IsTopologicalGroup.of_nhds_one {G : Type u} [Group G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 * x2) (nhds 1 ×ˢ nhds 1) (nhds 1)) (hinv : Filter.Tendsto (fun (x : G) => x⁻¹) (nhds 1) (nhds 1)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ * x) (nhds 1)) (hconj : ∀ (x₀ : G), Filter.Tendsto (fun (x : G) => x₀ * x * x₀⁻¹) (nhds 1) (nhds 1)) :
                              theorem IsTopologicalAddGroup.of_nhds_zero {G : Type u} [AddGroup G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 + x2) (nhds 0 ×ˢ nhds 0) (nhds 0)) (hinv : Filter.Tendsto (fun (x : G) => -x) (nhds 0) (nhds 0)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ + x) (nhds 0)) (hconj : ∀ (x₀ : G), Filter.Tendsto (fun (x : G) => x₀ + x + -x₀) (nhds 0) (nhds 0)) :
                              theorem IsTopologicalGroup.of_comm_of_nhds_one {G : Type u} [CommGroup G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 * x2) (nhds 1 ×ˢ nhds 1) (nhds 1)) (hinv : Filter.Tendsto (fun (x : G) => x⁻¹) (nhds 1) (nhds 1)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ * x) (nhds 1)) :
                              theorem IsTopologicalAddGroup.of_comm_of_nhds_zero {G : Type u} [AddCommGroup G] [TopologicalSpace G] (hmul : Filter.Tendsto (Function.uncurry fun (x1 x2 : G) => x1 + x2) (nhds 0 ×ˢ nhds 0) (nhds 0)) (hinv : Filter.Tendsto (fun (x : G) => -x) (nhds 0) (nhds 0)) (hleft : ∀ (x₀ : G), nhds x₀ = Filter.map (fun (x : G) => x₀ + x) (nhds 0)) :
                              theorem IsTopologicalGroup.exists_antitone_basis_nhds_one (G : Type w) [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [FirstCountableTopology G] :
                              ∃ (u : Set G), (nhds 1).HasAntitoneBasis u ∀ (n : ), u (n + 1) * u (n + 1) u n

                              Any first countable topological group has an antitone neighborhood basis u : ℕ → Set G for which (u (n + 1)) ^ 2 ⊆ u n. The existence of such a neighborhood basis is a key tool for QuotientGroup.completeSpace

                              Any first countable topological additive group has an antitone neighborhood basis u : ℕ → set G for which u (n + 1) + u (n + 1) ⊆ u n. The existence of such a neighborhood basis is a key tool for QuotientAddGroup.completeSpace

                              theorem Filter.Tendsto.const_div' {G : Type w} {α : Type u} [TopologicalSpace G] [Div G] [ContinuousDiv G] (b : G) {c : G} {f : αG} {l : Filter α} (h : Tendsto f l (nhds c)) :
                              Tendsto (fun (k : α) => b / f k) l (nhds (b / c))
                              theorem Filter.Tendsto.const_sub {G : Type w} {α : Type u} [TopologicalSpace G] [Sub G] [ContinuousSub G] (b : G) {c : G} {f : αG} {l : Filter α} (h : Tendsto f l (nhds c)) :
                              Tendsto (fun (k : α) => b - f k) l (nhds (b - c))
                              theorem Filter.tendsto_const_div_iff {α : Type u} {G : Type u_1} [CommGroup G] [TopologicalSpace G] [ContinuousDiv G] (b : G) {c : G} {f : αG} {l : Filter α} :
                              Tendsto (fun (k : α) => b / f k) l (nhds (b / c)) Tendsto f l (nhds c)
                              theorem Filter.tendsto_const_sub_iff {α : Type u} {G : Type u_1} [AddCommGroup G] [TopologicalSpace G] [ContinuousSub G] (b : G) {c : G} {f : αG} {l : Filter α} :
                              Tendsto (fun (k : α) => b - f k) l (nhds (b - c)) Tendsto f l (nhds c)
                              theorem Filter.Tendsto.div_const' {G : Type w} {α : Type u} [TopologicalSpace G] [Div G] [ContinuousDiv G] {c : G} {f : αG} {l : Filter α} (h : Tendsto f l (nhds c)) (b : G) :
                              Tendsto (fun (x : α) => f x / b) l (nhds (c / b))
                              theorem Filter.Tendsto.sub_const {G : Type w} {α : Type u} [TopologicalSpace G] [Sub G] [ContinuousSub G] {c : G} {f : αG} {l : Filter α} (h : Tendsto f l (nhds c)) (b : G) :
                              Tendsto (fun (x : α) => f x - b) l (nhds (c - b))
                              theorem Filter.tendsto_div_const_iff {α : Type u} {G : Type u_1} [CommGroupWithZero G] [TopologicalSpace G] [ContinuousDiv G] {b : G} (hb : b 0) {c : G} {f : αG} {l : Filter α} :
                              Tendsto (fun (x : α) => f x / b) l (nhds (c / b)) Tendsto f l (nhds c)
                              theorem Filter.tendsto_sub_const_iff {α : Type u} {G : Type u_1} [AddCommGroup G] [TopologicalSpace G] [ContinuousSub G] (b : G) {c : G} {f : αG} {l : Filter α} :
                              Tendsto (fun (x : α) => f x - b) l (nhds (c - b)) Tendsto f l (nhds c)
                              theorem continuous_div_left' {G : Type w} [TopologicalSpace G] [Div G] [ContinuousDiv G] (a : G) :
                              Continuous fun (x : G) => a / x
                              theorem continuous_sub_left {G : Type w} [TopologicalSpace G] [Sub G] [ContinuousSub G] (a : G) :
                              Continuous fun (x : G) => a - x
                              theorem continuous_div_right' {G : Type w} [TopologicalSpace G] [Div G] [ContinuousDiv G] (a : G) :
                              Continuous fun (x : G) => x / a
                              theorem continuous_sub_right {G : Type w} [TopologicalSpace G] [Sub G] [ContinuousSub G] (a : G) :
                              Continuous fun (x : G) => x - a

                              A version of Homeomorph.mulLeft a b⁻¹ that is defeq to a / b.

                              Equations
                              Instances For

                                A version of Homeomorph.addLeft a (-b) that is defeq to a - b.

                                Equations
                                Instances For
                                  @[simp]
                                  theorem Homeomorph.subLeft_apply {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (x b : G) :
                                  (subLeft x) b = x - b
                                  @[simp]
                                  @[simp]
                                  theorem Homeomorph.divLeft_apply {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (x b : G) :
                                  (divLeft x) b = x / b
                                  theorem isOpenMap_div_left {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (a : G) :
                                  IsOpenMap fun (x : G) => a / x
                                  theorem isOpenMap_sub_left {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (a : G) :
                                  IsOpenMap fun (x : G) => a - x
                                  theorem isClosedMap_div_left {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (a : G) :
                                  IsClosedMap fun (x : G) => a / x
                                  theorem isClosedMap_sub_left {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (a : G) :
                                  IsClosedMap fun (x : G) => a - x

                                  A version of Homeomorph.mulRight a⁻¹ b that is defeq to b / a.

                                  Equations
                                  Instances For

                                    A version of Homeomorph.addRight (-a) b that is defeq to b - a.

                                    Equations
                                    Instances For
                                      @[simp]
                                      @[simp]
                                      theorem Homeomorph.divRight_apply {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (x b : G) :
                                      (divRight x) b = b / x
                                      @[simp]
                                      theorem Homeomorph.subRight_apply {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (x b : G) :
                                      (subRight x) b = b - x
                                      theorem isOpenMap_div_right {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (a : G) :
                                      IsOpenMap fun (x : G) => x / a
                                      theorem isOpenMap_sub_right {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (a : G) :
                                      IsOpenMap fun (x : G) => x - a
                                      theorem isClosedMap_div_right {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (a : G) :
                                      IsClosedMap fun (x : G) => x / a
                                      theorem isClosedMap_sub_right {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (a : G) :
                                      IsClosedMap fun (x : G) => x - a
                                      theorem tendsto_div_nhds_one_iff {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] {α : Type u_1} {l : Filter α} {x : G} {u : αG} :
                                      Filter.Tendsto (fun (x_1 : α) => u x_1 / x) l (nhds 1) Filter.Tendsto u l (nhds x)
                                      theorem tendsto_sub_nhds_zero_iff {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] {α : Type u_1} {l : Filter α} {x : G} {u : αG} :
                                      Filter.Tendsto (fun (x_1 : α) => u x_1 - x) l (nhds 0) Filter.Tendsto u l (nhds x)
                                      theorem nhds_translation_div {G : Type w} [Group G] [TopologicalSpace G] [IsTopologicalGroup G] (x : G) :
                                      Filter.comap (fun (x_1 : G) => x_1 / x) (nhds 1) = nhds x
                                      theorem nhds_translation_sub {G : Type w} [AddGroup G] [TopologicalSpace G] [IsTopologicalAddGroup G] (x : G) :
                                      Filter.comap (fun (x_1 : G) => x_1 - x) (nhds 0) = nhds x

                                      A subgroup S of a topological group G acts on G properly discontinuously on the left, if it is discrete in the sense that S ∩ K is finite for all compact K. (See also DiscreteTopology.)

                                      A subgroup S of an additive topological group G acts on G properly discontinuously on the left, if it is discrete in the sense that S ∩ K is finite for all compact K. (See also DiscreteTopology.

                                      A subgroup S of a topological group G acts on G properly discontinuously on the right, if it is discrete in the sense that S ∩ K is finite for all compact K. (See also DiscreteTopology.)

                                      If G is Hausdorff, this can be combined with t2Space_of_properlyDiscontinuousSMul_of_t2Space to show that the quotient group G ⧸ S is Hausdorff.

                                      A subgroup S of an additive topological group G acts on G properly discontinuously on the right, if it is discrete in the sense that S ∩ K is finite for all compact K. (See also DiscreteTopology.)

                                      If G is Hausdorff, this can be combined with t2Space_of_properlyDiscontinuousVAdd_of_t2Space to show that the quotient group G ⧸ S is Hausdorff.

                                      Some results about an open set containing the product of two sets in a topological group.

                                      theorem compact_open_separated_mul_right {G : Type w} [TopologicalSpace G] [MulOneClass G] [ContinuousMul G] {K U : Set G} (hK : IsCompact K) (hU : IsOpen U) (hKU : K U) :
                                      Vnhds 1, K * V U

                                      Given a compact set K inside an open set U, there is an open neighborhood V of 1 such that K * V ⊆ U.

                                      theorem compact_open_separated_add_right {G : Type w} [TopologicalSpace G] [AddZeroClass G] [ContinuousAdd G] {K U : Set G} (hK : IsCompact K) (hU : IsOpen U) (hKU : K U) :
                                      Vnhds 0, K + V U

                                      Given a compact set K inside an open set U, there is an open neighborhood V of 0 such that K + V ⊆ U.

                                      theorem compact_open_separated_mul_left {G : Type w} [TopologicalSpace G] [MulOneClass G] [ContinuousMul G] {K U : Set G} (hK : IsCompact K) (hU : IsOpen U) (hKU : K U) :
                                      Vnhds 1, V * K U

                                      Given a compact set K inside an open set U, there is an open neighborhood V of 1 such that V * K ⊆ U.

                                      theorem compact_open_separated_add_left {G : Type w} [TopologicalSpace G] [AddZeroClass G] [ContinuousAdd G] {K U : Set G} (hK : IsCompact K) (hU : IsOpen U) (hKU : K U) :
                                      Vnhds 0, V + K U

                                      Given a compact set K inside an open set U, there is an open neighborhood V of 0 such that V + K ⊆ U.

                                      theorem compact_covered_by_mul_left_translates {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] {K V : Set G} (hK : IsCompact K) (hV : (interior V).Nonempty) :
                                      ∃ (t : Finset G), K gt, (fun (x : G) => g * x) ⁻¹' V

                                      A compact set is covered by finitely many left multiplicative translates of a set with non-empty interior.

                                      theorem compact_covered_by_add_left_translates {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] {K V : Set G} (hK : IsCompact K) (hV : (interior V).Nonempty) :
                                      ∃ (t : Finset G), K gt, (fun (x : G) => g + x) ⁻¹' V

                                      A compact set is covered by finitely many left additive translates of a set with non-empty interior.

                                      @[instance 100]

                                      Every weakly locally compact separable topological group is σ-compact. Note: this is not true if we drop the topological group hypothesis.

                                      @[instance 100]

                                      Every weakly locally compact separable topological additive group is σ-compact. Note: this is not true if we drop the topological group hypothesis.

                                      theorem exists_disjoint_smul_of_isCompact {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] [NoncompactSpace G] {K L : Set G} (hK : IsCompact K) (hL : IsCompact L) :
                                      ∃ (g : G), Disjoint K (g L)

                                      Given two compact sets in a noncompact topological group, there is a translate of the second one that is disjoint from the first one.

                                      theorem exists_disjoint_vadd_of_isCompact {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] [NoncompactSpace G] {K L : Set G} (hK : IsCompact K) (hL : IsCompact L) :
                                      ∃ (g : G), Disjoint K (g +ᵥ L)

                                      Given two compact sets in a noncompact additive topological group, there is a translate of the second one that is disjoint from the first one.

                                      theorem nhds_mul {G : Type w} [TopologicalSpace G] [Group G] [IsTopologicalGroup G] (x y : G) :
                                      nhds (x * y) = nhds x * nhds y
                                      theorem nhds_add {G : Type w} [TopologicalSpace G] [AddGroup G] [IsTopologicalAddGroup G] (x y : G) :
                                      nhds (x + y) = nhds x + nhds y

                                      On a topological group, 𝓝 : G → Filter G can be promoted to a MulHom.

                                      Equations
                                      Instances For

                                        On an additive topological group, 𝓝 : G → Filter G can be promoted to an AddHom.

                                        Equations
                                        Instances For
                                          @[simp]

                                          If G is a group with topological ⁻¹, then it is homeomorphic to its units.

                                          Equations
                                          Instances For

                                            If G is an additive group with topological negation, then it is homeomorphic to its additive units.

                                            Equations
                                            Instances For
                                              @[deprecated Units.isEmbedding_val (since := "2024-10-26")]

                                              Alias of Units.isEmbedding_val.

                                              def Homeomorph.prodUnits {α : Type u} {β : Type v} [Monoid α] [TopologicalSpace α] [Monoid β] [TopologicalSpace β] :
                                              (α × β)ˣ ≃ₜ αˣ × βˣ

                                              The topological group isomorphism between the units of a product of two monoids, and the product of the units of each monoid.

                                              Equations
                                              Instances For

                                                The topological group isomorphism between the additive units of a product of two additive monoids, and the product of the additive units of each additive monoid.

                                                Equations
                                                Instances For
                                                  @[deprecated Homeomorph.prodAddUnits (since := "2025-02-21")]

                                                  Alias of Homeomorph.prodAddUnits.


                                                  The topological group isomorphism between the additive units of a product of two additive monoids, and the product of the additive units of each additive monoid.

                                                  Equations
                                                  Instances For
                                                    @[deprecated Homeomorph.prodUnits (since := "2025-02-21")]
                                                    def Units.Homeomorph.prodUnits {α : Type u} {β : Type v} [Monoid α] [TopologicalSpace α] [Monoid β] [TopologicalSpace β] :
                                                    (α × β)ˣ ≃ₜ αˣ × βˣ

                                                    Alias of Homeomorph.prodUnits.


                                                    The topological group isomorphism between the units of a product of two monoids, and the product of the units of each monoid.

                                                    Equations
                                                    Instances For
                                                      theorem topologicalGroup_sInf {G : Type w} [Group G] {ts : Set (TopologicalSpace G)} (h : tts, IsTopologicalGroup G) :
                                                      theorem topologicalGroup_iInf {G : Type w} {ι : Sort u_1} [Group G] {ts' : ιTopologicalSpace G} (h' : ∀ (i : ι), IsTopologicalGroup G) :
                                                      theorem topologicalAddGroup_iInf {G : Type w} {ι : Sort u_1} [AddGroup G] {ts' : ιTopologicalSpace G} (h' : ∀ (i : ι), IsTopologicalAddGroup G) :