Strong topologies on the space of continuous linear maps #
In this file, we define the strong topologies on E →L[𝕜] F
associated with a family
𝔖 : Set (Set E)
to be the topology of uniform convergence on the elements of 𝔖
(also called
the topology of 𝔖
-convergence).
The lemma UniformOnFun.continuousSMul_of_image_bounded
tells us that this is a
vector space topology if the continuous linear image of any element of 𝔖
is bounded (in the sense
of Bornology.IsVonNBounded
).
We then declare an instance for the case where 𝔖
is exactly the set of all bounded subsets of
E
, giving us the so-called "topology of uniform convergence on bounded sets" (or "topology of
bounded convergence"), which coincides with the operator norm topology in the case of
NormedSpace
s.
Other useful examples include the weak-* topology (when 𝔖
is the set of finite sets or the set
of singletons) and the topology of compact convergence (when 𝔖
is the set of relatively compact
sets).
Main definitions #
UniformConvergenceCLM
is a type synonym forE →SL[σ] F
equipped with the𝔖
-topology.UniformConvergenceCLM.instTopologicalSpace
is the topology mentioned above for an arbitrary𝔖
.ContinuousLinearMap.topologicalSpace
is the topology of bounded convergence. This is declared as an instance.
Main statements #
UniformConvergenceCLM.instTopologicalAddGroup
andUniformConvergenceCLM.instContinuousSMul
show that the strong topology makesE →L[𝕜] F
a topological vector space, with the assumptions on𝔖
mentioned above.ContinuousLinearMap.topologicalAddGroup
andContinuousLinearMap.continuousSMul
register these facts as instances for the special case of bounded convergence.
References #
- [N. Bourbaki, Topological Vector Spaces][bourbaki1987]
TODO #
- Add convergence on compact subsets
Tags #
uniform convergence, bounded convergence
𝔖-Topologies #
Given E
and F
two topological vector spaces and 𝔖 : Set (Set E)
, then
UniformConvergenceCLM σ F 𝔖
is a type synonym of E →SL[σ] F
equipped with the "topology of
uniform convergence on the elements of 𝔖
".
If the continuous linear image of any element of 𝔖
is bounded, this makes E →SL[σ] F
a
topological vector space.
Equations
- UniformConvergenceCLM σ F x = (E →SL[σ] F)
Instances For
Equations
- UniformConvergenceCLM.instFunLike σ F 𝔖 = ContinuousLinearMap.funLike
Equations
- ⋯ = ⋯
Equations
- UniformConvergenceCLM.instTopologicalSpace σ F 𝔖 = TopologicalSpace.induced DFunLike.coe (UniformOnFun.topologicalSpace E F 𝔖)
The uniform structure associated with ContinuousLinearMap.strongTopology
. We make sure
that this has nice definitional properties.
Equations
- UniformConvergenceCLM.instUniformSpace σ F 𝔖 = (UniformSpace.comap DFunLike.coe (UniformOnFun.uniformSpace E F 𝔖)).replaceTopology ⋯
Equations
- UniformConvergenceCLM.instAddCommGroup σ F 𝔖 = ContinuousLinearMap.addCommGroup
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- UniformConvergenceCLM.instDistribMulAction σ F M 𝔖 = ContinuousLinearMap.distribMulAction
Equations
- UniformConvergenceCLM.instModule σ F R 𝔖 = ContinuousLinearMap.module
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Topology of bounded convergence #
The topology of bounded convergence on E →L[𝕜] F
. This coincides with the topology induced by
the operator norm when E
and F
are normed spaces.
Equations
- ContinuousLinearMap.topologicalSpace = UniformConvergenceCLM.instTopologicalSpace σ F {S : Set E | Bornology.IsVonNBounded 𝕜₁ S}
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ContinuousLinearMap.uniformSpace = UniformConvergenceCLM.instUniformSpace σ F {S : Set E | Bornology.IsVonNBounded 𝕜₁ S}
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Pre-composition by a fixed continuous linear map as a continuous linear map. Note that in non-normed space it is not always true that composition is continuous in both variables, so we have to fix one of them.
Equations
- ContinuousLinearMap.precomp G L = { toFun := fun (f : F →SL[τ] G) => f.comp L, map_add' := ⋯, map_smul' := ⋯, cont := ⋯ }
Instances For
Post-composition by a fixed continuous linear map as a continuous linear map. Note that in non-normed space it is not always true that composition is continuous in both variables, so we have to fix one of them.
Equations
- ContinuousLinearMap.postcomp E L = { toFun := fun (f : E →SL[σ] F) => L.comp f, map_add' := ⋯, map_smul' := ⋯, cont := ⋯ }
Instances For
Send a continuous bilinear map to an abstract bilinear map (forgetting continuity).
Equations
- L.toLinearMap₂ = ContinuousLinearMap.coeLM 𝕜 ∘ₗ ↑L
Instances For
Continuous linear equivalences #
A pair of continuous (semi)linear equivalences generates a (semi)linear equivalence between the spaces of continuous (semi)linear maps.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A pair of continuous linear equivalences generates a continuous linear equivalence between the spaces of continuous linear maps.
Equations
- e₁.arrowCongr e₂ = e₁.arrowCongrSL e₂