Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 1ffd04c

Browse files
committed
feat(analysis/locally_convex): add balanced hull and core (#12537)
1 parent 0f92307 commit 1ffd04c

File tree

1 file changed

+246
-0
lines changed

1 file changed

+246
-0
lines changed
Lines changed: 246 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,246 @@
1+
/-
2+
Copyright (c) 2022 Moritz Doll. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Moritz Doll
5+
-/
6+
import analysis.seminorm
7+
import order.closure
8+
9+
/-!
10+
# Balanced Core and Balanced Hull
11+
12+
## Main definitions
13+
14+
* `balanced_core`: the largest balanced subset of a set `s`.
15+
* `balanced_hull`: the smallest balanced superset of a set `s`.
16+
17+
## Main statements
18+
19+
* `balanced_core_eq_Inter`: Characterization of the balanced core as an intersection over subsets.
20+
21+
22+
## Implementation details
23+
24+
The balanced core and hull are implemented differently: for the core we take the obvious definition
25+
of the union over all balanced sets that are contained in `s`, whereas for the hull, we take the
26+
union over `r • s`, for `r` the scalars with `∥r∥ ≤ 1`. We show that `balanced_hull` has the
27+
defining properties of a hull in `balanced.hull_minimal` and `subset_balanced_hull`.
28+
For the core we need slightly stronger assumptions to obtain a characterization as an intersection,
29+
this is `balanced_core_eq_Inter`.
30+
31+
## References
32+
33+
* [Bourbaki, *Topological Vector Spaces*][bourbaki1987]
34+
35+
## Tags
36+
37+
balanced
38+
-/
39+
40+
41+
open set
42+
open_locale pointwise
43+
44+
variables {𝕜 E ι : Type*}
45+
46+
section balanced_hull
47+
48+
section semi_normed_ring
49+
variables [semi_normed_ring 𝕜]
50+
51+
section has_scalar
52+
variables [has_scalar 𝕜 E]
53+
54+
variables (𝕜)
55+
56+
/-- The largest balanced subset of `s`.-/
57+
def balanced_core (s : set E) := ⋃₀ {t : set E | balanced 𝕜 t ∧ t ⊆ s}
58+
59+
/-- Helper definition to prove `balanced_core_eq_Inter`-/
60+
def balanced_core_aux (s : set E) := ⋂ (r : 𝕜) (hr : 1 ≤ ∥r∥), r • s
61+
62+
/-- The smallest balanced superset of `s`.-/
63+
def balanced_hull (s : set E) := ⋃ (r : 𝕜) (hr : ∥r∥ ≤ 1), r • s
64+
65+
variables {𝕜}
66+
67+
lemma balanced_core_subset (s : set E) : balanced_core 𝕜 s ⊆ s :=
68+
begin
69+
refine sUnion_subset (λ t ht, _),
70+
simp only [mem_set_of_eq] at ht,
71+
exact ht.2,
72+
end
73+
74+
lemma balanced_core_mem_iff {s : set E} {x : E} : x ∈ balanced_core 𝕜 s ↔
75+
∃ t : set E, balanced 𝕜 t ∧ t ⊆ s ∧ x ∈ t :=
76+
by simp_rw [balanced_core, mem_sUnion, mem_set_of_eq, exists_prop, and_assoc]
77+
78+
lemma smul_balanced_core_subset (s : set E) {a : 𝕜} (ha : ∥a∥ ≤ 1) :
79+
a • balanced_core 𝕜 s ⊆ balanced_core 𝕜 s :=
80+
begin
81+
rw subset_def,
82+
intros x hx,
83+
rw mem_smul_set at hx,
84+
rcases hx with ⟨y, hy, hx⟩,
85+
rw balanced_core_mem_iff at hy,
86+
rcases hy with ⟨t, ht1, ht2, hy⟩,
87+
rw ←hx,
88+
refine ⟨t, _, ht1 a ha (smul_mem_smul_set hy)⟩,
89+
rw mem_set_of_eq,
90+
exact ⟨ht1, ht2⟩,
91+
end
92+
93+
lemma balanced_core_balanced (s : set E) : balanced 𝕜 (balanced_core 𝕜 s) :=
94+
λ _, smul_balanced_core_subset s
95+
96+
/-- The balanced core of `t` is maximal in the sense that it contains any balanced subset
97+
`s` of `t`.-/
98+
lemma balanced.subset_core_of_subset {s t : set E} (hs : balanced 𝕜 s) (h : s ⊆ t):
99+
s ⊆ balanced_core 𝕜 t :=
100+
begin
101+
refine subset_sUnion_of_mem _,
102+
rw [mem_set_of_eq],
103+
exact ⟨hs, h⟩,
104+
end
105+
106+
lemma balanced_core_aux_mem_iff (s : set E) (x : E) : x ∈ balanced_core_aux 𝕜 s ↔
107+
∀ (r : 𝕜) (hr : 1 ≤ ∥r∥), x ∈ r • s :=
108+
by rw [balanced_core_aux, set.mem_Inter₂]
109+
110+
lemma balanced_hull_mem_iff (s : set E) (x : E) : x ∈ balanced_hull 𝕜 s ↔
111+
∃ (r : 𝕜) (hr : ∥r∥ ≤ 1), x ∈ r • s :=
112+
by rw [balanced_hull, set.mem_Union₂]
113+
114+
/-- The balanced core of `s` is minimal in the sense that it is contained in any balanced superset
115+
`t` of `s`. -/
116+
lemma balanced.hull_subset_of_subset {s t : set E} (ht : balanced 𝕜 t) (h : s ⊆ t) :
117+
balanced_hull 𝕜 s ⊆ t :=
118+
begin
119+
intros x hx,
120+
rcases (balanced_hull_mem_iff _ _).mp hx with ⟨r, hr, hx⟩,
121+
rcases mem_smul_set.mp hx with ⟨y, hy, hx⟩,
122+
rw ←hx,
123+
exact balanced_mem ht (h hy) hr,
124+
end
125+
126+
end has_scalar
127+
128+
section add_comm_monoid
129+
130+
variables [add_comm_monoid E] [module 𝕜 E]
131+
132+
lemma balanced_core_nonempty_iff {s : set E} : (balanced_core 𝕜 s).nonempty ↔ (0 : E) ∈ s :=
133+
begin
134+
split; intro h,
135+
{ cases h with x hx,
136+
have h' : balanced 𝕜 (balanced_core 𝕜 s) := balanced_core_balanced s,
137+
have h'' := h' 0 (has_le.le.trans norm_zero.le zero_le_one),
138+
refine mem_of_subset_of_mem (subset.trans h'' (balanced_core_subset s)) _,
139+
exact mem_smul_set.mpr ⟨x, hx, zero_smul _ _⟩ },
140+
refine nonempty_of_mem (mem_of_subset_of_mem _ (mem_singleton 0)),
141+
exact balanced.subset_core_of_subset zero_singleton_balanced (singleton_subset_iff.mpr h),
142+
end
143+
144+
lemma balanced_core_zero_mem {s : set E} (hs: (0 : E) ∈ s) : (0 : E) ∈ balanced_core 𝕜 s :=
145+
balanced_core_mem_iff.mpr
146+
⟨{0}, zero_singleton_balanced, singleton_subset_iff.mpr hs, mem_singleton 0
147+
148+
variables (𝕜)
149+
150+
lemma subset_balanced_hull [norm_one_class 𝕜] {s : set E} : s ⊆ balanced_hull 𝕜 s :=
151+
λ _ hx, (balanced_hull_mem_iff _ _).mpr ⟨1, norm_one.le, mem_smul_set.mp ⟨_, hx, one_smul _ _⟩⟩
152+
153+
variables {𝕜}
154+
155+
lemma balanced_hull.balanced (s : set E) : balanced 𝕜 (balanced_hull 𝕜 s) :=
156+
begin
157+
intros a ha,
158+
simp_rw [balanced_hull, smul_set_Union₂, subset_def, mem_Union₂],
159+
intros x hx,
160+
rcases hx with ⟨r, hr, hx⟩,
161+
use [a • r],
162+
split,
163+
{ rw smul_eq_mul,
164+
refine has_le.le.trans (semi_normed_ring.norm_mul _ _) _,
165+
refine mul_le_one ha (norm_nonneg r) hr },
166+
rw smul_assoc,
167+
exact hx,
168+
end
169+
170+
end add_comm_monoid
171+
172+
end semi_normed_ring
173+
174+
section normed_field
175+
176+
variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E]
177+
178+
@[simp] lemma balanced_core_aux_empty : balanced_core_aux 𝕜 (∅ : set E) = ∅ :=
179+
begin
180+
rw [balanced_core_aux, set.Inter₂_eq_empty_iff],
181+
intros _,
182+
simp only [smul_set_empty, mem_empty_eq, not_false_iff, exists_prop, and_true],
183+
exact ⟨1, norm_one.ge⟩,
184+
end
185+
186+
lemma balanced_core_aux_subset (s : set E) : balanced_core_aux 𝕜 s ⊆ s :=
187+
begin
188+
rw subset_def,
189+
intros x hx,
190+
rw balanced_core_aux_mem_iff at hx,
191+
have h := hx 1 norm_one.ge,
192+
rw one_smul at h,
193+
exact h,
194+
end
195+
196+
lemma balanced_core_aux_balanced {s : set E} (h0 : (0 : E) ∈ balanced_core_aux 𝕜 s):
197+
balanced 𝕜 (balanced_core_aux 𝕜 s) :=
198+
begin
199+
intros a ha x hx,
200+
rcases mem_smul_set.mp hx with ⟨y, hy, hx⟩,
201+
by_cases (a = 0),
202+
{ simp[h] at hx,
203+
rw ←hx,
204+
exact h0 },
205+
rw [←hx, balanced_core_aux_mem_iff],
206+
rw balanced_core_aux_mem_iff at hy,
207+
intros r hr,
208+
have h'' : 1 ≤ ∥a⁻¹ • r∥ :=
209+
begin
210+
rw smul_eq_mul,
211+
simp only [norm_mul, norm_inv],
212+
exact one_le_mul_of_one_le_of_one_le (one_le_inv (norm_pos_iff.mpr h) ha) hr,
213+
end,
214+
have h' := hy (a⁻¹ • r) h'',
215+
rw smul_assoc at h',
216+
exact (mem_inv_smul_set_iff₀ h _ _).mp h',
217+
end
218+
219+
lemma balanced_core_aux_maximal {s t : set E} (h : t ⊆ s) (ht : balanced 𝕜 t) :
220+
t ⊆ balanced_core_aux 𝕜 s :=
221+
begin
222+
intros x hx,
223+
rw balanced_core_aux_mem_iff,
224+
intros r hr,
225+
rw mem_smul_set_iff_inv_smul_mem₀ (norm_pos_iff.mp (lt_of_lt_of_le zero_lt_one hr)),
226+
refine h (balanced_mem ht hx _),
227+
rw norm_inv,
228+
exact inv_le_one hr,
229+
end
230+
231+
lemma balanced_core_subset_balanced_core_aux {s : set E} :
232+
balanced_core 𝕜 s ⊆ balanced_core_aux 𝕜 s :=
233+
balanced_core_aux_maximal (balanced_core_subset s) (balanced_core_balanced s)
234+
235+
lemma balanced_core_eq_Inter {s : set E} (hs : (0 : E) ∈ s) :
236+
balanced_core 𝕜 s = ⋂ (r : 𝕜) (hr : 1 ≤ ∥r∥), r • s :=
237+
begin
238+
rw ←balanced_core_aux,
239+
refine subset_antisymm balanced_core_subset_balanced_core_aux _,
240+
refine balanced.subset_core_of_subset (balanced_core_aux_balanced _) (balanced_core_aux_subset s),
241+
refine mem_of_subset_of_mem balanced_core_subset_balanced_core_aux (balanced_core_zero_mem hs),
242+
end
243+
244+
end normed_field
245+
246+
end balanced_hull

0 commit comments

Comments
 (0)