Topological groups #
This file defines the following typeclasses:
- TopologicalGroup,- TopologicalAddGroup: multiplicative and additive topological groups, i.e., groups with continuous- (*)and- (⁻¹)/- (+)and- (-);
- ContinuousSub Gmeans that- Ghas a continuous subtraction operation.
There is an instance deducing ContinuousSub from TopologicalGroup 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 (*).
Addition from the left in a topological additive group as a homeomorphism.
Equations
- Homeomorph.addLeft a = let __src := Equiv.addLeft a; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Multiplication from the left in a topological group as a homeomorphism.
Equations
- Homeomorph.mulLeft a = let __src := Equiv.mulLeft a; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Addition from the right in a topological additive group as a homeomorphism.
Equations
- Homeomorph.addRight a = let __src := Equiv.addRight a; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Multiplication from the right in a topological group as a homeomorphism.
Equations
- Homeomorph.mulRight a = let __src := Equiv.mulRight a; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Basic hypothesis to talk about a topological additive group. A topological additive group
over M, for example, is obtained by requiring the instances AddGroup M and
ContinuousAdd M and ContinuousNeg M.
- continuous_neg : Continuous fun (a : G) => -a
Instances
Basic hypothesis to talk about a topological group. A topological group over M, for example,
is obtained by requiring the instances Group M and ContinuousMul M and
ContinuousInv M.
- continuous_inv : Continuous fun (a : G) => a⁻¹
Instances
Equations
- ⋯ = ⋯
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
If a function converges to a value in an additive topological group, then its negation converges to the negation of this value.
If a function converges to a value in a multiplicative topological group, then its inverse
converges to the inverse of this value. For the version in normed fields assuming additionally
that the limit is nonzero, use Tendsto.inv'.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A version of Pi.continuousNeg for non-dependent functions. It is needed
because sometimes Lean fails to use Pi.continuousNeg for non-dependent functions.
Equations
- ⋯ = ⋯
A version of Pi.continuousInv for non-dependent functions. It is needed because sometimes
Lean fails to use Pi.continuousInv for non-dependent functions.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Negation in a topological group as a homeomorphism.
Equations
- Homeomorph.neg G = let __src := Equiv.neg G; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Inversion in a topological group as a homeomorphism.
Equations
- Homeomorph.inv G = let __src := Equiv.inv G; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Alias of the forward direction of continuous_inv_iff.
Alias of the forward direction of continuousAt_inv_iff.
Alias of the forward direction of continuousOn_inv_iff.
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.
A topological (additive) group is a group in which the addition and negation operations are continuous.
Instances
A topological group is a group in which the multiplication and inversion operations are continuous.
When you declare an instance that does not already have a UniformSpace instance,
you should also provide an instance of UniformSpace and UniformGroup using
TopologicalGroup.toUniformSpace and topologicalCommGroup_isUniform.
Instances
Equations
- ⋯ = ⋯
Conjugation is jointly continuous on G × G when both add and neg are continuous.
Conjugation is jointly continuous on G × G when both mul and inv are continuous.
Conjugation by a fixed element is continuous when add is continuous.
Conjugation by a fixed element is continuous when mul is continuous.
Conjugation acting on fixed element of the additive group is continuous when both
add and neg are continuous.
Conjugation acting on fixed element of the group is continuous when both mul and
inv are continuous.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
If addition is continuous in α, then it also is in αᵃᵒᵖ.
Equations
- ⋯ = ⋯
If multiplication is continuous in α, then it also is in αᵐᵒᵖ.
Equations
- ⋯ = ⋯
The map (x, y) ↦ (x, x + y) as a homeomorphism. This is a shear mapping.
Equations
- Homeomorph.shearAddRight G = let __src := (Equiv.refl G).prodShear Equiv.addLeft; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
The map (x, y) ↦ (x, x * y) as a homeomorphism. This is a shear mapping.
Equations
- Homeomorph.shearMulRight G = let __src := (Equiv.refl G).prodShear Equiv.mulLeft; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
The (topological-space) closure of an additive subgroup of an additive topological group is itself an additive subgroup.
Equations
Instances For
The (topological-space) closure of a subgroup of a topological group is itself a subgroup.
Equations
Instances For
The topological closure of a normal additive subgroup is normal.
The topological closure of a normal subgroup is normal.
The connected component of 0 is a subgroup of G.
Equations
- AddSubgroup.connectedComponentOfZero G = { carrier := connectedComponent 0, add_mem' := ⋯, zero_mem' := ⋯, neg_mem' := ⋯ }
Instances For
The connected component of 1 is a subgroup of G.
Equations
- Subgroup.connectedComponentOfOne G = { carrier := connectedComponent 1, mul_mem' := ⋯, one_mem' := ⋯, inv_mem' := ⋯ }
Instances For
If a subgroup of an additive topological group is commutative, then so is its topological closure.
Equations
- s.addCommGroupTopologicalClosure hs = let __src := s.topologicalClosure.toAddGroup; let __src_1 := s.addCommMonoidTopologicalClosure hs; AddCommGroup.mk ⋯
Instances For
If a subgroup of a topological group is commutative, then so is its topological closure.
Equations
- s.commGroupTopologicalClosure hs = let __src := s.topologicalClosure.toGroup; let __src_1 := s.commMonoidTopologicalClosure hs; CommGroup.mk ⋯
Instances For
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.
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.
Equations
- QuotientAddGroup.Quotient.topologicalSpace N = instTopologicalSpaceQuotient
Equations
- QuotientGroup.Quotient.topologicalSpace N = instTopologicalSpaceQuotient
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient.
Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient.
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
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
In a first countable topological additive group G with normal additive subgroup
N, 0 : G ⧸ N has a countable neighborhood basis.
Equations
- ⋯ = ⋯
In a first countable topological group G with normal subgroup N, 1 : G ⧸ N has a
countable neighborhood basis.
Equations
- ⋯ = ⋯
A typeclass saying that p : G × G ↦ p.1 - p.2 is a continuous function. This property
automatically holds for topological additive groups but it also holds, e.g., for ℝ≥0.
- continuous_sub : Continuous fun (p : G × G) => p.1 - p.2
Instances
A typeclass saying that p : G × G ↦ p.1 / p.2 is a continuous function. This property
automatically holds for topological groups. Lemmas using this class have primes.
The unprimed version is for GroupWithZero.
- continuous_div' : Continuous fun (p : G × G) => p.1 / p.2
Instances
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A version of Homeomorph.addLeft a (-b) that is defeq to a - b.
Equations
- Homeomorph.subLeft x = let __src := Equiv.subLeft x; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A version of Homeomorph.mulLeft a b⁻¹ that is defeq to a / b.
Equations
- Homeomorph.divLeft x = let __src := Equiv.divLeft x; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A version of Homeomorph.addRight (-a) b that is defeq to b - a.
Equations
- Homeomorph.subRight x = let __src := Equiv.subRight x; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A version of Homeomorph.mulRight a⁻¹ b that is defeq to b / a.
Equations
- Homeomorph.divRight x = let __src := Equiv.divRight x; { toEquiv := __src, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Topological operations on pointwise sums and products #
A few results about interior and closure of the pointwise addition/multiplication of sets in groups
with continuous addition/multiplication. See also Submonoid.top_closure_mul_self_eq in
Topology.Algebra.Monoid.
One may expect a version of IsClosed.smul_left_of_isCompact where t is compact and s is
closed, but such a lemma can't be true in this level of generality. For a counterexample, consider
ℚ acting on ℝ by translation, and let s : Set ℚ := univ, t : set ℝ := {0}. Then s is
closed and t is compact, but s +ᵥ t is the set of all rationals, which is definitely not
closed in ℝ.
To fix the proof, we would need to make two additional assumptions:
- for any x ∈ t,s • {x}is closed
- for any x ∈ t, there is a continuous functiong : s • {x} → ssuch that, for ally ∈ s • {x}, we havey = (g y) • xThese are fairly specific hypotheses so we don't state this version of the lemmas, but an interesting fact is that these two assumptions are verified in the case of aNormedAddTorsor(or really, anyAddTorsorwith continuous-ᵥ). We prove this special case inIsClosed.vadd_right_of_isCompact.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Given a neighborhood U of the identity, one may find a neighborhood V of the
identity which is closed, symmetric, and satisfies V + V ⊆ U.
Given a neighborhood U of the identity, one may find a neighborhood V of the identity which
is closed, symmetric, and satisfies V * V ⊆ U.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
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 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 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.
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.
Some results about an open set containing the product of two sets in a topological group.
Given a compact set K inside an open set U, there is an open neighborhood V of
0 such that 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.
Given a compact set K inside an open set U, there is an open neighborhood V of
0 such that 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.
A compact set is covered by finitely many left additive translates of a set with non-empty interior.
A compact set is covered by finitely many left multiplicative translates of a set with non-empty interior.
Every weakly locally compact separable topological additive group is σ-compact. Note: this is not true if we drop the topological group hypothesis.
Equations
- ⋯ = ⋯
Every weakly locally compact separable topological group is σ-compact. Note: this is not true if we drop the topological group hypothesis.
Equations
- ⋯ = ⋯
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.
Given two compact sets in a noncompact topological group, there is a translate of the second one that is disjoint from the first one.
A compact neighborhood of 0 in a topological additive group
admits a closed compact subset that is a neighborhood of 0.
A compact neighborhood of 1 in a topological group admits a closed compact subset
that is a neighborhood of 1.
If a point in a topological group has a compact neighborhood, then the group is locally compact.
A topological group which is weakly locally compact is automatically locally compact.
If a function defined on a topological additive group has a support contained in a compact set, then either the function is trivial or the group is locally compact.
If a function defined on a topological group has a support contained in a compact set, then either the function is trivial or the group is locally compact.
If a function defined on a topological additive group has compact support, then either the function is trivial or the group is locally compact.
If a function defined on a topological group has compact support, then either the function is trivial or the group is locally compact.
In a locally compact additive group, any neighborhood of the identity contains a compact closed neighborhood of the identity, even without separation assumptions on the space.
In a locally compact group, any neighborhood of the identity contains a compact closed neighborhood of the identity, even without separation assumptions on the space.
Equations
- ⋯ = ⋯
Instances For
Equations
- ⋯ = ⋯
A quotient of a locally compact group is locally compact.
Equations
- ⋯ = ⋯
On an additive topological group, 𝓝 : G → Filter G can be promoted to an AddHom.
Equations
- nhdsAddHom = { toFun := nhds, map_add' := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
The quotient of a second countable additive topological group by a subgroup is second countable.
Equations
- ⋯ = ⋯
The quotient of a second countable topological group by a subgroup is second countable.
Equations
- ⋯ = ⋯
If G is an additive group with topological negation, then it is homeomorphic to
its additive units.
Equations
- toAddUnits_homeomorph = { toEquiv := toAddUnits.toEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
If G is a group with topological ⁻¹, then it is homeomorphic to its units.
Equations
- toUnits_homeomorph = { toEquiv := toUnits.toEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
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
- AddUnits.Homeomorph.sumAddUnits = { toEquiv := AddEquiv.prodAddUnits.toEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
The topological group isomorphism between the units of a product of two monoids, and the product of the units of each monoid.
Equations
- Units.Homeomorph.prodUnits = { toEquiv := MulEquiv.prodUnits.toEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Lattice of group topologies #
We define a type class GroupTopology α which endows a group α with a topology such that all
group operations are continuous.
Group topologies on a fixed group α are ordered, by reverse inclusion. They form a complete
lattice, with ⊥ the discrete topology and ⊤ the indiscrete topology.
Any function f : α → β induces coinduced f : TopologicalSpace α → GroupTopology β.
The additive version AddGroupTopology α and corresponding results are provided as well.
A group topology on a group α is a topology for which multiplication and inversion
are continuous.
Instances For
An additive group topology on an additive group α is a topology for which addition and
negation are continuous.
Instances For
A version of the global continuous_add suitable for dot notation.
A version of the global continuous_mul suitable for dot notation.
A version of the global continuous_neg suitable for dot notation.
A version of the global continuous_inv suitable for dot notation.
The ordering on group topologies on the group γ. t ≤ s if every set open in s
is also open in t (t is finer than s).
Equations
- AddGroupTopology.instPartialOrder = PartialOrder.lift AddGroupTopology.toTopologicalSpace ⋯
The ordering on group topologies on the group γ. t ≤ s if every set open in s is also open
in t (t is finer than s).
Equations
- GroupTopology.instPartialOrder = PartialOrder.lift GroupTopology.toTopologicalSpace ⋯
Equations
- AddGroupTopology.instBoundedOrder = BoundedOrder.mk
Equations
- GroupTopology.instBoundedOrder = BoundedOrder.mk
Equations
- AddGroupTopology.instInf = { inf := fun (x y : AddGroupTopology α) => { toTopologicalSpace := x.toTopologicalSpace ⊓ y.toTopologicalSpace, toTopologicalAddGroup := ⋯ } }
Equations
- GroupTopology.instInf = { inf := fun (x y : GroupTopology α) => { toTopologicalSpace := x.toTopologicalSpace ⊓ y.toTopologicalSpace, toTopologicalGroup := ⋯ } }
Equations
- AddGroupTopology.instSemilatticeInf = Function.Injective.semilatticeInf AddGroupTopology.toTopologicalSpace ⋯ ⋯
Equations
- GroupTopology.instSemilatticeInf = Function.Injective.semilatticeInf GroupTopology.toTopologicalSpace ⋯ ⋯
Infimum of a collection of additive group topologies
Equations
- AddGroupTopology.instInfSet = { sInf := fun (S : Set (AddGroupTopology α)) => { toTopologicalSpace := sInf (AddGroupTopology.toTopologicalSpace '' S), toTopologicalAddGroup := ⋯ } }
Infimum of a collection of group topologies.
Equations
- GroupTopology.instInfSet = { sInf := fun (S : Set (GroupTopology α)) => { toTopologicalSpace := sInf (GroupTopology.toTopologicalSpace '' S), toTopologicalGroup := ⋯ } }
Group topologies on γ form a complete lattice, with ⊥ the discrete topology and
⊤ the indiscrete topology.
The infimum of a collection of group topologies is the topology generated by all their open sets (which is a group topology).
The supremum of two group topologies s and t is the infimum of the family of all group
topologies contained in the intersection of s and t.
Equations
- One or more equations did not get rendered due to their size.
Group topologies on γ form a complete lattice, with ⊥ the discrete topology and ⊤ the
indiscrete topology.
The infimum of a collection of group topologies is the topology generated by all their open sets (which is a group topology).
The supremum of two group topologies s and t is the infimum of the family of all group
topologies contained in the intersection of s and t.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Given f : α → β and a topology on α, the coinduced additive group topology on β
is the finest topology such that f is continuous and β is a topological additive group.
Equations
- AddGroupTopology.coinduced f = sInf {b : AddGroupTopology β | TopologicalSpace.coinduced f t ≤ b.toTopologicalSpace}
Instances For
Given f : α → β and a topology on α, the coinduced group topology on β is the finest
topology such that f is continuous and β is a topological group.
Equations
- GroupTopology.coinduced f = sInf {b : GroupTopology β | TopologicalSpace.coinduced f t ≤ b.toTopologicalSpace}