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

[Merged by Bors] - feat(topology/metric_space): diameter of pointwise zero and addition #19028

Closed
wants to merge 10 commits into from
Closed
Show file tree
Hide file tree
Changes from 7 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
12 changes: 12 additions & 0 deletions src/analysis/normed/group/pointwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Yaël Dillies
-/
import analysis.normed.group.basic
import topology.metric_space.hausdorff_distance
import topology.metric_space.isometric_smul

/-!
# Properties of pointwise addition of sets in normed groups
Expand All @@ -24,6 +25,7 @@ variables {E : Type*}
section seminormed_group
variables [seminormed_group E] {ε δ : ℝ} {s t : set E} {x y : E}

-- note: we can't use `lipschitz_with.bounded_image2` here without adding `[isometric_smul E E]`
@[to_additive] lemma metric.bounded.mul (hs : bounded s) (ht : bounded t) : bounded (s * t) :=
Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

@YaelDillies, any ideas on whether I can make a better generalization that does work here?

Copy link
Collaborator

@YaelDillies YaelDillies Jul 14, 2023

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Yes, I believe you can instead use

lemma bounded_image2 (f : α → β → γ) {K₁ K₂ : ℝ≥0} {s : set α} {t : set β}
  (hf₁ : ∀ b ∈ t, lipschitz_with K₁ (λ a, f a b) s) (hf₂ : ∀ a ∈ s, lipschitz_on_with K₂ (f a) t)
  (hs : bounded s) (ht : bounded t) : bounded (set.image2 f s t) := sorry

Then you use the fact that multiplication is Lipschitz on bounded sets.

I'm not sure it results in a shorter proof, but it's at least conceptually more satisfying.

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I generalized bounded_image2, but haven't changed this proof

Then you use the fact that multiplication is Lipschitz on bounded sets.

Where is this result?

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Not in mathlib :/

Should be easy though. The bound on the set turns into the Lipschitz constant.

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Are you sure this doesn't need the same [isometric_smul E E] argument that I was trying to avoid?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Don't forget we're in weird normed_group land not normed_add_group.

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Then I'm confused: Isn't isometric_smul E E always true in that context?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Not without commutativity; only isometric_smul (MulOpposite E) E is true

Copy link
Collaborator

@YaelDillies YaelDillies Jul 16, 2023

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Do you care about noncommutative settings?

Also doesn't that mean you can do the proof along t instead of along s (whatever that means)?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Do you care about noncommutative settings?

No, but I don't want to un-generalize a lemma solely for the point of a golf.

begin
obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le',
Expand All @@ -33,6 +35,10 @@ begin
exact norm_mul_le_of_le (hRs x hx) (hRt y hy),
end

@[to_additive] lemma metric.bounded.of_mul (hst : bounded (s * t)) :
bounded s ∨ bounded t :=
antilipschitz_with.bounded_of_image2_left _ (λ x, (isometry_mul_right x).antilipschitz) hst

@[to_additive] lemma metric.bounded.inv : bounded s → bounded s⁻¹ :=
by { simp_rw [bounded_iff_forall_norm_le', ←image_inv, ball_image_iff, norm_inv'], exact id }

Expand All @@ -55,6 +61,12 @@ eq_of_forall_le_iff $ λ r, by simp_rw [le_inf_edist, ←image_inv, ball_image_i
lemma inf_edist_inv_inv (x : E) (s : set E) : inf_edist x⁻¹ s⁻¹ = inf_edist x s :=
by rw [inf_edist_inv, inv_inv]

@[to_additive] lemma ediam_mul_le (x y : set E) :
emetric.diam (x * y) ≤ emetric.diam x + emetric.diam y :=
eric-wieser marked this conversation as resolved.
Show resolved Hide resolved
(lipschitz_with.ediam_image2_le (*)
(λ _, (isometry_mul_right _).lipschitz) (λ _, (isometry_mul_left _).lipschitz) _ _).trans_eq $
by simp only [ennreal.coe_one, one_mul]

end emetric

variables (ε δ s t x y)
Expand Down
25 changes: 24 additions & 1 deletion src/topology/metric_space/antilipschitz.lean
Original file line number Diff line number Diff line change
Expand Up @@ -192,7 +192,8 @@ namespace antilipschitz_with

open metric

variables [pseudo_metric_space α] [pseudo_metric_space β] {K : ℝ≥0} {f : α → β}
variables [pseudo_metric_space α] [pseudo_metric_space β] [pseudo_metric_space γ]
variables {K : ℝ≥0} {f : α → β}

lemma bounded_preimage (hf : antilipschitz_with K f)
{s : set β} (hs : bounded s) :
Expand All @@ -219,6 +220,28 @@ begin
exact (hf.image_preimage _).symm
end

lemma bounded_of_image2_left (f : α → β → γ) {K₁ : ℝ≥0}
(hf : ∀ b, antilipschitz_with K₁ (λ a, f a b))
{s : set α} {t : set β} (hst : bounded (set.image2 f s t)) :
bounded s ∨ bounded t :=
begin
contrapose! hst,
obtain ⟨b, hb⟩ : t.nonempty := nonempty_of_unbounded hst.2,
have : ¬bounded (set.image2 f s {b}),
{ intro h,
apply hst.1,
rw set.image2_singleton_right at h,
replace h := (hf b).bounded_preimage h,
refine h.mono (subset_preimage_image _ _) },
exact mt (bounded.mono (image2_subset (subset.refl _) (singleton_subset_iff.mpr hb))) this,
eric-wieser marked this conversation as resolved.
Show resolved Hide resolved
end

lemma bounded_of_image2_right {f : α → β → γ} {K₂ : ℝ≥0}
(hf : ∀ a, antilipschitz_with K₂ (f a))
{s : set α} {t : set β} (hst : bounded (set.image2 f s t)) :
bounded s ∨ bounded t :=
or.symm $ bounded_of_image2_left (flip f) hf $ image2_swap f s t ▸ hst

end antilipschitz_with

lemma lipschitz_with.to_right_inverse [pseudo_emetric_space α] [pseudo_emetric_space β] {K : ℝ≥0}
Expand Down
11 changes: 10 additions & 1 deletion src/topology/metric_space/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -52,7 +52,7 @@ metric, pseudo_metric, dist

open set filter topological_space bornology

open_locale uniformity topology big_operators filter nnreal ennreal
open_locale uniformity topology big_operators filter nnreal ennreal pointwise

universes u v w
variables {α : Type u} {β : Type v} {X ι : Type*}
Expand Down Expand Up @@ -2107,6 +2107,13 @@ end
@[simp] lemma bounded_empty : bounded (∅ : set α) :=
⟨0, by simp⟩

lemma nonempty_of_unbounded (h : ¬ bounded s) : s.nonempty :=
urkud marked this conversation as resolved.
Show resolved Hide resolved
begin
rw nonempty_iff_ne_empty,
rintro rfl,
exact h bounded_empty
end

lemma bounded_iff_mem_bounded : bounded s ↔ ∀ x ∈ s, bounded s :=
⟨λ h _ _, h, λ H,
s.eq_empty_or_nonempty.elim
Expand Down Expand Up @@ -2389,6 +2396,8 @@ diam_subsingleton subsingleton_empty
@[simp] lemma diam_singleton : diam ({x} : set α) = 0 :=
diam_subsingleton subsingleton_singleton

@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton

-- Does not work as a simp-lemma, since {x, y} reduces to (insert y {x})
lemma diam_pair : diam ({x, y} : set α) = dist x y :=
by simp only [diam, emetric.diam_pair, dist_edist]
Expand Down
4 changes: 3 additions & 1 deletion src/topology/metric_space/emetric_space.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,7 +32,7 @@ to `emetric_space` at the end.

open set filter classical

open_locale uniformity topology big_operators filter nnreal ennreal
open_locale uniformity topology big_operators filter nnreal ennreal pointwise

universes u v w
variables {α : Type u} {β : Type v} {X : Type*}
Expand Down Expand Up @@ -793,6 +793,8 @@ diam_subsingleton subsingleton_empty
@[simp] lemma diam_singleton : diam ({x} : set α) = 0 :=
diam_subsingleton subsingleton_singleton

@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton

lemma diam_Union_mem_option {ι : Type*} (o : option ι) (s : ι → set α) :
diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) :=
by cases o; simp
Expand Down
23 changes: 23 additions & 0 deletions src/topology/metric_space/lipschitz.lean
Original file line number Diff line number Diff line change
Expand Up @@ -154,6 +154,19 @@ begin
exact hf.edist_le_mul_of_le (emetric.edist_le_diam_of_mem hx hy)
end

lemma ediam_image2_le (f : α → β → γ) {K₁ K₂ : ℝ≥0}
(hf₁ : ∀ b, lipschitz_with K₁ (λ a, f a b)) (hf₂ : ∀ a, lipschitz_with K₂ (f a))
(s : set α) (t : set β) :
emetric.diam (set.image2 f s t) ≤ ↑K₁ * diam s + ↑K₂ * diam t :=
begin
apply emetric.diam_le,
rintros _ ⟨a₁, b₁, ha₁, hb₁, rfl⟩ _ ⟨a₂, b₂, ha₂, hb₂, rfl⟩,
refine (edist_triangle _ (f a₂ b₁) _).trans _,
exact add_le_add
((hf₁ _).edist_le_mul_of_le (emetric.edist_le_diam_of_mem ha₁ ha₂))
((hf₂ _).edist_le_mul_of_le (emetric.edist_le_diam_of_mem hb₁ hb₂)),
end

lemma edist_lt_of_edist_lt_div (hf : lipschitz_with K f) {x y : α} {d : ℝ≥0∞}
(h : edist x y < d / K) : edist (f x) (f y) < d :=
calc edist (f x) (f y) ≤ K * edist x y : hf x y
Expand Down Expand Up @@ -355,6 +368,16 @@ lemma bounded_image (hf : lipschitz_with K f) {s : set α} (hs : metric.bounded
metric.bounded_iff_ediam_ne_top.2 $ ne_top_of_le_ne_top
(ennreal.mul_ne_top ennreal.coe_ne_top hs.ediam_ne_top) (hf.ediam_image_le s)

lemma bounded_image2 (f : α → β → γ) {K₁ K₂ : ℝ≥0}
(hf₁ : ∀ b, lipschitz_with K₁ (λ a, f a b)) (hf₂ : ∀ a, lipschitz_with K₂ (f a))
{s : set α} {t : set β} (hs : metric.bounded s) (ht : metric.bounded t) :
metric.bounded (set.image2 f s t) :=
metric.bounded_iff_ediam_ne_top.2 $ ne_top_of_le_ne_top
(ennreal.add_ne_top.mpr ⟨
ennreal.mul_ne_top ennreal.coe_ne_top hs.ediam_ne_top,
ennreal.mul_ne_top ennreal.coe_ne_top ht.ediam_ne_top⟩)
(ediam_image2_le _ hf₁ hf₂ _ _)

lemma diam_image_le (hf : lipschitz_with K f) (s : set α) (hs : metric.bounded s) :
metric.diam (f '' s) ≤ K * metric.diam s :=
metric.diam_le_of_forall_dist_le (mul_nonneg K.coe_nonneg metric.diam_nonneg) $
Expand Down