|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Yury Kudryashov. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Author: Yury Kudryashov |
| 5 | +-/ |
| 6 | +import topology.metric_space.emetric_space |
| 7 | + |
| 8 | +/-! |
| 9 | +# Metric separated pairs of sets |
| 10 | +
|
| 11 | +In this file we define the predicate `is_metric_separated`. We say that two sets in an (extended) |
| 12 | +metric space are *metric separated* if the (extended) distance between `x ∈ s` and `y ∈ t` is |
| 13 | +bounded from below by a positive constant. |
| 14 | +
|
| 15 | +This notion is useful, e.g., to define metric outer measures. |
| 16 | +-/ |
| 17 | + |
| 18 | +open emetric set |
| 19 | +noncomputable theory |
| 20 | + |
| 21 | +/-- Two sets in an (extended) metric space are called *metric separated* if the (extended) distance |
| 22 | +between `x ∈ s` and `y ∈ t` is bounded from below by a positive constant. -/ |
| 23 | +def is_metric_separated {X : Type*} [emetric_space X] (s t : set X) := |
| 24 | +∃ r ≠ 0, ∀ (x ∈ s) (y ∈ t), r ≤ edist x y |
| 25 | + |
| 26 | +namespace is_metric_separated |
| 27 | + |
| 28 | +variables {X : Type*} [emetric_space X] {s t : set X} {x y : X} |
| 29 | + |
| 30 | +@[symm] lemma symm (h : is_metric_separated s t) : is_metric_separated t s := |
| 31 | +let ⟨r, r0, hr⟩ := h in ⟨r, r0, λ y hy x hx, edist_comm x y ▸ hr x hx y hy⟩ |
| 32 | + |
| 33 | +lemma comm : is_metric_separated s t ↔ is_metric_separated t s := ⟨symm, symm⟩ |
| 34 | + |
| 35 | +@[simp] lemma empty_left (s : set X) : is_metric_separated ∅ s := |
| 36 | +⟨1, ennreal.zero_lt_one.ne', λ x, false.elim⟩ |
| 37 | + |
| 38 | +@[simp] lemma empty_right (s : set X) : is_metric_separated s ∅ := |
| 39 | +(empty_left s).symm |
| 40 | + |
| 41 | +protected lemma disjoint (h : is_metric_separated s t) : disjoint s t := |
| 42 | +let ⟨r, r0, hr⟩ := h in λ x hx, r0 $ by simpa using hr x hx.1 x hx.2 |
| 43 | + |
| 44 | +lemma subset_compl_right (h : is_metric_separated s t) : s ⊆ tᶜ := |
| 45 | +λ x hs ht, h.disjoint ⟨hs, ht⟩ |
| 46 | + |
| 47 | +@[mono] lemma mono {s' t'} (hs : s ⊆ s') (ht : t ⊆ t') : |
| 48 | + is_metric_separated s' t' → is_metric_separated s t := |
| 49 | +λ ⟨r, r0, hr⟩, ⟨r, r0, λ x hx y hy, hr x (hs hx) y (ht hy)⟩ |
| 50 | + |
| 51 | +lemma mono_left {s'} (h' : is_metric_separated s' t) (hs : s ⊆ s') : |
| 52 | + is_metric_separated s t := |
| 53 | +h'.mono hs subset.rfl |
| 54 | + |
| 55 | +lemma mono_right {t'} (h' : is_metric_separated s t') (ht : t ⊆ t') : |
| 56 | + is_metric_separated s t := |
| 57 | +h'.mono subset.rfl ht |
| 58 | + |
| 59 | +lemma union_left {s'} (h : is_metric_separated s t) (h' : is_metric_separated s' t) : |
| 60 | + is_metric_separated (s ∪ s') t := |
| 61 | +begin |
| 62 | + rcases ⟨h, h'⟩ with ⟨⟨r, r0, hr⟩, ⟨r', r0', hr'⟩⟩, |
| 63 | + refine ⟨min r r', _, λ x hx y hy, hx.elim _ _⟩, |
| 64 | + { rw [← pos_iff_ne_zero] at r0 r0' ⊢, |
| 65 | + exact lt_min r0 r0' }, |
| 66 | + { exact λ hx, (min_le_left _ _).trans (hr _ hx _ hy) }, |
| 67 | + { exact λ hx, (min_le_right _ _).trans (hr' _ hx _ hy) } |
| 68 | +end |
| 69 | + |
| 70 | +@[simp] lemma union_left_iff {s'} : |
| 71 | + is_metric_separated (s ∪ s') t ↔ is_metric_separated s t ∧ is_metric_separated s' t := |
| 72 | +⟨λ h, ⟨h.mono_left (subset_union_left _ _), h.mono_left (subset_union_right _ _)⟩, |
| 73 | + λ h, h.1.union_left h.2⟩ |
| 74 | + |
| 75 | +lemma union_right {t'} (h : is_metric_separated s t) (h' : is_metric_separated s t') : |
| 76 | + is_metric_separated s (t ∪ t') := |
| 77 | +(h.symm.union_left h'.symm).symm |
| 78 | + |
| 79 | +@[simp] lemma union_right_iff {t'} : |
| 80 | + is_metric_separated s (t ∪ t') ↔ is_metric_separated s t ∧ is_metric_separated s t' := |
| 81 | +comm.trans $ union_left_iff.trans $ and_congr comm comm |
| 82 | + |
| 83 | +lemma finite_Union_left_iff {ι : Type*} {I : set ι} (hI : finite I) {s : ι → set X} {t : set X} : |
| 84 | + is_metric_separated (⋃ i ∈ I, s i) t ↔ ∀ i ∈ I, is_metric_separated (s i) t := |
| 85 | +begin |
| 86 | + refine finite.induction_on hI (by simp) (λ i I hi _ hI, _), |
| 87 | + rw [bUnion_insert, ball_insert_iff, union_left_iff, hI] |
| 88 | +end |
| 89 | + |
| 90 | +alias finite_Union_left_iff ↔ _ is_metric_separated.finite_Union_left |
| 91 | + |
| 92 | +lemma finite_Union_right_iff {ι : Type*} {I : set ι} (hI : finite I) {s : set X} {t : ι → set X} : |
| 93 | + is_metric_separated s (⋃ i ∈ I, t i) ↔ ∀ i ∈ I, is_metric_separated s (t i) := |
| 94 | +by simpa only [@comm _ _ s] using finite_Union_left_iff hI |
| 95 | + |
| 96 | +@[simp] lemma finset_Union_left_iff {ι : Type*} {I : finset ι} {s : ι → set X} {t : set X} : |
| 97 | + is_metric_separated (⋃ i ∈ I, s i) t ↔ ∀ i ∈ I, is_metric_separated (s i) t := |
| 98 | +finite_Union_left_iff I.finite_to_set |
| 99 | + |
| 100 | +alias finset_Union_left_iff ↔ _ is_metric_separated.finset_Union_left |
| 101 | + |
| 102 | +@[simp] lemma finset_Union_right_iff {ι : Type*} {I : finset ι} {s : set X} {t : ι → set X} : |
| 103 | + is_metric_separated s (⋃ i ∈ I, t i) ↔ ∀ i ∈ I, is_metric_separated s (t i) := |
| 104 | +finite_Union_right_iff I.finite_to_set |
| 105 | + |
| 106 | +alias finset_Union_right_iff ↔ _ is_metric_separated.finset_Union_right |
| 107 | + |
| 108 | +end is_metric_separated |
0 commit comments