Inheritance of normed algebraic structures by bounded continuous functions #
For various types of normed algebraic structures β
, we show in this file that the space of
bounded continuous functions from α
to β
inherits the same normed structure, by using
pointwise operations and checking that they are compatible with the uniform distance.
Equations
- BoundedContinuousFunction.instNorm = { norm := fun (x : BoundedContinuousFunction α β) => dist x 0 }
The norm of a bounded continuous function is the supremum of ‖f x‖
.
We use sInf
to ensure that the definition works if α
has no elements.
When the domain is non-empty, we do not need the 0 ≤ C
condition in the formula for ‖f‖
as a
sInf
.
Distance between the images of any two points is at most twice the norm of the function.
The norm of a function is controlled by the supremum of the pointwise norms.
Norm of const α b
is less than or equal to ‖b‖
. If α
is nonempty,
then it is equal to ‖b‖
.
Constructing a bounded continuous function from a uniformly bounded continuous function taking values in a normed group.
Equations
- BoundedContinuousFunction.ofNormedAddCommGroup f Hf C H = { toFun := fun (n : α) => f n, continuous_toFun := Hf, map_bounded' := ⋯ }
Instances For
Constructing a bounded continuous function from a uniformly bounded function on a discrete space, taking values in a normed group.
Equations
Instances For
Taking the pointwise norm of a bounded continuous function with values in a
SeminormedAddCommGroup
yields a bounded continuous function with values in ℝ.
Equations
Instances For
If ‖(1 : β)‖ = 1
, then ‖(1 : α →ᵇ β)‖ = 1
if α
is nonempty.
The pointwise opposite of a bounded continuous function is again bounded continuous.
Equations
- BoundedContinuousFunction.instNeg = { neg := fun (f : BoundedContinuousFunction α β) => BoundedContinuousFunction.ofNormedAddCommGroup (-⇑f) ⋯ ‖f‖ ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- BoundedContinuousFunction.instAddCommGroup = Function.Injective.addCommGroup (fun (f : BoundedContinuousFunction α β) => ⇑f) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
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.
The nnnorm
of a function is controlled by the supremum of the pointwise nnnorm
s.
Equations
- BoundedContinuousFunction.instNormedSpace = { toModule := BoundedContinuousFunction.instModule, norm_smul_le := ⋯ }
Postcomposition of bounded continuous functions into a normed module by a continuous linear map
is a continuous linear map.
Upgraded version of ContinuousLinearMap.compLeftContinuous
, similar to LinearMap.compLeft
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- BoundedContinuousFunction.instNonUnitalRing = Function.Injective.nonUnitalRing (fun (f : BoundedContinuousFunction α R) => ⇑f) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- One or more equations did not get rendered due to their size.
Equations
- BoundedContinuousFunction.instNonUnitalSeminormedCommRing = { toNonUnitalSeminormedRing := BoundedContinuousFunction.instNonUnitalSeminormedRing, mul_comm := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- BoundedContinuousFunction.instNonUnitalNormedCommRing = { toNonUnitalNormedRing := BoundedContinuousFunction.instNonUnitalNormedRing, mul_comm := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- BoundedContinuousFunction.instNatCast_1 = { natCast := fun (n : ℕ) => BoundedContinuousFunction.const α ↑n }
Equations
- BoundedContinuousFunction.instIntCast_1 = { intCast := fun (n : ℤ) => BoundedContinuousFunction.const α ↑n }
Equations
- BoundedContinuousFunction.instRing = Function.Injective.ring (fun (f : BoundedContinuousFunction α R) => ⇑f) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- One or more equations did not get rendered due to their size.
Composition on the left by a (lipschitz-continuous) homomorphism of topological semirings, as a
RingHom
. Similar to RingHom.compLeftContinuous
.
Equations
- RingHom.compLeftContinuousBounded α g hg = { toMonoidHom := MonoidHom.compLeftContinuousBounded α (↑g) hg, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- BoundedContinuousFunction.instCommRing = { toRing := BoundedContinuousFunction.instRing, mul_comm := ⋯ }
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.
BoundedContinuousFunction.const
as a RingHom
.
Equations
- BoundedContinuousFunction.C = { toFun := fun (c : 𝕜) => BoundedContinuousFunction.const α ((algebraMap 𝕜 γ) c), map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Equations
- BoundedContinuousFunction.instAlgebra = { toSMul := BoundedContinuousFunction.instSMul, algebraMap := BoundedContinuousFunction.C, commutes' := ⋯, smul_def' := ⋯ }
Equations
- BoundedContinuousFunction.instNormedAlgebra = { toAlgebra := BoundedContinuousFunction.instAlgebra, norm_smul_le := ⋯ }
Composition on the left by a (lipschitz-continuous) homomorphism of topological R
-algebras,
as an AlgHom
. Similar to AlgHom.compLeftContinuous
.
Equations
- BoundedContinuousFunction.AlgHom.compLeftContinuousBounded 𝕜 g hg = { toRingHom := RingHom.compLeftContinuousBounded α g.toRingHom hg, commutes' := ⋯ }
Instances For
The algebra-homomorphism forgetting that a bounded continuous function is bounded.
Equations
- BoundedContinuousFunction.toContinuousMapₐ 𝕜 = { toFun := toContinuousMap, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯, commutes' := ⋯ }
Instances For
Structure as normed module over scalar functions #
If β
is a normed 𝕜
-space, then we show that the space of bounded continuous
functions from α
to β
is naturally a module over the algebra of bounded continuous
functions from α
to 𝕜
.
Equations
- One or more equations did not get rendered due to their size.
Equations
Equations
- BoundedContinuousFunction.instPartialOrder = PartialOrder.lift (fun (f : BoundedContinuousFunction α β) => f.toFun) ⋯
Equations
- BoundedContinuousFunction.instSup = { max := fun (f g : BoundedContinuousFunction α β) => { toFun := ⇑f ⊔ ⇑g, continuous_toFun := ⋯, map_bounded' := ⋯ } }
Equations
- BoundedContinuousFunction.instInf = { min := fun (f g : BoundedContinuousFunction α β) => { toFun := ⇑f ⊓ ⇑g, continuous_toFun := ⋯, map_bounded' := ⋯ } }
Equations
- BoundedContinuousFunction.instSemilatticeSup = Function.Injective.semilatticeSup (fun (f : BoundedContinuousFunction α β) => ⇑f) ⋯ ⋯
Equations
- BoundedContinuousFunction.instSemilatticeInf = Function.Injective.semilatticeInf (fun (f : BoundedContinuousFunction α β) => ⇑f) ⋯ ⋯
Equations
- BoundedContinuousFunction.instLattice = Function.Injective.lattice (fun (f : BoundedContinuousFunction α β) => ⇑f) ⋯ ⋯ ⋯
The nonnegative part of a bounded continuous ℝ
-valued function as a bounded
continuous ℝ≥0
-valued function.
Equations
Instances For
The absolute value of a bounded continuous ℝ
-valued function as a bounded
continuous ℝ≥0
-valued function.
Equations
- f.nnnorm = BoundedContinuousFunction.comp (fun (x : ℝ) => ‖x‖₊) BoundedContinuousFunction.nnnorm._proof_147 f
Instances For
Decompose a bounded continuous function to its positive and negative parts.
Express the absolute value of a bounded continuous function in terms of its positive and negative parts.