Documentation

Mathlib.Analysis.LocallyConvex.Basic

Local convexity #

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} [SeminormedRing 𝕜] [SMul 𝕜 E] (A : Set E) (B : Set E) :

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

Equations
Instances For
    @[simp]
    theorem absorbs_empty {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} :
    Absorbs 𝕜 s
    theorem Absorbs.mono {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {t : Set E} {u : Set E} {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} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {t : Set E} {u : Set E} (hs : Absorbs 𝕜 s u) (h : s t) :
    Absorbs 𝕜 t u
    theorem Absorbs.mono_right {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {u : Set E} {v : Set E} (hs : Absorbs 𝕜 s u) (h : v u) :
    Absorbs 𝕜 s v
    theorem Absorbs.union {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {u : Set E} {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} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {u : Set E} {v : Set E} :
    Absorbs 𝕜 s (u v) Absorbs 𝕜 s u Absorbs 𝕜 s v
    theorem absorbs_iUnion_finset {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} {ι : Type u_6} {t : Finset ι} {f : ιSet E} :
    Absorbs 𝕜 s (⋃ (i : ι) (_ : i t), f i) ∀ (i : ι), i tAbsorbs 𝕜 s (f i)
    theorem Set.Finite.absorbs_iUnion {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {ι : Type u_6} {s : Set E} {t : Set ι} {f : ιSet E} (hi : Set.Finite t) :
    Absorbs 𝕜 s (⋃ (i : ι) (_ : i t), f i) ∀ (i : ι), i tAbsorbs 𝕜 s (f i)
    def Absorbent (𝕜 : Type u_1) {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] (A : Set E) :

    A set is absorbent if it absorbs every singleton.

    Equations
    Instances For
      theorem Absorbent.subset {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {A : Set E} {B : Set E} (hA : Absorbent 𝕜 A) (hAB : A B) :
      Absorbent 𝕜 B
      theorem absorbent_iff_forall_absorbs_singleton {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {A : Set E} :
      Absorbent 𝕜 A ∀ (x : E), Absorbs 𝕜 A {x}
      theorem Absorbent.absorbs {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [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} [SeminormedRing 𝕜] [SMul 𝕜 E] {A : Set E} :
      Absorbent 𝕜 A ∀ (x : E), r, 0 r ∀ ⦃a : 𝕜⦄, r < ax a A
      theorem Absorbent.absorbs_finite {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} (hs : Absorbent 𝕜 s) {v : Set E} (hv : Set.Finite v) :
      Absorbs 𝕜 s v
      def Balanced (𝕜 : Type u_1) {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] (A : Set E) :

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

      Equations
      Instances For
        theorem balanced_iff_smul_mem {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} :
        Balanced 𝕜 s ∀ ⦃a : 𝕜⦄, a 1∀ ⦃x : E⦄, x sa x s
        theorem Balanced.smul_mem {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} :
        Balanced 𝕜 s∀ ⦃a : 𝕜⦄, a 1∀ ⦃x : E⦄, x sa x s

        Alias of the forward direction of balanced_iff_smul_mem.

        @[simp]
        theorem balanced_empty {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] :
        @[simp]
        theorem balanced_univ {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] :
        Balanced 𝕜 Set.univ
        theorem Balanced.union {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {A : Set E} {B : Set E} (hA : Balanced 𝕜 A) (hB : Balanced 𝕜 B) :
        Balanced 𝕜 (A B)
        theorem Balanced.inter {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [SMul 𝕜 E] {A : Set E} {B : Set E} (hA : Balanced 𝕜 A) (hB : Balanced 𝕜 B) :
        Balanced 𝕜 (A B)
        theorem balanced_iUnion {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} [SeminormedRing 𝕜] [SMul 𝕜 E] {f : ιSet E} (h : ∀ (i : ι), Balanced 𝕜 (f i)) :
        Balanced 𝕜 (⋃ (i : ι), f i)
        theorem balanced_iUnion₂ {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} {κ : ιSort u_5} [SeminormedRing 𝕜] [SMul 𝕜 E] {f : (i : ι) → κ iSet E} (h : ∀ (i : ι) (j : κ i), Balanced 𝕜 (f i j)) :
        Balanced 𝕜 (⋃ (i : ι) (j : κ i), f i j)
        theorem balanced_iInter {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} [SeminormedRing 𝕜] [SMul 𝕜 E] {f : ιSet E} (h : ∀ (i : ι), Balanced 𝕜 (f i)) :
        Balanced 𝕜 (⋂ (i : ι), f i)
        theorem balanced_iInter₂ {𝕜 : Type u_1} {E : Type u_3} {ι : Sort u_4} {κ : ιSort u_5} [SeminormedRing 𝕜] [SMul 𝕜 E] {f : (i : ι) → κ iSet 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} [SeminormedRing 𝕜] [SMul 𝕜 E] {s : Set E} [SMul 𝕝 E] [SMulCommClass 𝕜 𝕝 E] (a : 𝕝) (hs : Balanced 𝕜 s) :
        Balanced 𝕜 (a s)
        theorem Absorbs.neg {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {t : Set E} :
        Absorbs 𝕜 s tAbsorbs 𝕜 (-s) (-t)
        theorem Balanced.neg {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} :
        Balanced 𝕜 sBalanced 𝕜 (-s)
        theorem Absorbs.add {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s₁ : Set E} {s₂ : Set E} {t₁ : Set E} {t₂ : Set E} :
        Absorbs 𝕜 s₁ t₁Absorbs 𝕜 s₂ t₂Absorbs 𝕜 (s₁ + s₂) (t₁ + t₂)
        theorem Balanced.add {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {t : Set E} (hs : Balanced 𝕜 s) (ht : Balanced 𝕜 t) :
        Balanced 𝕜 (s + t)
        theorem Absorbs.sub {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s₁ : Set E} {s₂ : Set E} {t₁ : Set E} {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} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {t : Set E} (hs : Balanced 𝕜 s) (ht : Balanced 𝕜 t) :
        Balanced 𝕜 (s - t)
        theorem balanced_zero {𝕜 : Type u_1} {E : Type u_3} [SeminormedRing 𝕜] [AddCommGroup E] [Module 𝕜 E] :
        Balanced 𝕜 0
        theorem Balanced.smul_mono {𝕜 : Type u_1} {𝕝 : Type u_2} {E : Type u_3} [NormedField 𝕜] [NormedRing 𝕝] [NormedSpace 𝕜 𝕝] [AddCommGroup E] [Module 𝕜 E] [SMulWithZero 𝕝 E] [IsScalarTower 𝕜 𝕝 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} [NormedField 𝕜] [AddCommGroup 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} [NormedField 𝕜] [AddCommGroup 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} [NormedField 𝕜] [AddCommGroup 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} [NormedField 𝕜] [AddCommGroup 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} [NormedField 𝕜] [AddCommGroup 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} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {t : Set E} {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} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} {t : Set E} {u : Set E} :
        Absorbs 𝕜 (s t) u Absorbs 𝕜 s u Absorbs 𝕜 t u
        theorem absorbent_univ {𝕜 : Type u_1} {E : Type u_3} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] :
        Absorbent 𝕜 Set.univ
        theorem absorbent_nhds_zero {𝕜 : Type u_1} {E : Type u_3} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {A : Set E} [TopologicalSpace E] [ContinuousSMul 𝕜 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} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {A : Set E} [TopologicalSpace E] [ContinuousSMul 𝕜 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} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {A : Set E} [TopologicalSpace E] [ContinuousSMul 𝕜 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} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {A : Set E} [TopologicalSpace E] [ContinuousSMul 𝕜 E] (hA : Balanced 𝕜 A) :
        Balanced 𝕜 (closure A)
        theorem absorbs_zero_iff {𝕜 : Type u_1} {E : Type u_3} [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} :
        Absorbs 𝕜 s 0 0 s
        theorem Absorbent.zero_mem {𝕜 : Type u_1} {E : Type u_3} [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} (hs : Absorbent 𝕜 s) :
        0 s
        theorem balanced_convexHull_of_balanced {𝕜 : Type u_1} {E : Type u_3} [NontriviallyNormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] {s : Set E} [Module E] [SMulCommClass 𝕜 E] (hs : Balanced 𝕜 s) :
        Balanced 𝕜 (↑(convexHull ) s)
        theorem balanced_iff_neg_mem {E : Type u_3} [AddCommGroup E] [Module E] {s : Set E} (hs : Convex s) :
        Balanced s ∀ ⦃x : E⦄, x s-x s