Skip to content

Commit 5115198

Browse files
committed
chore: last localEquiv -> partialEquiv renames in lemma names (#9620)
1 parent db84fa8 commit 5115198

File tree

4 files changed

+8
-8
lines changed

4 files changed

+8
-8
lines changed

Mathlib/Geometry/Manifold/ContMDiff/Defs.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -583,7 +583,7 @@ theorem contMDiffOn_iff_target :
583583
∀ y : M',
584584
ContMDiffOn I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) (s ∩ f ⁻¹' (extChartAt I' y).source) := by
585585
simp only [contMDiffOn_iff, ModelWithCorners.source_eq, chartAt_self_eq,
586-
PartialHomeomorph.refl_localEquiv, PartialEquiv.refl_trans, extChartAt,
586+
PartialHomeomorph.refl_partialEquiv, PartialEquiv.refl_trans, extChartAt,
587587
PartialHomeomorph.extend, Set.preimage_univ, Set.inter_univ, and_congr_right_iff]
588588
intro h
589589
constructor

Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -374,9 +374,9 @@ variable (𝕜 E)
374374

375375
/-- In the trivial model with corners, the associated `PartialEquiv` is the identity. -/
376376
@[simp, mfld_simps]
377-
theorem modelWithCornersSelf_localEquiv : 𝓘(𝕜, E).toPartialEquiv = PartialEquiv.refl E :=
377+
theorem modelWithCornersSelf_partialEquiv : 𝓘(𝕜, E).toPartialEquiv = PartialEquiv.refl E :=
378378
rfl
379-
#align model_with_corners_self_local_equiv modelWithCornersSelf_localEquiv
379+
#align model_with_corners_self_local_equiv modelWithCornersSelf_partialEquiv
380380

381381
@[simp, mfld_simps]
382382
theorem modelWithCornersSelf_coe : (𝓘(𝕜, E) : E → E) = id :=

Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -279,7 +279,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where
279279
-/
280280
refine mem_iUnion.2fun _ ↦ .refl 𝕜 F, mem_iUnion.2 ⟨univ, mem_iUnion.2 ⟨isOpen_univ, ?_⟩⟩⟩
281281
refine mem_iUnion.2 ⟨contMDiffOn_const, mem_iUnion.2 ⟨contMDiffOn_const, ?_, ?_⟩⟩
282-
· simp only [FiberwiseLinear.partialHomeomorph, PartialHomeomorph.refl_localEquiv,
282+
· simp only [FiberwiseLinear.partialHomeomorph, PartialHomeomorph.refl_partialEquiv,
283283
PartialEquiv.refl_source, univ_prod_univ]
284284
· exact eqOn_refl id _
285285
locality' := by

Mathlib/Topology/PartialHomeomorph.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -240,10 +240,10 @@ theorem source_preimage_target : e.source ⊆ e ⁻¹' e.target :=
240240
#align local_homeomorph.source_preimage_target PartialHomeomorph.source_preimage_target
241241

242242
@[deprecated toPartialEquiv_injective]
243-
theorem eq_of_localEquiv_eq {e e' : PartialHomeomorph α β}
243+
theorem eq_of_partialEquiv_eq {e e' : PartialHomeomorph α β}
244244
(h : e.toPartialEquiv = e'.toPartialEquiv) : e = e' :=
245245
toPartialEquiv_injective h
246-
#align local_homeomorph.eq_of_local_equiv_eq PartialHomeomorph.eq_of_localEquiv_eq
246+
#align local_homeomorph.eq_of_local_equiv_eq PartialHomeomorph.eq_of_partialEquiv_eq
247247

248248
theorem eventually_left_inverse {x} (hx : x ∈ e.source) :
249249
∀ᶠ y in 𝓝 x, e.symm (e y) = y :=
@@ -773,9 +773,9 @@ protected def refl (α : Type*) [TopologicalSpace α] : PartialHomeomorph α α
773773
#align local_homeomorph.refl PartialHomeomorph.refl
774774

775775
@[simp, mfld_simps]
776-
theorem refl_localEquiv : (PartialHomeomorph.refl α).toPartialEquiv = PartialEquiv.refl α :=
776+
theorem refl_partialEquiv : (PartialHomeomorph.refl α).toPartialEquiv = PartialEquiv.refl α :=
777777
rfl
778-
#align local_homeomorph.refl_local_equiv PartialHomeomorph.refl_localEquiv
778+
#align local_homeomorph.refl_local_equiv PartialHomeomorph.refl_partialEquiv
779779

780780
@[simp, mfld_simps]
781781
theorem refl_symm : (PartialHomeomorph.refl α).symm = PartialHomeomorph.refl α :=

0 commit comments

Comments
 (0)