|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Joseph Myers. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Joseph Myers |
| 5 | +-/ |
| 6 | +import analysis.convex.strict_convex_between |
| 7 | +import geometry.euclidean.basic |
| 8 | + |
| 9 | +/-! |
| 10 | +# Spheres |
| 11 | +
|
| 12 | +This file defines and proves basic results about spheres and cospherical sets of points in |
| 13 | +Euclidean affine spaces. |
| 14 | +
|
| 15 | +## Main definitions |
| 16 | +
|
| 17 | +* `euclidean_geometry.sphere` bundles a `center` and a `radius`. |
| 18 | +
|
| 19 | +* `euclidean_geometry.cospherical` is the property of a set of points being equidistant from some |
| 20 | + point. |
| 21 | +
|
| 22 | +* `euclidean_geometry.concyclic` is the property of a set of points being cospherical and |
| 23 | + coplanar. |
| 24 | +
|
| 25 | +-/ |
| 26 | + |
| 27 | +noncomputable theory |
| 28 | +open_locale real_inner_product_space |
| 29 | + |
| 30 | +namespace euclidean_geometry |
| 31 | + |
| 32 | +variables {V : Type*} (P : Type*) [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] |
| 33 | + |
| 34 | +open finite_dimensional |
| 35 | + |
| 36 | +/-- A `sphere P` bundles a `center` and `radius`. This definition does not require the radius to |
| 37 | +be positive; that should be given as a hypothesis to lemmas that require it. -/ |
| 38 | +@[ext] structure sphere := |
| 39 | +(center : P) |
| 40 | +(radius : ℝ) |
| 41 | + |
| 42 | +variables {P} |
| 43 | + |
| 44 | +instance [nonempty P] : nonempty (sphere P) := ⟨⟨classical.arbitrary P, 0⟩⟩ |
| 45 | + |
| 46 | +instance : has_coe (sphere P) (set P) := ⟨λ s, metric.sphere s.center s.radius⟩ |
| 47 | +instance : has_mem P (sphere P) := ⟨λ p s, p ∈ (s : set P)⟩ |
| 48 | + |
| 49 | +lemma sphere.mk_center (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).center = c := rfl |
| 50 | + |
| 51 | +lemma sphere.mk_radius (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).radius = r := rfl |
| 52 | + |
| 53 | +@[simp] lemma sphere.mk_center_radius (s : sphere P) : (⟨s.center, s.radius⟩ : sphere P) = s := |
| 54 | +by ext; refl |
| 55 | + |
| 56 | +lemma sphere.coe_def (s : sphere P) : (s : set P) = metric.sphere s.center s.radius := rfl |
| 57 | + |
| 58 | +@[simp] lemma sphere.coe_mk (c : P) (r : ℝ) : ↑(⟨c, r⟩ : sphere P) = metric.sphere c r := rfl |
| 59 | + |
| 60 | +@[simp] lemma sphere.mem_coe {p : P} {s : sphere P} : p ∈ (s : set P) ↔ p ∈ s := iff.rfl |
| 61 | + |
| 62 | +lemma mem_sphere {p : P} {s : sphere P} : p ∈ s ↔ dist p s.center = s.radius := iff.rfl |
| 63 | + |
| 64 | +lemma mem_sphere' {p : P} {s : sphere P} : p ∈ s ↔ dist s.center p = s.radius := |
| 65 | +metric.mem_sphere' |
| 66 | + |
| 67 | +lemma subset_sphere {ps : set P} {s : sphere P} : ps ⊆ s ↔ ∀ p ∈ ps, p ∈ s := iff.rfl |
| 68 | + |
| 69 | +lemma dist_of_mem_subset_sphere {p : P} {ps : set P} {s : sphere P} (hp : p ∈ ps) |
| 70 | + (hps : ps ⊆ (s : set P)) : dist p s.center = s.radius := |
| 71 | +mem_sphere.1 (sphere.mem_coe.1 (set.mem_of_mem_of_subset hp hps)) |
| 72 | + |
| 73 | +lemma dist_of_mem_subset_mk_sphere {p c : P} {ps : set P} {r : ℝ} (hp : p ∈ ps) |
| 74 | + (hps : ps ⊆ ↑(⟨c, r⟩ : sphere P)) : dist p c = r := |
| 75 | +dist_of_mem_subset_sphere hp hps |
| 76 | + |
| 77 | +lemma sphere.ne_iff {s₁ s₂ : sphere P} : |
| 78 | + s₁ ≠ s₂ ↔ s₁.center ≠ s₂.center ∨ s₁.radius ≠ s₂.radius := |
| 79 | +by rw [←not_and_distrib, ←sphere.ext_iff] |
| 80 | + |
| 81 | +lemma sphere.center_eq_iff_eq_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : |
| 82 | + s₁.center = s₂.center ↔ s₁ = s₂ := |
| 83 | +begin |
| 84 | + refine ⟨λ h, sphere.ext _ _ h _, λ h, h ▸ rfl⟩, |
| 85 | + rw mem_sphere at hs₁ hs₂, |
| 86 | + rw [←hs₁, ←hs₂, h] |
| 87 | +end |
| 88 | + |
| 89 | +lemma sphere.center_ne_iff_ne_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : |
| 90 | + s₁.center ≠ s₂.center ↔ s₁ ≠ s₂ := |
| 91 | +(sphere.center_eq_iff_eq_of_mem hs₁ hs₂).not |
| 92 | + |
| 93 | +lemma dist_center_eq_dist_center_of_mem_sphere {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) |
| 94 | + (hp₂ : p₂ ∈ s) : dist p₁ s.center = dist p₂ s.center := |
| 95 | +by rw [mem_sphere.1 hp₁, mem_sphere.1 hp₂] |
| 96 | + |
| 97 | +lemma dist_center_eq_dist_center_of_mem_sphere' {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) |
| 98 | + (hp₂ : p₂ ∈ s) : dist s.center p₁ = dist s.center p₂ := |
| 99 | +by rw [mem_sphere'.1 hp₁, mem_sphere'.1 hp₂] |
| 100 | + |
| 101 | +/-- A set of points is cospherical if they are equidistant from some |
| 102 | +point. In two dimensions, this is the same thing as being |
| 103 | +concyclic. -/ |
| 104 | +def cospherical (ps : set P) : Prop := |
| 105 | +∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius |
| 106 | + |
| 107 | +/-- The definition of `cospherical`. -/ |
| 108 | +lemma cospherical_def (ps : set P) : |
| 109 | + cospherical ps ↔ ∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius := |
| 110 | +iff.rfl |
| 111 | + |
| 112 | +/-- A set of points is cospherical if and only if they lie in some sphere. -/ |
| 113 | +lemma cospherical_iff_exists_sphere {ps : set P} : |
| 114 | + cospherical ps ↔ ∃ s : sphere P, ps ⊆ (s : set P) := |
| 115 | +begin |
| 116 | + refine ⟨λ h, _, λ h, _⟩, |
| 117 | + { rcases h with ⟨c, r, h⟩, |
| 118 | + exact ⟨⟨c, r⟩, h⟩ }, |
| 119 | + { rcases h with ⟨s, h⟩, |
| 120 | + exact ⟨s.center, s.radius, h⟩ } |
| 121 | +end |
| 122 | + |
| 123 | +/-- The set of points in a sphere is cospherical. -/ |
| 124 | +lemma sphere.cospherical (s : sphere P) : cospherical (s : set P) := |
| 125 | +cospherical_iff_exists_sphere.2 ⟨s, set.subset.rfl⟩ |
| 126 | + |
| 127 | +/-- A subset of a cospherical set is cospherical. -/ |
| 128 | +lemma cospherical.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : |
| 129 | + cospherical ps₁ := |
| 130 | +begin |
| 131 | + rcases hc with ⟨c, r, hcr⟩, |
| 132 | + exact ⟨c, r, λ p hp, hcr p (hs hp)⟩ |
| 133 | +end |
| 134 | + |
| 135 | +include V |
| 136 | + |
| 137 | +/-- The empty set is cospherical. -/ |
| 138 | +lemma cospherical_empty : cospherical (∅ : set P) := |
| 139 | +begin |
| 140 | + use add_torsor.nonempty.some, |
| 141 | + simp, |
| 142 | +end |
| 143 | + |
| 144 | +omit V |
| 145 | + |
| 146 | +/-- A single point is cospherical. -/ |
| 147 | +lemma cospherical_singleton (p : P) : cospherical ({p} : set P) := |
| 148 | +begin |
| 149 | + use p, |
| 150 | + simp |
| 151 | +end |
| 152 | + |
| 153 | +include V |
| 154 | + |
| 155 | +/-- Two points are cospherical. -/ |
| 156 | +lemma cospherical_pair (p₁ p₂ : P) : cospherical ({p₁, p₂} : set P) := |
| 157 | +begin |
| 158 | + use [(2⁻¹ : ℝ) • (p₂ -ᵥ p₁) +ᵥ p₁, (2⁻¹ : ℝ) * (dist p₂ p₁)], |
| 159 | + intro p, |
| 160 | + rw [set.mem_insert_iff, set.mem_singleton_iff], |
| 161 | + rintro ⟨_|_⟩, |
| 162 | + { rw [dist_eq_norm_vsub V p₁, vsub_vadd_eq_vsub_sub, vsub_self, zero_sub, norm_neg, norm_smul, |
| 163 | + dist_eq_norm_vsub V p₂], |
| 164 | + simp }, |
| 165 | + { rw [H, dist_eq_norm_vsub V p₂, vsub_vadd_eq_vsub_sub, dist_eq_norm_vsub V p₂], |
| 166 | + conv_lhs { congr, congr, rw ←one_smul ℝ (p₂ -ᵥ p₁ : V) }, |
| 167 | + rw [←sub_smul, norm_smul], |
| 168 | + norm_num } |
| 169 | +end |
| 170 | + |
| 171 | +/-- Any three points in a cospherical set are affinely independent. -/ |
| 172 | +lemma cospherical.affine_independent {s : set P} (hs : cospherical s) {p : fin 3 → P} |
| 173 | + (hps : set.range p ⊆ s) (hpi : function.injective p) : |
| 174 | + affine_independent ℝ p := |
| 175 | +begin |
| 176 | + rw affine_independent_iff_not_collinear, |
| 177 | + intro hc, |
| 178 | + rw collinear_iff_of_mem (set.mem_range_self (0 : fin 3)) at hc, |
| 179 | + rcases hc with ⟨v, hv⟩, |
| 180 | + rw set.forall_range_iff at hv, |
| 181 | + have hv0 : v ≠ 0, |
| 182 | + { intro h, |
| 183 | + have he : p 1 = p 0, by simpa [h] using hv 1, |
| 184 | + exact (dec_trivial : (1 : fin 3) ≠ 0) (hpi he) }, |
| 185 | + rcases hs with ⟨c, r, hs⟩, |
| 186 | + have hs' := λ i, hs (p i) (set.mem_of_mem_of_subset (set.mem_range_self _) hps), |
| 187 | + choose f hf using hv, |
| 188 | + have hsd : ∀ i, dist ((f i • v) +ᵥ p 0) c = r, |
| 189 | + { intro i, |
| 190 | + rw ←hf, |
| 191 | + exact hs' i }, |
| 192 | + have hf0 : f 0 = 0, |
| 193 | + { have hf0' := hf 0, |
| 194 | + rw [eq_comm, ←@vsub_eq_zero_iff_eq V, vadd_vsub, smul_eq_zero] at hf0', |
| 195 | + simpa [hv0] using hf0' }, |
| 196 | + have hfi : function.injective f, |
| 197 | + { intros i j h, |
| 198 | + have hi := hf i, |
| 199 | + rw [h, ←hf j] at hi, |
| 200 | + exact hpi hi }, |
| 201 | + simp_rw [←hsd 0, hf0, zero_smul, zero_vadd, dist_smul_vadd_eq_dist (p 0) c hv0] at hsd, |
| 202 | + have hfn0 : ∀ i, i ≠ 0 → f i ≠ 0 := λ i, (hfi.ne_iff' hf0).2, |
| 203 | + have hfn0' : ∀ i, i ≠ 0 → f i = (-2) * ⟪v, (p 0 -ᵥ c)⟫ / ⟪v, v⟫, |
| 204 | + { intros i hi, |
| 205 | + have hsdi := hsd i, |
| 206 | + simpa [hfn0, hi] using hsdi }, |
| 207 | + have hf12 : f 1 = f 2, { rw [hfn0' 1 dec_trivial, hfn0' 2 dec_trivial] }, |
| 208 | + exact (dec_trivial : (1 : fin 3) ≠ 2) (hfi hf12) |
| 209 | +end |
| 210 | + |
| 211 | +/-- Any three points in a cospherical set are affinely independent. -/ |
| 212 | +lemma cospherical.affine_independent_of_mem_of_ne {s : set P} (hs : cospherical s) {p₁ p₂ p₃ : P} |
| 213 | + (h₁ : p₁ ∈ s) (h₂ : p₂ ∈ s) (h₃ : p₃ ∈ s) (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : |
| 214 | + affine_independent ℝ ![p₁, p₂, p₃] := |
| 215 | +begin |
| 216 | + refine hs.affine_independent _ _, |
| 217 | + { simp [h₁, h₂, h₃, set.insert_subset] }, |
| 218 | + { erw [fin.cons_injective_iff, fin.cons_injective_iff], |
| 219 | + simp [h₁₂, h₁₃, h₂₃, function.injective] } |
| 220 | +end |
| 221 | + |
| 222 | +/-- The three points of a cospherical set are affinely independent. -/ |
| 223 | +lemma cospherical.affine_independent_of_ne {p₁ p₂ p₃ : P} (hs : cospherical ({p₁, p₂, p₃} : set P)) |
| 224 | + (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : |
| 225 | + affine_independent ℝ ![p₁, p₂, p₃] := |
| 226 | +hs.affine_independent_of_mem_of_ne (set.mem_insert _ _) |
| 227 | + (set.mem_insert_of_mem _ (set.mem_insert _ _)) |
| 228 | + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_singleton _))) h₁₂ h₁₃ h₂₃ |
| 229 | + |
| 230 | +/-- Suppose that `p₁` and `p₂` lie in spheres `s₁` and `s₂`. Then the vector between the centers |
| 231 | +of those spheres is orthogonal to that between `p₁` and `p₂`; this is a version of |
| 232 | +`inner_vsub_vsub_of_dist_eq_of_dist_eq` for bundled spheres. (In two dimensions, this says that |
| 233 | +the diagonals of a kite are orthogonal.) -/ |
| 234 | +lemma inner_vsub_vsub_of_mem_sphere_of_mem_sphere {p₁ p₂ : P} {s₁ s₂ : sphere P} |
| 235 | + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) : |
| 236 | + ⟪s₂.center -ᵥ s₁.center, p₂ -ᵥ p₁⟫ = 0 := |
| 237 | +inner_vsub_vsub_of_dist_eq_of_dist_eq (dist_center_eq_dist_center_of_mem_sphere hp₁s₁ hp₂s₁) |
| 238 | + (dist_center_eq_dist_center_of_mem_sphere hp₁s₂ hp₂s₂) |
| 239 | + |
| 240 | +/-- Two spheres intersect in at most two points in a two-dimensional subspace containing their |
| 241 | +centers; this is a version of `eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two` for bundled |
| 242 | +spheres. -/ |
| 243 | +lemma eq_of_mem_sphere_of_mem_sphere_of_mem_of_finrank_eq_two {s : affine_subspace ℝ P} |
| 244 | + [finite_dimensional ℝ s.direction] (hd : finrank ℝ s.direction = 2) {s₁ s₂ : sphere P} |
| 245 | + {p₁ p₂ p : P} (hs₁ : s₁.center ∈ s) (hs₂ : s₂.center ∈ s) (hp₁s : p₁ ∈ s) (hp₂s : p₂ ∈ s) |
| 246 | + (hps : p ∈ s) (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) |
| 247 | + (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := |
| 248 | +eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two hd hs₁ hs₂ hp₁s hp₂s hps |
| 249 | + ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ |
| 250 | + |
| 251 | +/-- Two spheres intersect in at most two points in two-dimensional space; this is a version of |
| 252 | +`eq_of_dist_eq_of_dist_eq_of_finrank_eq_two` for bundled spheres. -/ |
| 253 | +lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V] |
| 254 | + (hd : finrank ℝ V = 2) {s₁ s₂ : sphere P} {p₁ p₂ p : P} (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) |
| 255 | + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) |
| 256 | + (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := |
| 257 | +eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) |
| 258 | + hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ |
| 259 | + |
| 260 | +/-- Given a point on a sphere and a point not outside it, the inner product between the |
| 261 | +difference of those points and the radius vector is positive unless the points are equal. -/ |
| 262 | +lemma inner_pos_or_eq_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) |
| 263 | + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ ∨ p₁ = p₂ := |
| 264 | +begin |
| 265 | + by_cases h : p₁ = p₂, { exact or.inr h }, |
| 266 | + refine or.inl _, |
| 267 | + rw mem_sphere at hp₁, |
| 268 | + rw [←vsub_sub_vsub_cancel_right p₁ p₂ s.center, inner_sub_left, |
| 269 | + real_inner_self_eq_norm_mul_norm/-, ←dist_eq_norm_vsub, hp₁-/, sub_pos], |
| 270 | + refine lt_of_le_of_ne |
| 271 | + ((real_inner_le_norm _ _).trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _))) |
| 272 | + _, |
| 273 | + { rwa [←dist_eq_norm_vsub, ←dist_eq_norm_vsub, hp₁] }, |
| 274 | + { rcases hp₂.lt_or_eq with hp₂' | hp₂', |
| 275 | + { refine ((real_inner_le_norm _ _).trans_lt (mul_lt_mul_of_pos_right _ _)).ne, |
| 276 | + { rwa [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂' }, |
| 277 | + { rw [norm_pos_iff, vsub_ne_zero], |
| 278 | + rintro rfl, |
| 279 | + rw ←hp₁ at hp₂', |
| 280 | + refine (dist_nonneg.not_lt : ¬dist p₂ s.center < 0) _, |
| 281 | + simpa using hp₂' } }, |
| 282 | + { rw [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂', |
| 283 | + nth_rewrite 0 ←hp₂', |
| 284 | + rw [ne.def, inner_eq_norm_mul_iff_real, hp₂', ←sub_eq_zero, ←smul_sub, |
| 285 | + vsub_sub_vsub_cancel_right, ←ne.def, smul_ne_zero_iff, vsub_ne_zero, |
| 286 | + and_iff_left (ne.symm h), norm_ne_zero_iff, vsub_ne_zero], |
| 287 | + rintro rfl, |
| 288 | + refine h (eq.symm _), |
| 289 | + simpa using hp₂' } } |
| 290 | +end |
| 291 | + |
| 292 | +/-- Given a point on a sphere and a point not outside it, the inner product between the |
| 293 | +difference of those points and the radius vector is nonnegative. -/ |
| 294 | +lemma inner_nonneg_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) |
| 295 | + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 ≤ ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := |
| 296 | +begin |
| 297 | + rcases inner_pos_or_eq_of_dist_le_radius hp₁ hp₂ with h | rfl, |
| 298 | + { exact h.le }, |
| 299 | + { simp } |
| 300 | +end |
| 301 | + |
| 302 | +/-- Given a point on a sphere and a point inside it, the inner product between the difference of |
| 303 | +those points and the radius vector is positive. -/ |
| 304 | +lemma inner_pos_of_dist_lt_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) |
| 305 | + (hp₂ : dist p₂ s.center < s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := |
| 306 | +begin |
| 307 | + by_cases h : p₁ = p₂, |
| 308 | + { rw [h, mem_sphere] at hp₁, |
| 309 | + exact false.elim (hp₂.ne hp₁) }, |
| 310 | + exact (inner_pos_or_eq_of_dist_le_radius hp₁ hp₂.le).resolve_right h |
| 311 | +end |
| 312 | + |
| 313 | +/-- Given three collinear points, two on a sphere and one not outside it, the one not outside it |
| 314 | +is weakly between the other two points. -/ |
| 315 | +lemma wbtw_of_collinear_of_dist_center_le_radius {s : sphere P} {p₁ p₂ p₃ : P} |
| 316 | + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center ≤ s.radius) |
| 317 | + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : wbtw ℝ p₁ p₂ p₃ := |
| 318 | +h.wbtw_of_dist_eq_of_dist_le hp₁ hp₂ hp₃ hp₁p₃ |
| 319 | + |
| 320 | +/-- Given three collinear points, two on a sphere and one inside it, the one inside it is |
| 321 | +strictly between the other two points. -/ |
| 322 | +lemma sbtw_of_collinear_of_dist_center_lt_radius {s : sphere P} {p₁ p₂ p₃ : P} |
| 323 | + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center < s.radius) |
| 324 | + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : sbtw ℝ p₁ p₂ p₃ := |
| 325 | +h.sbtw_of_dist_eq_of_dist_lt hp₁ hp₂ hp₃ hp₁p₃ |
| 326 | + |
| 327 | +/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated |
| 328 | +directly in terms of `cospherical` instead of using `concyclic`.) -/ |
| 329 | +structure concyclic (ps : set P) : Prop := |
| 330 | +(cospherical : cospherical ps) |
| 331 | +(coplanar : coplanar ℝ ps) |
| 332 | + |
| 333 | +/-- A subset of a concyclic set is concyclic. -/ |
| 334 | +lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := |
| 335 | +⟨h.1.subset hs, h.2.subset hs⟩ |
| 336 | + |
| 337 | +/-- The empty set is concyclic. -/ |
| 338 | +lemma concyclic_empty : concyclic (∅ : set P) := |
| 339 | +⟨cospherical_empty, coplanar_empty ℝ P⟩ |
| 340 | + |
| 341 | +/-- A single point is concyclic. -/ |
| 342 | +lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := |
| 343 | +⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ |
| 344 | + |
| 345 | +/-- Two points are concyclic. -/ |
| 346 | +lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := |
| 347 | +⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ |
| 348 | + |
| 349 | +end euclidean_geometry |
0 commit comments