Normed spaces #
In this file we define (semi)normed spaces and algebras. We also prove some theorems about these definitions.
- smul : α → β → β
A normed space over a normed field is a vector space endowed with a norm which satisfies the equality
‖c • x‖ = ‖c‖ ‖x‖
. We require only‖c • x‖ ≤ ‖c‖ ‖x‖
in the definition, then prove‖c • x‖ = ‖c‖ ‖x‖
innorm_smul
.Note that since this requires
SeminormedAddCommGroup
and notNormedAddCommGroup
, this typeclass can be used for "semi normed spaces" too, just asModule
can be used for "semi modules".
A normed space over a normed field is a vector space endowed with a norm which satisfies the
equality ‖c • x‖ = ‖c‖ ‖x‖
. We require only ‖c • x‖ ≤ ‖c‖ ‖x‖
in the definition, then prove
‖c • x‖ = ‖c‖ ‖x‖
in norm_smul
.
Note that since this requires SeminormedAddCommGroup
and not NormedAddCommGroup
, this
typeclass can be used for "semi normed spaces" too, just as Module
can be used for
"semi modules".
Instances
The product of two normed spaces is a normed space, with the sup norm.
The product of finitely many normed spaces is a normed space, with the sup norm.
A subspace of a normed space is also a normed space, with the restriction of the norm.
A linear map from a Module
to a NormedSpace
induces a NormedSpace
structure on the
domain, using the SeminormedAddCommGroup.induced
norm.
See note [reducible non-instances]
Instances For
While this may appear identical to NormedSpace.toModule
, it contains an implicit argument
involving NormedAddCommGroup.toSeminormedAddCommGroup
that typeclass inference has trouble
inferring.
Specifically, the following instance cannot be found without this NormedSpace.toModule'
:
example
(𝕜 ι : Type*) (E : ι → Type*)
[NormedField 𝕜] [Π i, NormedAddCommGroup (E i)] [Π i, NormedSpace 𝕜 (E i)] :
Π i, Module 𝕜 (E i) := by infer_instance
This Zulip thread gives some more context.
If E
is a nontrivial topological module over ℝ
, then E
has no isolated points.
This is a particular case of Module.punctured_nhds_neBot
.
If E
is a nontrivial normed space over a nontrivially normed field 𝕜
, then E
is unbounded:
for any c : ℝ
, there exists a vector x : E
with norm strictly greater than c
.
A normed vector space over a nontrivially normed field is a noncompact space. This cannot be
an instance because in order to apply it, Lean would have to search for NormedSpace 𝕜 E
with
unknown 𝕜
. We register this as an instance in two cases: 𝕜 = E
and 𝕜 = ℝ
.
- smul : 𝕜 → 𝕜' → 𝕜'
- toFun : 𝕜 → 𝕜'
- map_one' : OneHom.toFun (↑↑Algebra.toRingHom) 1 = 1
- map_mul' : ∀ (x y : 𝕜), OneHom.toFun (↑↑Algebra.toRingHom) (x * y) = OneHom.toFun (↑↑Algebra.toRingHom) x * OneHom.toFun (↑↑Algebra.toRingHom) y
- map_zero' : OneHom.toFun (↑↑Algebra.toRingHom) 0 = 0
- map_add' : ∀ (x y : 𝕜), OneHom.toFun (↑↑Algebra.toRingHom) (x + y) = OneHom.toFun (↑↑Algebra.toRingHom) x + OneHom.toFun (↑↑Algebra.toRingHom) y
A normed algebra
𝕜'
over𝕜
is normed module that is also an algebra.See the implementation notes for
Algebra
for a discussion about non-unital algebras. Following the strategy there, a non-unital normed algebra can be written as:variables [NormedField 𝕜] [NonunitalSeminormedRing 𝕜'] variables [NormedModule 𝕜 𝕜'] [SMulCommClass 𝕜 𝕜' 𝕜'] [IsScalarTower 𝕜 𝕜' 𝕜']
A normed algebra 𝕜'
over 𝕜
is normed module that is also an algebra.
See the implementation notes for Algebra
for a discussion about non-unital algebras. Following
the strategy there, a non-unital normed algebra can be written as:
variables [NormedField 𝕜] [NonunitalSeminormedRing 𝕜']
variables [NormedModule 𝕜 𝕜'] [SMulCommClass 𝕜 𝕜' 𝕜'] [IsScalarTower 𝕜 𝕜' 𝕜']
Instances
While this may appear identical to NormedAlgebra.toNormedSpace
, it contains an implicit
argument involving NormedRing.toSeminormedRing
that typeclass inference has trouble inferring.
Specifically, the following instance cannot be found without this NormedSpace.toModule'
:
example
(𝕜 ι : Type*) (E : ι → Type*)
[NormedField 𝕜] [Π i, NormedRing (E i)] [Π i, NormedAlgebra 𝕜 (E i)] :
Π i, Module 𝕜 (E i) := by infer_instance
See NormedSpace.toModule'
for a similar situation.
In a normed algebra, the inclusion of the base field in the extended field is an isometry.
Any normed characteristic-zero division ring that is a normed algebra over the reals is also a normed algebra over the rationals.
Phrased another way, if 𝕜
is a normed algebra over the reals, then AlgebraRat
respects that
norm.
The product of two normed algebras is a normed algebra, with the sup norm.
The product of finitely many normed algebras is a normed algebra, with the sup norm.
A non-unital algebra homomorphism from an Algebra
to a NormedAlgebra
induces a
NormedAlgebra
structure on the domain, using the SeminormedRing.induced
norm.
See note [reducible non-instances]
Instances For
If E
is a normed space over 𝕜'
and 𝕜
is a normed algebra over 𝕜'
, then
RestrictScalars.module
is additionally a NormedSpace
.
The action of the original normed_field on RestrictScalars 𝕜 𝕜' E
.
This is not an instance as it would be contrary to the purpose of RestrictScalars
.
Instances For
Warning: This declaration should be used judiciously.
Please consider using IsScalarTower
and/or RestrictScalars 𝕜 𝕜' E
instead.
This definition allows the RestrictScalars.normedSpace
instance to be put directly on E
rather on RestrictScalars 𝕜 𝕜' E
. This would be a very bad instance; both because 𝕜'
cannot be
inferred, and because it is likely to create instance diamonds.