Documentation

Mathlib.Topology.Algebra.Module.UniformConvergence

Algebraic facts about the topology of uniform convergence #

This file contains algebraic compatibility results about the uniform structure of uniform convergence / ๐”–-convergence. They will mostly be useful for defining strong topologies on the space of continuous linear maps between two topological vector spaces.

Main statements #

Implementation notes #

Like in Mathlib.Topology.UniformSpace.UniformConvergenceTopology, we use the type aliases UniformFun (denoted ฮฑ โ†’แตค ฮฒ) and UniformOnFun (denoted ฮฑ โ†’แตค[๐”–] ฮฒ) for functions from ฮฑ to ฮฒ endowed with the structures of uniform convergence and ๐”–-convergence.

References #

Tags #

uniform convergence, strong dual

theorem UniformFun.continuousSMul_induced_of_range_bounded (๐•œ : Type u_1) (ฮฑ : Type u_2) (E : Type u_3) (H : Type u_4) {hom : Type u_5} [NormedField ๐•œ] [AddCommGroup H] [Module ๐•œ H] [AddCommGroup E] [Module ๐•œ E] [TopologicalSpace H] [UniformSpace E] [UniformAddGroup E] [ContinuousSMul ๐•œ E] [FunLike hom H (ฮฑ โ†’ E)] [LinearMapClass hom ๐•œ H (ฮฑ โ†’ E)] (ฯ† : hom) (hฯ† : Inducing (โ‡‘UniformFun.ofFun โˆ˜ โ‡‘ฯ†)) (h : โˆ€ (u : H), Bornology.IsVonNBounded ๐•œ (Set.range (ฯ† u))) :
ContinuousSMul ๐•œ H

Let E be a topological vector space over a normed field ๐•œ, let ฮฑ be any type. Let H be a submodule of ฮฑ โ†’แตค E such that the range of each f โˆˆ H is von Neumann bounded. Then H is a topological vector space over ๐•œ, i.e., the pointwise scalar multiplication is continuous in both variables.

For convenience we require that H is a vector space over ๐•œ with a topology induced by UniformFun.ofFun โˆ˜ ฯ†, where ฯ† : H โ†’โ‚—[๐•œ] (ฮฑ โ†’ E).

theorem UniformOnFun.continuousSMul_induced_of_image_bounded (๐•œ : Type u_1) (ฮฑ : Type u_2) (E : Type u_3) (H : Type u_4) {hom : Type u_5} [NormedField ๐•œ] [AddCommGroup H] [Module ๐•œ H] [AddCommGroup E] [Module ๐•œ E] [TopologicalSpace H] [UniformSpace E] [UniformAddGroup E] [ContinuousSMul ๐•œ E] {๐”– : Set (Set ฮฑ)} [FunLike hom H (ฮฑ โ†’ E)] [LinearMapClass hom ๐•œ H (ฮฑ โ†’ E)] (ฯ† : hom) (hฯ† : Inducing (โ‡‘(UniformOnFun.ofFun ๐”–) โˆ˜ โ‡‘ฯ†)) (h : โˆ€ (u : H), โˆ€ s โˆˆ ๐”–, Bornology.IsVonNBounded ๐•œ (ฯ† u '' s)) :
ContinuousSMul ๐•œ H

Let E be a TVS, ๐”– : Set (Set ฮฑ) and H a submodule of ฮฑ โ†’แตค[๐”–] E. If the image of any S โˆˆ ๐”– by any u โˆˆ H is bounded (in the sense of Bornology.IsVonNBounded), then H, equipped with the topology of ๐”–-convergence, is a TVS.

For convenience, we don't literally ask for H : Submodule (ฮฑ โ†’แตค[๐”–] E). Instead, we prove the result for any vector space H equipped with a linear inducing to ฮฑ โ†’แตค[๐”–] E, which is often easier to use. We also state the Submodule version as UniformOnFun.continuousSMul_submodule_of_image_bounded.

theorem UniformOnFun.continuousSMul_submodule_of_image_bounded (๐•œ : Type u_1) (ฮฑ : Type u_2) (E : Type u_3) [NormedField ๐•œ] [AddCommGroup E] [Module ๐•œ E] [UniformSpace E] [UniformAddGroup E] [ContinuousSMul ๐•œ E] {๐”– : Set (Set ฮฑ)} (H : Submodule ๐•œ (UniformOnFun ฮฑ E ๐”–)) (h : โˆ€ u โˆˆ H, โˆ€ s โˆˆ ๐”–, Bornology.IsVonNBounded ๐•œ (u '' s)) :
ContinuousSMul ๐•œ โ†ฅH

Let E be a TVS, ๐”– : Set (Set ฮฑ) and H a submodule of ฮฑ โ†’แตค[๐”–] E. If the image of any S โˆˆ ๐”– by any u โˆˆ H is bounded (in the sense of Bornology.IsVonNBounded), then H, equipped with the topology of ๐”–-convergence, is a TVS.

If you have a hard time using this lemma, try the one above instead.