Operator norm on the space of continuous linear maps #
Define the operator norm on the space of continuous (semi)linear maps between normed spaces, and prove its basic properties. In particular, show that this space is itself a normed space.
Since a lot of elementary properties don't require ‖x‖ = 0 → x = 0
we start setting up the
theory for SeminormedAddCommGroup
and we specialize to NormedAddCommGroup
at the end.
Note that most of statements that apply to semilinear maps only hold when the ring homomorphism
is isometric, as expressed by the typeclass [RingHomIsometric σ]
.
If ‖x‖ = 0
and f
is continuous then ‖f x‖ = 0
.
A continuous linear map between seminormed spaces is bounded when the field is nontrivially
normed. The continuity ensures boundedness on a ball of some radius ε
. The nontriviality of the
norm is then used to rescale any element into an element of norm in [ε/C, ε]
, whose image has a
controlled norm. The norm control for the original element follows by rescaling.
Given a unit-length element x
of a normed space E
over a field 𝕜
, the natural linear
isometry map from 𝕜
to E
by taking multiples of x
.
Instances For
The operator norm of a continuous linear map is the inf of all its bounds.
Instances For
If one controls the norm of every A x
, then one controls the norm of A
.
If one controls the norm of every A x
, ‖x‖ ≠ 0
, then one controls the norm of A
.
The norm of the 0
operator is 0
.
The norm of the identity is at most 1
. It is in fact 1
, except when the space is trivial
where it is 0
. It means that one can not do better than an inequality in general.
The fundamental property of the operator norm: ‖f x‖ ≤ ‖f‖ * ‖x‖
.
The image of the unit ball under a continuous linear map is bounded.
For a continuous real linear map f
, if one controls the norm of every f x
, ‖x‖ = 1
, then
one controls the norm of f
.
The operator norm satisfies the triangle inequality.
If there is an element with norm different from 0
, then the norm of the identity equals 1
.
(Since we are working with seminorms supposing that the space is non-trivial is not enough.)
Continuous linear maps themselves form a seminormed space with respect to
the operator norm. This is only a temporary definition because we want to replace the topology
with ContinuousLinearMap.topologicalSpace
to avoid diamond issues.
See Note [forgetful inheritance]
Instances For
The PseudoMetricSpace
structure on E →SL[σ₁₂] F
coming from
ContinuousLinearMap.tmpSeminormedAddCommGroup
.
See Note [forgetful inheritance]
Instances For
The UniformSpace
structure on E →SL[σ₁₂] F
coming from
ContinuousLinearMap.tmpSeminormedAddCommGroup
.
See Note [forgetful inheritance]
Instances For
The TopologicalSpace
structure on E →SL[σ₁₂] F
coming from
ContinuousLinearMap.tmpSeminormedAddCommGroup
.
See Note [forgetful inheritance]
Instances For
Continuous linear maps themselves form a seminormed space with respect to the operator norm.
If one controls the norm of every A x
, then one controls the norm of A
.
If one controls the norm of every A x
, ‖x‖₊ ≠ 0
, then one controls the norm of A
.
For a continuous real linear map f
, if one controls the norm of every f x
, ‖x‖₊ = 1
, then
one controls the norm of f
.
The operator norm is submultiplicative.
The operator norm is submultiplicative.
Continuous linear maps form a seminormed ring with respect to the operator norm.
For a normed space E
, continuous linear endomorphisms form a normed algebra with
respect to the operator norm.
continuous linear maps are Lipschitz continuous.
Evaluation of a continuous linear map f
at a point is Lipschitz continuous in f
.
ContinuousLinearMap.prod
as a LinearIsometryEquiv
.
Instances For
If a continuous linear map is constructed from a linear map via the constructor mkContinuous
,
then its norm is bounded by the bound given to the constructor if it is nonnegative.
If a continuous linear map is constructed from a linear map via the constructor mkContinuous
,
then its norm is bounded by the bound or zero if bound is negative.
Create a bilinear map (represented as a map E →L[𝕜] F →L[𝕜] G
) from the corresponding linear
map and existence of a bound on the norm of the image. The linear map can be constructed using
LinearMap.mk₂
.
If you have an explicit bound, use LinearMap.mkContinuous₂
instead, as a norm estimate will
follow automatically in LinearMap.mkContinuous₂_norm_le
.
Instances For
Create a bilinear map (represented as a map E →L[𝕜] F →L[𝕜] G
) from the corresponding linear
map and a bound on the norm of the image. The linear map can be constructed using
LinearMap.mk₂
. Lemmas LinearMap.mkContinuous₂_norm_le'
and LinearMap.mkContinuous₂_norm_le
provide estimates on the norm of an operator constructed using this function.
Instances For
Flip the order of arguments of a continuous bilinear map.
For a version bundled as LinearIsometryEquiv
, see
ContinuousLinearMap.flipL
.
Instances For
Flip the order of arguments of a continuous bilinear map.
This is a version bundled as a LinearIsometryEquiv
.
For an unbundled version see ContinuousLinearMap.flip
.
Instances For
Flip the order of arguments of a continuous bilinear map.
This is a version bundled as a LinearIsometryEquiv
.
For an unbundled version see ContinuousLinearMap.flip
.
Instances For
The continuous semilinear map obtained by applying a continuous semilinear map at a given vector.
This is the continuous version of LinearMap.applyₗ
.
Instances For
The continuous semilinear map obtained by applying a continuous semilinear map at a given vector.
This is the continuous version of LinearMap.applyₗ
.
Instances For
Composition of continuous semilinear maps as a continuous semibilinear map.
Instances For
Porting note: Local instance for norm_compSL_le
.
Should be by inferInstance
, and indeed not be needed.
Instances For
Composition of continuous linear maps as a continuous bilinear map.
Instances For
Porting note: Local instance for norm_compL_le
.
Should be by inferInstance
, and indeed not be needed.
Instances For
Apply L(x,-)
pointwise to bilinear maps, as a continuous bilinear map
Instances For
Apply L(-,y)
pointwise to bilinear maps, as a continuous bilinear map
Instances For
Porting note: Local instances for norm_precompR_le
.
Should be by inferInstance
, and indeed not be needed.
Instances For
Instances For
Porting note: Local instance for norm_precompL_le
.
Should be by inferInstance
, and indeed not be needed.
Instances For
ContinuousLinearMap.prodMap
as a continuous linear map.
Instances For
Multiplication in a non-unital normed algebra as a continuous bilinear map.
Instances For
Multiplication on the left in a non-unital normed algebra 𝕜'
as a non-unital algebra
homomorphism into the algebra of continuous linear maps. This is the left regular representation
of A
acting on itself.
This has more algebraic structure than ContinuousLinearMap.mul
, but there is no longer continuity
bundled in the first coordinate. An alternative viewpoint is that this upgrades
NonUnitalAlgHom.lmul
from a homomorphism into linear maps to a homomorphism into continuous
linear maps.
Instances For
Simultaneous left- and right-multiplication in a non-unital normed algebra, considered as a
continuous trilinear map. This is akin to its non-continuous version LinearMap.mulLeftRight
,
but there is a minor difference: LinearMap.mulLeftRight
is uncurried.
Instances For
Porting note: Local instance for op_norm_mulLeftRight_le
.
Should be by inferInstance
, and indeed not be needed.
Instances For
- isometry_mul' : Isometry ↑(ContinuousLinearMap.mul 𝕜 𝕜')
The left regular representation of the algebra on itself is an isometry.
This is a mixin class for non-unital normed algebras which states that the left-regular
representation of the algebra on itself is isometric. Every unital normed algebra with ‖1‖ = 1
is
a regular normed algebra (see NormedAlgebra.instRegularNormedAlgebra
). In addition, so is every
C⋆-algebra, non-unital included (see CstarRing.instRegularNormedAlgebra
), but there are yet other
examples. Any algebra with an approximate identity (e.g., $$L^1$$) is also regular.
This is a useful class because it gives rise to a nice norm on the unitization; in particular it is
a C⋆-norm when the norm on A
is a C⋆-norm.
Instances
Every (unital) normed algebra such that ‖1‖ = 1
is a RegularNormedAlgebra
.
Multiplication in a normed algebra as a linear isometry to the space of continuous linear maps.
Instances For
Scalar multiplication as a continuous bilinear map.
Instances For
The norm of lsmul
is at most 1 in any semi-normed group.
ContinuousLinearMap.restrictScalars
as a LinearIsometry
.
Instances For
ContinuousLinearMap.restrictScalars
as a ContinuousLinearMap
.
Instances For
Compose a bilinear map E →SL[σ₁₃] F →SL[σ₂₃] G
with two linear maps
E' →SL[σ₁'] E
and F' →SL[σ₂'] F
.
Instances For
Derivative of a continuous bilinear map f : E →L[𝕜] F →L[𝕜] G
interpreted as a map E × F → G
at point p : E × F
evaluated at q : E × F
, as a continuous bilinear map.
Instances For
LinearMap.bound_of_ball_bound'
is a version of this lemma over a field satisfying IsROrC
that produces a concrete bound.
An operator is zero iff its norm vanishes.
If a normed space is non-trivial, then the norm of the identity equals 1
.
Continuous linear maps themselves form a normed space with respect to the operator norm.
Continuous linear maps form a normed ring with respect to the operator norm.
If a continuous linear map is a topology embedding, then it is expands the distances by a positive factor.
Construct a bundled continuous (semi)linear map from a map f : E → F
and a proof of the fact
that it belongs to the closure of the image of a bounded set s : Set (E →SL[σ₁₂] F)
under coercion
to function. Coercion to function of the result is definitionally equal to f
.