mathlib3 documentation

analysis.locally_convex.basic

Local convexity #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This file defines absorbent and balanced sets.

An absorbent set is one that "surrounds" the origin. The idea is made precise by requiring that any point belongs to all large enough scalings of the set. This is the vector world analog of a topological neighborhood of the origin.

A balanced set is one that is everywhere around the origin. This means that a • s ⊆ s for all a of norm less than 1.

Main declarations #

For a module over a normed ring:

References #

Tags #

absorbent, balanced, locally convex, LCTVS

def absorbs (𝕜 : Type u_1) {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] (A B : set E) :
Prop

A set A absorbs another set B if B is contained in all scalings of A by elements of sufficiently large norm.

Equations
@[simp]
theorem absorbs_empty {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} :
absorbs 𝕜 s
theorem absorbs.mono {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s t u v : set E} (hs : absorbs 𝕜 s u) (hst : s t) (hvu : v u) :
absorbs 𝕜 t v
theorem absorbs.mono_left {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s t u : set E} (hs : absorbs 𝕜 s u) (h : s t) :
absorbs 𝕜 t u
theorem absorbs.mono_right {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s u v : set E} (hs : absorbs 𝕜 s u) (h : v u) :
absorbs 𝕜 s v
theorem absorbs.union {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s u v : set E} (hu : absorbs 𝕜 s u) (hv : absorbs 𝕜 s v) :
absorbs 𝕜 s (u v)
@[simp]
theorem absorbs_union {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s u v : set E} :
absorbs 𝕜 s (u v) absorbs 𝕜 s u absorbs 𝕜 s v
theorem absorbs_Union_finset {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} {ι : Type u_2} {t : finset ι} {f : ι set E} :
absorbs 𝕜 s ( (i : ι) (H : i t), f i) (i : ι), i t absorbs 𝕜 s (f i)
theorem set.finite.absorbs_Union {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {ι : Type u_2} {s : set E} {t : set ι} {f : ι set E} (hi : t.finite) :
absorbs 𝕜 s ( (i : ι) (H : i t), f i) (i : ι), i t absorbs 𝕜 s (f i)
def absorbent (𝕜 : Type u_1) {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] (A : set E) :
Prop

A set is absorbent if it absorbs every singleton.

Equations
theorem absorbent.subset {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {A B : set E} (hA : absorbent 𝕜 A) (hAB : A B) :
absorbent 𝕜 B
theorem absorbent_iff_forall_absorbs_singleton {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {A : set E} :
absorbent 𝕜 A (x : E), absorbs 𝕜 A {x}
theorem absorbent.absorbs {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} (hs : absorbent 𝕜 s) {x : E} :
absorbs 𝕜 s {x}
theorem absorbent_iff_nonneg_lt {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {A : set E} :
absorbent 𝕜 A (x : E), (r : ), 0 r ⦃a : 𝕜⦄, r < a x a A
theorem absorbent.absorbs_finite {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} (hs : absorbent 𝕜 s) {v : set E} (hv : v.finite) :
absorbs 𝕜 s v
def balanced (𝕜 : Type u_1) {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] (A : set E) :
Prop

A set A is balanced if a • A is contained in A whenever a has norm at most 1.

Equations
theorem balanced_iff_smul_mem {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} :
balanced 𝕜 s ⦃a : 𝕜⦄, a 1 ⦃x : E⦄, x s a x s
theorem balanced.smul_mem {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} :
balanced 𝕜 s ⦃a : 𝕜⦄, a 1 ⦃x : E⦄, x s a x s

Alias of the forward direction of balanced_iff_smul_mem.

@[simp]
theorem balanced_empty {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] :
@[simp]
theorem balanced_univ {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] :
theorem balanced.union {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {A B : set E} (hA : balanced 𝕜 A) (hB : balanced 𝕜 B) :
balanced 𝕜 (A B)
theorem balanced.inter {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {A B : set E} (hA : balanced 𝕜 A) (hB : balanced 𝕜 B) :
balanced 𝕜 (A B)
theorem balanced_Union {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {f : ι set E} (h : (i : ι), balanced 𝕜 (f i)) :
balanced 𝕜 ( (i : ι), f i)
theorem balanced_Union₂ {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} {κ : ι Sort u_5} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {f : Π (i : ι), κ i set E} (h : (i : ι) (j : κ i), balanced 𝕜 (f i j)) :
balanced 𝕜 ( (i : ι) (j : κ i), f i j)
theorem balanced_Inter {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {f : ι set E} (h : (i : ι), balanced 𝕜 (f i)) :
balanced 𝕜 ( (i : ι), f i)
theorem balanced_Inter₂ {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} {κ : ι Sort u_5} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {f : Π (i : ι), κ i set E} (h : (i : ι) (j : κ i), balanced 𝕜 (f i j)) :
balanced 𝕜 ( (i : ι) (j : κ i), f i j)
theorem balanced.smul {𝕜 : Type u_1} {𝕝 : Type u_2} {E : Type u_3} [semi_normed_ring 𝕜] [has_smul 𝕜 E] {s : set E} [has_smul 𝕝 E] [smul_comm_class 𝕜 𝕝 E] (a : 𝕝) (hs : balanced 𝕜 s) :
balanced 𝕜 (a s)
theorem absorbs.neg {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s t : set E} :
absorbs 𝕜 s t absorbs 𝕜 (-s) (-t)
theorem balanced.neg {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} :
balanced 𝕜 s balanced 𝕜 (-s)
theorem absorbs.add {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s₁ s₂ t₁ t₂ : set E} :
absorbs 𝕜 s₁ t₁ absorbs 𝕜 s₂ t₂ absorbs 𝕜 (s₁ + s₂) (t₁ + t₂)
theorem balanced.add {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s t : set E} (hs : balanced 𝕜 s) (ht : balanced 𝕜 t) :
balanced 𝕜 (s + t)
theorem absorbs.sub {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s₁ s₂ t₁ t₂ : set E} (h₁ : absorbs 𝕜 s₁ t₁) (h₂ : absorbs 𝕜 s₂ t₂) :
absorbs 𝕜 (s₁ - s₂) (t₁ - t₂)
theorem balanced.sub {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] {s t : set E} (hs : balanced 𝕜 s) (ht : balanced 𝕜 t) :
balanced 𝕜 (s - t)
theorem balanced_zero {𝕜 : Type u_1} {E : Type u_3} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] :
balanced 𝕜 0
theorem balanced.smul_mono {𝕜 : Type u_1} {𝕝 : Type u_2} {E : Type u_3} [normed_field 𝕜] [normed_ring 𝕝] [normed_space 𝕜 𝕝] [add_comm_group E] [module 𝕜 E] [smul_with_zero 𝕝 E] [is_scalar_tower 𝕜 𝕝 E] {s : set E} (hs : balanced 𝕝 s) {a : 𝕝} {b : 𝕜} (h : a b) :
a s b s

Scalar multiplication (by possibly different types) of a balanced set is monotone.

theorem balanced.absorbs_self {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} (hA : balanced 𝕜 A) :
absorbs 𝕜 A A

A balanced set absorbs itself.

theorem balanced.subset_smul {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} {a : 𝕜} (hA : balanced 𝕜 A) (ha : 1 a) :
A a A
theorem balanced.smul_eq {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} {a : 𝕜} (hA : balanced 𝕜 A) (ha : a = 1) :
a A = A
theorem balanced.mem_smul_iff {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} {x : E} {a b : 𝕜} (hs : balanced 𝕜 s) (h : a = b) :
a x s b x s
theorem balanced.neg_mem_iff {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} {x : E} (hs : balanced 𝕜 s) :
-x s x s
theorem absorbs.inter {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s t u : set E} (hs : absorbs 𝕜 s u) (ht : absorbs 𝕜 t u) :
absorbs 𝕜 (s t) u
@[simp]
theorem absorbs_inter {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s t u : set E} :
absorbs 𝕜 (s t) u absorbs 𝕜 s u absorbs 𝕜 t u
theorem absorbent_univ {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] :
theorem absorbent_nhds_zero {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} [topological_space E] [has_continuous_smul 𝕜 E] (hA : A nhds 0) :
absorbent 𝕜 A

Every neighbourhood of the origin is absorbent.

theorem balanced_zero_union_interior {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} [topological_space E] [has_continuous_smul 𝕜 E] (hA : balanced 𝕜 A) :
balanced 𝕜 (0 interior A)

The union of {0} with the interior of a balanced set is balanced.

theorem balanced.interior {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} [topological_space E] [has_continuous_smul 𝕜 E] (hA : balanced 𝕜 A) (h : 0 interior A) :

The interior of a balanced set is balanced if it contains the origin.

theorem balanced.closure {𝕜 : Type u_1} {E : Type u_3} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {A : set E} [topological_space E] [has_continuous_smul 𝕜 E] (hA : balanced 𝕜 A) :
balanced 𝕜 (closure A)
theorem absorbs_zero_iff {𝕜 : Type u_1} {E : Type u_3} [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} :
absorbs 𝕜 s 0 0 s
theorem absorbent.zero_mem {𝕜 : Type u_1} {E : Type u_3} [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} (hs : absorbent 𝕜 s) :
0 s
theorem balanced_convex_hull_of_balanced {𝕜 : Type u_1} {E : Type u_3} [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {s : set E} [module E] [smul_comm_class 𝕜 E] (hs : balanced 𝕜 s) :
theorem balanced_iff_neg_mem {E : Type u_3} [add_comm_group E] [module E] {s : set E} (hs : convex s) :
balanced s ⦃x : E⦄, x s -x s