Documentation

Mathlib.Analysis.Normed.Group.Basic

Normed (semi)groups #

In this file we define 10 classes:

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

class NNNorm (E : Type u_8) :
Type u_8

Auxiliary class, endowing a type α with a function nnnorm : α → ℝ≥0 with notation ‖x‖₊.

  • nnnorm : ENNReal

    the ℝ≥0-valued norm function.

Instances
class ENorm (E : Type u_8) :
Type u_8

Auxiliary class, endowing a type α with a function enorm : α → ℝ≥0∞ with notation ‖x‖ₑ.

  • enorm : EENNReal

    the ℝ≥0∞-valued norm function.

Instances

the -valued norm function.

Equations
  • One or more equations did not get rendered due to their size.

the ℝ≥0-valued norm function.

Equations
  • One or more equations did not get rendered due to their size.

the ℝ≥0∞-valued norm function.

Equations
  • One or more equations did not get rendered due to their size.
instance NNNorm.toENorm {E : Type u_8} [NNNorm E] :
Equations
theorem enorm_eq_nnnorm {E : Type u_8} [NNNorm E] (x : E) :
@[simp]
theorem toNNReal_enorm {E : Type u_8} [NNNorm E] (x : E) :
@[simp]
theorem coe_le_enorm {E : Type u_8} [NNNorm E] {x : E} {r : NNReal} :
@[simp]
theorem enorm_le_coe {E : Type u_8} [NNNorm E] {x : E} {r : NNReal} :
@[simp]
theorem coe_lt_enorm {E : Type u_8} [NNNorm E] {x : E} {r : NNReal} :
@[simp]
theorem enorm_lt_coe {E : Type u_8} [NNNorm E] {x : E} {r : NNReal} :
@[simp]
theorem enorm_ne_top {E : Type u_8} [NNNorm E] {x : E} :
@[simp]
theorem enorm_lt_top {E : Type u_8} [NNNorm E] {x : E} :
class SeminormedAddGroup (E : Type u_8) extends Norm E, AddGroup E, PseudoMetricSpace E :
Type u_8

A seminormed group is an additive group endowed with a norm for which dist x y = ‖x - y‖ defines a pseudometric space structure.

Instances
class SeminormedGroup (E : Type u_8) extends Norm E, Group E, PseudoMetricSpace E :
Type u_8

A seminormed group is a group endowed with a norm for which dist x y = ‖x / y‖ defines a pseudometric space structure.

Instances
class NormedAddGroup (E : Type u_8) extends Norm E, AddGroup E, MetricSpace E :
Type u_8

A normed group is an additive group endowed with a norm for which dist x y = ‖x - y‖ defines a metric space structure.

Instances
class NormedGroup (E : Type u_8) extends Norm E, Group E, MetricSpace E :
Type u_8

A normed group is a group endowed with a norm for which dist x y = ‖x / y‖ defines a metric space structure.

Instances
class SeminormedAddCommGroup (E : Type u_8) extends Norm E, AddCommGroup E, PseudoMetricSpace E :
Type u_8

A seminormed group is an additive group endowed with a norm for which dist x y = ‖x - y‖ defines a pseudometric space structure.

Instances
class SeminormedCommGroup (E : Type u_8) extends Norm E, CommGroup E, PseudoMetricSpace E :
Type u_8

A seminormed group is a group endowed with a norm for which dist x y = ‖x / y‖ defines a pseudometric space structure.

Instances
class NormedAddCommGroup (E : Type u_8) extends Norm E, AddCommGroup E, MetricSpace E :
Type u_8

A normed group is an additive group endowed with a norm for which dist x y = ‖x - y‖ defines a metric space structure.

Instances
class NormedCommGroup (E : Type u_8) extends Norm E, CommGroup E, MetricSpace E :
Type u_8

A normed group is a group endowed with a norm for which dist x y = ‖x / y‖ defines a metric space structure.

Instances
@[reducible, inline]
abbrev NormedGroup.ofSeparation {E : Type u_5} [SeminormedGroup E] (h : ∀ (x : E), x = 0x = 1) :

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
@[reducible, inline]
abbrev NormedAddGroup.ofSeparation {E : Type u_5} [SeminormedAddGroup E] (h : ∀ (x : E), x = 0x = 0) :

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
@[reducible, inline]
abbrev NormedCommGroup.ofSeparation {E : Type u_5} [SeminormedCommGroup E] (h : ∀ (x : E), x = 0x = 1) :

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
@[reducible, inline]
abbrev NormedAddCommGroup.ofSeparation {E : Type u_5} [SeminormedAddCommGroup E] (h : ∀ (x : E), x = 0x = 0) :

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
@[reducible, inline]
abbrev SeminormedGroup.ofMulDist {E : Type u_5} [Norm E] [Group E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist x y dist (x * z) (y * z)) :

Construct a seminormed group from a multiplication-invariant distance.

Equations
@[reducible, inline]
abbrev SeminormedAddGroup.ofAddDist {E : Type u_5} [Norm E] [AddGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist x y dist (x + z) (y + z)) :

Construct a seminormed group from a translation-invariant distance.

Equations
@[reducible, inline]
abbrev SeminormedGroup.ofMulDist' {E : Type u_5} [Norm E] [Group E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist (x * z) (y * z) dist x y) :

Construct a seminormed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev SeminormedAddGroup.ofAddDist' {E : Type u_5} [Norm E] [AddGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist (x + z) (y + z) dist x y) :

Construct a seminormed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev SeminormedCommGroup.ofMulDist {E : Type u_5} [Norm E] [CommGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist x y dist (x * z) (y * z)) :

Construct a seminormed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev SeminormedAddCommGroup.ofAddDist {E : Type u_5} [Norm E] [AddCommGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist x y dist (x + z) (y + z)) :

Construct a seminormed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev SeminormedCommGroup.ofMulDist' {E : Type u_5} [Norm E] [CommGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist (x * z) (y * z) dist x y) :

Construct a seminormed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev SeminormedAddCommGroup.ofAddDist' {E : Type u_5} [Norm E] [AddCommGroup E] [PseudoMetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist (x + z) (y + z) dist x y) :

Construct a seminormed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedGroup.ofMulDist {E : Type u_5} [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist x y dist (x * z) (y * z)) :

Construct a normed group from a multiplication-invariant distance.

Equations
@[reducible, inline]
abbrev NormedAddGroup.ofAddDist {E : Type u_5} [Norm E] [AddGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist x y dist (x + z) (y + z)) :

Construct a normed group from a translation-invariant distance.

Equations
@[reducible, inline]
abbrev NormedGroup.ofMulDist' {E : Type u_5} [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist (x * z) (y * z) dist x y) :

Construct a normed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedAddGroup.ofAddDist' {E : Type u_5} [Norm E] [AddGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist (x + z) (y + z) dist x y) :

Construct a normed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedCommGroup.ofMulDist {E : Type u_5} [Norm E] [CommGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist x y dist (x * z) (y * z)) :

Construct a normed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedAddCommGroup.ofAddDist {E : Type u_5} [Norm E] [AddCommGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist x y dist (x + z) (y + z)) :

Construct a normed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedCommGroup.ofMulDist' {E : Type u_5} [Norm E] [CommGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 1) (h₂ : ∀ (x y z : E), dist (x * z) (y * z) dist x y) :

Construct a normed group from a multiplication-invariant pseudodistance.

Equations
@[reducible, inline]
abbrev NormedAddCommGroup.ofAddDist' {E : Type u_5} [Norm E] [AddCommGroup E] [MetricSpace E] (h₁ : ∀ (x : E), x = dist x 0) (h₂ : ∀ (x y z : E), dist (x + z) (y + z) dist x y) :

Construct a normed group from a translation-invariant pseudodistance.

Equations
@[reducible, inline]

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
@[reducible, inline]

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
@[reducible, inline]

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
@[reducible, inline]

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
@[reducible, inline]
abbrev GroupNorm.toNormedGroup {E : Type u_5} [Group E] (f : GroupNorm E) :

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
@[reducible, inline]

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
@[reducible, inline]

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
@[reducible, inline]

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
theorem dist_eq_norm_div {E : Type u_5} [SeminormedGroup E] (a b : E) :
dist a b = a / b
theorem dist_eq_norm_sub {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
dist a b = a - b
theorem dist_eq_norm_div' {E : Type u_5} [SeminormedGroup E] (a b : E) :
dist a b = b / a
theorem dist_eq_norm_sub' {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
dist a b = b - a
theorem dist_eq_norm {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
dist a b = a - b

Alias of dist_eq_norm_sub.

theorem dist_eq_norm' {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
dist a b = b - a

Alias of dist_eq_norm_sub'.

theorem DiscreteTopology.of_forall_le_norm' {E : Type u_5} [SeminormedGroup E] {r : } (hpos : 0 < r) (hr : ∀ (x : E), x 1r x) :
theorem DiscreteTopology.of_forall_le_norm {E : Type u_5} [SeminormedAddGroup E] {r : } (hpos : 0 < r) (hr : ∀ (x : E), x 0r x) :
@[simp]
theorem dist_one_right {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem dist_zero_right {E : Type u_5} [SeminormedAddGroup E] (a : E) :
theorem dist_one_left {E : Type u_5} [SeminormedGroup E] (a : E) :
theorem dist_zero_left {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem dist_one {E : Type u_5} [SeminormedGroup E] :
@[simp]
theorem dist_zero {E : Type u_5} [SeminormedAddGroup E] :
theorem norm_div_rev {E : Type u_5} [SeminormedGroup E] (a b : E) :
a / b = b / a
theorem norm_sub_rev {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
a - b = b - a
@[simp]
theorem norm_inv' {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem norm_neg {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem norm_zpow_abs {E : Type u_5} [SeminormedGroup E] (a : E) (n : ) :
@[simp]
theorem norm_abs_zsmul {E : Type u_5} [SeminormedAddGroup E] (a : E) (n : ) :
@[simp]
theorem norm_pow_natAbs {E : Type u_5} [SeminormedGroup E] (a : E) (n : ) :
@[simp]
theorem norm_natAbs_smul {E : Type u_5} [SeminormedAddGroup E] (a : E) (n : ) :
theorem norm_zpow_isUnit {E : Type u_5} [SeminormedGroup E] (a : E) {n : } (hn : IsUnit n) :
theorem norm_isUnit_zsmul {E : Type u_5} [SeminormedAddGroup E] (a : E) {n : } (hn : IsUnit n) :
@[simp]
theorem norm_units_zsmul {E : Type u_8} [SeminormedAddGroup E] (n : ˣ) (a : E) :
theorem dist_mulIndicator {α : Type u_2} {E : Type u_5} [SeminormedGroup E] (s t : Set α) (f : αE) (x : α) :
theorem dist_indicator {α : Type u_2} {E : Type u_5} [SeminormedAddGroup E] (s t : Set α) (f : αE) (x : α) :
dist (s.indicator f x) (t.indicator f x) = (symmDiff s t).indicator f x
theorem norm_mul_le' {E : Type u_5} [SeminormedGroup E] (a b : E) :

Triangle inequality for the norm.

theorem norm_add_le {E : Type u_5} [SeminormedAddGroup E] (a b : E) :

Triangle inequality for the norm.

theorem norm_mul_le_of_le' {E : Type u_5} [SeminormedGroup E] {a₁ a₂ : E} {r₁ r₂ : } (h₁ : a₁ r₁) (h₂ : a₂ r₂) :
a₁ * a₂ r₁ + r₂

Triangle inequality for the norm.

theorem norm_add_le_of_le {E : Type u_5} [SeminormedAddGroup E] {a₁ a₂ : E} {r₁ r₂ : } (h₁ : a₁ r₁) (h₂ : a₂ r₂) :
a₁ + a₂ r₁ + r₂

Triangle inequality for the norm.

theorem norm_mul₃_le' {E : Type u_5} [SeminormedGroup E] {a b c : E} :

Triangle inequality for the norm.

theorem norm_add₃_le {E : Type u_5} [SeminormedAddGroup E] {a b c : E} :

Triangle inequality for the norm.

@[simp]
theorem norm_nonneg' {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem norm_nonneg {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem abs_norm' {E : Type u_5} [SeminormedGroup E] (z : E) :
@[simp]
theorem abs_norm {E : Type u_5} [SeminormedAddGroup E] (z : E) :
@[simp]
theorem norm_one' {E : Type u_5} [SeminormedGroup E] :
@[simp]
theorem norm_zero {E : Type u_5} [SeminormedAddGroup E] :
theorem ne_one_of_norm_ne_zero {E : Type u_5} [SeminormedGroup E] {a : E} :
a 0a 1
theorem ne_zero_of_norm_ne_zero {E : Type u_5} [SeminormedAddGroup E] {a : E} :
a 0a 0
theorem zero_lt_one_add_norm_sq' {E : Type u_5} [SeminormedGroup E] (x : E) :
0 < 1 + x ^ 2
theorem zero_lt_one_add_norm_sq {E : Type u_5} [SeminormedAddGroup E] (x : E) :
0 < 1 + x ^ 2
theorem norm_div_le {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem norm_sub_le {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
theorem norm_div_le_of_le {E : Type u_5} [SeminormedGroup E] {a₁ a₂ : E} {r₁ r₂ : } (H₁ : a₁ r₁) (H₂ : a₂ r₂) :
a₁ / a₂ r₁ + r₂
theorem norm_sub_le_of_le {E : Type u_5} [SeminormedAddGroup E] {a₁ a₂ : E} {r₁ r₂ : } (H₁ : a₁ r₁) (H₂ : a₂ r₂) :
a₁ - a₂ r₁ + r₂
theorem dist_le_norm_add_norm' {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem dist_le_norm_add_norm {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
theorem norm_sub_norm_le' {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem norm_sub_norm_le {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
theorem norm_sub_le_norm_mul {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem dist_norm_norm_le' {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem norm_le_mul_norm_add {E : Type u_5} [SeminormedGroup E] (u v : E) :
theorem norm_le_mul_norm_add' {E : Type u_5} [SeminormedGroup E] (u v : E) :

An analogue of norm_le_mul_norm_add for the multiplication from the left.

theorem norm_le_add_norm_add' {E : Type u_5} [SeminormedAddGroup E] (u v : E) :

An analogue of norm_le_add_norm_add for the addition from the left.

theorem norm_mul_eq_norm_right {E : Type u_5} [SeminormedGroup E] {x : E} (y : E) (h : x = 0) :
theorem norm_add_eq_norm_right {E : Type u_5} [SeminormedAddGroup E] {x : E} (y : E) (h : x = 0) :
theorem norm_mul_eq_norm_left {E : Type u_5} [SeminormedGroup E] (x : E) {y : E} (h : y = 0) :
theorem norm_add_eq_norm_left {E : Type u_5} [SeminormedAddGroup E] (x : E) {y : E} (h : y = 0) :
theorem norm_div_eq_norm_right {E : Type u_5} [SeminormedGroup E] {x : E} (y : E) (h : x = 0) :
theorem norm_sub_eq_norm_right {E : Type u_5} [SeminormedAddGroup E] {x : E} (y : E) (h : x = 0) :
theorem norm_div_eq_norm_left {E : Type u_5} [SeminormedGroup E] (x : E) {y : E} (h : y = 0) :
theorem norm_sub_eq_norm_left {E : Type u_5} [SeminormedAddGroup E] (x : E) {y : E} (h : y = 0) :
theorem ball_eq' {E : Type u_5} [SeminormedGroup E] (y : E) (ε : ) :
Metric.ball y ε = {x : E | x / y < ε}
theorem ball_eq {E : Type u_5} [SeminormedAddGroup E] (y : E) (ε : ) :
Metric.ball y ε = {x : E | x - y < ε}
theorem ball_one_eq {E : Type u_5} [SeminormedGroup E] (r : ) :
Metric.ball 1 r = {x : E | x < r}
theorem ball_zero_eq {E : Type u_5} [SeminormedAddGroup E] (r : ) :
Metric.ball 0 r = {x : E | x < r}
theorem mem_ball_iff_norm'' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
theorem mem_ball_iff_norm {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
theorem mem_ball_iff_norm''' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
theorem mem_ball_iff_norm' {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
theorem mem_ball_one_iff {E : Type u_5} [SeminormedGroup E] {a : E} {r : } :
theorem mem_ball_zero_iff {E : Type u_5} [SeminormedAddGroup E] {a : E} {r : } :
theorem mem_closedBall_iff_norm'' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
theorem mem_closedBall_iff_norm {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
theorem mem_closedBall_iff_norm''' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
theorem mem_closedBall_iff_norm' {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
theorem norm_le_of_mem_closedBall' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } (h : b Metric.closedBall a r) :
theorem norm_le_of_mem_closedBall {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } (h : b Metric.closedBall a r) :
theorem norm_le_norm_add_const_of_dist_le' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
dist a b ra b + r
theorem norm_le_norm_add_const_of_dist_le {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
dist a b ra b + r
theorem norm_lt_of_mem_ball' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } (h : b Metric.ball a r) :
theorem norm_lt_of_mem_ball {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } (h : b Metric.ball a r) :
@[simp]
theorem mem_sphere_iff_norm' {E : Type u_5} [SeminormedGroup E] {a b : E} {r : } :
@[simp]
theorem mem_sphere_iff_norm {E : Type u_5} [SeminormedAddGroup E] {a b : E} {r : } :
theorem mem_sphere_one_iff_norm {E : Type u_5} [SeminormedGroup E] {a : E} {r : } :
theorem mem_sphere_zero_iff_norm {E : Type u_5} [SeminormedAddGroup E] {a : E} {r : } :
@[simp]
theorem norm_eq_of_mem_sphere' {E : Type u_5} [SeminormedGroup E] {r : } (x : (Metric.sphere 1 r)) :
x = r
@[simp]
theorem norm_eq_of_mem_sphere {E : Type u_5} [SeminormedAddGroup E] {r : } (x : (Metric.sphere 0 r)) :
x = r
theorem ne_one_of_mem_sphere {E : Type u_5} [SeminormedGroup E] {r : } (hr : r 0) (x : (Metric.sphere 1 r)) :
x 1
theorem ne_zero_of_mem_sphere {E : Type u_5} [SeminormedAddGroup E] {r : } (hr : r 0) (x : (Metric.sphere 0 r)) :
x 0
theorem ne_one_of_mem_unit_sphere {E : Type u_5} [SeminormedGroup E] (x : (Metric.sphere 1 1)) :
x 1
theorem ne_zero_of_mem_unit_sphere {E : Type u_5} [SeminormedAddGroup E] (x : (Metric.sphere 0 1)) :
x 0

The norm of a seminormed group as a group seminorm.

Equations

The norm of a seminormed group as an additive group seminorm.

Equations
theorem NormedCommGroup.tendsto_nhds_one {α : Type u_2} {E : Type u_5} [SeminormedGroup E] {f : αE} {l : Filter α} :
Filter.Tendsto f l (nhds 1) ε > 0, ∀ᶠ (x : α) in l, f x < ε
theorem NormedAddCommGroup.tendsto_nhds_zero {α : Type u_2} {E : Type u_5} [SeminormedAddGroup E] {f : αE} {l : Filter α} :
Filter.Tendsto f l (nhds 0) ε > 0, ∀ᶠ (x : α) in l, f x < ε
theorem NormedCommGroup.tendsto_nhds_nhds {E : Type u_5} {F : Type u_6} [SeminormedGroup E] [SeminormedGroup F] {f : EF} {x : E} {y : F} :
Filter.Tendsto f (nhds x) (nhds y) ε > 0, δ > 0, ∀ (x' : E), x' / x < δf x' / y < ε
theorem NormedAddCommGroup.tendsto_nhds_nhds {E : Type u_5} {F : Type u_6} [SeminormedAddGroup E] [SeminormedAddGroup F] {f : EF} {x : E} {y : F} :
Filter.Tendsto f (nhds x) (nhds y) ε > 0, δ > 0, ∀ (x' : E), x' - x < δf x' - y < ε
theorem NormedCommGroup.nhds_basis_norm_lt {E : Type u_5} [SeminormedGroup E] (x : E) :
(nhds x).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {y : E | y / x < ε}
theorem NormedAddCommGroup.nhds_basis_norm_lt {E : Type u_5} [SeminormedAddGroup E] (x : E) :
(nhds x).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {y : E | y - x < ε}
theorem NormedCommGroup.nhds_one_basis_norm_lt {E : Type u_5} [SeminormedGroup E] :
(nhds 1).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {y : E | y < ε}
theorem NormedAddCommGroup.nhds_zero_basis_norm_lt {E : Type u_5} [SeminormedAddGroup E] :
(nhds 0).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {y : E | y < ε}
theorem NormedCommGroup.uniformity_basis_dist {E : Type u_5} [SeminormedGroup E] :
(uniformity E).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {p : E × E | p.1 / p.2 < ε}
theorem NormedAddCommGroup.uniformity_basis_dist {E : Type u_5} [SeminormedAddGroup E] :
(uniformity E).HasBasis (fun (ε : ) => 0 < ε) fun (ε : ) => {p : E × E | p.1 - p.2 < ε}
@[instance 100]
Equations
@[instance 100]
Equations
@[simp]
theorem coe_nnnorm' {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem coe_nnnorm {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem norm_toNNReal' {E : Type u_5} [SeminormedGroup E] {a : E} :
@[simp]
@[simp]
theorem toReal_enorm' {E : Type u_5} [SeminormedGroup E] (x : E) :
@[simp]
theorem toReal_enorm {E : Type u_5} [SeminormedAddGroup E] (x : E) :
@[simp]
@[simp]
theorem nndist_eq_nnnorm_div {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem nndist_eq_nnnorm_sub {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
theorem nndist_eq_nnnorm {E : Type u_5} [SeminormedAddGroup E] (a b : E) :

Alias of nndist_eq_nnnorm_sub.

@[simp]
theorem nndist_one_right {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem nndist_zero_right {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem edist_one_right {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem edist_zero_right {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem nnnorm_one' {E : Type u_5} [SeminormedGroup E] :
@[simp]
theorem nnnorm_zero {E : Type u_5} [SeminormedAddGroup E] :
theorem ne_one_of_nnnorm_ne_zero {E : Type u_5} [SeminormedGroup E] {a : E} :
a‖₊ 0a 1
theorem ne_zero_of_nnnorm_ne_zero {E : Type u_5} [SeminormedAddGroup E] {a : E} :
a‖₊ 0a 0
theorem norm_pow_le_mul_norm {E : Type u_5} [SeminormedGroup E] {a : E} {n : } :
a ^ n n * a
theorem norm_nsmul_le {E : Type u_5} [SeminormedAddGroup E] {a : E} {n : } :
n a n * a
theorem nnnorm_pow_le_mul_norm {E : Type u_5} [SeminormedGroup E] {a : E} {n : } :
theorem nnnorm_nsmul_le {E : Type u_5} [SeminormedAddGroup E] {a : E} {n : } :
@[simp]
theorem nnnorm_inv' {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem nnnorm_neg {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem enorm_inv' {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem enorm_neg {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem nnnorm_zpow_abs {E : Type u_5} [SeminormedGroup E] (a : E) (n : ) :
@[simp]
theorem nnnorm_abs_zsmul {E : Type u_5} [SeminormedAddGroup E] (a : E) (n : ) :
@[simp]
theorem nnnorm_pow_natAbs {E : Type u_5} [SeminormedGroup E] (a : E) (n : ) :
@[simp]
theorem nnnorm_natAbs_smul {E : Type u_5} [SeminormedAddGroup E] (a : E) (n : ) :
theorem nnnorm_zpow_isUnit {E : Type u_5} [SeminormedGroup E] (a : E) {n : } (hn : IsUnit n) :
theorem nnnorm_isUnit_zsmul {E : Type u_5} [SeminormedAddGroup E] (a : E) {n : } (hn : IsUnit n) :
@[simp]
theorem nnnorm_units_zsmul {E : Type u_8} [SeminormedAddGroup E] (n : ˣ) (a : E) :
@[simp]
theorem nndist_one_left {E : Type u_5} [SeminormedGroup E] (a : E) :
@[simp]
theorem nndist_zero_left {E : Type u_5} [SeminormedAddGroup E] (a : E) :
@[simp]
theorem edist_one_left {E : Type u_5} [SeminormedGroup E] (a : E) :
edist 1 a = a‖₊
@[simp]
theorem edist_zero_left {E : Type u_5} [SeminormedAddGroup E] (a : E) :
edist 0 a = a‖₊
theorem nndist_mulIndicator {α : Type u_2} {E : Type u_5} [SeminormedGroup E] (s t : Set α) (f : αE) (x : α) :
theorem nndist_indicator {α : Type u_2} {E : Type u_5} [SeminormedAddGroup E] (s t : Set α) (f : αE) (x : α) :
theorem enorm_div_le {E : Type u_5} [SeminormedGroup E] {a b : E} :

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.

theorem nnnorm_mul_eq_nnnorm_right {E : Type u_5} [SeminormedGroup E] {x : E} (y : E) (h : x‖₊ = 0) :
theorem nnnorm_add_eq_nnnorm_right {E : Type u_5} [SeminormedAddGroup E] {x : E} (y : E) (h : x‖₊ = 0) :
theorem nnnorm_mul_eq_nnnorm_left {E : Type u_5} [SeminormedGroup E] (x : E) {y : E} (h : y‖₊ = 0) :
theorem nnnorm_add_eq_nnnorm_left {E : Type u_5} [SeminormedAddGroup E] (x : E) {y : E} (h : y‖₊ = 0) :
theorem nnnorm_div_eq_nnnorm_right {E : Type u_5} [SeminormedGroup E] {x : E} (y : E) (h : x‖₊ = 0) :
theorem nnnorm_sub_eq_nnnorm_right {E : Type u_5} [SeminormedAddGroup E] {x : E} (y : E) (h : x‖₊ = 0) :
theorem nnnorm_div_eq_nnnorm_left {E : Type u_5} [SeminormedGroup E] (x : E) {y : E} (h : y‖₊ = 0) :
theorem nnnorm_sub_eq_nnnorm_left {E : Type u_5} [SeminormedAddGroup E] (x : E) {y : E} (h : y‖₊ = 0) :
theorem toReal_coe_nnnorm' {E : Type u_5} [SeminormedGroup E] (a : E) :

The non negative norm seen as an ENNReal and then as a Real is equal to the norm.

theorem toReal_coe_nnnorm {E : Type u_5} [SeminormedAddGroup E] (a : E) :

The non negative norm seen as an ENNReal and then as a Real is equal to the norm.

theorem edist_mulIndicator {α : Type u_2} {E : Type u_5} [SeminormedGroup E] (s t : Set α) (f : αE) (x : α) :
theorem edist_indicator {α : Type u_2} {E : Type u_5} [SeminormedAddGroup E] (s t : Set α) (f : αE) (x : α) :
edist (s.indicator f x) (t.indicator f x) = (symmDiff s t).indicator f x‖₊
@[simp]
theorem enorm_one' {E : Type u_5} [SeminormedGroup E] :
@[simp]
theorem enorm_zero {E : Type u_5} [SeminormedAddGroup E] :
@[deprecated ofReal_norm_eq_enorm (since := "2025-01-17")]

Alias of ofReal_norm_eq_enorm.

@[deprecated ofReal_norm_eq_enorm' (since := "2025-01-17")]

Alias of ofReal_norm_eq_enorm'.

Equations
@[simp]
theorem enorm_eq_self (x : ENNReal) :
theorem edist_eq_enorm_div {E : Type u_5} [SeminormedGroup E] (a b : E) :
theorem edist_eq_enorm_sub {E : Type u_5} [SeminormedAddGroup E] (a b : E) :
@[deprecated edist_eq_enorm_sub (since := "2025-01-17")]
theorem edist_eq_coe_nnnorm_sub {E : Type u_5} [SeminormedAddGroup E] (a b : E) :

Alias of edist_eq_enorm_sub.

@[deprecated edist_eq_enorm_div (since := "2025-01-17")]
theorem edist_eq_coe_nnnorm_div {E : Type u_5} [SeminormedGroup E] (a b : E) :

Alias of edist_eq_enorm_div.

theorem edist_one_eq_enorm {E : Type u_5} [SeminormedGroup E] (x : E) :
@[deprecated edist_zero_eq_enorm (since := "2025-01-17")]
theorem edist_eq_coe_nnnorm {E : Type u_5} [SeminormedAddGroup E] (x : E) :

Alias of edist_zero_eq_enorm.

@[deprecated edist_one_eq_enorm (since := "2025-01-17")]
theorem edist_eq_coe_nnnorm' {E : Type u_5} [SeminormedGroup E] (x : E) :

Alias of edist_one_eq_enorm.

theorem SeminormedGroup.disjoint_nhds {E : Type u_5} [SeminormedGroup E] (x : E) (f : Filter E) :
Disjoint (nhds x) f δ > 0, ∀ᶠ (y : E) in f, δ y / x
theorem SeminormedAddGroup.disjoint_nhds {E : Type u_5} [SeminormedAddGroup E] (x : E) (f : Filter E) :
Disjoint (nhds x) f δ > 0, ∀ᶠ (y : E) in f, δ y - x
theorem SeminormedGroup.disjoint_nhds_one {E : Type u_5} [SeminormedGroup E] (f : Filter E) :
Disjoint (nhds 1) f δ > 0, ∀ᶠ (y : E) in f, δ y
theorem SeminormedAddGroup.disjoint_nhds_zero {E : Type u_5} [SeminormedAddGroup E] (f : Filter E) :
Disjoint (nhds 0) f δ > 0, ∀ᶠ (y : E) in f, δ y
@[reducible, inline]
abbrev SeminormedGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [Group E] [SeminormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) :

A group homomorphism from a Group to a SeminormedGroup induces a SeminormedGroup structure on the domain.

Equations
@[reducible, inline]
abbrev SeminormedAddGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [AddGroup E] [SeminormedAddGroup F] [AddMonoidHomClass 𝓕 E F] (f : 𝓕) :

A group homomorphism from an AddGroup to a SeminormedAddGroup induces a SeminormedAddGroup structure on the domain.

Equations
@[reducible, inline]
abbrev SeminormedCommGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [CommGroup E] [SeminormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) :

A group homomorphism from a CommGroup to a SeminormedGroup induces a SeminormedCommGroup structure on the domain.

Equations
@[reducible, inline]
abbrev SeminormedAddCommGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [AddCommGroup E] [SeminormedAddGroup F] [AddMonoidHomClass 𝓕 E F] (f : 𝓕) :

A group homomorphism from an AddCommGroup to a SeminormedAddGroup induces a SeminormedAddCommGroup structure on the domain.

Equations
@[reducible, inline]
abbrev NormedGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [Group E] [NormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) (h : Function.Injective f) :

An injective group homomorphism from a Group to a NormedGroup induces a NormedGroup structure on the domain.

Equations
@[reducible, inline]
abbrev NormedAddGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [AddGroup E] [NormedAddGroup F] [AddMonoidHomClass 𝓕 E F] (f : 𝓕) (h : Function.Injective f) :

An injective group homomorphism from an AddGroup to a NormedAddGroup induces a NormedAddGroup structure on the domain.

Equations
@[reducible, inline]
abbrev NormedCommGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [CommGroup E] [NormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) (h : Function.Injective f) :

An injective group homomorphism from a CommGroup to a NormedGroup induces a NormedCommGroup structure on the domain.

Equations
@[reducible, inline]
abbrev NormedAddCommGroup.induced {𝓕 : Type u_1} (E : Type u_5) (F : Type u_6) [FunLike 𝓕 E F] [AddCommGroup E] [NormedAddGroup F] [AddMonoidHomClass 𝓕 E F] (f : 𝓕) (h : Function.Injective f) :

An injective group homomorphism from a CommGroup to a NormedCommGroup induces a NormedCommGroup structure on the domain.

Equations
instance Real.norm :
Equations
@[simp]
theorem Real.norm_eq_abs (r : ) :
theorem Real.norm_of_nonneg {r : } (hr : 0 r) :
theorem Real.norm_of_nonpos {r : } (hr : r 0) :
@[simp]
theorem Real.norm_natCast (n : ) :
n = n
@[simp]
theorem Real.nnnorm_natCast (n : ) :
n‖₊ = n
@[simp]
theorem Real.enorm_natCast (n : ) :
n‖ₑ = n
@[simp]
theorem Real.norm_nnratCast (q : ℚ≥0) :
q = q
@[simp]
theorem Real.nnnorm_nnratCast (q : ℚ≥0) :
q‖₊ = q
theorem Real.nnnorm_of_nonneg {r : } (hr : 0 r) :
r‖₊ = r, hr
@[deprecated Real.enorm_eq_ofReal (since := "2025-01-17")]

Alias of Real.enorm_eq_ofReal.

@[deprecated Real.enorm_eq_ofReal_abs (since := "2025-01-17")]

Alias of Real.enorm_eq_ofReal_abs.

@[deprecated Real.ofReal_le_enorm (since := "2025-01-17")]

Alias of Real.ofReal_le_enorm.

Equations
@[simp]
theorem dist_inv {E : Type u_5} [SeminormedCommGroup E] (x y : E) :
theorem dist_neg {E : Type u_5} [SeminormedAddCommGroup E] (x y : E) :
dist (-x) y = dist x (-y)
theorem norm_multiset_sum_le {E : Type u_8} [SeminormedAddCommGroup E] (m : Multiset E) :
m.sum (Multiset.map (fun (x : E) => x) m).sum
theorem norm_multiset_prod_le {E : Type u_5} [SeminormedCommGroup E] (m : Multiset E) :
m.prod (Multiset.map (fun (x : E) => x) m).sum
theorem norm_sum_le {ι : Type u_8} {E : Type u_9} [SeminormedAddCommGroup E] (s : Finset ι) (f : ιE) :
is, f i is, f i
theorem norm_prod_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) (f : ιE) :
is, f i is, f i
theorem norm_prod_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) {f : ιE} {n : ι} (h : bs, f b n b) :
bs, f b bs, n b
theorem norm_sum_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedAddCommGroup E] (s : Finset ι) {f : ιE} {n : ι} (h : bs, f b n b) :
bs, f b bs, n b
theorem dist_prod_prod_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) {f a : ιE} {d : ι} (h : bs, dist (f b) (a b) d b) :
dist (∏ bs, f b) (∏ bs, a b) bs, d b
theorem dist_sum_sum_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedAddCommGroup E] (s : Finset ι) {f a : ιE} {d : ι} (h : bs, dist (f b) (a b) d b) :
dist (∑ bs, f b) (∑ bs, a b) bs, d b
theorem dist_prod_prod_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) (f a : ιE) :
dist (∏ bs, f b) (∏ bs, a b) bs, dist (f b) (a b)
theorem dist_sum_sum_le {ι : Type u_3} {E : Type u_5} [SeminormedAddCommGroup E] (s : Finset ι) (f a : ιE) :
dist (∑ bs, f b) (∑ bs, a b) bs, dist (f b) (a b)
theorem mul_mem_ball_iff_norm {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } :
theorem add_mem_ball_iff_norm {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } :
@[simp]
theorem preimage_mul_ball {E : Type u_5} [SeminormedCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b * x) ⁻¹' Metric.ball a r = Metric.ball (a / b) r
@[simp]
theorem preimage_add_ball {E : Type u_5} [SeminormedAddCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b + x) ⁻¹' Metric.ball a r = Metric.ball (a - b) r
@[simp]
theorem preimage_mul_closedBall {E : Type u_5} [SeminormedCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b * x) ⁻¹' Metric.closedBall a r = Metric.closedBall (a / b) r
@[simp]
theorem preimage_add_closedBall {E : Type u_5} [SeminormedAddCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b + x) ⁻¹' Metric.closedBall a r = Metric.closedBall (a - b) r
@[simp]
theorem preimage_mul_sphere {E : Type u_5} [SeminormedCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b * x) ⁻¹' Metric.sphere a r = Metric.sphere (a / b) r
@[simp]
theorem preimage_add_sphere {E : Type u_5} [SeminormedAddCommGroup E] (a b : E) (r : ) :
(fun (x : E) => b + x) ⁻¹' Metric.sphere a r = Metric.sphere (a - b) r
theorem pow_mem_closedBall {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } {n : } (h : a Metric.closedBall b r) :
a ^ n Metric.closedBall (b ^ n) (n r)
theorem nsmul_mem_closedBall {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } {n : } (h : a Metric.closedBall b r) :
n a Metric.closedBall (n b) (n r)
theorem pow_mem_ball {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } {n : } (hn : 0 < n) (h : a Metric.ball b r) :
a ^ n Metric.ball (b ^ n) (n r)
theorem nsmul_mem_ball {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } {n : } (hn : 0 < n) (h : a Metric.ball b r) :
n a Metric.ball (n b) (n r)
theorem mul_mem_closedBall_mul_iff {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } {c : E} :
theorem add_mem_closedBall_add_iff {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } {c : E} :
theorem mul_mem_ball_mul_iff {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } {c : E} :
a * c Metric.ball (b * c) r a Metric.ball b r
theorem add_mem_ball_add_iff {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } {c : E} :
a + c Metric.ball (b + c) r a Metric.ball b r
theorem smul_closedBall'' {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } :
theorem smul_ball'' {E : Type u_5} [SeminormedCommGroup E] {a b : E} {r : } :
theorem vadd_ball'' {E : Type u_5} [SeminormedAddCommGroup E] {a b : E} {r : } :
theorem nnnorm_prod_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) (f : ιE) :
as, f a‖₊ as, f a‖₊
theorem nnnorm_sum_le {ι : Type u_3} {E : Type u_5} [SeminormedAddCommGroup E] (s : Finset ι) (f : ιE) :
as, f a‖₊ as, f a‖₊
theorem nnnorm_prod_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedCommGroup E] (s : Finset ι) {f : ιE} {n : ιNNReal} (h : bs, f b‖₊ n b) :
bs, f b‖₊ bs, n b
theorem nnnorm_sum_le_of_le {ι : Type u_3} {E : Type u_5} [SeminormedAddCommGroup E] (s : Finset ι) {f : ιE} {n : ιNNReal} (h : bs, f b‖₊ n b) :
bs, f b‖₊ bs, n b
@[simp]
theorem norm_norm' {E : Type u_5} [SeminormedCommGroup E] (x : E) :
@[simp]
theorem norm_norm {E : Type u_5} [SeminormedAddCommGroup E] (x : E) :
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem norm_le_zero_iff' {E : Type u_5} [NormedGroup E] {a : E} :
a 0 a = 1
@[simp]
theorem norm_le_zero_iff {E : Type u_5} [NormedAddGroup E] {a : E} :
a 0 a = 0
@[simp]
theorem norm_pos_iff' {E : Type u_5} [NormedGroup E] {a : E} :
0 < a a 1
@[simp]
theorem norm_pos_iff {E : Type u_5} [NormedAddGroup E] {a : E} :
0 < a a 0
@[simp]
theorem norm_eq_zero' {E : Type u_5} [NormedGroup E] {a : E} :
a = 0 a = 1
@[simp]
theorem norm_eq_zero {E : Type u_5} [NormedAddGroup E] {a : E} :
a = 0 a = 0
theorem norm_ne_zero_iff' {E : Type u_5} [NormedGroup E] {a : E} :
a 0 a 1
theorem norm_ne_zero_iff {E : Type u_5} [NormedAddGroup E] {a : E} :
a 0 a 0
@[deprecated norm_le_zero_iff' (since := "2024-11-24")]
theorem norm_le_zero_iff'' {E : Type u_5} [NormedGroup E] {a : E} :
a 0 a = 1

Alias of norm_le_zero_iff'.

@[deprecated norm_le_zero_iff' (since := "2024-11-24")]
theorem norm_le_zero_iff''' {E : Type u_5} [NormedGroup E] {a : E} :
a 0 a = 1

Alias of norm_le_zero_iff'.

@[deprecated norm_pos_iff' (since := "2024-11-24")]
theorem norm_pos_iff'' {E : Type u_5} [NormedGroup E] {a : E} :
0 < a a 1

Alias of norm_pos_iff'.

@[deprecated norm_eq_zero' (since := "2024-11-24")]
theorem norm_eq_zero'' {E : Type u_5} [NormedGroup E] {a : E} :
a = 0 a = 1

Alias of norm_eq_zero'.

@[deprecated norm_eq_zero' (since := "2024-11-24")]
theorem norm_eq_zero''' {E : Type u_5} [NormedGroup E] {a : E} :
a = 0 a = 1

Alias of norm_eq_zero'.

theorem norm_div_eq_zero_iff {E : Type u_5} [NormedGroup E] {a b : E} :
a / b = 0 a = b
theorem norm_sub_eq_zero_iff {E : Type u_5} [NormedAddGroup E] {a b : E} :
a - b = 0 a = b
theorem norm_div_pos_iff {E : Type u_5} [NormedGroup E] {a b : E} :
0 < a / b a b
theorem norm_sub_pos_iff {E : Type u_5} [NormedAddGroup E] {a b : E} :
0 < a - b a b
theorem eq_of_norm_div_le_zero {E : Type u_5} [NormedGroup E] {a b : E} (h : a / b 0) :
a = b
theorem eq_of_norm_sub_le_zero {E : Type u_5} [NormedAddGroup E] {a b : E} (h : a - b 0) :
a = b
theorem eq_of_norm_div_eq_zero {E : Type u_5} [NormedGroup E] {a b : E} :
a / b = 0a = b

Alias of the forward direction of norm_div_eq_zero_iff.

theorem eq_of_norm_sub_eq_zero {E : Type u_5} [NormedAddGroup E] {a b : E} :
a - b = 0a = b
theorem eq_one_or_norm_pos {E : Type u_5} [NormedGroup E] (a : E) :
a = 1 0 < a
theorem eq_zero_or_norm_pos {E : Type u_5} [NormedAddGroup E] (a : E) :
a = 0 0 < a
theorem eq_one_or_nnnorm_pos {E : Type u_5} [NormedGroup E] (a : E) :
a = 1 0 < a‖₊
theorem eq_zero_or_nnnorm_pos {E : Type u_5} [NormedAddGroup E] (a : E) :
a = 0 0 < a‖₊
@[simp]
theorem nnnorm_eq_zero' {E : Type u_5} [NormedGroup E] {a : E} :
a‖₊ = 0 a = 1
@[simp]
theorem nnnorm_eq_zero {E : Type u_5} [NormedAddGroup E] {a : E} :
a‖₊ = 0 a = 0
@[simp]
theorem enorm_eq_zero' {E : Type u_5} [NormedGroup E] {a : E} :
a‖ₑ = 0 a = 1
@[simp]
theorem enorm_eq_zero {E : Type u_5} [NormedAddGroup E] {a : E} :
a‖ₑ = 0 a = 0
theorem nnnorm_ne_zero_iff' {E : Type u_5} [NormedGroup E] {a : E} :
theorem nnnorm_ne_zero_iff {E : Type u_5} [NormedAddGroup E] {a : E} :
theorem enorm_ne_zero' {E : Type u_5} [NormedGroup E] {a : E} :
theorem enorm_ne_zero {E : Type u_5} [NormedAddGroup E] {a : E} :
@[simp]
theorem nnnorm_pos' {E : Type u_5} [NormedGroup E] {a : E} :
@[simp]
theorem nnnorm_pos {E : Type u_5} [NormedAddGroup E] {a : E} :
@[simp]
theorem enorm_pos' {E : Type u_5} [NormedGroup E] {a : E} :
@[simp]
theorem enorm_pos {E : Type u_5} [NormedAddGroup E] {a : E} :

The norm of a normed group as a group norm.

Equations

The norm of a normed group as an additive group norm.

Equations
@[simp]
theorem coe_normGroupNorm (E : Type u_5) [NormedGroup E] :

Some relations with HasCompactSupport

theorem hasCompactSupport_norm_iff {α : Type u_2} {E : Type u_5} [NormedAddGroup E] [TopologicalSpace α] {f : αE} :
theorem HasCompactSupport.norm {α : Type u_2} {E : Type u_5} [NormedAddGroup E] [TopologicalSpace α] {f : αE} :
HasCompactSupport fHasCompactSupport fun (x : α) => f x

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.

Extension for the positivity tactic: additive norms are always nonnegative, and positive on non-zero inputs.