Skip to content

Commit

Permalink
chore: split Mathlib.Geometry.Euclidean.Inversion (#5868)
Browse files Browse the repository at this point in the history
Move theorems about image of a sphere passing through the center to a new file.
  • Loading branch information
urkud committed Jul 15, 2023
1 parent 2f7b995 commit 3acb5b8
Show file tree
Hide file tree
Showing 4 changed files with 89 additions and 55 deletions.
3 changes: 2 additions & 1 deletion Mathlib.lean
Expand Up @@ -1879,7 +1879,8 @@ import Mathlib.Geometry.Euclidean.Angle.Unoriented.Conformal
import Mathlib.Geometry.Euclidean.Angle.Unoriented.RightAngle
import Mathlib.Geometry.Euclidean.Basic
import Mathlib.Geometry.Euclidean.Circumcenter
import Mathlib.Geometry.Euclidean.Inversion
import Mathlib.Geometry.Euclidean.Inversion.Basic
import Mathlib.Geometry.Euclidean.Inversion.ImageHyperplane
import Mathlib.Geometry.Euclidean.MongePoint
import Mathlib.Geometry.Euclidean.PerpBisector
import Mathlib.Geometry.Euclidean.Sphere.Basic
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean
Expand Up @@ -10,7 +10,7 @@ Authors: Yury Kudryashov
-/
import Mathlib.Analysis.Complex.UpperHalfPlane.Topology
import Mathlib.Analysis.SpecialFunctions.Arsinh
import Mathlib.Geometry.Euclidean.Inversion
import Mathlib.Geometry.Euclidean.Inversion.Basic

/-!
# Metric on the upper half-plane
Expand Down
Expand Up @@ -9,7 +9,6 @@ Authors: Yury G. Kudryashov
! if you have ported upstream changes.
-/
import Mathlib.Analysis.InnerProductSpace.Basic
import Mathlib.Geometry.Euclidean.PerpBisector

/-!
# Inversion in an affine space
Expand Down Expand Up @@ -200,58 +199,6 @@ theorem mul_dist_le_mul_dist_add_mul_dist (a b c d : P) :
convert H using 1 <;> (field_simp [hb.ne', hc.ne', hd.ne', dist_comm a]; ring)
#align euclidean_geometry.mul_dist_le_mul_dist_add_mul_dist EuclideanGeometry.mul_dist_le_mul_dist_add_mul_dist

/-!
### Images of spheres and hyperplanes
In this section we prove that the inversion with center `c` and radius `R ≠ 0` maps a sphere passing
through the center to a hyperplane. More precisely, it maps a sphere with center `y ≠ c` and radius
`dist y c` to the hyperplane `AffineSubspace.perpBisector c (EuclideanGeometry.inversion c R y)`.
-/

/-- The inversion with center `c` and radius `R` maps a sphere passing through the center to a
hyperplane. -/
theorem inversion_mem_perpBisector_inversion_iff (hR : R ≠ 0) (hx : x ≠ c) (hy : y ≠ c) :
inversion c R x ∈ perpBisector c (inversion c R y) ↔ dist x y = dist y c := by
rw [mem_perpBisector_iff_dist_eq, dist_inversion_inversion hx hy, dist_inversion_center]
have hx' := dist_ne_zero.2 hx
have hy' := dist_ne_zero.2 hy
field_simp [mul_assoc, mul_comm, hx, hx.symm, eq_comm]

/-- The inversion with center `c` and radius `R` maps a sphere passing through the center to a
hyperplane. -/
theorem inversion_mem_perpBisector_inversion_iff' (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R x ∈ perpBisector c (inversion c R y) ↔ dist x y = dist y c ∧ x ≠ c := by
rcases eq_or_ne x c with rfl | hx
· simp [*]
· simp [inversion_mem_perpBisector_inversion_iff hR hx hy, hx]

theorem preimage_inversion_perpBisector_inversion (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' perpBisector c (inversion c R y) = sphere y (dist y c) \ {c} :=
Set.ext fun _ ↦ inversion_mem_perpBisector_inversion_iff' hR hy

theorem preimage_inversion_perpBisector (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' perpBisector c y = sphere (inversion c R y) (R ^ 2 / dist y c) \ {c} := by
rw [← dist_inversion_center, ← preimage_inversion_perpBisector_inversion hR,
inversion_inversion] <;> simp [*]

theorem image_inversion_perpBisector (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R '' perpBisector c y = sphere (inversion c R y) (R ^ 2 / dist y c) \ {c} := by
rw [image_eq_preimage_of_inverse (inversion_involutive _ hR) (inversion_involutive _ hR),
preimage_inversion_perpBisector hR hy]

theorem preimage_inversion_sphere_dist_center (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' sphere y (dist y c) =
insert c (perpBisector c (inversion c R y) : Set P) := by
ext x
rcases eq_or_ne x c with rfl | hx; · simp [dist_comm]
rw [mem_preimage, mem_sphere, ← inversion_mem_perpBisector_inversion_iff hR] <;> simp [*]

theorem image_inversion_sphere_dist_center (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R '' sphere y (dist y c) = insert c (perpBisector c (inversion c R y) : Set P) := by
rw [image_eq_preimage_of_inverse (inversion_involutive _ hR) (inversion_involutive _ hR),
preimage_inversion_sphere_dist_center hR hy]

end EuclideanGeometry

open EuclideanGeometry
Expand Down
86 changes: 86 additions & 0 deletions Mathlib/Geometry/Euclidean/Inversion/ImageHyperplane.lean
@@ -0,0 +1,86 @@
/-
Copyright (c) 2023 Yury Kudryashov. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yury Kudryashov
-/
import Mathlib.Geometry.Euclidean.Inversion.Basic
import Mathlib.Geometry.Euclidean.PerpBisector

/-!
# Image of a hyperplane under inversion
In this file we prove that the inversion with center `c` and radius `R ≠ 0` maps a sphere passing
through the center to a hyperplane, and vice versa. More precisely, it maps a sphere with center
`y ≠ c` and radius `dist y c` to the hyperplane
`AffineSubspace.perpBisector c (EuclideanGeometry.inversion c R y)`.
The exact statements are a little more complicated because `EuclideanGeometry.inversion c R` sends
the center to itself, not to a point at infinity.
We also prove that the inversion sends an affine subspace passing through the center to itself.
## Keywords
inversion
-/

open Metric Function AffineMap Set AffineSubspace
open scoped Topology

variable {V P : Type _} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P]
[NormedAddTorsor V P] {c x y : P} {R : ℝ}

namespace EuclideanGeometry

/-- The inversion with center `c` and radius `R` maps a sphere passing through the center to a
hyperplane. -/
theorem inversion_mem_perpBisector_inversion_iff (hR : R ≠ 0) (hx : x ≠ c) (hy : y ≠ c) :
inversion c R x ∈ perpBisector c (inversion c R y) ↔ dist x y = dist y c := by
rw [mem_perpBisector_iff_dist_eq, dist_inversion_inversion hx hy, dist_inversion_center]
have hx' := dist_ne_zero.2 hx
have hy' := dist_ne_zero.2 hy
field_simp [mul_assoc, mul_comm, hx, hx.symm, eq_comm]

/-- The inversion with center `c` and radius `R` maps a sphere passing through the center to a
hyperplane. -/
theorem inversion_mem_perpBisector_inversion_iff' (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R x ∈ perpBisector c (inversion c R y) ↔ dist x y = dist y c ∧ x ≠ c := by
rcases eq_or_ne x c with rfl | hx
· simp [*]
· simp [inversion_mem_perpBisector_inversion_iff hR hx hy, hx]

theorem preimage_inversion_perpBisector_inversion (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' perpBisector c (inversion c R y) = sphere y (dist y c) \ {c} :=
Set.ext fun _ ↦ inversion_mem_perpBisector_inversion_iff' hR hy

theorem preimage_inversion_perpBisector (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' perpBisector c y = sphere (inversion c R y) (R ^ 2 / dist y c) \ {c} := by
rw [← dist_inversion_center, ← preimage_inversion_perpBisector_inversion hR,
inversion_inversion] <;> simp [*]

theorem image_inversion_perpBisector (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R '' perpBisector c y = sphere (inversion c R y) (R ^ 2 / dist y c) \ {c} := by
rw [image_eq_preimage_of_inverse (inversion_involutive _ hR) (inversion_involutive _ hR),
preimage_inversion_perpBisector hR hy]

theorem preimage_inversion_sphere_dist_center (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R ⁻¹' sphere y (dist y c) =
insert c (perpBisector c (inversion c R y) : Set P) := by
ext x
rcases eq_or_ne x c with rfl | hx; · simp [dist_comm]
rw [mem_preimage, mem_sphere, ← inversion_mem_perpBisector_inversion_iff hR] <;> simp [*]

theorem image_inversion_sphere_dist_center (hR : R ≠ 0) (hy : y ≠ c) :
inversion c R '' sphere y (dist y c) = insert c (perpBisector c (inversion c R y) : Set P) := by
rw [image_eq_preimage_of_inverse (inversion_involutive _ hR) (inversion_involutive _ hR),
preimage_inversion_sphere_dist_center hR hy]

/-- Inversion sends an affine subspace passing through the center to itself. -/
theorem mapsTo_inversion_affineSubspace_of_mem {p : AffineSubspace ℝ P} (hp : c ∈ p) :
MapsTo (inversion c R) p p := fun _ ↦ AffineMap.lineMap_mem _ hp

/-- Inversion sends an affine subspace passing through the center to itself. -/
theorem image_inversion_affineSubspace_of_mem {p : AffineSubspace ℝ P} (hR : R ≠ 0) (hp : c ∈ p) :
inversion c R '' p = p :=
(mapsTo_inversion_affineSubspace_of_mem hp).image_subset.antisymm <| fun x hx ↦
⟨inversion c R x, mapsTo_inversion_affineSubspace_of_mem hp hx, inversion_inversion _ hR _⟩

0 comments on commit 3acb5b8

Please sign in to comment.