Documentation

Mathlib.Analysis.Convex.Gauge

The Minkowski functional #

This file defines the Minkowski functional, aka gauge.

The Minkowski functional of a set s is the function which associates each point to how much you need to scale s for x to be inside it. When s is symmetric, convex and absorbent, its gauge is a seminorm. Reciprocally, any seminorm arises as the gauge of some set, namely its unit ball. This induces the equivalence of seminorms and locally convex topological vector spaces.

Main declarations #

For a real vector space,

References #

Tags #

Minkowski functional, gauge

def gauge {E : Type u_2} [AddCommGroup E] [Module E] (s : Set E) (x : E) :

The Minkowski functional. Given a set s in a real vector space, gauge s is the functional which sends x : E to the smallest r : ℝ such that x is in s scaled by r.

Equations
theorem gauge_def {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} :
theorem gauge_def' {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} :

An alternative definition of the gauge using scalar multiplication on the element rather than on the set.

theorem Absorbent.gauge_set_nonempty {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} (absorbs : Absorbent s) :

If the given subset is Absorbent then the set we take an infimum over in gauge is nonempty, which is useful for proving many properties about the gauge.

theorem gauge_mono {E : Type u_2} [AddCommGroup E] [Module E] {s t : Set E} (hs : Absorbent s) (h : s t) :
theorem exists_lt_of_gauge_lt {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} {a : } (absorbs : Absorbent s) (h : gauge s x < a) :
∃ (b : ), 0 < b b < a x b s
@[simp]
theorem gauge_zero {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} :
gauge s 0 = 0

The gauge evaluated at 0 is always zero (mathematically this requires 0 to be in the set s but, the real infimum of the empty set in Lean being defined as 0, it holds unconditionally).

@[simp]
theorem gauge_zero' {E : Type u_2} [AddCommGroup E] [Module E] :
gauge 0 = 0
@[simp]
theorem gauge_empty {E : Type u_2} [AddCommGroup E] [Module E] :
theorem gauge_of_subset_zero {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (h : s 0) :
gauge s = 0
theorem gauge_nonneg {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (x : E) :

The gauge is always nonnegative.

theorem gauge_neg {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (symmetric : xs, -x s) (x : E) :
gauge s (-x) = gauge s x
theorem gauge_neg_set_neg {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (x : E) :
gauge (-s) (-x) = gauge s x
theorem gauge_neg_set_eq_gauge_neg {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (x : E) :
gauge (-s) x = gauge s (-x)
theorem gauge_le_of_mem {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} {a : } (ha : 0 a) (hx : x a s) :
gauge s x a
theorem gauge_le_eq {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {a : } (hs₁ : Convex s) (hs₀ : 0 s) (hs₂ : Absorbent s) (ha : 0 a) :
{x : E | gauge s x a} = ⋂ (r : ), ⋂ (_ : a < r), r s
theorem gauge_lt_eq' {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (absorbs : Absorbent s) (a : ) :
{x : E | gauge s x < a} = ⋃ (r : ), ⋃ (_ : 0 < r), ⋃ (_ : r < a), r s
theorem gauge_lt_eq {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (absorbs : Absorbent s) (a : ) :
{x : E | gauge s x < a} = rSet.Ioo 0 a, r s
theorem mem_openSegment_of_gauge_lt_one {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} (absorbs : Absorbent s) (hgauge : gauge s x < 1) :
ys, x openSegment 0 y
theorem gauge_lt_one_subset_self {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (hs : Convex s) (h₀ : 0 s) (absorbs : Absorbent s) :
{x : E | gauge s x < 1} s
theorem gauge_le_one_of_mem {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} (hx : x s) :
gauge s x 1
theorem gauge_add_le {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (hs : Convex s) (absorbs : Absorbent s) (x y : E) :
gauge s (x + y) gauge s x + gauge s y

Gauge is subadditive.

theorem self_subset_gauge_le_one {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} :
s {x : E | gauge s x 1}
theorem Convex.gauge_le {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (hs : Convex s) (h₀ : 0 s) (absorbs : Absorbent s) (a : ) :
Convex {x : E | gauge s x a}
theorem Balanced.starConvex {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} (hs : Balanced s) :
theorem le_gauge_of_not_mem {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} {a : } (hs₀ : StarConvex 0 s) (hs₂ : Absorbs s {x}) (hx : xa s) :
theorem one_le_gauge_of_not_mem {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} (hs₁ : StarConvex 0 s) (hs₂ : Absorbs s {x}) (hx : xs) :
theorem gauge_smul_of_nonneg {E : Type u_2} [AddCommGroup E] [Module E] {α : Type u_3} [LinearOrderedField α] [MulActionWithZero α ] [OrderedSMul α ] [MulActionWithZero α E] [IsScalarTower α (Set E)] {s : Set E} {a : α} (ha : 0 a) (x : E) :
theorem gauge_smul_left {E : Type u_2} [AddCommGroup E] [Module E] {α : Type u_3} [LinearOrderedField α] [MulActionWithZero α ] [OrderedSMul α ] [Module α E] [SMulCommClass α ] [IsScalarTower α ] [IsScalarTower α E] {s : Set E} (symmetric : xs, -x s) (a : α) :
theorem gauge_norm_smul {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] (hs : Balanced 𝕜 s) (r : 𝕜) (x : E) :
theorem gauge_smul {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] (hs : Balanced 𝕜 s) (r : 𝕜) (x : E) :

If s is balanced, then the Minkowski functional is ℂ-homogeneous.

theorem gauge_eq_zero {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} [TopologicalSpace E] [T1Space E] (hs : Absorbent s) (hb : Bornology.IsVonNBounded s) :
gauge s x = 0 x = 0
theorem gauge_pos {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} [TopologicalSpace E] [T1Space E] (hs : Absorbent s) (hb : Bornology.IsVonNBounded s) :
theorem gauge_lt_one_eq_self_of_isOpen {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [TopologicalSpace E] [ContinuousSMul E] (hs₁ : Convex s) (hs₀ : 0 s) (hs₂ : IsOpen s) :
{x : E | gauge s x < 1} = s
theorem gauge_lt_one_of_mem_of_isOpen {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [TopologicalSpace E] [ContinuousSMul E] (hs₂ : IsOpen s) {x : E} (hx : x s) :
gauge s x < 1
theorem gauge_lt_of_mem_smul {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [TopologicalSpace E] [ContinuousSMul E] (x : E) (ε : ) (hε : 0 < ε) (hs₂ : IsOpen s) (hx : x ε s) :
gauge s x < ε
theorem mem_closure_of_gauge_le_one {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} [TopologicalSpace E] [ContinuousSMul E] (hc : Convex s) (hs₀ : 0 s) (ha : Absorbent s) (h : gauge s x 1) :
theorem mem_frontier_of_gauge_eq_one {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} [TopologicalSpace E] [ContinuousSMul E] (hc : Convex s) (hs₀ : 0 s) (ha : Absorbent s) (h : gauge s x = 1) :

If s is a neighborhood of the origin, then gauge s is continuous at the origin. See also continuousAt_gauge.

theorem continuousAt_gauge {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} {x : E} [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul E] (hc : Convex s) (hs₀ : s nhds 0) :

If s is a convex neighborhood of the origin in a topological real vector space, then gauge s is continuous. If the ambient space is a normed space, then gauge s is Lipschitz continuous, see Convex.lipschitz_gauge.

theorem continuous_gauge {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul E] (hc : Convex s) (hs₀ : s nhds 0) :

If s is a convex neighborhood of the origin in a topological real vector space, then gauge s is continuous. If the ambient space is a normed space, then gauge s is Lipschitz continuous, see Convex.lipschitz_gauge.

theorem gauge_lt_one_eq_interior {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul E] (hc : Convex s) (hs₀ : s nhds 0) :
{x : E | gauge s x < 1} = interior s
def gaugeSeminorm {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] (hs₀ : Balanced 𝕜 s) (hs₁ : Convex s) (hs₂ : Absorbent s) :
Seminorm 𝕜 E

gauge s as a seminorm when s is balanced, convex and absorbent.

Equations
@[simp]
theorem gaugeSeminorm_toFun {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] (hs₀ : Balanced 𝕜 s) (hs₁ : Convex s) (hs₂ : Absorbent s) (x : E) :
(gaugeSeminorm hs₀ hs₁ hs₂) x = gauge s x
theorem gaugeSeminorm_lt_one_of_isOpen {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] {hs₀ : Balanced 𝕜 s} {hs₁ : Convex s} {hs₂ : Absorbent s} [TopologicalSpace E] [ContinuousSMul E] (hs : IsOpen s) {x : E} (hx : x s) :
(gaugeSeminorm hs₀ hs₁ hs₂) x < 1
theorem gaugeSeminorm_ball_one {𝕜 : Type u_1} {E : Type u_2} [AddCommGroup E] [Module E] {s : Set E} [RCLike 𝕜] [Module 𝕜 E] [IsScalarTower 𝕜 E] {hs₀ : Balanced 𝕜 s} {hs₁ : Convex s} {hs₂ : Absorbent s} [TopologicalSpace E] [ContinuousSMul E] (hs : IsOpen s) :
(gaugeSeminorm hs₀ hs₁ hs₂).ball 0 1 = s
@[simp]
theorem Seminorm.gauge_ball {E : Type u_2} [AddCommGroup E] [Module E] (p : Seminorm E) :
gauge (p.ball 0 1) = p

Any seminorm arises as the gauge of its unit ball.

theorem Seminorm.gaugeSeminorm_ball {E : Type u_2} [AddCommGroup E] [Module E] (p : Seminorm E) :
gaugeSeminorm = p
theorem gauge_ball {E : Type u_2} [SeminormedAddCommGroup E] [NormedSpace E] {r : } (hr : 0 r) (x : E) :
@[simp]
theorem gauge_closedBall {E : Type u_2} [SeminormedAddCommGroup E] [NormedSpace E] {r : } (hr : 0 r) (x : E) :
theorem mul_gauge_le_norm {E : Type u_2} [SeminormedAddCommGroup E] [NormedSpace E] {s : Set E} {r : } {x : E} (hs : Metric.ball 0 r s) :
theorem Convex.lipschitzWith_gauge {E : Type u_2} [SeminormedAddCommGroup E] [NormedSpace E] {s : Set E} {r : NNReal} (hc : Convex s) (hr : 0 < r) (hs : Metric.ball 0 r s) :
theorem Convex.lipschitz_gauge {E : Type u_2} [SeminormedAddCommGroup E] [NormedSpace E] {s : Set E} (hc : Convex s) (h₀ : s nhds 0) :
∃ (K : NNReal), LipschitzWith K (gauge s)
theorem le_gauge_of_subset_closedBall {E : Type u_2} [NormedAddCommGroup E] [NormedSpace E] {s : Set E} {r : } {x : E} (hs : Absorbent s) (hr : 0 r) (hsr : s Metric.closedBall 0 r) :
x / r gauge s x