Normed (semi)groups #
In this file we define 10 classes:
Norm
,NNNorm
: auxiliary classes endowing a typeα
with a functionnorm : α → ℝ
(notation:‖x‖
) andnnnorm : α → ℝ≥0
(notation:‖x‖₊
), respectively;Seminormed...Group
: A seminormed (additive) (commutative) group is an (additive) (commutative) group with a norm and a compatible pseudometric space structure:∀ x y, dist x y = ‖x / y‖
or∀ x y, dist x y = ‖x - y‖
, depending on the group operation.Normed...Group
: A normed (additive) (commutative) group is an (additive) (commutative) group with a norm and a compatible metric space structure.
We also prove basic properties of (semi)normed groups and provide some instances.
Notes #
The current convention dist x y = ‖x - y‖
means that the distance is invariant under right
addition, but actions in mathlib are usually from the left. This means we might want to change it to
dist x y = ‖-x + y‖
.
The normed group hierarchy would lend itself well to a mixin design (that is, having
SeminormedGroup
and SeminormedAddGroup
not extend Group
and AddGroup
), but we choose not
to for performance concerns.
Tags #
normed group
the ℝ
-valued norm function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
the ℝ≥0
-valued norm function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
the ℝ≥0∞
-valued norm function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- NNNorm.toENorm = { enorm := fun (x : E) => ↑‖x‖₊ }
A seminormed group is an additive group endowed with a norm for which dist x y = ‖x - y‖
defines a pseudometric space structure.
- add : E → E → E
- zero : E
- neg : E → E
- sub : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A seminormed group is a group endowed with a norm for which dist x y = ‖x / y‖
defines a
pseudometric space structure.
- mul : E → E → E
- one : E
- inv : E → E
- div : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A normed group is an additive group endowed with a norm for which dist x y = ‖x - y‖
defines a
metric space structure.
- add : E → E → E
- zero : E
- neg : E → E
- sub : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A normed group is a group endowed with a norm for which dist x y = ‖x / y‖
defines a metric
space structure.
- mul : E → E → E
- one : E
- inv : E → E
- div : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A seminormed group is an additive group endowed with a norm for which dist x y = ‖x - y‖
defines a pseudometric space structure.
- add : E → E → E
- zero : E
- neg : E → E
- sub : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A seminormed group is a group endowed with a norm for which dist x y = ‖x / y‖
defines a pseudometric space structure.
- mul : E → E → E
- one : E
- inv : E → E
- div : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A normed group is an additive group endowed with a norm for which dist x y = ‖x - y‖
defines a
metric space structure.
- add : E → E → E
- zero : E
- neg : E → E
- sub : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
A normed group is a group endowed with a norm for which dist x y = ‖x / y‖
defines a metric
space structure.
- mul : E → E → E
- one : E
- inv : E → E
- div : E → E → E
- uniformity_dist : uniformity E = ⨅ (ε : ℝ), ⨅ (_ : ε > 0), Filter.principal {p : E × E | dist p.1 p.2 < ε}
- toBornology : Bornology E
The distance function is induced by the norm.
Instances
Equations
Equations
Construct a NormedGroup
from a SeminormedGroup
satisfying ∀ x, ‖x‖ = 0 → x = 1
. This
avoids having to go back to the (Pseudo)MetricSpace
level when declaring a NormedGroup
instance as a special case of a more general SeminormedGroup
instance.
Equations
Instances For
Construct a NormedAddGroup
from a SeminormedAddGroup
satisfying ∀ x, ‖x‖ = 0 → x = 0
. This avoids having to go back to the (Pseudo)MetricSpace
level when declaring a NormedAddGroup
instance as a special case of a more general
SeminormedAddGroup
instance.
Equations
Instances For
Construct a NormedCommGroup
from a SeminormedCommGroup
satisfying
∀ x, ‖x‖ = 0 → x = 1
. This avoids having to go back to the (Pseudo)MetricSpace
level when
declaring a NormedCommGroup
instance as a special case of a more general SeminormedCommGroup
instance.
Equations
Instances For
Construct a NormedAddCommGroup
from a
SeminormedAddCommGroup
satisfying ∀ x, ‖x‖ = 0 → x = 0
. This avoids having to go back to the
(Pseudo)MetricSpace
level when declaring a NormedAddCommGroup
instance as a special case
of a more general SeminormedAddCommGroup
instance.
Equations
Instances For
Construct a seminormed group from a translation-invariant pseudodistance.
Equations
Instances For
Construct a seminormed group from a multiplication-invariant pseudodistance.
Equations
Instances For
Construct a seminormed group from a translation-invariant pseudodistance.
Equations
Instances For
Construct a seminormed group from a multiplication-invariant pseudodistance.
Equations
Instances For
Construct a seminormed group from a translation-invariant pseudodistance.
Equations
Instances For
Construct a normed group from a multiplication-invariant distance.
Equations
- NormedGroup.ofMulDist h₁ h₂ = NormedGroup.mk ⋯
Instances For
Construct a normed group from a multiplication-invariant pseudodistance.
Equations
- NormedGroup.ofMulDist' h₁ h₂ = NormedGroup.mk ⋯
Instances For
Construct a normed group from a translation-invariant pseudodistance.
Equations
Instances For
Construct a normed group from a translation-invariant pseudodistance.
Equations
Instances For
Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the
pseudometric space structure from the seminorm properties. Note that in most cases this instance
creates bad definitional equalities (e.g., it does not take into account a possibly existing
UniformSpace
instance on E
).
Equations
Instances For
Construct a seminormed group from a seminorm, i.e., registering the pseudodistance
and the pseudometric space structure from the seminorm properties. Note that in most cases this
instance creates bad definitional equalities (e.g., it does not take into account a possibly
existing UniformSpace
instance on E
).
Equations
Instances For
Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the
pseudometric space structure from the seminorm properties. Note that in most cases this instance
creates bad definitional equalities (e.g., it does not take into account a possibly existing
UniformSpace
instance on E
).
Equations
Instances For
Construct a seminormed group from a seminorm, i.e., registering the pseudodistance
and the pseudometric space structure from the seminorm properties. Note that in most cases this
instance creates bad definitional equalities (e.g., it does not take into account a possibly
existing UniformSpace
instance on E
).
Equations
Instances For
Construct a normed group from a norm, i.e., registering the distance and the metric space
structure from the norm properties. Note that in most cases this instance creates bad definitional
equalities (e.g., it does not take into account a possibly existing UniformSpace
instance on
E
).
Equations
Instances For
Construct a normed group from a norm, i.e., registering the distance and the metric
space structure from the norm properties. Note that in most cases this instance creates bad
definitional equalities (e.g., it does not take into account a possibly existing UniformSpace
instance on E
).
Equations
Instances For
Construct a normed group from a norm, i.e., registering the distance and the metric space
structure from the norm properties. Note that in most cases this instance creates bad definitional
equalities (e.g., it does not take into account a possibly existing UniformSpace
instance on
E
).
Equations
Instances For
Construct a normed group from a norm, i.e., registering the distance and the metric
space structure from the norm properties. Note that in most cases this instance creates bad
definitional equalities (e.g., it does not take into account a possibly existing UniformSpace
instance on E
).
Equations
Instances For
Alias of dist_eq_norm_sub
.
Alias of dist_eq_norm_sub'
.
Alias of norm_le_norm_add_norm_sub'
.
Alias of norm_le_norm_add_norm_sub
.
An analogue of norm_le_mul_norm_add
for the multiplication from the left.
An analogue of norm_le_add_norm_add
for the addition from the left.
The norm of a seminormed group as a group seminorm.
Equations
- normGroupSeminorm E = { toFun := norm, map_one' := ⋯, mul_le' := ⋯, inv' := ⋯ }
Instances For
The norm of a seminormed group as an additive group seminorm.
Equations
- normAddGroupSeminorm E = { toFun := norm, map_zero' := ⋯, add_le' := ⋯, neg' := ⋯ }
Instances For
Equations
- SeminormedGroup.toNNNorm = { nnnorm := fun (a : E) => ⟨‖a‖, ⋯⟩ }
Equations
- SeminormedAddGroup.toNNNorm = { nnnorm := fun (a : E) => ⟨‖a‖, ⋯⟩ }
Alias of nndist_eq_nnnorm_sub
.
Alias of nnnorm_le_nnnorm_add_nnnorm_sub'
.
Alias of nnnorm_le_nnnorm_add_nnnorm_sub
.
An analogue of nnnorm_le_mul_nnnorm_add
for the multiplication from the left.
An analogue of nnnorm_le_add_nnnorm_add
for the addition from the left.
Alias of ofReal_norm_eq_enorm
.
Alias of ofReal_norm_eq_enorm'
.
Equations
- instENormENNReal = { enorm := fun (x : ENNReal) => x }
Alias of edist_eq_enorm_sub
.
Alias of edist_eq_enorm_div
.
Alias of edist_zero_eq_enorm
.
Alias of edist_one_eq_enorm
.
A group homomorphism from a Group
to a SeminormedGroup
induces a SeminormedGroup
structure on the domain.
Equations
Instances For
A group homomorphism from an AddGroup
to a
SeminormedAddGroup
induces a SeminormedAddGroup
structure on the domain.
Equations
Instances For
A group homomorphism from a CommGroup
to a SeminormedGroup
induces a
SeminormedCommGroup
structure on the domain.
Equations
Instances For
A group homomorphism from an AddCommGroup
to a
SeminormedAddGroup
induces a SeminormedAddCommGroup
structure on the domain.
Equations
Instances For
An injective group homomorphism from a Group
to a NormedGroup
induces a NormedGroup
structure on the domain.
Equations
- NormedGroup.induced E F f h = NormedGroup.mk ⋯
Instances For
An injective group homomorphism from an AddGroup
to a
NormedAddGroup
induces a NormedAddGroup
structure on the domain.
Equations
- NormedAddGroup.induced E F f h = NormedAddGroup.mk ⋯
Instances For
An injective group homomorphism from a CommGroup
to a NormedGroup
induces a
NormedCommGroup
structure on the domain.
Equations
- NormedCommGroup.induced E F f h = NormedCommGroup.mk ⋯
Instances For
An injective group homomorphism from a CommGroup
to a
NormedCommGroup
induces a NormedCommGroup
structure on the domain.
Equations
- NormedAddCommGroup.induced E F f h = NormedAddCommGroup.mk ⋯
Instances For
Alias of Real.enorm_eq_ofReal
.
Alias of Real.enorm_eq_ofReal_abs
.
Alias of Real.ofReal_le_enorm
.
Equations
- NNReal.instNNNorm = { nnnorm := fun (x : NNReal) => x }
Alias of norm_le_zero_iff'
.
Alias of norm_le_zero_iff'
.
Alias of norm_pos_iff'
.
Alias of norm_eq_zero'
.
Alias of norm_eq_zero'
.
Alias of the forward direction of norm_div_eq_zero_iff
.
The norm of a normed group as a group norm.
Equations
- normGroupNorm E = { toGroupSeminorm := normGroupSeminorm E, eq_one_of_map_eq_zero' := ⋯ }
Instances For
The norm of a normed group as an additive group norm.
Equations
- normAddGroupNorm E = { toAddGroupSeminorm := normAddGroupSeminorm E, eq_zero_of_map_eq_zero' := ⋯ }
Instances For
Some relations with HasCompactSupport
Alias of the reverse direction of hasCompactSupport_norm_iff
.
positivity
extensions #
Extension for the positivity
tactic: multiplicative norms are always nonnegative, and positive
on non-one inputs.
Instances For
Extension for the positivity
tactic: additive norms are always nonnegative, and positive
on non-zero inputs.