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

Commit 0b1e039

Browse files
committed
feat(topology/metric_space/isometry): use namespace, add lemmas (#15591)
* Use `namespace isometry`. * Add lemmas like `isometry.preimage_ball`.
1 parent b93a64d commit 0b1e039

File tree

11 files changed

+118
-97
lines changed

11 files changed

+118
-97
lines changed

src/analysis/normed/group/add_torsor.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -74,7 +74,7 @@ by rw [dist_comm, dist_vadd_left]
7474
addition/subtraction of `x : P`. -/
7575
@[simps] def isometric.vadd_const (x : P) : V ≃ᵢ P :=
7676
{ to_equiv := equiv.vadd_const x,
77-
isometry_to_fun := isometry_emetric_iff_metric.2 $ λ _ _, dist_vadd_cancel_right _ _ _ }
77+
isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_right _ _ _ }
7878

7979
section
8080

@@ -83,7 +83,7 @@ variable (P)
8383
/-- Self-isometry of a (semi)normed add torsor given by addition of a constant vector `x`. -/
8484
@[simps] def isometric.const_vadd (x : V) : P ≃ᵢ P :=
8585
{ to_equiv := equiv.const_vadd P x,
86-
isometry_to_fun := isometry_emetric_iff_metric.2 $ λ _ _, dist_vadd_cancel_left _ _ _ }
86+
isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_left _ _ _ }
8787

8888
end
8989

@@ -94,7 +94,7 @@ by rw [dist_eq_norm, vsub_sub_vsub_cancel_left, dist_comm, dist_eq_norm_vsub V]
9494
subtraction from `x : P`. -/
9595
@[simps] def isometric.const_vsub (x : P) : P ≃ᵢ V :=
9696
{ to_equiv := equiv.const_vsub x,
97-
isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_vsub_cancel_left _ _ _ }
97+
isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_vsub_cancel_left _ _ _ }
9898

9999
@[simp] lemma dist_vsub_cancel_right (x y z : P) : dist (x -ᵥ z) (y -ᵥ z) = dist x y :=
100100
(isometric.vadd_const z).symm.dist_eq x y

src/analysis/normed/group/basic.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -378,7 +378,7 @@ namespace isometric
378378

379379
/-- Addition `y ↦ y + x` as an `isometry`. -/
380380
protected def add_right (x : E) : E ≃ᵢ E :=
381-
{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_add_right _ _ _,
381+
{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_right _ _ _,
382382
.. equiv.add_right x }
383383

384384
@[simp] lemma add_right_to_equiv (x : E) :
@@ -394,7 +394,7 @@ ext $ λ y, rfl
394394

395395
/-- Addition `y ↦ x + y` as an `isometry`. -/
396396
protected def add_left (x : E) : E ≃ᵢ E :=
397-
{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ y z, dist_add_left _ _ _,
397+
{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_left _ _ _,
398398
to_equiv := equiv.add_left x }
399399

400400
@[simp] lemma add_left_to_equiv (x : E) :
@@ -410,7 +410,7 @@ variable (E)
410410

411411
/-- Negation `x ↦ -x` as an `isometry`. -/
412412
protected def neg : E ≃ᵢ E :=
413-
{ isometry_to_fun := isometry_emetric_iff_metric.2 $ λ x y, dist_neg_neg _ _,
413+
{ isometry_to_fun := isometry.of_dist_eq $ λ x y, dist_neg_neg _ _,
414414
to_equiv := equiv.neg E }
415415

416416
variable {E}
@@ -511,7 +511,7 @@ end
511511
lemma add_monoid_hom_class.isometry_iff_norm {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F]
512512
(f : 𝓕) : isometry f ↔ ∀ x, ∥f x∥ = ∥x∥ :=
513513
begin
514-
simp only [isometry_emetric_iff_metric, dist_eq_norm, ←map_sub],
514+
simp only [isometry_iff_dist_eq, dist_eq_norm, ←map_sub],
515515
refine ⟨λ h x, _, λ h x y, h _⟩,
516516
simpa using h x 0
517517
end

src/analysis/normed_space/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -411,7 +411,7 @@ variables (𝕜 𝕜')
411411
/-- In a normed algebra, the inclusion of the base field in the extended field is an isometry. -/
412412
lemma algebra_map_isometry [norm_one_class 𝕜'] : isometry (algebra_map 𝕜 𝕜') :=
413413
begin
414-
refine isometry_emetric_iff_metric.2 (λx y, _),
414+
refine isometry.of_dist_eq (λx y, _),
415415
rw [dist_eq_norm, dist_eq_norm, ← ring_hom.map_sub, norm_algebra_map'],
416416
end
417417

src/measure_theory/function/conditional_expectation.lean

Lines changed: 2 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -475,12 +475,8 @@ end
475475

476476
lemma isometry_Lp_meas_subgroup_to_Lp_trim [hp : fact (1 ≤ p)] (hm : m ≤ m0) :
477477
isometry (Lp_meas_subgroup_to_Lp_trim F p μ hm) :=
478-
begin
479-
rw isometry_emetric_iff_metric,
480-
intros f g,
481-
rw [dist_eq_norm, ← Lp_meas_subgroup_to_Lp_trim_sub, Lp_meas_subgroup_to_Lp_trim_norm_map,
482-
dist_eq_norm],
483-
end
478+
isometry.of_dist_eq $ λ f g, by rw [dist_eq_norm, ← Lp_meas_subgroup_to_Lp_trim_sub,
479+
Lp_meas_subgroup_to_Lp_trim_norm_map, dist_eq_norm]
484480

485481
variables (F p μ)
486482
/-- `Lp_meas_subgroup` and `Lp F p (μ.trim hm)` are isometric. -/

src/topology/continuous_function/bounded.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -392,7 +392,7 @@ end
392392

393393
lemma isometry_extend (f : α ↪ δ) (h : δ →ᵇ β) :
394394
isometry (λ g : α →ᵇ β, extend f g h) :=
395-
isometry_emetric_iff_metric.2 $ λ g₁ g₂, by simp [dist_nonneg]
395+
isometry.of_dist_eq $ λ g₁ g₂, by simp [dist_nonneg]
396396

397397
end extend
398398

src/topology/metric_space/completion.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -173,7 +173,7 @@ instance : metric_space (completion α) :=
173173

174174
/-- The embedding of a metric space in its completion is an isometry. -/
175175
lemma coe_isometry : isometry (coe : α → completion α) :=
176-
isometry_emetric_iff_metric.2 completion.dist_eq
176+
isometry.of_dist_eq completion.dist_eq
177177

178178
@[simp] protected lemma edist_eq (x y : α) : edist (x : completion α) y = edist x y :=
179179
coe_isometry x y

src/topology/metric_space/gluing.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -338,11 +338,11 @@ lemma sum.dist_eq {x y : X ⊕ Y} : dist x y = sum.dist x y := rfl
338338

339339
/-- The left injection of a space in a disjoint union is an isometry -/
340340
lemma isometry_inl : isometry (sum.inl : X → (X ⊕ Y)) :=
341-
isometry_emetric_iff_metric.2 $ λx y, rfl
341+
isometry.of_dist_eq $ λ x y, rfl
342342

343343
/-- The right injection of a space in a disjoint union is an isometry -/
344344
lemma isometry_inr : isometry (sum.inr : Y → (X ⊕ Y)) :=
345-
isometry_emetric_iff_metric.2 $ λx y, rfl
345+
isometry.of_dist_eq $ λ x y, rfl
346346

347347
end sum
348348

@@ -487,7 +487,7 @@ open filter
487487

488488
/-- The injection of a space in a disjoint union is an isometry -/
489489
lemma isometry_mk (i : ι) : isometry (sigma.mk i : E i → Σ k, E k) :=
490-
isometry_emetric_iff_metric.2 (by simp)
490+
isometry.of_dist_eq (λ x y, by simp)
491491

492492
/-- A disjoint union of complete metric spaces is complete. -/
493493
protected lemma complete_space [∀ i, complete_space (E i)] : complete_space (Σ i, E i) :=
@@ -559,10 +559,10 @@ begin
559559
end
560560

561561
lemma to_glue_l_isometry (hΦ : isometry Φ) (hΨ : isometry Ψ) : isometry (to_glue_l hΦ hΨ) :=
562-
isometry_emetric_iff_metric.2 $ λ_ _, rfl
562+
isometry.of_dist_eq $ λ_ _, rfl
563563

564564
lemma to_glue_r_isometry (hΦ : isometry Φ) (hΨ : isometry Ψ) : isometry (to_glue_r hΦ hΨ) :=
565-
isometry_emetric_iff_metric.2 $ λ_ _, rfl
565+
isometry.of_dist_eq $ λ_ _, rfl
566566

567567
end gluing --section
568568

@@ -656,7 +656,7 @@ instance (I : ∀ n, isometry (f n)) [inhabited (X 0)] : inhabited (inductive_li
656656

657657
/-- The map `to_inductive_limit n` mapping `X n` to the inductive limit is an isometry. -/
658658
lemma to_inductive_limit_isometry (I : ∀ n, isometry (f n)) (n : ℕ) :
659-
isometry (to_inductive_limit I n) := isometry_emetric_iff_metric.2 $ λx y,
659+
isometry (to_inductive_limit I n) := isometry.of_dist_eq $ λ x y,
660660
begin
661661
change inductive_limit_dist f ⟨n, x⟩ ⟨n, y⟩ = dist x y,
662662
rw [inductive_limit_dist_eq_dist I ⟨n, x⟩ ⟨n, y⟩ n (le_refl n) (le_refl n),

src/topology/metric_space/gromov_hausdorff.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -534,8 +534,8 @@ begin
534534
glue_metric_approx (λ x:s, (x:X)) (λ x, Φ x) (ε₂/2 + δ) (by linarith) this,
535535
let Fl := @sum.inl X Y,
536536
let Fr := @sum.inr X Y,
537-
have Il : isometry Fl := isometry_emetric_iff_metric.2 (λ x y, rfl),
538-
have Ir : isometry Fr := isometry_emetric_iff_metric.2 (λ x y, rfl),
537+
have Il : isometry Fl := isometry.of_dist_eq (λ x y, rfl),
538+
have Ir : isometry Fr := isometry.of_dist_eq (λ x y, rfl),
539539
/- The proof goes as follows : the `GH_dist` is bounded by the Hausdorff distance of the images
540540
in the coupling, which is bounded (using the triangular inequality) by the sum of the Hausdorff
541541
distances of `X` and `s` (in the coupling or, equivalently in the original space), of `s` and

src/topology/metric_space/gromov_hausdorff_realized.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -461,7 +461,7 @@ def optimal_GH_injl (x : X) : optimal_GH_coupling X Y := ⟦inl x⟧
461461
/-- The injection of `X` in the optimal coupling between `X` and `Y` is an isometry. -/
462462
lemma isometry_optimal_GH_injl : isometry (optimal_GH_injl X Y) :=
463463
begin
464-
refine isometry_emetric_iff_metric.2 (λx y, _),
464+
refine isometry.of_dist_eq (λx y, _),
465465
change dist ⟦inl x⟧ ⟦inl y⟧ = dist x y,
466466
exact candidates_dist_inl (optimal_GH_dist_mem_candidates_b X Y) _ _,
467467
end
@@ -472,7 +472,7 @@ def optimal_GH_injr (y : Y) : optimal_GH_coupling X Y := ⟦inr y⟧
472472
/-- The injection of `Y` in the optimal coupling between `X` and `Y` is an isometry. -/
473473
lemma isometry_optimal_GH_injr : isometry (optimal_GH_injr X Y) :=
474474
begin
475-
refine isometry_emetric_iff_metric.2 (λx y, _),
475+
refine isometry.of_dist_eq (λx y, _),
476476
change dist ⟦inr x⟧ ⟦inr y⟧ = dist x y,
477477
exact candidates_dist_inr (optimal_GH_dist_mem_candidates_b X Y) _ _,
478478
end

0 commit comments

Comments
 (0)