From bbc6e56d76341e236fa4e15748bebff612b5cb4f Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 11 Dec 2023 20:54:43 +0000 Subject: [PATCH] chore: rename LocalHomeomorph to PartialHomeomorph (#8982) `LocalHomeomorph` evokes a "local homeomorphism": this is *not* what this means. Instead, this is a homeomorphism on an open set of the domain (extended to the whole space, by the junk value pattern). Hence, partial homeomorphism is more appropriate, and avoids confusion with `IsLocallyHomeomorph`. A future PR will rename `LocalEquiv` to `PartialEquiv`. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Naming.20around.20local.20homeomorphisms/near/407090631) --- Mathlib.lean | 2 +- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 28 +- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 22 +- Mathlib/Analysis/Calculus/Deriv/Inverse.lean | 8 +- Mathlib/Analysis/Calculus/FDeriv/Equiv.lean | 10 +- Mathlib/Analysis/Calculus/Implicit.lean | 154 ++-- .../ApproximatesLinearOn.lean | 42 +- .../InverseFunctionTheorem/ContDiff.lean | 38 +- .../InverseFunctionTheorem/FDeriv.lean | 60 +- .../Analysis/InnerProductSpace/Calculus.lean | 16 +- .../Analysis/NormedSpace/HomeomorphBall.lean | 38 +- .../Analysis/NormedSpace/SphereNormEquiv.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Arsinh.lean | 2 +- .../SpecialFunctions/Complex/LogDeriv.lean | 13 +- .../Analysis/SpecialFunctions/PolarCoord.lean | 6 +- Mathlib/Analysis/SpecialFunctions/Sqrt.lean | 8 +- .../Trigonometric/Arctan.lean | 16 +- .../Trigonometric/ArctanDeriv.lean | 4 +- .../Trigonometric/Inverse.lean | 6 +- .../Trigonometric/InverseDeriv.lean | 4 +- Mathlib/Geometry/Manifold/ChartedSpace.lean | 180 ++-- .../Geometry/Manifold/ContMDiff/Atlas.lean | 14 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 4 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 12 +- .../Geometry/Manifold/ContMDiff/Product.lean | 4 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 2 +- Mathlib/Geometry/Manifold/Diffeomorph.lean | 8 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 4 +- .../Geometry/Manifold/Instances/Sphere.lean | 10 +- .../Manifold/LocalInvariantProperties.lean | 64 +- Mathlib/Geometry/Manifold/MFDeriv.lean | 52 +- .../Geometry/Manifold/PartitionOfUnity.lean | 2 +- .../Manifold/SmoothManifoldWithCorners.lean | 140 +-- .../Geometry/Manifold/VectorBundle/Basic.lean | 26 +- .../VectorBundle/FiberwiseLinear.lean | 16 +- .../Geometry/Manifold/VectorBundle/Hom.lean | 7 +- .../Manifold/VectorBundle/Tangent.lean | 16 +- Mathlib/Logic/Equiv/LocalEquiv.lean | 2 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 2 +- Mathlib/Topology/Covering.lean | 4 +- Mathlib/Topology/FiberBundle/Basic.lean | 8 +- .../Topology/FiberBundle/Constructions.lean | 4 +- .../Topology/FiberBundle/Trivialization.lean | 77 +- Mathlib/Topology/IsLocalHomeomorph.lean | 24 +- ...Homeomorph.lean => PartialHomeomorph.lean} | 828 +++++++++--------- Mathlib/Topology/Sheaves/LocalPredicate.lean | 2 +- Mathlib/Topology/VectorBundle/Basic.lean | 10 +- scripts/nolints.json | 8 +- test/MfldSetTac.lean | 24 +- 49 files changed, 1019 insertions(+), 1014 deletions(-) rename Mathlib/Topology/{LocalHomeomorph.lean => PartialHomeomorph.lean} (57%) diff --git a/Mathlib.lean b/Mathlib.lean index 2ae177da1d1b4..603b2eb62f174 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3585,7 +3585,6 @@ import Mathlib.Topology.IsLocalHomeomorph import Mathlib.Topology.List import Mathlib.Topology.LocalAtTarget import Mathlib.Topology.LocalExtr -import Mathlib.Topology.LocalHomeomorph import Mathlib.Topology.LocallyConstant.Algebra import Mathlib.Topology.LocallyConstant.Basic import Mathlib.Topology.LocallyFinite @@ -3643,6 +3642,7 @@ import Mathlib.Topology.Order.Priestley import Mathlib.Topology.Order.ScottTopology import Mathlib.Topology.Order.UpperLowerSetTopology import Mathlib.Topology.Partial +import Mathlib.Topology.PartialHomeomorph import Mathlib.Topology.PartitionOfUnity import Mathlib.Topology.Perfect import Mathlib.Topology.ProperMap diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 4a4494c843e03..4a6ecf09d8d13 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Yury Kudryashov import Mathlib.Analysis.Normed.Group.InfiniteSum import Mathlib.Analysis.NormedSpace.Basic import Mathlib.Topology.Algebra.Order.LiminfLimsup -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph #align_import analysis.asymptotics.asymptotics from "leanprover-community/mathlib"@"f2ce6086713c78a7f880485f7917ea547a215982" @@ -2233,14 +2233,14 @@ lemma Asymptotics.IsBigO.comp_summable_norm {ι E F : Type*} summable_of_isBigO hg <| hf.norm_norm.comp_tendsto <| tendsto_zero_iff_norm_tendsto_zero.2 hg.tendsto_cofinite_zero -namespace LocalHomeomorph +namespace PartialHomeomorph variable {α : Type*} {β : Type*} [TopologicalSpace α] [TopologicalSpace β] variable {E : Type*} [Norm E] {F : Type*} [Norm F] -/-- Transfer `IsBigOWith` over a `LocalHomeomorph`. -/ -theorem isBigOWith_congr (e : LocalHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} +/-- Transfer `IsBigOWith` over a `PartialHomeomorph`. -/ +theorem isBigOWith_congr (e : PartialHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} {g : β → F} {C : ℝ} : IsBigOWith C (𝓝 b) f g ↔ IsBigOWith C (𝓝 (e.symm b)) (f ∘ e) (g ∘ e) := ⟨fun h => h.comp_tendsto <| by @@ -2251,25 +2251,25 @@ theorem isBigOWith_congr (e : LocalHomeomorph α β) {b : β} (hb : b ∈ e.targ ((e.eventually_right_inverse hb).mono fun x hx => congr_arg f hx) ((e.eventually_right_inverse hb).mono fun x hx => congr_arg g hx)⟩ set_option linter.uppercaseLean3 false in -#align local_homeomorph.is_O_with_congr LocalHomeomorph.isBigOWith_congr +#align local_homeomorph.is_O_with_congr PartialHomeomorph.isBigOWith_congr -/-- Transfer `IsBigO` over a `LocalHomeomorph`. -/ -theorem isBigO_congr (e : LocalHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} {g : β → F} : - f =O[𝓝 b] g ↔ (f ∘ e) =O[𝓝 (e.symm b)] (g ∘ e) := by +/-- Transfer `IsBigO` over a `PartialHomeomorph`. -/ +theorem isBigO_congr (e : PartialHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} + {g : β → F} : f =O[𝓝 b] g ↔ (f ∘ e) =O[𝓝 (e.symm b)] (g ∘ e) := by simp only [IsBigO_def] exact exists_congr fun C => e.isBigOWith_congr hb set_option linter.uppercaseLean3 false in -#align local_homeomorph.is_O_congr LocalHomeomorph.isBigO_congr +#align local_homeomorph.is_O_congr PartialHomeomorph.isBigO_congr -/-- Transfer `IsLittleO` over a `LocalHomeomorph`. -/ -theorem isLittleO_congr (e : LocalHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} +/-- Transfer `IsLittleO` over a `PartialHomeomorph`. -/ +theorem isLittleO_congr (e : PartialHomeomorph α β) {b : β} (hb : b ∈ e.target) {f : β → E} {g : β → F} : f =o[𝓝 b] g ↔ (f ∘ e) =o[𝓝 (e.symm b)] (g ∘ e) := by simp only [IsLittleO_def] exact forall₂_congr fun c _hc => e.isBigOWith_congr hb set_option linter.uppercaseLean3 false in -#align local_homeomorph.is_o_congr LocalHomeomorph.isLittleO_congr +#align local_homeomorph.is_o_congr PartialHomeomorph.isLittleO_congr -end LocalHomeomorph +end PartialHomeomorph namespace Homeomorph @@ -2282,7 +2282,7 @@ open Asymptotics /-- Transfer `IsBigOWith` over a `Homeomorph`. -/ theorem isBigOWith_congr (e : α ≃ₜ β) {b : β} {f : β → E} {g : β → F} {C : ℝ} : IsBigOWith C (𝓝 b) f g ↔ IsBigOWith C (𝓝 (e.symm b)) (f ∘ e) (g ∘ e) := - e.toLocalHomeomorph.isBigOWith_congr trivial + e.toPartialHomeomorph.isBigOWith_congr trivial set_option linter.uppercaseLean3 false in #align homeomorph.is_O_with_congr Homeomorph.isBigOWith_congr diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 4bb803f9be188..c26ff10cf451b 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1810,7 +1810,7 @@ then `f.symm` is `n` times continuously differentiable at the point `a`. This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.contDiffAt_symm [CompleteSpace E] (f : LocalHomeomorph E F) +theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomorph E F) {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : ContDiffAt 𝕜 n f.symm a := by @@ -1855,7 +1855,7 @@ theorem LocalHomeomorph.contDiffAt_symm [CompleteSpace E] (f : LocalHomeomorph E · refine' contDiffAt_top.mpr _ intro n exact Itop n (contDiffAt_top.mp hf n) -#align local_homeomorph.cont_diff_at_symm LocalHomeomorph.contDiffAt_symm +#align local_homeomorph.cont_diff_at_symm PartialHomeomorph.contDiffAt_symm /-- If `f` is an `n` times continuously differentiable homeomorphism, and if the derivative of `f` at each point is a continuous linear equivalence, @@ -1867,7 +1867,7 @@ theorem Homeomorph.contDiff_symm [CompleteSpace E] (f : E ≃ₜ F) {f₀' : E (hf₀' : ∀ a, HasFDerivAt f (f₀' a : E →L[𝕜] F) a) (hf : ContDiff 𝕜 n (f : E → F)) : ContDiff 𝕜 n (f.symm : F → E) := contDiff_iff_contDiffAt.2 fun x => - f.toLocalHomeomorph.contDiffAt_symm (mem_univ x) (hf₀' _) hf.contDiffAt + f.toPartialHomeomorph.contDiffAt_symm (mem_univ x) (hf₀' _) hf.contDiffAt #align homeomorph.cont_diff_symm Homeomorph.contDiff_symm /-- Let `f` be a local homeomorphism of a nontrivially normed field, let `a` be a point in its @@ -1876,11 +1876,11 @@ target. if `f` is `n` times continuously differentiable at `f.symm a`, and if th This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.contDiffAt_symm_deriv [CompleteSpace 𝕜] (f : LocalHomeomorph 𝕜 𝕜) +theorem PartialHomeomorph.contDiffAt_symm_deriv [CompleteSpace 𝕜] (f : PartialHomeomorph 𝕜 𝕜) {f₀' a : 𝕜} (h₀ : f₀' ≠ 0) (ha : a ∈ f.target) (hf₀' : HasDerivAt f f₀' (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : ContDiffAt 𝕜 n f.symm a := f.contDiffAt_symm ha (hf₀'.hasFDerivAt_equiv h₀) hf -#align local_homeomorph.cont_diff_at_symm_deriv LocalHomeomorph.contDiffAt_symm_deriv +#align local_homeomorph.cont_diff_at_symm_deriv PartialHomeomorph.contDiffAt_symm_deriv /-- Let `f` be an `n` times continuously differentiable homeomorphism of a nontrivially normed field. Suppose that the derivative of `f` is never equal to zero. Then `f.symm` is `n` times @@ -1892,10 +1892,10 @@ theorem Homeomorph.contDiff_symm_deriv [CompleteSpace 𝕜] (f : 𝕜 ≃ₜ (h₀ : ∀ x, f' x ≠ 0) (hf' : ∀ x, HasDerivAt f (f' x) x) (hf : ContDiff 𝕜 n (f : 𝕜 → 𝕜)) : ContDiff 𝕜 n (f.symm : 𝕜 → 𝕜) := contDiff_iff_contDiffAt.2 fun x => - f.toLocalHomeomorph.contDiffAt_symm_deriv (h₀ _) (mem_univ x) (hf' _) hf.contDiffAt + f.toPartialHomeomorph.contDiffAt_symm_deriv (h₀ _) (mem_univ x) (hf' _) hf.contDiffAt #align homeomorph.cont_diff_symm_deriv Homeomorph.contDiff_symm_deriv -namespace LocalHomeomorph +namespace PartialHomeomorph variable (𝕜) @@ -1906,7 +1906,7 @@ such that `f` is `C^n` at `x` and `f.symm` is `C^n` at `y`. Note that `n` is a natural number, not `∞`, because the set of points of `C^∞`-smoothness of `f` is not guaranteed to be open. -/ @[simps! apply symm_apply source target] -def restrContDiff (f : LocalHomeomorph E F) (n : ℕ) : LocalHomeomorph E F := +def restrContDiff (f : PartialHomeomorph E F) (n : ℕ) : PartialHomeomorph E F := haveI H : f.IsImage {x | ContDiffAt 𝕜 n f x ∧ ContDiffAt 𝕜 n f.symm (f x)} {y | ContDiffAt 𝕜 n f.symm y ∧ ContDiffAt 𝕜 n f (f.symm y)} := fun x hx ↦ by simp [hx, and_comm] @@ -1914,13 +1914,13 @@ def restrContDiff (f : LocalHomeomorph E F) (n : ℕ) : LocalHomeomorph E F := inter_mem (f.open_source.mem_nhds hxs) <| hxf.eventually.and <| f.continuousAt hxs hxf'.eventually -lemma contDiffOn_restrContDiff_source (f : LocalHomeomorph E F) (n : ℕ) : +lemma contDiffOn_restrContDiff_source (f : PartialHomeomorph E F) (n : ℕ) : ContDiffOn 𝕜 n f (f.restrContDiff 𝕜 n).source := fun _x hx ↦ hx.2.1.contDiffWithinAt -lemma contDiffOn_restrContDiff_target (f : LocalHomeomorph E F) (n : ℕ) : +lemma contDiffOn_restrContDiff_target (f : PartialHomeomorph E F) (n : ℕ) : ContDiffOn 𝕜 n f.symm (f.restrContDiff 𝕜 n).target := fun _x hx ↦ hx.2.1.contDiffWithinAt -end LocalHomeomorph +end PartialHomeomorph end FunctionInverse diff --git a/Mathlib/Analysis/Calculus/Deriv/Inverse.lean b/Mathlib/Analysis/Calculus/Deriv/Inverse.lean index 71ad5b48850ef..aefe0e88c592e 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Inverse.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Inverse.lean @@ -76,11 +76,11 @@ at `a` in the strict sense. This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.hasStrictDerivAt_symm (f : LocalHomeomorph 𝕜 𝕜) {a f' : 𝕜} +theorem PartialHomeomorph.hasStrictDerivAt_symm (f : PartialHomeomorph 𝕜 𝕜) {a f' : 𝕜} (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : HasStrictDerivAt f f' (f.symm a)) : HasStrictDerivAt f.symm f'⁻¹ a := htff'.of_local_left_inverse (f.symm.continuousAt ha) hf' (f.eventually_right_inverse ha) -#align local_homeomorph.has_strict_deriv_at_symm LocalHomeomorph.hasStrictDerivAt_symm +#align local_homeomorph.has_strict_deriv_at_symm PartialHomeomorph.hasStrictDerivAt_symm /-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. @@ -98,10 +98,10 @@ nonzero derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹ This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.hasDerivAt_symm (f : LocalHomeomorph 𝕜 𝕜) {a f' : 𝕜} (ha : a ∈ f.target) +theorem PartialHomeomorph.hasDerivAt_symm (f : PartialHomeomorph 𝕜 𝕜) {a f' : 𝕜} (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : HasDerivAt f f' (f.symm a)) : HasDerivAt f.symm f'⁻¹ a := htff'.of_local_left_inverse (f.symm.continuousAt ha) hf' (f.eventually_right_inverse ha) -#align local_homeomorph.has_deriv_at_symm LocalHomeomorph.hasDerivAt_symm +#align local_homeomorph.has_deriv_at_symm PartialHomeomorph.hasDerivAt_symm theorem HasDerivAt.eventually_ne (h : HasDerivAt f f' x) (hf' : f' ≠ 0) : ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index 14255476d5640..257012e267825 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -410,22 +410,22 @@ the derivative `f'⁻¹` at `a`. This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.hasStrictFDerivAt_symm (f : LocalHomeomorph E F) {f' : E ≃L[𝕜] F} {a : F} - (ha : a ∈ f.target) (htff' : HasStrictFDerivAt f (f' : E →L[𝕜] F) (f.symm a)) : +theorem PartialHomeomorph.hasStrictFDerivAt_symm (f : PartialHomeomorph E F) {f' : E ≃L[𝕜] F} + {a : F} (ha : a ∈ f.target) (htff' : HasStrictFDerivAt f (f' : E →L[𝕜] F) (f.symm a)) : HasStrictFDerivAt f.symm (f'.symm : F →L[𝕜] E) a := htff'.of_local_left_inverse (f.symm.continuousAt ha) (f.eventually_right_inverse ha) -#align local_homeomorph.has_strict_fderiv_at_symm LocalHomeomorph.hasStrictFDerivAt_symm +#align local_homeomorph.has_strict_fderiv_at_symm PartialHomeomorph.hasStrictFDerivAt_symm /-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an invertible derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹` at `a`. This is one of the easy parts of the inverse function theorem: it assumes that we already have an inverse function. -/ -theorem LocalHomeomorph.hasFDerivAt_symm (f : LocalHomeomorph E F) {f' : E ≃L[𝕜] F} {a : F} +theorem PartialHomeomorph.hasFDerivAt_symm (f : PartialHomeomorph E F) {f' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) (htff' : HasFDerivAt f (f' : E →L[𝕜] F) (f.symm a)) : HasFDerivAt f.symm (f'.symm : F →L[𝕜] E) a := htff'.of_local_left_inverse (f.symm.continuousAt ha) (f.eventually_right_inverse ha) -#align local_homeomorph.has_fderiv_at_symm LocalHomeomorph.hasFDerivAt_symm +#align local_homeomorph.has_fderiv_at_symm PartialHomeomorph.hasFDerivAt_symm theorem HasFDerivWithinAt.eventually_ne (h : HasFDerivWithinAt f f' s x) (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : ∀ᶠ z in 𝓝[s \ {x}] x, f z ≠ f x := by diff --git a/Mathlib/Analysis/Calculus/Implicit.lean b/Mathlib/Analysis/Calculus/Implicit.lean index e8a14eb50237d..42f6618ccd1a1 100644 --- a/Mathlib/Analysis/Calculus/Implicit.lean +++ b/Mathlib/Analysis/Calculus/Implicit.lean @@ -143,35 +143,35 @@ protected theorem hasStrictFDerivAt : at `a`, their derivatives `f'`, `g'` are surjective, and the kernels of these derivatives are complementary subspaces of `E`, then `x ↦ (f x, g x)` defines a local homeomorphism between `E` and `F × G`. In particular, `{x | f x = f a}` is locally homeomorphic to `G`. -/ -def toLocalHomeomorph : LocalHomeomorph E (F × G) := - φ.hasStrictFDerivAt.toLocalHomeomorph _ -#align implicit_function_data.to_local_homeomorph ImplicitFunctionData.toLocalHomeomorph +def toPartialHomeomorph : PartialHomeomorph E (F × G) := + φ.hasStrictFDerivAt.toPartialHomeomorph _ +#align implicit_function_data.to_local_homeomorph ImplicitFunctionData.toPartialHomeomorph /-- Implicit function theorem. If `f : E → F` and `g : E → G` are two maps strictly differentiable at `a`, their derivatives `f'`, `g'` are surjective, and the kernels of these derivatives are complementary subspaces of `E`, then `implicitFunction` is the unique (germ of a) map `φ : F → G → E` such that `f (φ y z) = y` and `g (φ y z) = z`. -/ def implicitFunction : F → G → E := - Function.curry <| φ.toLocalHomeomorph.symm + Function.curry <| φ.toPartialHomeomorph.symm #align implicit_function_data.implicit_function ImplicitFunctionData.implicitFunction @[simp] -theorem toLocalHomeomorph_coe : ⇑φ.toLocalHomeomorph = φ.prodFun := +theorem toPartialHomeomorph_coe : ⇑φ.toPartialHomeomorph = φ.prodFun := rfl -#align implicit_function_data.to_local_homeomorph_coe ImplicitFunctionData.toLocalHomeomorph_coe +#align implicit_function_data.to_local_homeomorph_coe ImplicitFunctionData.toPartialHomeomorph_coe -theorem toLocalHomeomorph_apply (x : E) : φ.toLocalHomeomorph x = (φ.leftFun x, φ.rightFun x) := +theorem toPartialHomeomorph_apply (x : E) : φ.toPartialHomeomorph x = (φ.leftFun x, φ.rightFun x) := rfl -#align implicit_function_data.to_local_homeomorph_apply ImplicitFunctionData.toLocalHomeomorph_apply +#align implicit_function_data.to_local_homeomorph_apply ImplicitFunctionData.toPartialHomeomorph_apply -theorem pt_mem_toLocalHomeomorph_source : φ.pt ∈ φ.toLocalHomeomorph.source := - φ.hasStrictFDerivAt.mem_toLocalHomeomorph_source -#align implicit_function_data.pt_mem_to_local_homeomorph_source ImplicitFunctionData.pt_mem_toLocalHomeomorph_source +theorem pt_mem_toPartialHomeomorph_source : φ.pt ∈ φ.toPartialHomeomorph.source := + φ.hasStrictFDerivAt.mem_toPartialHomeomorph_source +#align implicit_function_data.pt_mem_to_local_homeomorph_source ImplicitFunctionData.pt_mem_toPartialHomeomorph_source -theorem map_pt_mem_toLocalHomeomorph_target : - (φ.leftFun φ.pt, φ.rightFun φ.pt) ∈ φ.toLocalHomeomorph.target := - φ.toLocalHomeomorph.map_source <| φ.pt_mem_toLocalHomeomorph_source -#align implicit_function_data.map_pt_mem_to_local_homeomorph_target ImplicitFunctionData.map_pt_mem_toLocalHomeomorph_target +theorem map_pt_mem_toPartialHomeomorph_target : + (φ.leftFun φ.pt, φ.rightFun φ.pt) ∈ φ.toPartialHomeomorph.target := + φ.toPartialHomeomorph.map_source <| φ.pt_mem_toPartialHomeomorph_source +#align implicit_function_data.map_pt_mem_to_local_homeomorph_target ImplicitFunctionData.map_pt_mem_toPartialHomeomorph_target theorem prod_map_implicitFunction : ∀ᶠ p : F × G in 𝓝 (φ.prodFun φ.pt), φ.prodFun (φ.implicitFunction p.1 p.2) = p := @@ -262,10 +262,10 @@ def implicitFunctionDataOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : ra /-- A local homeomorphism between `E` and `F × f'.ker` sending level surfaces of `f` to vertical subspaces. -/ -def implicitToLocalHomeomorphOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) - (hker : (ker f').ClosedComplemented) : LocalHomeomorph E (F × ker f') := - (implicitFunctionDataOfComplemented f f' hf hf' hker).toLocalHomeomorph -#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented HasStrictFDerivAt.implicitToLocalHomeomorphOfComplemented +def implicitToPartialHomeomorphOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) + (hker : (ker f').ClosedComplemented) : PartialHomeomorph E (F × ker f') := + (implicitFunctionDataOfComplemented f f' hf hf' hker).toPartialHomeomorph +#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented HasStrictFDerivAt.implicitToPartialHomeomorphOfComplemented /-- Implicit function `g` defined by `f (g z y) = z`. -/ def implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) @@ -276,55 +276,55 @@ def implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range end Defs @[simp] -theorem implicitToLocalHomeomorphOfComplemented_fst (hf : HasStrictFDerivAt f f' a) +theorem implicitToPartialHomeomorphOfComplemented_fst (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) (x : E) : - (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker x).fst = f x := + (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker x).fst = f x := rfl -#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_fst HasStrictFDerivAt.implicitToLocalHomeomorphOfComplemented_fst +#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_fst HasStrictFDerivAt.implicitToPartialHomeomorphOfComplemented_fst -theorem implicitToLocalHomeomorphOfComplemented_apply (hf : HasStrictFDerivAt f f' a) +theorem implicitToPartialHomeomorphOfComplemented_apply (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) (y : E) : - hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker y = + hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker y = (f y, Classical.choose hker (y - a)) := rfl -#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_apply HasStrictFDerivAt.implicitToLocalHomeomorphOfComplemented_apply +#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_apply HasStrictFDerivAt.implicitToPartialHomeomorphOfComplemented_apply @[simp] -theorem implicitToLocalHomeomorphOfComplemented_apply_ker (hf : HasStrictFDerivAt f f' a) +theorem implicitToPartialHomeomorphOfComplemented_apply_ker (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) (y : ker f') : - hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker (y + a) = (f (y + a), y) := by - simp only [implicitToLocalHomeomorphOfComplemented_apply, add_sub_cancel, + hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker (y + a) = (f (y + a), y) := by + simp only [implicitToPartialHomeomorphOfComplemented_apply, add_sub_cancel, Classical.choose_spec hker] -#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_apply_ker HasStrictFDerivAt.implicitToLocalHomeomorphOfComplemented_apply_ker +#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_apply_ker HasStrictFDerivAt.implicitToPartialHomeomorphOfComplemented_apply_ker @[simp] -theorem implicitToLocalHomeomorphOfComplemented_self (hf : HasStrictFDerivAt f f' a) +theorem implicitToPartialHomeomorphOfComplemented_self (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : - hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker a = (f a, 0) := by - simp [hf.implicitToLocalHomeomorphOfComplemented_apply] -#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_self HasStrictFDerivAt.implicitToLocalHomeomorphOfComplemented_self + hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker a = (f a, 0) := by + simp [hf.implicitToPartialHomeomorphOfComplemented_apply] +#align has_strict_fderiv_at.implicit_to_local_homeomorph_of_complemented_self HasStrictFDerivAt.implicitToPartialHomeomorphOfComplemented_self -theorem mem_implicitToLocalHomeomorphOfComplemented_source (hf : HasStrictFDerivAt f f' a) +theorem mem_implicitToPartialHomeomorphOfComplemented_source (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : - a ∈ (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker).source := - ImplicitFunctionData.pt_mem_toLocalHomeomorph_source _ -#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_of_complemented_source HasStrictFDerivAt.mem_implicitToLocalHomeomorphOfComplemented_source + a ∈ (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker).source := + ImplicitFunctionData.pt_mem_toPartialHomeomorph_source _ +#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_of_complemented_source HasStrictFDerivAt.mem_implicitToPartialHomeomorphOfComplemented_source -theorem mem_implicitToLocalHomeomorphOfComplemented_target (hf : HasStrictFDerivAt f f' a) +theorem mem_implicitToPartialHomeomorphOfComplemented_target (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : - (f a, (0 : ker f')) ∈ (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker).target := by - simpa only [implicitToLocalHomeomorphOfComplemented_self] using - (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker).map_source <| - hf.mem_implicitToLocalHomeomorphOfComplemented_source hf' hker -#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_of_complemented_target HasStrictFDerivAt.mem_implicitToLocalHomeomorphOfComplemented_target + (f a, (0 : ker f')) ∈ (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker).target := by + simpa only [implicitToPartialHomeomorphOfComplemented_self] using + (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker).map_source <| + hf.mem_implicitToPartialHomeomorphOfComplemented_source hf' hker +#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_of_complemented_target HasStrictFDerivAt.mem_implicitToPartialHomeomorphOfComplemented_target /-- `HasStrictFDerivAt.implicitFunctionOfComplemented` sends `(z, y)` to a point in `f ⁻¹' z`. -/ theorem map_implicitFunctionOfComplemented_eq (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : ∀ᶠ p : F × ker f' in 𝓝 (f a, 0), f (hf.implicitFunctionOfComplemented f f' hf' hker p.1 p.2) = p.1 := - ((hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker).eventually_right_inverse <| - hf.mem_implicitToLocalHomeomorphOfComplemented_target hf' hker).mono + ((hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker).eventually_right_inverse <| + hf.mem_implicitToPartialHomeomorphOfComplemented_target hf' hker).mono fun ⟨_, _⟩ h => congr_arg Prod.fst h #align has_strict_fderiv_at.map_implicit_function_of_complemented_eq HasStrictFDerivAt.map_implicitFunctionOfComplemented_eq @@ -333,7 +333,7 @@ theorem map_implicitFunctionOfComplemented_eq (hf : HasStrictFDerivAt f f' a) (h theorem eq_implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : ∀ᶠ x in 𝓝 a, hf.implicitFunctionOfComplemented f f' hf' hker (f x) - (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker x).snd = x := + (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker x).snd = x := (implicitFunctionDataOfComplemented f f' hf hf' hker).implicitFunction_apply_image #align has_strict_fderiv_at.eq_implicit_function_of_complemented HasStrictFDerivAt.eq_implicitFunctionOfComplemented @@ -341,9 +341,9 @@ theorem eq_implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : theorem implicitFunctionOfComplemented_apply_image (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (hker : (ker f').ClosedComplemented) : hf.implicitFunctionOfComplemented f f' hf' hker (f a) 0 = a := by - simpa only [implicitToLocalHomeomorphOfComplemented_self] using - (hf.implicitToLocalHomeomorphOfComplemented f f' hf' hker).left_inv - (hf.mem_implicitToLocalHomeomorphOfComplemented_source hf' hker) + simpa only [implicitToPartialHomeomorphOfComplemented_self] using + (hf.implicitToPartialHomeomorphOfComplemented f f' hf' hker).left_inv + (hf.mem_implicitToPartialHomeomorphOfComplemented_source hf' hker) #align has_strict_fderiv_at.implicit_function_of_complemented_apply_image HasStrictFDerivAt.implicitFunctionOfComplemented_apply_image theorem to_implicitFunctionOfComplemented (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) @@ -393,60 +393,60 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] {E : /-- Given a map `f : E → F` to a finite dimensional space with a surjective derivative `f'`, returns a local homeomorphism between `E` and `F × ker f'`. -/ -def implicitToLocalHomeomorph (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : - LocalHomeomorph E (F × ker f') := +def implicitToPartialHomeomorph (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : + PartialHomeomorph E (F × ker f') := haveI := FiniteDimensional.complete 𝕜 F - hf.implicitToLocalHomeomorphOfComplemented f f' hf' + hf.implicitToPartialHomeomorphOfComplemented f f' hf' f'.ker_closedComplemented_of_finiteDimensional_range -#align has_strict_fderiv_at.implicit_to_local_homeomorph HasStrictFDerivAt.implicitToLocalHomeomorph +#align has_strict_fderiv_at.implicit_to_local_homeomorph HasStrictFDerivAt.implicitToPartialHomeomorph /-- Implicit function `g` defined by `f (g z y) = z`. -/ def implicitFunction (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : F → ker f' → E := - Function.curry <| (hf.implicitToLocalHomeomorph f f' hf').symm + Function.curry <| (hf.implicitToPartialHomeomorph f f' hf').symm #align has_strict_fderiv_at.implicit_function HasStrictFDerivAt.implicitFunction variable {f f'} @[simp] -theorem implicitToLocalHomeomorph_fst (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) (x : E) : - (hf.implicitToLocalHomeomorph f f' hf' x).fst = f x := +theorem implicitToPartialHomeomorph_fst (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) + (x : E) : (hf.implicitToPartialHomeomorph f f' hf' x).fst = f x := rfl -#align has_strict_fderiv_at.implicit_to_local_homeomorph_fst HasStrictFDerivAt.implicitToLocalHomeomorph_fst +#align has_strict_fderiv_at.implicit_to_local_homeomorph_fst HasStrictFDerivAt.implicitToPartialHomeomorph_fst @[simp] -theorem implicitToLocalHomeomorph_apply_ker (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) - (y : ker f') : hf.implicitToLocalHomeomorph f f' hf' (y + a) = (f (y + a), y) := +theorem implicitToPartialHomeomorph_apply_ker (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) + (y : ker f') : hf.implicitToPartialHomeomorph f f' hf' (y + a) = (f (y + a), y) := -- porting note: had to add `haveI` (here and below) haveI := FiniteDimensional.complete 𝕜 F - implicitToLocalHomeomorphOfComplemented_apply_ker .. -#align has_strict_fderiv_at.implicit_to_local_homeomorph_apply_ker HasStrictFDerivAt.implicitToLocalHomeomorph_apply_ker + implicitToPartialHomeomorphOfComplemented_apply_ker .. +#align has_strict_fderiv_at.implicit_to_local_homeomorph_apply_ker HasStrictFDerivAt.implicitToPartialHomeomorph_apply_ker @[simp] -theorem implicitToLocalHomeomorph_self (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : - hf.implicitToLocalHomeomorph f f' hf' a = (f a, 0) := +theorem implicitToPartialHomeomorph_self (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : + hf.implicitToPartialHomeomorph f f' hf' a = (f a, 0) := haveI := FiniteDimensional.complete 𝕜 F - implicitToLocalHomeomorphOfComplemented_self .. -#align has_strict_fderiv_at.implicit_to_local_homeomorph_self HasStrictFDerivAt.implicitToLocalHomeomorph_self + implicitToPartialHomeomorphOfComplemented_self .. +#align has_strict_fderiv_at.implicit_to_local_homeomorph_self HasStrictFDerivAt.implicitToPartialHomeomorph_self -theorem mem_implicitToLocalHomeomorph_source (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : - a ∈ (hf.implicitToLocalHomeomorph f f' hf').source := +theorem mem_implicitToPartialHomeomorph_source (hf : HasStrictFDerivAt f f' a) + (hf' : range f' = ⊤) : a ∈ (hf.implicitToPartialHomeomorph f f' hf').source := haveI := FiniteDimensional.complete 𝕜 F - ImplicitFunctionData.pt_mem_toLocalHomeomorph_source _ -#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_source HasStrictFDerivAt.mem_implicitToLocalHomeomorph_source + ImplicitFunctionData.pt_mem_toPartialHomeomorph_source _ +#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_source HasStrictFDerivAt.mem_implicitToPartialHomeomorph_source -theorem mem_implicitToLocalHomeomorph_target (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : - (f a, (0 : ker f')) ∈ (hf.implicitToLocalHomeomorph f f' hf').target := +theorem mem_implicitToPartialHomeomorph_target (hf : HasStrictFDerivAt f f' a) + (hf' : range f' = ⊤) : (f a, (0 : ker f')) ∈ (hf.implicitToPartialHomeomorph f f' hf').target := haveI := FiniteDimensional.complete 𝕜 F - mem_implicitToLocalHomeomorphOfComplemented_target .. -#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_target HasStrictFDerivAt.mem_implicitToLocalHomeomorph_target + mem_implicitToPartialHomeomorphOfComplemented_target .. +#align has_strict_fderiv_at.mem_implicit_to_local_homeomorph_target HasStrictFDerivAt.mem_implicitToPartialHomeomorph_target theorem tendsto_implicitFunction (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) {α : Type*} {l : Filter α} {g₁ : α → F} {g₂ : α → ker f'} (h₁ : Tendsto g₁ l (𝓝 <| f a)) (h₂ : Tendsto g₂ l (𝓝 0)) : Tendsto (fun t => hf.implicitFunction f f' hf' (g₁ t) (g₂ t)) l (𝓝 a) := by - refine' ((hf.implicitToLocalHomeomorph f f' hf').tendsto_symm - (hf.mem_implicitToLocalHomeomorph_source hf')).comp _ - rw [implicitToLocalHomeomorph_self] + refine' ((hf.implicitToPartialHomeomorph f f' hf').tendsto_symm + (hf.mem_implicitToPartialHomeomorph_source hf')).comp _ + rw [implicitToPartialHomeomorph_self] exact h₁.prod_mk_nhds h₂ #align has_strict_fderiv_at.tendsto_implicit_function HasStrictFDerivAt.tendsto_implicitFunction @@ -471,7 +471,7 @@ theorem implicitFunction_apply_image (hf : HasStrictFDerivAt f f' a) (hf' : rang of some point. -/ theorem eq_implicitFunction (hf : HasStrictFDerivAt f f' a) (hf' : range f' = ⊤) : ∀ᶠ x in 𝓝 a, - hf.implicitFunction f f' hf' (f x) (hf.implicitToLocalHomeomorph f f' hf' x).snd = x := + hf.implicitFunction f f' hf' (f x) (hf.implicitToPartialHomeomorph f f' hf' x).snd = x := haveI := FiniteDimensional.complete 𝕜 F eq_implicitFunctionOfComplemented .. #align has_strict_fderiv_at.eq_implicit_function HasStrictFDerivAt.eq_implicitFunction diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 73d03a40bf415..bdd3c5f7fd21f 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -16,11 +16,11 @@ behave like `f a + f' (x - a)` near each `a ∈ s`. When `f'` is onto, we show that `f` is locally onto. When `f'` is a continuous linear equiv, we show that `f` is a homeomorphism -between `s` and `f '' s`. More precisely, we define `ApproximatesLinearOn.toLocalHomeomorph` to -be a `LocalHomeomorph` with `toFun = f`, `source = s`, and `target = f '' s`. +between `s` and `f '' s`. More precisely, we define `ApproximatesLinearOn.toPartialHomeomorph` to +be a `PartialHomeomorph` with `toFun = f`, `source = s`, and `target = f '' s`. Maps of this type naturally appear in the proof of the inverse function theorem (see next section), -and `ApproximatesLinearOn.toLocalHomeomorph` will imply that the locally inverse function +and `ApproximatesLinearOn.toPartialHomeomorph` will imply that the locally inverse function exists. We define this auxiliary notion to split the proof of the inverse function theorem into small @@ -353,7 +353,7 @@ protected theorem surjective [CompleteSpace E] (hf : ApproximatesLinearOn f (f' #align approximates_linear_on.surjective ApproximatesLinearOn.surjective /-- A map approximating a linear equivalence on a set defines a local equivalence on this set. -Should not be used outside of this file, because it is superseded by `toLocalHomeomorph` below. +Should not be used outside of this file, because it is superseded by `toPartialHomeomorph` below. This is a first step towards the inverse function. -/ def toLocalEquiv (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) @@ -361,7 +361,8 @@ def toLocalEquiv (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hf.injOn hc).toLocalEquiv _ _ #align approximates_linear_on.to_local_equiv ApproximatesLinearOn.toLocalEquiv -/-- The inverse function is continuous on `f '' s`. Use properties of `LocalHomeomorph` instead. -/ +/-- The inverse function is continuous on `f '' s`. +Use properties of `PartialHomeomorph` instead. -/ theorem inverse_continuousOn (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hc : Subsingleton E ∨ c < N⁻¹) : ContinuousOn (hf.toLocalEquiv hc).symm (f '' s) := by apply continuousOn_iff_continuous_restrict.2 @@ -400,41 +401,42 @@ variable (f s) /-- Given a function `f` that approximates a linear equivalence on an open set `s`, returns a local homeomorph with `toFun = f` and `source = s`. -/ -def toLocalHomeomorph (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) - (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : LocalHomeomorph E F where +def toPartialHomeomorph (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) + (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : PartialHomeomorph E F where toLocalEquiv := hf.toLocalEquiv hc open_source := hs open_target := hf.open_image f'.toNonlinearRightInverse hs <| by rwa [f'.toEquiv.subsingleton_congr] at hc continuousOn_toFun := hf.continuousOn continuousOn_invFun := hf.inverse_continuousOn hc -#align approximates_linear_on.to_local_homeomorph ApproximatesLinearOn.toLocalHomeomorph +#align approximates_linear_on.to_local_homeomorph ApproximatesLinearOn.toPartialHomeomorph @[simp] -theorem toLocalHomeomorph_coe (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) +theorem toPartialHomeomorph_coe (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : - (hf.toLocalHomeomorph f s hc hs : E → F) = f := + (hf.toPartialHomeomorph f s hc hs : E → F) = f := rfl -#align approximates_linear_on.to_local_homeomorph_coe ApproximatesLinearOn.toLocalHomeomorph_coe +#align approximates_linear_on.to_local_homeomorph_coe ApproximatesLinearOn.toPartialHomeomorph_coe @[simp] -theorem toLocalHomeomorph_source (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) - (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : (hf.toLocalHomeomorph f s hc hs).source = s := +theorem toPartialHomeomorph_source (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) + (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : + (hf.toPartialHomeomorph f s hc hs).source = s := rfl -#align approximates_linear_on.to_local_homeomorph_source ApproximatesLinearOn.toLocalHomeomorph_source +#align approximates_linear_on.to_local_homeomorph_source ApproximatesLinearOn.toPartialHomeomorph_source @[simp] -theorem toLocalHomeomorph_target (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) +theorem toPartialHomeomorph_target (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) : - (hf.toLocalHomeomorph f s hc hs).target = f '' s := + (hf.toPartialHomeomorph f s hc hs).target = f '' s := rfl -#align approximates_linear_on.to_local_homeomorph_target ApproximatesLinearOn.toLocalHomeomorph_target +#align approximates_linear_on.to_local_homeomorph_target ApproximatesLinearOn.toPartialHomeomorph_target /-- A function `f` that approximates a linear equivalence on the whole space is a homeomorphism. -/ def toHomeomorph (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) univ c) (hc : Subsingleton E ∨ c < N⁻¹) : E ≃ₜ F := by - refine' (hf.toLocalHomeomorph _ _ hc isOpen_univ).toHomeomorphOfSourceEqUnivTargetEqUniv rfl _ - rw [toLocalHomeomorph_target, image_univ, range_iff_surjective] + refine' (hf.toPartialHomeomorph _ _ hc isOpen_univ).toHomeomorphOfSourceEqUnivTargetEqUniv rfl _ + rw [toPartialHomeomorph_target, image_univ, range_iff_surjective] exact hf.surjective hc #align approximates_linear_on.to_homeomorph ApproximatesLinearOn.toHomeomorph @@ -442,7 +444,7 @@ end theorem closedBall_subset_target (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hc : Subsingleton E ∨ c < N⁻¹) (hs : IsOpen s) {b : E} (ε0 : 0 ≤ ε) (hε : closedBall b ε ⊆ s) : - closedBall (f b) ((N⁻¹ - c) * ε) ⊆ (hf.toLocalHomeomorph f s hc hs).target := + closedBall (f b) ((N⁻¹ - c) * ε) ⊆ (hf.toPartialHomeomorph f s hc hs).target := (hf.surjOn_closedBall_of_nonlinearRightInverse f'.toNonlinearRightInverse ε0 hε).mono hε Subset.rfl #align approximates_linear_on.closed_ball_subset_target ApproximatesLinearOn.closedBall_subset_target diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean index d3d57f77ac0c3..5561a84ba4f0d 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean @@ -25,32 +25,32 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕂 F] variable [CompleteSpace E] (f : E → F) {f' : E ≃L[𝕂] F} {a : E} /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible -derivative at `a`, returns a `LocalHomeomorph` with `to_fun = f` and `a ∈ source`. -/ -def toLocalHomeomorph {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) - (hn : 1 ≤ n) : LocalHomeomorph E F := - (hf.hasStrictFDerivAt' hf' hn).toLocalHomeomorph f -#align cont_diff_at.to_local_homeomorph ContDiffAt.toLocalHomeomorph +derivative at `a`, returns a `PartialHomeomorph` with `to_fun = f` and `a ∈ source`. -/ +def toPartialHomeomorph {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) + (hn : 1 ≤ n) : PartialHomeomorph E F := + (hf.hasStrictFDerivAt' hf' hn).toPartialHomeomorph f +#align cont_diff_at.to_local_homeomorph ContDiffAt.toPartialHomeomorph variable {f} @[simp] -theorem toLocalHomeomorph_coe {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem toPartialHomeomorph_coe {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : - (hf.toLocalHomeomorph f hf' hn : E → F) = f := + (hf.toPartialHomeomorph f hf' hn : E → F) = f := rfl -#align cont_diff_at.to_local_homeomorph_coe ContDiffAt.toLocalHomeomorph_coe +#align cont_diff_at.to_local_homeomorph_coe ContDiffAt.toPartialHomeomorph_coe -theorem mem_toLocalHomeomorph_source {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem mem_toPartialHomeomorph_source {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : - a ∈ (hf.toLocalHomeomorph f hf' hn).source := - (hf.hasStrictFDerivAt' hf' hn).mem_toLocalHomeomorph_source -#align cont_diff_at.mem_to_local_homeomorph_source ContDiffAt.mem_toLocalHomeomorph_source + a ∈ (hf.toPartialHomeomorph f hf' hn).source := + (hf.hasStrictFDerivAt' hf' hn).mem_toPartialHomeomorph_source +#align cont_diff_at.mem_to_local_homeomorph_source ContDiffAt.mem_toPartialHomeomorph_source -theorem image_mem_toLocalHomeomorph_target {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem image_mem_toPartialHomeomorph_target {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : - f a ∈ (hf.toLocalHomeomorph f hf' hn).target := - (hf.hasStrictFDerivAt' hf' hn).image_mem_toLocalHomeomorph_target -#align cont_diff_at.image_mem_to_local_homeomorph_target ContDiffAt.image_mem_toLocalHomeomorph_target + f a ∈ (hf.toPartialHomeomorph f hf' hn).target := + (hf.hasStrictFDerivAt' hf' hn).image_mem_toPartialHomeomorph_target +#align cont_diff_at.image_mem_to_local_homeomorph_target ContDiffAt.image_mem_toPartialHomeomorph_target /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a function that is locally inverse to `f`. -/ @@ -65,14 +65,14 @@ theorem localInverse_apply_image {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) #align cont_diff_at.local_inverse_apply_image ContDiffAt.localInverse_apply_image /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative -at `a`, the inverse function (produced by `ContDiff.toLocalHomeomorph`) is +at `a`, the inverse function (produced by `ContDiff.toPartialHomeomorph`) is also `ContDiff`. -/ theorem to_localInverse {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : ContDiffAt 𝕂 n (hf.localInverse hf' hn) (f a) := by have := hf.localInverse_apply_image hf' hn - apply (hf.toLocalHomeomorph f hf' hn).contDiffAt_symm - (image_mem_toLocalHomeomorph_target hf hf' hn) + apply (hf.toPartialHomeomorph f hf' hn).contDiffAt_symm + (image_mem_toPartialHomeomorph_target hf hf' hn) · convert hf' · convert hf #align cont_diff_at.to_local_inverse ContDiffAt.to_localInverse diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean index 890db86a1b653..247015e779ac1 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean @@ -15,12 +15,12 @@ In this file we prove the inverse function theorem. It says that if a map `f : E has an invertible strict derivative `f'` at `a`, then it is locally invertible, and the inverse function has derivative `f' ⁻¹`. -We define `HasStrictFDerivAt.toLocalHomeomorph` that repacks a function `f` -with a `hf : HasStrictFDerivAt f f' a`, `f' : E ≃L[𝕜] F`, into a `LocalHomeomorph`. -The `toFun` of this `LocalHomeomorph` is defeq to `f`, so one can apply theorems -about `LocalHomeomorph` to `hf.toLocalHomeomorph f`, and get statements about `f`. +We define `HasStrictFDerivAt.toPartialHomeomorph` that repacks a function `f` +with a `hf : HasStrictFDerivAt f f' a`, `f' : E ≃L[𝕜] F`, into a `PartialHomeomorph`. +The `toFun` of this `PartialHomeomorph` is defeq to `f`, so one can apply theorems +about `PartialHomeomorph` to `hf.toPartialHomeomorph f`, and get statements about `f`. -Then we define `HasStrictFDerivAt.localInverse` to be the `invFun` of this `LocalHomeomorph`, +Then we define `HasStrictFDerivAt.localInverse` to be the `invFun` of this `PartialHomeomorph`, and prove two versions of the inverse function theorem: * `HasStrictFDerivAt.to_localInverse`: if `f` has an invertible derivative `f'` at `a` in the @@ -69,7 +69,7 @@ open ContinuousLinearMap (id) Let `f : E → F` be a map defined on a complete vector space `E`. Assume that `f` has an invertible derivative `f' : E ≃L[𝕜] F` at `a : E` in the strict sense. Then `f` approximates `f'` in the sense of `ApproximatesLinearOn` on an open neighborhood -of `a`, and we can apply `ApproximatesLinearOn.toLocalHomeomorph` to construct the inverse +of `a`, and we can apply `ApproximatesLinearOn.toPartialHomeomorph` to construct the inverse function. -/ namespace HasStrictFDerivAt @@ -115,39 +115,39 @@ theorem approximates_deriv_on_open_nhds (hf : HasStrictFDerivAt f (f' : E →L[ variable (f) -/-- Given a function with an invertible strict derivative at `a`, returns a `LocalHomeomorph` +/-- Given a function with an invertible strict derivative at `a`, returns a `PartialHomeomorph` with `to_fun = f` and `a ∈ source`. This is a part of the inverse function theorem. The other part `HasStrictFDerivAt.to_localInverse` states that the inverse function -of this `LocalHomeomorph` has derivative `f'.symm`. -/ -def toLocalHomeomorph (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : LocalHomeomorph E F := - ApproximatesLinearOn.toLocalHomeomorph f (Classical.choose hf.approximates_deriv_on_open_nhds) +of this `PartialHomeomorph` has derivative `f'.symm`. -/ +def toPartialHomeomorph (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : PartialHomeomorph E F := + ApproximatesLinearOn.toPartialHomeomorph f (Classical.choose hf.approximates_deriv_on_open_nhds) (Classical.choose_spec hf.approximates_deriv_on_open_nhds).2.2 (f'.subsingleton_or_nnnorm_symm_pos.imp id fun hf' => NNReal.half_lt_self <| ne_of_gt <| inv_pos.2 hf') (Classical.choose_spec hf.approximates_deriv_on_open_nhds).2.1 -#align has_strict_fderiv_at.to_local_homeomorph HasStrictFDerivAt.toLocalHomeomorph +#align has_strict_fderiv_at.to_local_homeomorph HasStrictFDerivAt.toPartialHomeomorph variable {f} @[simp] -theorem toLocalHomeomorph_coe (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : - (hf.toLocalHomeomorph f : E → F) = f := +theorem toPartialHomeomorph_coe (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : + (hf.toPartialHomeomorph f : E → F) = f := rfl -#align has_strict_fderiv_at.to_local_homeomorph_coe HasStrictFDerivAt.toLocalHomeomorph_coe +#align has_strict_fderiv_at.to_local_homeomorph_coe HasStrictFDerivAt.toPartialHomeomorph_coe -theorem mem_toLocalHomeomorph_source (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : - a ∈ (hf.toLocalHomeomorph f).source := +theorem mem_toPartialHomeomorph_source (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : + a ∈ (hf.toPartialHomeomorph f).source := (Classical.choose_spec hf.approximates_deriv_on_open_nhds).1 -#align has_strict_fderiv_at.mem_to_local_homeomorph_source HasStrictFDerivAt.mem_toLocalHomeomorph_source +#align has_strict_fderiv_at.mem_to_local_homeomorph_source HasStrictFDerivAt.mem_toPartialHomeomorph_source -theorem image_mem_toLocalHomeomorph_target (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : - f a ∈ (hf.toLocalHomeomorph f).target := - (hf.toLocalHomeomorph f).map_source hf.mem_toLocalHomeomorph_source -#align has_strict_fderiv_at.image_mem_to_local_homeomorph_target HasStrictFDerivAt.image_mem_toLocalHomeomorph_target +theorem image_mem_toPartialHomeomorph_target (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : + f a ∈ (hf.toPartialHomeomorph f).target := + (hf.toPartialHomeomorph f).map_source hf.mem_toPartialHomeomorph_source +#align has_strict_fderiv_at.image_mem_to_local_homeomorph_target HasStrictFDerivAt.image_mem_toPartialHomeomorph_target theorem map_nhds_eq_of_equiv (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : map f (𝓝 a) = 𝓝 (f a) := - (hf.toLocalHomeomorph f).map_nhds_eq hf.mem_toLocalHomeomorph_source + (hf.toPartialHomeomorph f).map_nhds_eq hf.mem_toPartialHomeomorph_source #align has_strict_fderiv_at.map_nhds_eq_of_equiv HasStrictFDerivAt.map_nhds_eq_of_equiv variable (f f' a) @@ -155,19 +155,19 @@ variable (f f' a) /-- Given a function `f` with an invertible derivative, returns a function that is locally inverse to `f`. -/ def localInverse (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : F → E := - (hf.toLocalHomeomorph f).symm + (hf.toPartialHomeomorph f).symm #align has_strict_fderiv_at.local_inverse HasStrictFDerivAt.localInverse variable {f f' a} theorem localInverse_def (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : - hf.localInverse f _ _ = (hf.toLocalHomeomorph f).symm := + hf.localInverse f _ _ = (hf.toPartialHomeomorph f).symm := rfl #align has_strict_fderiv_at.local_inverse_def HasStrictFDerivAt.localInverse_def theorem eventually_left_inverse (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : ∀ᶠ x in 𝓝 a, hf.localInverse f f' a (f x) = x := - (hf.toLocalHomeomorph f).eventually_left_inverse hf.mem_toLocalHomeomorph_source + (hf.toPartialHomeomorph f).eventually_left_inverse hf.mem_toPartialHomeomorph_source #align has_strict_fderiv_at.eventually_left_inverse HasStrictFDerivAt.eventually_left_inverse @[simp] @@ -178,30 +178,30 @@ theorem localInverse_apply_image (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) theorem eventually_right_inverse (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : ∀ᶠ y in 𝓝 (f a), f (hf.localInverse f f' a y) = y := - (hf.toLocalHomeomorph f).eventually_right_inverse' hf.mem_toLocalHomeomorph_source + (hf.toPartialHomeomorph f).eventually_right_inverse' hf.mem_toPartialHomeomorph_source #align has_strict_fderiv_at.eventually_right_inverse HasStrictFDerivAt.eventually_right_inverse theorem localInverse_continuousAt (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : ContinuousAt (hf.localInverse f f' a) (f a) := - (hf.toLocalHomeomorph f).continuousAt_symm hf.image_mem_toLocalHomeomorph_target + (hf.toPartialHomeomorph f).continuousAt_symm hf.image_mem_toPartialHomeomorph_target #align has_strict_fderiv_at.local_inverse_continuous_at HasStrictFDerivAt.localInverse_continuousAt theorem localInverse_tendsto (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : Tendsto (hf.localInverse f f' a) (𝓝 <| f a) (𝓝 a) := - (hf.toLocalHomeomorph f).tendsto_symm hf.mem_toLocalHomeomorph_source + (hf.toPartialHomeomorph f).tendsto_symm hf.mem_toPartialHomeomorph_source #align has_strict_fderiv_at.local_inverse_tendsto HasStrictFDerivAt.localInverse_tendsto theorem localInverse_unique (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) {g : F → E} (hg : ∀ᶠ x in 𝓝 a, g (f x) = x) : ∀ᶠ y in 𝓝 (f a), g y = localInverse f f' a hf y := eventuallyEq_of_left_inv_of_right_inv hg hf.eventually_right_inverse <| - (hf.toLocalHomeomorph f).tendsto_symm hf.mem_toLocalHomeomorph_source + (hf.toPartialHomeomorph f).tendsto_symm hf.mem_toPartialHomeomorph_source #align has_strict_fderiv_at.local_inverse_unique HasStrictFDerivAt.localInverse_unique /-- If `f` has an invertible derivative `f'` at `a` in the sense of strict differentiability `(hf)`, then the inverse function `hf.localInverse f` has derivative `f'.symm` at `f a`. -/ theorem to_localInverse (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : HasStrictFDerivAt (hf.localInverse f f' a) (f'.symm : F →L[𝕜] E) (f a) := - (hf.toLocalHomeomorph f).hasStrictFDerivAt_symm hf.image_mem_toLocalHomeomorph_target <| by + (hf.toPartialHomeomorph f).hasStrictFDerivAt_symm hf.image_mem_toPartialHomeomorph_target <| by simpa [← localInverse_def] using hf #align has_strict_fderiv_at.to_local_inverse HasStrictFDerivAt.to_localInverse diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index 6aac288bcb92f..f50f72439632c 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -377,13 +377,13 @@ open Metric hiding mem_nhds_iff variable {n : ℕ∞} {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] -theorem LocalHomeomorph.contDiff_univUnitBall : ContDiff ℝ n (univUnitBall : E → E) := by +theorem PartialHomeomorph.contDiff_univUnitBall : ContDiff ℝ n (univUnitBall : E → E) := by suffices ContDiff ℝ n fun x : E => (1 + ‖x‖ ^ 2 : ℝ).sqrt⁻¹ from this.smul contDiff_id have h : ∀ x : E, (0 : ℝ) < (1 : ℝ) + ‖x‖ ^ 2 := fun x => by positivity refine' ContDiff.inv _ fun x => Real.sqrt_ne_zero'.mpr (h x) exact (contDiff_const.add <| contDiff_norm_sq ℝ).sqrt fun x => (h x).ne' -theorem LocalHomeomorph.contDiffOn_univUnitBall_symm : +theorem PartialHomeomorph.contDiffOn_univUnitBall_symm : ContDiffOn ℝ n univUnitBall.symm (ball (0 : E) 1) := fun y hy ↦ by apply ContDiffAt.contDiffWithinAt suffices ContDiffAt ℝ n (fun y : E => (1 - ‖y‖ ^ 2 : ℝ).sqrt⁻¹) y from this.smul contDiffAt_id @@ -394,17 +394,17 @@ theorem LocalHomeomorph.contDiffOn_univUnitBall_symm : exact contDiffAt_const.sub (contDiff_norm_sq ℝ).contDiffAt theorem Homeomorph.contDiff_unitBall : ContDiff ℝ n fun x : E => (unitBall x : E) := - LocalHomeomorph.contDiff_univUnitBall + PartialHomeomorph.contDiff_univUnitBall #align cont_diff_homeomorph_unit_ball Homeomorph.contDiff_unitBall -@[deprecated LocalHomeomorph.contDiffOn_univUnitBall_symm] +@[deprecated PartialHomeomorph.contDiffOn_univUnitBall_symm] theorem Homeomorph.contDiffOn_unitBall_symm {f : E → E} (h : ∀ (y) (hy : y ∈ ball (0 : E) 1), f y = Homeomorph.unitBall.symm ⟨y, hy⟩) : ContDiffOn ℝ n f <| ball 0 1 := - LocalHomeomorph.contDiffOn_univUnitBall_symm.congr h + PartialHomeomorph.contDiffOn_univUnitBall_symm.congr h #align cont_diff_on_homeomorph_unit_ball_symm Homeomorph.contDiffOn_unitBall_symm -namespace LocalHomeomorph +namespace PartialHomeomorph variable {c : E} {r : ℝ} @@ -424,9 +424,9 @@ theorem contDiffOn_univBall_symm : unfold univBall; split_ifs with h · refine contDiffOn_univUnitBall_symm.comp (contDiff_unitBallBall_symm h).contDiffOn ?_ rw [← unitBallBall_source c r h, ← unitBallBall_target c r h] - apply LocalHomeomorph.symm_mapsTo + apply PartialHomeomorph.symm_mapsTo · exact contDiffOn_id.sub contDiffOn_const -end LocalHomeomorph +end PartialHomeomorph end DiffeomorphUnitBall diff --git a/Mathlib/Analysis/NormedSpace/HomeomorphBall.lean b/Mathlib/Analysis/NormedSpace/HomeomorphBall.lean index 3c4d74839ae28..eaf3d6f95a0ab 100644 --- a/Mathlib/Analysis/NormedSpace/HomeomorphBall.lean +++ b/Mathlib/Analysis/NormedSpace/HomeomorphBall.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Oliver Nash -/ -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph import Mathlib.Analysis.NormedSpace.AddTorsor import Mathlib.Analysis.NormedSpace.Pointwise @@ -17,14 +17,14 @@ In this file we show that a real (semi)normed vector space is homeomorphic to th We formalize it in two ways: - as a `Homeomorph`, see `Homeomorph.unitBall`; -- as a `LocalHomeomorph` with `source = Set.univ` and `target = Metric.ball (0 : E) 1`. +- as a `PartialHomeomorph` with `source = Set.univ` and `target = Metric.ball (0 : E) 1`. While the former approach is more natural, the latter approach provides us with a globally defined inverse function which makes it easier to say that this homeomorphism is in fact a diffeomorphism. We also show that the unit ball `Metric.ball (0 : E) 1` is homeomorphic -to a ball of positive radius in an affine space over `E`, see `LocalHomeomorph.unitBallBall`. +to a ball of positive radius in an affine space over `E`, see `PartialHomeomorph.unitBallBall`. ## Tags @@ -39,7 +39,7 @@ noncomputable section /-- Local homeomorphism between a real (semi)normed space and the unit ball. See also `Homeomorph.unitBall`. -/ @[simps (config := .lemmasOnly)] -def LocalHomeomorph.univUnitBall : LocalHomeomorph E E where +def PartialHomeomorph.univUnitBall : PartialHomeomorph E E where toFun x := (1 + ‖x‖ ^ 2).sqrt⁻¹ • x invFun y := (1 - ‖(y : E)‖ ^ 2).sqrt⁻¹ • (y : E) source := univ @@ -73,12 +73,12 @@ def LocalHomeomorph.univUnitBall : LocalHomeomorph E E where (continuousOn_const.sub (continuous_norm.continuousOn.pow _)).sqrt this) continuousOn_id @[simp] -theorem LocalHomeomorph.univUnitBall_apply_zero : univUnitBall (0 : E) = 0 := by - simp [LocalHomeomorph.univUnitBall_apply] +theorem PartialHomeomorph.univUnitBall_apply_zero : univUnitBall (0 : E) = 0 := by + simp [PartialHomeomorph.univUnitBall_apply] @[simp] -theorem LocalHomeomorph.univUnitBall_symm_apply_zero : univUnitBall.symm (0 : E) = 0 := by - simp [LocalHomeomorph.univUnitBall_symm_apply] +theorem PartialHomeomorph.univUnitBall_symm_apply_zero : univUnitBall.symm (0 : E) = 0 := by + simp [PartialHomeomorph.univUnitBall_symm_apply] /-- A (semi) normed real vector space is homeomorphic to the unit ball in the same space. This homeomorphism sends `x : E` to `(1 + ‖x‖²)^(- ½) • x`. @@ -86,41 +86,41 @@ This homeomorphism sends `x : E` to `(1 + ‖x‖²)^(- ½) • x`. In many cases the actual implementation is not important, so we don't mark the projection lemmas `Homeomorph.unitBall_apply_coe` and `Homeomorph.unitBall_symm_apply` as `@[simp]`. -See also `Homeomorph.contDiff_unitBall` and `LocalHomeomorph.contDiffOn_unitBall_symm` +See also `Homeomorph.contDiff_unitBall` and `PartialHomeomorph.contDiffOn_unitBall_symm` for smoothness properties that hold when `E` is an inner-product space. -/ @[simps! (config := .lemmasOnly)] def Homeomorph.unitBall : E ≃ₜ ball (0 : E) 1 := - (Homeomorph.Set.univ _).symm.trans LocalHomeomorph.univUnitBall.toHomeomorphSourceTarget + (Homeomorph.Set.univ _).symm.trans PartialHomeomorph.univUnitBall.toHomeomorphSourceTarget #align homeomorph_unit_ball Homeomorph.unitBall @[simp] theorem Homeomorph.coe_unitBall_apply_zero : (Homeomorph.unitBall (0 : E) : E) = 0 := - LocalHomeomorph.univUnitBall_apply_zero + PartialHomeomorph.univUnitBall_apply_zero #align coe_homeomorph_unit_ball_apply_zero Homeomorph.coe_unitBall_apply_zero variable {P : Type*} [PseudoMetricSpace P] [NormedAddTorsor E P] -namespace LocalHomeomorph +namespace PartialHomeomorph /-- Affine homeomorphism `(r • · +ᵥ c)` between a normed space and an add torsor over this space, -interpreted as a `LocalHomeomorph` between `Metric.ball 0 1` and `Metric.ball c r`. -/ +interpreted as a `PartialHomeomorph` between `Metric.ball 0 1` and `Metric.ball c r`. -/ @[simps!] -def unitBallBall (c : P) (r : ℝ) (hr : 0 < r) : LocalHomeomorph E P := +def unitBallBall (c : P) (r : ℝ) (hr : 0 < r) : PartialHomeomorph E P := ((Homeomorph.smulOfNeZero r hr.ne').trans - (IsometryEquiv.vaddConst c).toHomeomorph).toLocalHomeomorphOfImageEq + (IsometryEquiv.vaddConst c).toHomeomorph).toPartialHomeomorphOfImageEq (ball 0 1) isOpen_ball (ball c r) <| by change (IsometryEquiv.vaddConst c) ∘ (r • ·) '' ball (0 : E) 1 = ball c r rw [image_comp, image_smul, smul_unitBall hr.ne', IsometryEquiv.image_ball] simp [abs_of_pos hr] -/-- If `r > 0`, then `LocalHomeomorph.univBall c r` is a smooth local homeomorphism +/-- If `r > 0`, then `PartialHomeomorph.univBall c r` is a smooth local homeomorphism with `source = Set.univ` and `target = Metric.ball c r`. Otherwise, it is the translation by `c`. -Thus in all cases, it sends `0` to `c`, see `LocalHomeomorph.univBall_apply_zero`. -/ -def univBall (c : P) (r : ℝ) : LocalHomeomorph E P := +Thus in all cases, it sends `0` to `c`, see `PartialHomeomorph.univBall_apply_zero`. -/ +def univBall (c : P) (r : ℝ) : PartialHomeomorph E P := if h : 0 < r then univUnitBall.trans' (unitBallBall c r h) rfl - else (IsometryEquiv.vaddConst c).toHomeomorph.toLocalHomeomorph + else (IsometryEquiv.vaddConst c).toHomeomorph.toPartialHomeomorph @[simp] theorem univBall_source (c : P) (r : ℝ) : (univBall c r).source = univ := by diff --git a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean index 3f60ef42c5c43..a259dbaa0ee8b 100644 --- a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean +++ b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.NormedSpace.Basic -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph /-! # Homeomorphism between a normed space and sphere times `(0, +∞)` diff --git a/Mathlib/Analysis/SpecialFunctions/Arsinh.lean b/Mathlib/Analysis/SpecialFunctions/Arsinh.lean index 46ff37e5dd4db..6dc3283b72965 100644 --- a/Mathlib/Analysis/SpecialFunctions/Arsinh.lean +++ b/Mathlib/Analysis/SpecialFunctions/Arsinh.lean @@ -177,7 +177,7 @@ theorem arsinh_neg_iff : arsinh x < 0 ↔ x < 0 := #align real.arsinh_neg_iff Real.arsinh_neg_iff theorem hasStrictDerivAt_arsinh (x : ℝ) : HasStrictDerivAt arsinh (sqrt (1 + x ^ 2))⁻¹ x := by - convert sinhHomeomorph.toLocalHomeomorph.hasStrictDerivAt_symm (mem_univ x) (cosh_pos _).ne' + convert sinhHomeomorph.toPartialHomeomorph.hasStrictDerivAt_symm (mem_univ x) (cosh_pos _).ne' (hasStrictDerivAt_sinh _) using 2 exact (cosh_arsinh _).symm #align real.has_strict_deriv_at_arsinh Real.hasStrictDerivAt_arsinh diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean index b9384e448ab32..fc9706d537e53 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogDeriv.lean @@ -25,11 +25,11 @@ theorem isOpenMap_exp : IsOpenMap exp := open_map_of_strict_deriv hasStrictDerivAt_exp exp_ne_zero #align complex.is_open_map_exp Complex.isOpenMap_exp -/-- `Complex.exp` as a `LocalHomeomorph` with `source = {z | -π < im z < π}` and +/-- `Complex.exp` as a `PartialHomeomorph` with `source = {z | -π < im z < π}` and `target = {z | 0 < re z} ∪ {z | im z ≠ 0}`. This definition is used to prove that `Complex.log` is complex differentiable at all points but the negative real semi-axis. -/ -noncomputable def expLocalHomeomorph : LocalHomeomorph ℂ ℂ := - LocalHomeomorph.ofContinuousOpen +noncomputable def expPartialHomeomorph : PartialHomeomorph ℂ ℂ := + PartialHomeomorph.ofContinuousOpen { toFun := exp invFun := log source := {z : ℂ | z.im ∈ Ioo (-π) π} @@ -48,11 +48,11 @@ noncomputable def expLocalHomeomorph : LocalHomeomorph ℂ ℂ := left_inv' := fun x hx => log_exp hx.1 (le_of_lt hx.2) right_inv' := fun x hx => exp_log <| by rintro rfl; simp [lt_irrefl] at hx } continuous_exp.continuousOn isOpenMap_exp (isOpen_Ioo.preimage continuous_im) -#align complex.exp_local_homeomorph Complex.expLocalHomeomorph +#align complex.exp_local_homeomorph Complex.expPartialHomeomorph theorem hasStrictDerivAt_log {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) : HasStrictDerivAt log x⁻¹ x := have h0 : x ≠ 0 := by rintro rfl; simp [lt_irrefl] at h - expLocalHomeomorph.hasStrictDerivAt_symm h h0 <| by + expPartialHomeomorph.hasStrictDerivAt_symm h h0 <| by simpa [exp_log h0] using hasStrictDerivAt_exp (log x) #align complex.has_strict_deriv_at_log Complex.hasStrictDerivAt_log @@ -62,7 +62,7 @@ theorem hasStrictFDerivAt_log_real {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) : #align complex.has_strict_fderiv_at_log_real Complex.hasStrictFDerivAt_log_real theorem contDiffAt_log {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) {n : ℕ∞} : ContDiffAt ℂ n log x := - expLocalHomeomorph.contDiffAt_symm_deriv (exp_ne_zero <| log x) h (hasDerivAt_exp _) + expPartialHomeomorph.contDiffAt_symm_deriv (exp_ne_zero <| log x) h (hasDerivAt_exp _) contDiff_exp.contDiffAt #align complex.cont_diff_at_log Complex.contDiffAt_log @@ -148,4 +148,3 @@ theorem Differentiable.clog {f : E → ℂ} (h₁ : Differentiable ℂ f) #align differentiable.clog Differentiable.clog end LogDeriv - diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index dadb276935031..80db011c91c89 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -28,7 +28,7 @@ open scoped Real Topology /-- The polar coordinates local homeomorphism in `ℝ^2`, mapping `(r cos θ, r sin θ)` to `(r, θ)`. It is a homeomorphism between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. -/ @[simps] -def polarCoord : LocalHomeomorph (ℝ × ℝ) (ℝ × ℝ) where +def polarCoord : PartialHomeomorph (ℝ × ℝ) (ℝ × ℝ) where toFun q := (Real.sqrt (q.1 ^ 2 + q.2 ^ 2), Complex.arg (Complex.equivRealProd.symm q)) invFun p := (p.1 * cos p.2, p.1 * sin p.2) source := {q | 0 < q.1} ∪ {q | q.2 ≠ 0} @@ -160,8 +160,8 @@ open scoped Real /-- The polar coordinates local homeomorphism in `ℂ`, mapping `r (cos θ + I * sin θ)` to `(r, θ)`. It is a homeomorphism between `ℂ - ℝ≤0` and `(0, +∞) × (-π, π)`. -/ -protected noncomputable def polarCoord : LocalHomeomorph ℂ (ℝ × ℝ) := - equivRealProdClm.toHomeomorph.toLocalHomeomorph.trans polarCoord +protected noncomputable def polarCoord : PartialHomeomorph ℂ (ℝ × ℝ) := + equivRealProdClm.toHomeomorph.toPartialHomeomorph.trans polarCoord protected theorem polarCoord_apply (a : ℂ) : Complex.polarCoord a = (Complex.abs a, Complex.arg a) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean index 23fb378c841d7..dda00db3ac55d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean +++ b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean @@ -28,7 +28,7 @@ namespace Real /-- Local homeomorph between `(0, +∞)` and `(0, +∞)` with `toFun = (· ^ 2)` and `invFun = Real.sqrt`. -/ -noncomputable def sqLocalHomeomorph : LocalHomeomorph ℝ ℝ where +noncomputable def sqPartialHomeomorph : PartialHomeomorph ℝ ℝ where toFun x := x ^ 2 invFun := sqrt source := Ioi 0 @@ -41,7 +41,7 @@ noncomputable def sqLocalHomeomorph : LocalHomeomorph ℝ ℝ where open_target := isOpen_Ioi continuousOn_toFun := (continuous_pow 2).continuousOn continuousOn_invFun := continuousOn_id.sqrt -#align real.sq_local_homeomorph Real.sqLocalHomeomorph +#align real.sq_local_homeomorph Real.sqPartialHomeomorph theorem deriv_sqrt_aux {x : ℝ} (hx : x ≠ 0) : HasStrictDerivAt sqrt (1 / (2 * sqrt x)) x ∧ ∀ n, ContDiffAt ℝ n sqrt x := by @@ -53,8 +53,8 @@ theorem deriv_sqrt_aux {x : ℝ} (hx : x ≠ 0) : contDiffAt_const.congr_of_eventuallyEq this⟩ · have : ↑2 * sqrt x ^ (2 - 1) ≠ 0 := by simp [(sqrt_pos.2 hx).ne', @two_ne_zero ℝ] constructor - · simpa using sqLocalHomeomorph.hasStrictDerivAt_symm hx this (hasStrictDerivAt_pow 2 _) - · exact fun n => sqLocalHomeomorph.contDiffAt_symm_deriv this hx (hasDerivAt_pow 2 (sqrt x)) + · simpa using sqPartialHomeomorph.hasStrictDerivAt_symm hx this (hasStrictDerivAt_pow 2 _) + · exact fun n => sqPartialHomeomorph.contDiffAt_symm_deriv this hx (hasDerivAt_pow 2 (sqrt x)) (contDiffAt_id.pow 2) #align real.deriv_sqrt_aux Real.deriv_sqrt_aux diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index bd7d6df676705..eedb24944e24c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -11,7 +11,7 @@ import Mathlib.Analysis.SpecialFunctions.Trigonometric.Complex # The `arctan` function. Inequalities, derivatives, -and `Real.tan` as a `LocalHomeomorph` between `(-(π / 2), π / 2)` and the whole line. +and `Real.tan` as a `PartialHomeomorph` between `(-(π / 2), π / 2)` and the whole line. -/ @@ -209,8 +209,8 @@ theorem continuousAt_arctan {x : ℝ} : ContinuousAt arctan x := continuous_arctan.continuousAt #align real.continuous_at_arctan Real.continuousAt_arctan -/-- `Real.tan` as a `LocalHomeomorph` between `(-(π / 2), π / 2)` and the whole line. -/ -def tanLocalHomeomorph : LocalHomeomorph ℝ ℝ where +/-- `Real.tan` as a `PartialHomeomorph` between `(-(π / 2), π / 2)` and the whole line. -/ +def tanPartialHomeomorph : PartialHomeomorph ℝ ℝ where toFun := tan invFun := arctan source := Ioo (-(π / 2)) (π / 2) @@ -223,16 +223,16 @@ def tanLocalHomeomorph : LocalHomeomorph ℝ ℝ where open_target := isOpen_univ continuousOn_toFun := continuousOn_tan_Ioo continuousOn_invFun := continuous_arctan.continuousOn -#align real.tan_local_homeomorph Real.tanLocalHomeomorph +#align real.tan_local_homeomorph Real.tanPartialHomeomorph @[simp] -theorem coe_tanLocalHomeomorph : ⇑tanLocalHomeomorph = tan := +theorem coe_tanPartialHomeomorph : ⇑tanPartialHomeomorph = tan := rfl -#align real.coe_tan_local_homeomorph Real.coe_tanLocalHomeomorph +#align real.coe_tan_local_homeomorph Real.coe_tanPartialHomeomorph @[simp] -theorem coe_tanLocalHomeomorph_symm : ⇑tanLocalHomeomorph.symm = arctan := +theorem coe_tanPartialHomeomorph_symm : ⇑tanPartialHomeomorph.symm = arctan := rfl -#align real.coe_tan_local_homeomorph_symm Real.coe_tanLocalHomeomorph_symm +#align real.coe_tan_local_homeomorph_symm Real.coe_tanPartialHomeomorph_symm end Real diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean index d2798e45a2854..4f017ca41a53e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean @@ -82,7 +82,7 @@ theorem differentiableAt_tan_of_mem_Ioo {x : ℝ} (h : x ∈ Ioo (-(π / 2) : theorem hasStrictDerivAt_arctan (x : ℝ) : HasStrictDerivAt arctan (1 / (1 + x ^ 2)) x := by have A : cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne' simpa [cos_sq_arctan] using - tanLocalHomeomorph.hasStrictDerivAt_symm trivial (by simpa) (hasStrictDerivAt_tan A) + tanPartialHomeomorph.hasStrictDerivAt_symm trivial (by simpa) (hasStrictDerivAt_tan A) #align real.has_strict_deriv_at_arctan Real.hasStrictDerivAt_arctan theorem hasDerivAt_arctan (x : ℝ) : HasDerivAt arctan (1 / (1 + x ^ 2)) x := @@ -105,7 +105,7 @@ theorem deriv_arctan : deriv arctan = fun (x : ℝ) => 1 / (1 + x ^ 2) := theorem contDiff_arctan {n : ℕ∞} : ContDiff ℝ n arctan := contDiff_iff_contDiffAt.2 fun x => have : cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne' - tanLocalHomeomorph.contDiffAt_symm_deriv (by simpa) trivial (hasDerivAt_tan this) + tanPartialHomeomorph.contDiffAt_symm_deriv (by simpa) trivial (hasDerivAt_tan this) (contDiffAt_tan.2 this) #align real.cont_diff_arctan Real.contDiff_arctan diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean index 7a57e3523fd64..a13666b448a01 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean @@ -278,9 +278,9 @@ theorem mapsTo_sin_Ioo : MapsTo sin (Ioo (-(π / 2)) (π / 2)) (Ioo (-1) 1) := f rwa [mem_Ioo, ← arcsin_lt_pi_div_two, ← neg_pi_div_two_lt_arcsin, arcsin_sin h.1.le h.2.le] #align real.maps_to_sin_Ioo Real.mapsTo_sin_Ioo -/-- `Real.sin` as a `LocalHomeomorph` between `(-π / 2, π / 2)` and `(-1, 1)`. -/ +/-- `Real.sin` as a `PartialHomeomorph` between `(-π / 2, π / 2)` and `(-1, 1)`. -/ @[simp] -def sinLocalHomeomorph : LocalHomeomorph ℝ ℝ where +def sinPartialHomeomorph : PartialHomeomorph ℝ ℝ where toFun := sin invFun := arcsin source := Ioo (-(π / 2)) (π / 2) @@ -293,7 +293,7 @@ def sinLocalHomeomorph : LocalHomeomorph ℝ ℝ where open_target := isOpen_Ioo continuousOn_toFun := continuous_sin.continuousOn continuousOn_invFun := continuous_arcsin.continuousOn -#align real.sin_local_homeomorph Real.sinLocalHomeomorph +#align real.sin_local_homeomorph Real.sinPartialHomeomorph theorem cos_arcsin_nonneg (x : ℝ) : 0 ≤ cos (arcsin x) := cos_nonneg_of_mem_Icc ⟨neg_pi_div_two_le_arcsin _, arcsin_le_pi_div_two _⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean index 3617da9898c2d..a4c3b44502c5e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean @@ -39,8 +39,8 @@ theorem deriv_arcsin_aux {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : cases' h₂.lt_or_lt with h₂ h₂ · have : 0 < sqrt (1 - x ^ 2) := sqrt_pos.2 (by nlinarith [h₁, h₂]) simp only [← cos_arcsin, one_div] at this ⊢ - exact ⟨sinLocalHomeomorph.hasStrictDerivAt_symm ⟨h₁, h₂⟩ this.ne' (hasStrictDerivAt_sin _), - sinLocalHomeomorph.contDiffAt_symm_deriv this.ne' ⟨h₁, h₂⟩ (hasDerivAt_sin _) + exact ⟨sinPartialHomeomorph.hasStrictDerivAt_symm ⟨h₁, h₂⟩ this.ne' (hasStrictDerivAt_sin _), + sinPartialHomeomorph.contDiffAt_symm_deriv this.ne' ⟨h₁, h₂⟩ (hasDerivAt_sin _) contDiff_sin.contDiffAt⟩ · have : 1 - x ^ 2 < 0 := by nlinarith [h₂] rw [sqrt_eq_zero'.2 this.le, div_zero] diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 64f744464d426..71716e955d3d2 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Init.Align -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph #align_import geometry.manifold.charted_space from "leanprover-community/mathlib"@"431589bce478b2229eba14b14a283250428217db" @@ -122,14 +122,14 @@ universe u variable {H : Type u} {H' : Type*} {M : Type*} {M' : Type*} {M'' : Type*} /- Notational shortcut for the composition of local homeomorphisms and local equivs, i.e., -`LocalHomeomorph.trans` and `LocalEquiv.trans`. +`PartialHomeomorph.trans` and `LocalEquiv.trans`. Note that, as is usual for equivs, the composition is from left to right, hence the direction of the arrow. -/ -scoped[Manifold] infixr:100 " ≫ₕ " => LocalHomeomorph.trans +scoped[Manifold] infixr:100 " ≫ₕ " => PartialHomeomorph.trans scoped[Manifold] infixr:100 " ≫ " => LocalEquiv.trans -open Set LocalHomeomorph Manifold -- Porting note: Added `Manifold` +open Set PartialHomeomorph Manifold -- Porting note: Added `Manifold` /-! ### Structure groupoids -/ @@ -164,21 +164,21 @@ We use primes in the structure names as we will reformulate them below (without /-- A structure groupoid is a set of local homeomorphisms of a topological space stable under composition and inverse. They appear in the definition of the smoothness class of a manifold. -/ structure StructureGroupoid (H : Type u) [TopologicalSpace H] where - members : Set (LocalHomeomorph H H) - trans' : ∀ e e' : LocalHomeomorph H H, e ∈ members → e' ∈ members → e ≫ₕ e' ∈ members - symm' : ∀ e : LocalHomeomorph H H, e ∈ members → e.symm ∈ members - id_mem' : LocalHomeomorph.refl H ∈ members - locality' : ∀ e : LocalHomeomorph H H, + members : Set (PartialHomeomorph H H) + trans' : ∀ e e' : PartialHomeomorph H H, e ∈ members → e' ∈ members → e ≫ₕ e' ∈ members + symm' : ∀ e : PartialHomeomorph H H, e ∈ members → e.symm ∈ members + id_mem' : PartialHomeomorph.refl H ∈ members + locality' : ∀ e : PartialHomeomorph H H, (∀ x ∈ e.source, ∃ s, IsOpen s ∧ x ∈ s ∧ e.restr s ∈ members) → e ∈ members - eq_on_source' : ∀ e e' : LocalHomeomorph H H, e ∈ members → e' ≈ e → e' ∈ members + eq_on_source' : ∀ e e' : PartialHomeomorph H H, e ∈ members → e' ≈ e → e' ∈ members #align structure_groupoid StructureGroupoid variable [TopologicalSpace H] -instance : Membership (LocalHomeomorph H H) (StructureGroupoid H) := - ⟨fun (e : LocalHomeomorph H H) (G : StructureGroupoid H) ↦ e ∈ G.members⟩ +instance : Membership (PartialHomeomorph H H) (StructureGroupoid H) := + ⟨fun (e : PartialHomeomorph H H) (G : StructureGroupoid H) ↦ e ∈ G.members⟩ -instance (H : Type u) [TopologicalSpace H] : SetLike (StructureGroupoid H) (LocalHomeomorph H H) +instance (H : Type u) [TopologicalSpace H] : SetLike (StructureGroupoid H) (PartialHomeomorph H H) where coe s := s.members coe_injective' N O h := by cases N; cases O; congr @@ -231,26 +231,26 @@ instance : InfSet (StructureGroupoid H) := intro e e' he he'e exact fun i hi => i.eq_on_source' e e' (he i hi) he'e)⟩ -theorem StructureGroupoid.trans (G : StructureGroupoid H) {e e' : LocalHomeomorph H H} (he : e ∈ G) - (he' : e' ∈ G) : e ≫ₕ e' ∈ G := +theorem StructureGroupoid.trans (G : StructureGroupoid H) {e e' : PartialHomeomorph H H} + (he : e ∈ G) (he' : e' ∈ G) : e ≫ₕ e' ∈ G := G.trans' e e' he he' #align structure_groupoid.trans StructureGroupoid.trans -theorem StructureGroupoid.symm (G : StructureGroupoid H) {e : LocalHomeomorph H H} (he : e ∈ G) : +theorem StructureGroupoid.symm (G : StructureGroupoid H) {e : PartialHomeomorph H H} (he : e ∈ G) : e.symm ∈ G := G.symm' e he #align structure_groupoid.symm StructureGroupoid.symm -theorem StructureGroupoid.id_mem (G : StructureGroupoid H) : LocalHomeomorph.refl H ∈ G := +theorem StructureGroupoid.id_mem (G : StructureGroupoid H) : PartialHomeomorph.refl H ∈ G := G.id_mem' #align structure_groupoid.id_mem StructureGroupoid.id_mem -theorem StructureGroupoid.locality (G : StructureGroupoid H) {e : LocalHomeomorph H H} +theorem StructureGroupoid.locality (G : StructureGroupoid H) {e : PartialHomeomorph H H} (h : ∀ x ∈ e.source, ∃ s, IsOpen s ∧ x ∈ s ∧ e.restr s ∈ G) : e ∈ G := G.locality' e h #align structure_groupoid.locality StructureGroupoid.locality -theorem StructureGroupoid.eq_on_source (G : StructureGroupoid H) {e e' : LocalHomeomorph H H} +theorem StructureGroupoid.eq_on_source (G : StructureGroupoid H) {e e' : PartialHomeomorph H H} (he : e ∈ G) (h : e' ≈ e) : e' ∈ G := G.eq_on_source' e e' he h #align structure_groupoid.eq_on_source StructureGroupoid.eq_on_source @@ -272,13 +272,13 @@ theorem StructureGroupoid.le_iff {G₁ G₂ : StructureGroupoid H} : G₁ ≤ G /-- The trivial groupoid, containing only the identity (and maps with empty source, as this is necessary from the definition). -/ def idGroupoid (H : Type u) [TopologicalSpace H] : StructureGroupoid H where - members := {LocalHomeomorph.refl H} ∪ { e : LocalHomeomorph H H | e.source = ∅ } + members := {PartialHomeomorph.refl H} ∪ { e : PartialHomeomorph H H | e.source = ∅ } trans' e e' he he' := by cases' he with he he · simpa only [mem_singleton_iff.1 he, refl_trans] · have : (e ≫ₕ e').source ⊆ e.source := sep_subset _ _ rw [he] at this - have : e ≫ₕ e' ∈ { e : LocalHomeomorph H H | e.source = ∅ } := eq_bot_iff.2 this + have : e ≫ₕ e' ∈ { e : PartialHomeomorph H H | e.source = ∅ } := eq_bot_iff.2 this exact (mem_union _ _ _).2 (Or.inr this) symm' e he := by cases' (mem_union _ _ _).1 he with E E @@ -297,7 +297,7 @@ def idGroupoid (H : Type u) [TopologicalSpace H] : StructureGroupoid H where rw [restr_source, open_s.interior_eq] exact ⟨hx, xs⟩ cases' hs with hs hs - · replace hs : LocalHomeomorph.restr e s = LocalHomeomorph.refl H + · replace hs : PartialHomeomorph.restr e s = PartialHomeomorph.refl H · simpa only using hs have : (e.restr s).source = univ := by rw [hs] @@ -328,7 +328,7 @@ instance instStructureGroupoidOrderBot : OrderBot (StructureGroupoid H) where bot := idGroupoid H bot_le := by intro u f hf - have hf : f ∈ {LocalHomeomorph.refl H} ∪ { e : LocalHomeomorph H H | e.source = ∅ } := hf + have hf : f ∈ {PartialHomeomorph.refl H} ∪ { e : PartialHomeomorph H H | e.source = ∅ } := hf simp only [singleton_union, mem_setOf_eq, mem_insert_iff] at hf cases' hf with hf hf · rw [hf] @@ -357,7 +357,7 @@ structure Pregroupoid (H : Type*) [TopologicalSpace H] where /-- Construct a groupoid of local homeos for which the map and its inverse have some property, from a pregroupoid asserting that this property is stable under composition. -/ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where - members := { e : LocalHomeomorph H H | PG.property e e.source ∧ PG.property e.symm e.target } + members := { e : PartialHomeomorph H H | PG.property e e.source ∧ PG.property e.symm e.target } trans' e e' he he' := by constructor · apply PG.comp he.1 he'.1 e.open_source e'.open_source @@ -372,7 +372,7 @@ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where rcases he x xu with ⟨s, s_open, xs, hs⟩ refine' ⟨s, s_open, xs, _⟩ convert hs.1 using 1 - dsimp [LocalHomeomorph.restr] + dsimp [PartialHomeomorph.restr] rw [s_open.interior_eq] · refine' PG.locality e.open_target fun x xu ↦ _ rcases he (e.symm x) (e.map_target xu) with ⟨s, s_open, xs, hs⟩ @@ -380,7 +380,7 @@ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where · exact ContinuousOn.isOpen_inter_preimage e.continuousOn_invFun e.open_target s_open · rw [← inter_assoc, inter_self] convert hs.2 using 1 - dsimp [LocalHomeomorph.restr] + dsimp [PartialHomeomorph.restr] rw [s_open.interior_eq] eq_on_source' e e' he ee' := by constructor @@ -396,7 +396,7 @@ def Pregroupoid.groupoid (PG : Pregroupoid H) : StructureGroupoid H where exact he.2 #align pregroupoid.groupoid Pregroupoid.groupoid -theorem mem_groupoid_of_pregroupoid {PG : Pregroupoid H} {e : LocalHomeomorph H H} : +theorem mem_groupoid_of_pregroupoid {PG : Pregroupoid H} {e : PartialHomeomorph H H} : e ∈ PG.groupoid ↔ PG.property e e.source ∧ PG.property e.symm e.target := Iff.rfl #align mem_groupoid_of_pregroupoid mem_groupoid_of_pregroupoid @@ -408,7 +408,7 @@ theorem groupoid_of_pregroupoid_le (PG₁ PG₂ : Pregroupoid H) exact ⟨h _ _ he.1, h _ _ he.2⟩ #align groupoid_of_pregroupoid_le groupoid_of_pregroupoid_le -theorem mem_pregroupoid_of_eq_on_source (PG : Pregroupoid H) {e e' : LocalHomeomorph H H} +theorem mem_pregroupoid_of_eq_on_source (PG : Pregroupoid H) {e e' : PartialHomeomorph H H} (he' : e ≈ e') (he : PG.property e e.source) : PG.property e' e'.source := by rw [← he'.1] exact PG.congr e.open_source he'.eqOn.symm he @@ -458,28 +458,28 @@ instance : CompleteLattice (StructureGroupoid H) := homeomorphisms to open subsets of the source. -/ class ClosedUnderRestriction (G : StructureGroupoid H) : Prop where closedUnderRestriction : - ∀ {e : LocalHomeomorph H H}, e ∈ G → ∀ s : Set H, IsOpen s → e.restr s ∈ G + ∀ {e : PartialHomeomorph H H}, e ∈ G → ∀ s : Set H, IsOpen s → e.restr s ∈ G #align closed_under_restriction ClosedUnderRestriction theorem closedUnderRestriction' {G : StructureGroupoid H} [ClosedUnderRestriction G] - {e : LocalHomeomorph H H} (he : e ∈ G) {s : Set H} (hs : IsOpen s) : e.restr s ∈ G := + {e : PartialHomeomorph H H} (he : e ∈ G) {s : Set H} (hs : IsOpen s) : e.restr s ∈ G := ClosedUnderRestriction.closedUnderRestriction he s hs #align closed_under_restriction' closedUnderRestriction' /-- The trivial restriction-closed groupoid, containing only local homeomorphisms equivalent to the restriction of the identity to the various open subsets. -/ def idRestrGroupoid : StructureGroupoid H where - members := { e | ∃ (s : Set H) (h : IsOpen s), e ≈ LocalHomeomorph.ofSet s h } + members := { e | ∃ (s : Set H) (h : IsOpen s), e ≈ PartialHomeomorph.ofSet s h } trans' := by rintro e e' ⟨s, hs, hse⟩ ⟨s', hs', hse'⟩ refine' ⟨s ∩ s', IsOpen.inter hs hs', _⟩ - have := LocalHomeomorph.EqOnSource.trans' hse hse' - rwa [LocalHomeomorph.ofSet_trans_ofSet] at this + have := PartialHomeomorph.EqOnSource.trans' hse hse' + rwa [PartialHomeomorph.ofSet_trans_ofSet] at this symm' := by rintro e ⟨s, hs, hse⟩ refine' ⟨s, hs, _⟩ rw [← ofSet_symm] - exact LocalHomeomorph.EqOnSource.symm' hse + exact PartialHomeomorph.EqOnSource.symm' hse id_mem' := ⟨univ, isOpen_univ, by simp only [mfld_simps, refl]⟩ locality' := by intro e h @@ -491,7 +491,7 @@ def idRestrGroupoid : StructureGroupoid H where refine' ⟨hx, _⟩ rw [hs.interior_eq] exact hxs - simpa only [mfld_simps] using LocalHomeomorph.EqOnSource.eqOn hes' hes + simpa only [mfld_simps] using PartialHomeomorph.EqOnSource.eqOn hes' hes eq_on_source' := by rintro e e' ⟨s, hs, hse⟩ hee' exact ⟨s, hs, Setoid.trans hee' hse⟩ @@ -506,7 +506,7 @@ instance closedUnderRestriction_idRestrGroupoid : ClosedUnderRestriction (@idRes ⟨by rintro e ⟨s', hs', he⟩ s hs use s' ∩ s, IsOpen.inter hs' hs - refine' Setoid.trans (LocalHomeomorph.EqOnSource.restr he s) _ + refine' Setoid.trans (PartialHomeomorph.EqOnSource.restr he s) _ exact ⟨by simp only [hs.interior_eq, mfld_simps], by simp only [mfld_simps, eqOn_refl]⟩⟩ #align closed_under_restriction_id_restr_groupoid closedUnderRestriction_idRestrGroupoid @@ -525,13 +525,13 @@ theorem closedUnderRestriction_iff_id_le (G : StructureGroupoid H) : -- rw [hs.interior_eq] -- simp only [mfld_simps] ext - · rw [LocalHomeomorph.restr_apply, LocalHomeomorph.refl_apply, id, ofSet_apply, id_eq] + · rw [PartialHomeomorph.restr_apply, PartialHomeomorph.refl_apply, id, ofSet_apply, id_eq] · simp [hs] · simp [hs.interior_eq] · intro h constructor intro e he s hs - rw [← ofSet_trans (e : LocalHomeomorph H H) hs] + rw [← ofSet_trans (e : PartialHomeomorph H H) hs] refine' G.trans _ he apply StructureGroupoid.le_iff.mp h exact idRestrGroupoid_mem hs @@ -558,18 +558,18 @@ sometimes as a real manifold over `ℝ^(2n)`. -/ @[ext] class ChartedSpace (H : Type*) [TopologicalSpace H] (M : Type*) [TopologicalSpace M] where - protected atlas : Set (LocalHomeomorph M H) - protected chartAt : M → LocalHomeomorph M H + protected atlas : Set (PartialHomeomorph M H) + protected chartAt : M → PartialHomeomorph M H protected mem_chart_source : ∀ x, x ∈ (chartAt x).source protected chart_mem_atlas : ∀ x, chartAt x ∈ atlas #align charted_space ChartedSpace abbrev atlas (H : Type*) [TopologicalSpace H] (M : Type*) [TopologicalSpace M] - [ChartedSpace H M] : Set (LocalHomeomorph M H) := + [ChartedSpace H M] : Set (PartialHomeomorph M H) := ChartedSpace.atlas abbrev chartAt (H : Type*) [TopologicalSpace H] {M : Type*} [TopologicalSpace M] - [ChartedSpace H M] (x : M) : LocalHomeomorph M H := + [ChartedSpace H M] (x : M) : PartialHomeomorph M H := ChartedSpace.chartAt x @[simp, mfld_simps] @@ -586,8 +586,8 @@ section ChartedSpace /-- Any space is a `ChartedSpace` modelled over itself, by just using the identity chart. -/ instance chartedSpaceSelf (H : Type*) [TopologicalSpace H] : ChartedSpace H H where - atlas := {LocalHomeomorph.refl H} - chartAt _ := LocalHomeomorph.refl H + atlas := {PartialHomeomorph.refl H} + chartAt _ := PartialHomeomorph.refl H mem_chart_source x := mem_univ x chart_mem_atlas _ := mem_singleton _ #align charted_space_self chartedSpaceSelf @@ -595,14 +595,14 @@ instance chartedSpaceSelf (H : Type*) [TopologicalSpace H] : ChartedSpace H H wh /-- In the trivial `ChartedSpace` structure of a space modelled over itself through the identity, the atlas members are just the identity. -/ @[simp, mfld_simps] -theorem chartedSpaceSelf_atlas {H : Type*} [TopologicalSpace H] {e : LocalHomeomorph H H} : - e ∈ atlas H H ↔ e = LocalHomeomorph.refl H := +theorem chartedSpaceSelf_atlas {H : Type*} [TopologicalSpace H] {e : PartialHomeomorph H H} : + e ∈ atlas H H ↔ e = PartialHomeomorph.refl H := Iff.rfl #align charted_space_self_atlas chartedSpaceSelf_atlas /-- In the model space, `chartAt` is always the identity. -/ theorem chartAt_self_eq {H : Type*} [TopologicalSpace H] {x : H} : - chartAt H x = LocalHomeomorph.refl H := rfl + chartAt H x = PartialHomeomorph.refl H := rfl #align chart_at_self_eq chartAt_self_eq section @@ -633,7 +633,7 @@ theorem achart_def (x : M) : achart H x = ⟨chartAt H x, chart_mem_atlas H x⟩ #align achart_def achart_def @[simp, mfld_simps] -theorem coe_achart (x : M) : (achart H x : LocalHomeomorph M H) = chartAt H x := +theorem coe_achart (x : M) : (achart H x : PartialHomeomorph M H) = chartAt H x := rfl #align coe_achart coe_achart @@ -684,11 +684,11 @@ theorem ChartedSpace.locallyCompactSpace [LocallyCompactSpace H] : LocallyCompac /-- If a topological space admits an atlas with locally connected charts, then the space itself is locally connected. -/ theorem ChartedSpace.locallyConnectedSpace [LocallyConnectedSpace H] : LocallyConnectedSpace M := by - let e : M → LocalHomeomorph M H := chartAt H + let e : M → PartialHomeomorph M H := chartAt H refine' locallyConnectedSpace_of_connected_bases (fun x s ↦ (e x).symm '' s) (fun x s ↦ (IsOpen s ∧ e x x ∈ s ∧ IsConnected s) ∧ s ⊆ (e x).target) _ _ · intro x - simpa only [LocalHomeomorph.symm_map_nhds_eq, mem_chart_source] using + simpa only [PartialHomeomorph.symm_map_nhds_eq, mem_chart_source] using ((LocallyConnectedSpace.open_connected_basis (e x x)).restrict_subset ((e x).open_target.mem_nhds (mem_chart_target H x))).map (e x).symm · rintro x s ⟨⟨-, -, hsconn⟩, hssubset⟩ @@ -700,7 +700,7 @@ modelled on `H`. -/ def ChartedSpace.comp (H : Type*) [TopologicalSpace H] (H' : Type*) [TopologicalSpace H'] (M : Type*) [TopologicalSpace M] [ChartedSpace H H'] [ChartedSpace H' M] : ChartedSpace H M where - atlas := image2 LocalHomeomorph.trans (atlas H' M) (atlas H H') + atlas := image2 PartialHomeomorph.trans (atlas H' M) (atlas H H') chartAt p := (chartAt H' p).trans (chartAt H (chartAt H' p p)) mem_chart_source p := by simp only [mfld_simps] chart_mem_atlas p := ⟨chartAt _ p, chartAt _ _, chart_mem_atlas _ p, chart_mem_atlas _ _, rfl⟩ @@ -783,7 +783,7 @@ construction of the atlas of product maps. -/ instance prodChartedSpace (H : Type*) [TopologicalSpace H] (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (H' : Type*) [TopologicalSpace H'] (M' : Type*) [TopologicalSpace M'] [ChartedSpace H' M'] : ChartedSpace (ModelProd H H') (M × M') where - atlas := image2 LocalHomeomorph.prod (atlas H M) (atlas H' M') + atlas := image2 PartialHomeomorph.prod (atlas H M) (atlas H' M') chartAt x := (chartAt H x.1).prod (chartAt H' x.2) mem_chart_source x := ⟨mem_chart_source H x.1, mem_chart_source H' x.2⟩ chart_mem_atlas x := mem_image2_of_mem (chart_mem_atlas H x.1) (chart_mem_atlas H' x.2) @@ -819,8 +819,8 @@ canonical construction of the atlas of finite product maps. -/ instance piChartedSpace {ι : Type*} [Fintype ι] (H : ι → Type*) [∀ i, TopologicalSpace (H i)] (M : ι → Type*) [∀ i, TopologicalSpace (M i)] [∀ i, ChartedSpace (H i) (M i)] : ChartedSpace (ModelPi H) (∀ i, M i) where - atlas := LocalHomeomorph.pi '' Set.pi univ fun _ ↦ atlas (H _) (M _) - chartAt f := LocalHomeomorph.pi fun i ↦ chartAt (H i) (f i) + atlas := PartialHomeomorph.pi '' Set.pi univ fun _ ↦ atlas (H _) (M _) + chartAt f := PartialHomeomorph.pi fun i ↦ chartAt (H i) (f i) mem_chart_source f i _ := mem_chart_source (H i) (f i) chart_mem_atlas f := mem_image_of_mem _ fun i _ ↦ chart_mem_atlas (H i) (f i) #align pi_charted_space piChartedSpace @@ -829,7 +829,7 @@ instance piChartedSpace {ι : Type*} [Fintype ι] (H : ι → Type*) [∀ i, Top theorem piChartedSpace_chartAt {ι : Type*} [Fintype ι] (H : ι → Type*) [∀ i, TopologicalSpace (H i)] (M : ι → Type*) [∀ i, TopologicalSpace (M i)] [∀ i, ChartedSpace (H i) (M i)] (f : ∀ i, M i) : - chartAt (H := ModelPi H) f = LocalHomeomorph.pi fun i ↦ chartAt (H i) (f i) := + chartAt (H := ModelPi H) f = PartialHomeomorph.pi fun i ↦ chartAt (H i) (f i) := rfl #align pi_charted_space_chart_at piChartedSpace_chartAt @@ -882,7 +882,7 @@ theorem open_target (he : e ∈ c.atlas) : IsOpen e.target := by for the topology constructed from this atlas. The `localHomeomorph` version is given in this definition. -/ protected def localHomeomorph (e : LocalEquiv M H) (he : e ∈ c.atlas) : - @LocalHomeomorph M H c.toTopologicalSpace _ := + @PartialHomeomorph M H c.toTopologicalSpace _ := { c.toTopologicalSpace, e with open_source := by convert c.open_source' he open_target := by convert c.open_target he @@ -938,7 +938,7 @@ variable [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] groupoid. -/ class HasGroupoid {H : Type*} [TopologicalSpace H] (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (G : StructureGroupoid H) : Prop where - compatible : ∀ {e e' : LocalHomeomorph M H}, e ∈ atlas H M → e' ∈ atlas H M → e.symm ≫ₕ e' ∈ G + compatible : ∀ {e e' : PartialHomeomorph M H}, e ∈ atlas H M → e' ∈ atlas H M → e.symm ≫ₕ e' ∈ G #align has_groupoid HasGroupoid /-- Reformulate in the `StructureGroupoid` namespace the compatibility condition of charts in a @@ -946,7 +946,7 @@ charted space admitting a structure groupoid, to make it more easily accessible notation. -/ theorem StructureGroupoid.compatible {H : Type*} [TopologicalSpace H] (G : StructureGroupoid H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [HasGroupoid M G] - {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) : e.symm ≫ₕ e' ∈ G := + {e e' : PartialHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) : e.symm ≫ₕ e' ∈ G := HasGroupoid.compatible he he' #align structure_groupoid.compatible StructureGroupoid.compatible @@ -960,7 +960,7 @@ theorem hasGroupoid_inf_iff {G₁ G₂ : StructureGroupoid H} : HasGroupoid M (G ⟨(fun h ↦ ⟨hasGroupoid_of_le h inf_le_left, hasGroupoid_of_le h inf_le_right⟩), fun ⟨h1, h2⟩ ↦ { compatible := fun he he' ↦ ⟨h1.compatible he he', h2.compatible he he'⟩ }⟩ -theorem hasGroupoid_of_pregroupoid (PG : Pregroupoid H) (h : ∀ {e e' : LocalHomeomorph M H}, +theorem hasGroupoid_of_pregroupoid (PG : Pregroupoid H) (h : ∀ {e e' : PartialHomeomorph M H}, e ∈ atlas H M → e' ∈ atlas H M → PG.property (e.symm ≫ₕ e') (e.symm ≫ₕ e').source) : HasGroupoid M PG.groupoid := ⟨fun he he' ↦ mem_groupoid_of_pregroupoid.mpr ⟨h he he', h he' he⟩⟩ @@ -988,7 +988,7 @@ variable (M) (G : StructureGroupoid H) /-- Given a charted space admitting a structure groupoid, the maximal atlas associated to this structure groupoid is the set of all local charts that are compatible with the atlas, i.e., such that changing coordinates with an atlas member gives an element of the groupoid. -/ -def StructureGroupoid.maximalAtlas : Set (LocalHomeomorph M H) := +def StructureGroupoid.maximalAtlas : Set (PartialHomeomorph M H) := { e | ∀ e' ∈ atlas H M, e.symm ≫ₕ e' ∈ G ∧ e'.symm ≫ₕ e ∈ G } #align structure_groupoid.maximal_atlas StructureGroupoid.maximalAtlas @@ -1006,14 +1006,14 @@ theorem StructureGroupoid.chart_mem_maximalAtlas [HasGroupoid M G] (x : M) : variable {G} -theorem mem_maximalAtlas_iff {e : LocalHomeomorph M H} : +theorem mem_maximalAtlas_iff {e : PartialHomeomorph M H} : e ∈ G.maximalAtlas M ↔ ∀ e' ∈ atlas H M, e.symm ≫ₕ e' ∈ G ∧ e'.symm ≫ₕ e ∈ G := Iff.rfl #align mem_maximal_atlas_iff mem_maximalAtlas_iff /-- Changing coordinates between two elements of the maximal atlas gives rise to an element of the structure groupoid. -/ -theorem StructureGroupoid.compatible_of_mem_maximalAtlas {e e' : LocalHomeomorph M H} +theorem StructureGroupoid.compatible_of_mem_maximalAtlas {e e' : PartialHomeomorph M H} (he : e ∈ G.maximalAtlas M) (he' : e' ∈ G.maximalAtlas M) : e.symm ≫ₕ e' ∈ G := by refine' G.locality fun x hx ↦ _ set f := chartAt (H := H) (e.symm x) @@ -1040,14 +1040,14 @@ theorem StructureGroupoid.compatible_of_mem_maximalAtlas {e e' : LocalHomeomorph variable (G) /-- In the model space, the identity is in any maximal atlas. -/ -theorem StructureGroupoid.id_mem_maximalAtlas : LocalHomeomorph.refl H ∈ G.maximalAtlas H := +theorem StructureGroupoid.id_mem_maximalAtlas : PartialHomeomorph.refl H ∈ G.maximalAtlas H := G.subset_maximalAtlas <| by simp #align structure_groupoid.id_mem_maximal_atlas StructureGroupoid.id_mem_maximalAtlas /-- In the model space, any element of the groupoid is in the maximal atlas. -/ -theorem StructureGroupoid.mem_maximalAtlas_of_mem_groupoid {f : LocalHomeomorph H H} (hf : f ∈ G) : - f ∈ G.maximalAtlas H := by - rintro e (rfl : e = LocalHomeomorph.refl H) +theorem StructureGroupoid.mem_maximalAtlas_of_mem_groupoid {f : PartialHomeomorph H H} + (hf : f ∈ G) : f ∈ G.maximalAtlas H := by + rintro e (rfl : e = PartialHomeomorph.refl H) exact ⟨G.trans (G.symm hf) G.id_mem, G.trans (G.symm G.id_mem) hf⟩ #align structure_groupoid.mem_maximal_atlas_of_mem_groupoid StructureGroupoid.mem_maximalAtlas_of_mem_groupoid @@ -1057,9 +1057,9 @@ section Singleton variable {α : Type*} [TopologicalSpace α] -namespace LocalHomeomorph +namespace PartialHomeomorph -variable (e : LocalHomeomorph α H) +variable (e : PartialHomeomorph α H) /-- If a single local homeomorphism `e` from a space `α` into `H` has source covering the whole space `α`, then that local homeomorphism induces an `H`-charted space structure on `α`. @@ -1070,23 +1070,23 @@ def singletonChartedSpace (h : e.source = Set.univ) : ChartedSpace H α where chartAt _ := e mem_chart_source _ := by rw [h]; apply mem_univ chart_mem_atlas _ := by tauto -#align local_homeomorph.singleton_charted_space LocalHomeomorph.singletonChartedSpace +#align local_homeomorph.singleton_charted_space PartialHomeomorph.singletonChartedSpace @[simp, mfld_simps] theorem singletonChartedSpace_chartAt_eq (h : e.source = Set.univ) {x : α} : @chartAt H _ α _ (e.singletonChartedSpace h) x = e := rfl -#align local_homeomorph.singleton_charted_space_chart_at_eq LocalHomeomorph.singletonChartedSpace_chartAt_eq +#align local_homeomorph.singleton_charted_space_chart_at_eq PartialHomeomorph.singletonChartedSpace_chartAt_eq theorem singletonChartedSpace_chartAt_source (h : e.source = Set.univ) {x : α} : (@chartAt H _ α _ (e.singletonChartedSpace h) x).source = Set.univ := h -#align local_homeomorph.singleton_charted_space_chart_at_source LocalHomeomorph.singletonChartedSpace_chartAt_source +#align local_homeomorph.singleton_charted_space_chart_at_source PartialHomeomorph.singletonChartedSpace_chartAt_source -theorem singletonChartedSpace_mem_atlas_eq (h : e.source = Set.univ) (e' : LocalHomeomorph α H) +theorem singletonChartedSpace_mem_atlas_eq (h : e.source = Set.univ) (e' : PartialHomeomorph α H) (h' : e' ∈ (e.singletonChartedSpace h).atlas) : e' = e := h' -#align local_homeomorph.singleton_charted_space_mem_atlas_eq LocalHomeomorph.singletonChartedSpace_mem_atlas_eq +#align local_homeomorph.singleton_charted_space_mem_atlas_eq PartialHomeomorph.singletonChartedSpace_mem_atlas_eq /-- Given a local homeomorphism `e` from a space `α` into `H`, if its source covers the whole space `α`, then the induced charted space structure on `α` is `HasGroupoid G` for any structure @@ -1101,18 +1101,18 @@ theorem singleton_hasGroupoid (h : e.source = Set.univ) (G : StructureGroupoid H refine' G.eq_on_source _ e.trans_symm_self have hle : idRestrGroupoid ≤ G := (closedUnderRestriction_iff_id_le G).mp (by assumption) exact StructureGroupoid.le_iff.mp hle _ (idRestrGroupoid_mem _) } -#align local_homeomorph.singleton_has_groupoid LocalHomeomorph.singleton_hasGroupoid +#align local_homeomorph.singleton_has_groupoid PartialHomeomorph.singleton_hasGroupoid -end LocalHomeomorph +end PartialHomeomorph namespace OpenEmbedding variable [Nonempty α] /-- An open embedding of `α` into `H` induces an `H`-charted space structure on `α`. -See `LocalHomeomorph.singletonChartedSpace`. -/ +See `PartialHomeomorph.singletonChartedSpace`. -/ def singletonChartedSpace {f : α → H} (h : OpenEmbedding f) : ChartedSpace H α := - (h.toLocalHomeomorph f).singletonChartedSpace (toLocalHomeomorph_source _ _) + (h.toPartialHomeomorph f).singletonChartedSpace (toPartialHomeomorph_source _ _) #align open_embedding.singleton_charted_space OpenEmbedding.singletonChartedSpace theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : OpenEmbedding f) {x : α} : @@ -1122,7 +1122,7 @@ theorem singletonChartedSpace_chartAt_eq {f : α → H} (h : OpenEmbedding f) {x theorem singleton_hasGroupoid {f : α → H} (h : OpenEmbedding f) (G : StructureGroupoid H) [ClosedUnderRestriction G] : @HasGroupoid _ _ _ _ h.singletonChartedSpace G := - (h.toLocalHomeomorph f).singleton_hasGroupoid (toLocalHomeomorph_source _ _) G + (h.toPartialHomeomorph f).singleton_hasGroupoid (toPartialHomeomorph_source _ _) G #align open_embedding.singleton_has_groupoid OpenEmbedding.singleton_hasGroupoid end OpenEmbedding @@ -1139,8 +1139,8 @@ variable (s : Opens M) /-- An open subset of a charted space is naturally a charted space. -/ protected instance instChartedSpace : ChartedSpace H s where - atlas := ⋃ x : s, {@LocalHomeomorph.subtypeRestr _ _ _ _ (chartAt H x.1) s ⟨x⟩} - chartAt x := @LocalHomeomorph.subtypeRestr _ _ _ _ (chartAt H x.1) s ⟨x⟩ + atlas := ⋃ x : s, {@PartialHomeomorph.subtypeRestr _ _ _ _ (chartAt H x.1) s ⟨x⟩} + chartAt x := @PartialHomeomorph.subtypeRestr _ _ _ _ (chartAt H x.1) s ⟨x⟩ mem_chart_source x := ⟨trivial, mem_chart_source H x.1⟩ chart_mem_atlas x := by simp only [mem_iUnion, mem_singleton_iff] @@ -1197,8 +1197,8 @@ and use structomorph instead. -/ -- @[nolint has_nonempty_instance] -- Porting note: commented out structure Structomorph (G : StructureGroupoid H) (M : Type*) (M' : Type*) [TopologicalSpace M] [TopologicalSpace M'] [ChartedSpace H M] [ChartedSpace H M'] extends Homeomorph M M' where - mem_groupoid : ∀ c : LocalHomeomorph M H, ∀ c' : LocalHomeomorph M' H, c ∈ atlas H M → - c' ∈ atlas H M' → c.symm ≫ₕ toHomeomorph.toLocalHomeomorph ≫ₕ c' ∈ G + mem_groupoid : ∀ c : PartialHomeomorph M H, ∀ c' : PartialHomeomorph M' H, c ∈ atlas H M → + c' ∈ atlas H M' → c.symm ≫ₕ toHomeomorph.toPartialHomeomorph ≫ₕ c' ∈ G #align structomorph Structomorph variable [TopologicalSpace M'] [TopologicalSpace M''] {G : StructureGroupoid H} [ChartedSpace H M'] @@ -1209,8 +1209,8 @@ def Structomorph.refl (M : Type*) [TopologicalSpace M] [ChartedSpace H M] [HasGr Structomorph G M M := { Homeomorph.refl M with mem_groupoid := fun c c' hc hc' ↦ by - change LocalHomeomorph.symm c ≫ₕ LocalHomeomorph.refl M ≫ₕ c' ∈ G - rw [LocalHomeomorph.refl_trans] + change PartialHomeomorph.symm c ≫ₕ PartialHomeomorph.refl M ≫ₕ c' ∈ G + rw [PartialHomeomorph.refl_trans] exact G.compatible hc hc' } #align structomorph.refl Structomorph.refl @@ -1219,7 +1219,7 @@ def Structomorph.symm (e : Structomorph G M M') : Structomorph G M' M := { e.toHomeomorph.symm with mem_groupoid := by intro c c' hc hc' - have : (c'.symm ≫ₕ e.toHomeomorph.toLocalHomeomorph ≫ₕ c).symm ∈ G := + have : (c'.symm ≫ₕ e.toHomeomorph.toPartialHomeomorph ≫ₕ c).symm ∈ G := G.symm (e.mem_groupoid c' c hc' hc) rwa [trans_symm_eq_symm_trans_symm, trans_symm_eq_symm_trans_symm, symm_symm, trans_assoc] at this } @@ -1236,10 +1236,10 @@ def Structomorph.trans (e : Structomorph G M M') (e' : Structomorph G M' M'') : their composition is smooth, and it coincides with c' ∘ e' ∘ e ∘ c⁻¹ around x. -/ intro c c' hc hc' refine' G.locality fun x hx ↦ _ - let f₁ := e.toHomeomorph.toLocalHomeomorph - let f₂ := e'.toHomeomorph.toLocalHomeomorph - let f := (e.toHomeomorph.trans e'.toHomeomorph).toLocalHomeomorph - have feq : f = f₁ ≫ₕ f₂ := Homeomorph.trans_toLocalHomeomorph _ _ + let f₁ := e.toHomeomorph.toPartialHomeomorph + let f₂ := e'.toHomeomorph.toPartialHomeomorph + let f := (e.toHomeomorph.trans e'.toHomeomorph).toPartialHomeomorph + have feq : f = f₁ ≫ₕ f₂ := Homeomorph.trans_toPartialHomeomorph _ _ -- define the atlas g around y let y := (c.symm ≫ₕ f₁) x let g := chartAt (H := H) y @@ -1250,7 +1250,7 @@ def Structomorph.trans (e : Structomorph G M M') (e' : Structomorph G M' M'') : apply (c.symm ≫ₕ f₁).continuousOn_toFun.isOpen_inter_preimage <;> apply open_source have : x ∈ s := by constructor - · simp only [trans_source, preimage_univ, inter_univ, Homeomorph.toLocalHomeomorph_source] + · simp only [trans_source, preimage_univ, inter_univ, Homeomorph.toPartialHomeomorph_source] rw [trans_source] at hx exact hx.1 · exact hg₂ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 3bb94c4a4194f..e709286238f70 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -27,7 +27,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -- declare functions, sets, points and smoothness indices - {e : LocalHomeomorph M H} {x : M} {m n : ℕ∞} + {e : PartialHomeomorph M H} {x : M} {m n : ℕ∞} /-! ### Atlas members are smooth -/ @@ -117,7 +117,7 @@ theorem contMDiffOn_extChartAt_symm (x : M) : #align cont_mdiff_on_ext_chart_at_symm contMDiffOn_extChartAt_symm /-- An element of `contDiffGroupoid ⊤ I` is `C^n` for any `n`. -/ -theorem contMDiffOn_of_mem_contDiffGroupoid {e' : LocalHomeomorph H H} +theorem contMDiffOn_of_mem_contDiffGroupoid {e' : PartialHomeomorph H H} (h : e' ∈ contDiffGroupoid ⊤ I) : ContMDiffOn I I n e' e'.source := (contDiffWithinAt_localInvariantProp I I n).liftPropOn_of_mem_groupoid (contDiffWithinAtProp_id I) h @@ -131,7 +131,7 @@ section IsLocalStructomorph variable [ChartedSpace H M'] [IsM' : SmoothManifoldWithCorners I M'] -theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : LocalHomeomorph M M'} +theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M M'} (hf : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source) : SmoothOn I I f f.source := by -- It suffices to show smoothness near each `x` @@ -177,7 +177,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : LocalHomeomorph M M' /-- Let `M` and `M'` be smooth manifolds with the same model-with-corners, `I`. Then `f : M → M'` is a local structomorphism for `I`, if and only if it is manifold-smooth on the domain of definition in both directions. -/ -theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : LocalHomeomorph M M') : +theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source ↔ SmoothOn I I f f.source ∧ SmoothOn I I f.symm f.target := by constructor @@ -193,7 +193,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : LocalHomeomorph M M') : obtain ⟨-, hxf⟩ := h x hx refine' ⟨(f.symm.continuousAt hX).continuousWithinAt, fun h2x => _⟩ obtain ⟨e, he, h2e, hef, hex⟩ : - ∃ e : LocalHomeomorph H H, + ∃ e : PartialHomeomorph H H, e ∈ contDiffGroupoid ⊤ I ∧ e.source ⊆ (c.symm ≫ₕ f ≫ₕ c').source ∧ EqOn (c' ∘ f ∘ c.symm) e e.source ∧ c x ∈ e.source := by @@ -201,7 +201,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : LocalHomeomorph M M') : have h2 : c' ∘ f ∘ c.symm = ⇑(c.symm ≫ₕ f ≫ₕ c') := rfl have hcx : c x ∈ c.symm ⁻¹' f.source := by simp only [hx, mfld_simps] rw [h2] - rw [← h1, h2, LocalHomeomorph.isLocalStructomorphWithinAt_iff'] at hxf + rw [← h1, h2, PartialHomeomorph.isLocalStructomorphWithinAt_iff'] at hxf · exact hxf hcx · mfld_set_tac · apply Or.inl @@ -219,7 +219,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : LocalHomeomorph M M') : · rw [inter_self]; exact hef.symm have h2 : e.target ⊆ (c.symm ≫ₕ f ≫ₕ c').target := by intro x hx; rw [← e.right_inv hx, ← hef (e.symm.mapsTo hx)] - exact LocalHomeomorph.mapsTo _ (h2e <| e.symm.mapsTo hx) + exact PartialHomeomorph.mapsTo _ (h2e <| e.symm.mapsTo hx) rw [inter_self] at h1 rwa [inter_eq_right.mpr] refine' h2.trans _ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index 7abad04c84dd4..b0bd432410d7e 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -53,8 +53,8 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] -- declare functions, sets, points and smoothness indices - {e : LocalHomeomorph M H} - {e' : LocalHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} + {e : PartialHomeomorph M H} + {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} variable {I I'} /-! ### Smoothness of the composition of smooth functions between manifolds -/ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index cefea4f4b6fd5..056396fd3da2c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -84,8 +84,8 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] -- declare functions, sets, points and smoothness indices - {e : LocalHomeomorph M H} - {e' : LocalHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} + {e : PartialHomeomorph M H} + {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} /-- Property in the model space of a model with corners of being `C^n` within at set at a point, when read in the model vector space. This property will be lifted to manifolds to define smooth @@ -359,9 +359,9 @@ theorem contMDiffWithinAt_iff_target : ContinuousWithinAt f s x ∧ ContinuousWithinAt (extChartAt I' (f x) ∘ f) s x ↔ ContinuousWithinAt f s x := and_iff_left_of_imp <| (continuousAt_extChartAt _ _).comp_continuousWithinAt - simp_rw [cont, ContDiffWithinAtProp, extChartAt, LocalHomeomorph.extend, LocalEquiv.coe_trans, - ModelWithCorners.toLocalEquiv_coe, LocalHomeomorph.coe_coe, modelWithCornersSelf_coe, - chartAt_self_eq, LocalHomeomorph.refl_apply, comp.left_id] + simp_rw [cont, ContDiffWithinAtProp, extChartAt, PartialHomeomorph.extend, LocalEquiv.coe_trans, + ModelWithCorners.toLocalEquiv_coe, PartialHomeomorph.coe_coe, modelWithCornersSelf_coe, + chartAt_self_eq, PartialHomeomorph.refl_apply, comp.left_id] rfl #align cont_mdiff_within_at_iff_target contMDiffWithinAt_iff_target @@ -582,7 +582,7 @@ theorem contMDiffOn_iff_target : ∀ y : M', ContMDiffOn I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) (s ∩ f ⁻¹' (extChartAt I' y).source) := by simp only [contMDiffOn_iff, ModelWithCorners.source_eq, chartAt_self_eq, - LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_trans, extChartAt, LocalHomeomorph.extend, + PartialHomeomorph.refl_localEquiv, LocalEquiv.refl_trans, extChartAt, PartialHomeomorph.extend, Set.preimage_univ, Set.inter_univ, and_congr_right_iff] intro h constructor diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 74eb5da9a250f..911ccddd110b8 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -50,8 +50,8 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] -- declare functions, sets, points and smoothness indices - {e : LocalHomeomorph M H} - {e' : LocalHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} + {e : PartialHomeomorph M H} + {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} variable {I I'} section ProdMk diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index f12513c1cf22c..c5a559558e690 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -567,7 +567,7 @@ theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : rcases p with ⟨x, v⟩ have N : I.symm ⁻¹' (chartAt H x).target ∈ 𝓝 (I ((chartAt H x) x)) := by apply IsOpen.mem_nhds - apply (LocalHomeomorph.open_target _).preimage I.continuous_invFun + apply (PartialHomeomorph.open_target _).preimage I.continuous_invFun simp only [mfld_simps] have A : MDifferentiableAt I I.tangent (fun x => @TotalSpace.mk M E (TangentSpace I) x 0) x := haveI : Smooth I (I.prod 𝓘(𝕜, E)) (zeroSection E (TangentSpace I : M → Type _)) := diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 5de132745d0d8..c89faec79e7cb 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -374,10 +374,10 @@ theorem contMDiff_diffeomorph_comp_iff {m} (h : M ≃ₘ^n⟮I, J⟯ N) {f : M' forall_congr' fun _ => h.contMDiffWithinAt_diffeomorph_comp_iff hm #align diffeomorph.cont_mdiff_diffeomorph_comp_iff Diffeomorph.contMDiff_diffeomorph_comp_iff -theorem toLocalHomeomorph_mdifferentiable (h : M ≃ₘ^n⟮I, J⟯ N) (hn : 1 ≤ n) : - h.toHomeomorph.toLocalHomeomorph.MDifferentiable I J := +theorem toPartialHomeomorph_mdifferentiable (h : M ≃ₘ^n⟮I, J⟯ N) (hn : 1 ≤ n) : + h.toHomeomorph.toPartialHomeomorph.MDifferentiable I J := ⟨h.mdifferentiableOn _ hn, h.symm.mdifferentiableOn _ hn⟩ -#align diffeomorph.to_local_homeomorph_mdifferentiable Diffeomorph.toLocalHomeomorph_mdifferentiable +#align diffeomorph.to_local_homeomorph_mdifferentiable Diffeomorph.toPartialHomeomorph_mdifferentiable section Constructions @@ -442,7 +442,7 @@ variable [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners J N] theorem uniqueMDiffOn_image_aux (h : M ≃ₘ^n⟮I, J⟯ N) (hn : 1 ≤ n) {s : Set M} (hs : UniqueMDiffOn I s) : UniqueMDiffOn J (h '' s) := by - convert hs.uniqueMDiffOn_preimage (h.toLocalHomeomorph_mdifferentiable hn) + convert hs.uniqueMDiffOn_preimage (h.toPartialHomeomorph_mdifferentiable hn) simp [h.image_eq_preimage] #align diffeomorph.unique_mdiff_on_image_aux Diffeomorph.uniqueMDiffOn_image_aux diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index adfa58088fd11..462fe126222ec 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -166,7 +166,7 @@ scoped[Manifold] `EuclideanHalfSpace 1`. -/ def IccLeftChart (x y : ℝ) [h : Fact (x < y)] : - LocalHomeomorph (Icc x y) (EuclideanHalfSpace 1) where + PartialHomeomorph (Icc x y) (EuclideanHalfSpace 1) where source := { z : Icc x y | z.val < y } target := { z : EuclideanHalfSpace 1 | z.val 0 < y - x } toFun := fun z : Icc x y => ⟨fun _ => z.val - x, sub_nonneg.mpr z.property.1⟩ @@ -214,7 +214,7 @@ def IccLeftChart (x y : ℝ) [h : Fact (x < y)] : `EuclideanHalfSpace 1`. -/ def IccRightChart (x y : ℝ) [h : Fact (x < y)] : - LocalHomeomorph (Icc x y) (EuclideanHalfSpace 1) where + PartialHomeomorph (Icc x y) (EuclideanHalfSpace 1) where source := { z : Icc x y | x < z.val } target := { z : EuclideanHalfSpace 1 | z.val 0 < y - x } toFun z := ⟨fun _ => y - z.val, sub_nonneg.mpr z.property.2⟩ diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index c39df393555b8..3cfb5ef76fcc1 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -281,7 +281,7 @@ theorem stereo_right_inv (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : stereoToFun v /-- Stereographic projection from the unit sphere in `E`, centred at a unit vector `v` in `E`; this is the version as a local homeomorphism. -/ -def stereographic (hv : ‖v‖ = 1) : LocalHomeomorph (sphere (0 : E) 1) (ℝ ∙ v)ᗮ where +def stereographic (hv : ‖v‖ = 1) : PartialHomeomorph (sphere (0 : E) 1) (ℝ ∙ v)ᗮ where toFun := stereoToFun v ∘ (↑) invFun := stereoInvFun hv source := {⟨v, by simp [hv]⟩}ᶜ @@ -364,10 +364,10 @@ space `E`. This version has codomain the Euclidean space of dimension `n`, and composing the original sterographic projection (`stereographic`) with an arbitrary linear isometry from `(ℝ ∙ v)ᗮ` to the Euclidean space. -/ def stereographic' (n : ℕ) [Fact (finrank ℝ E = n + 1)] (v : sphere (0 : E) 1) : - LocalHomeomorph (sphere (0 : E) 1) (EuclideanSpace ℝ (Fin n)) := + PartialHomeomorph (sphere (0 : E) 1) (EuclideanSpace ℝ (Fin n)) := stereographic (norm_eq_of_mem_sphere v) ≫ₕ (OrthonormalBasis.fromOrthogonalSpanSingleton n - (ne_zero_of_mem_unit_sphere v)).repr.toHomeomorph.toLocalHomeomorph + (ne_zero_of_mem_unit_sphere v)).repr.toHomeomorph.toPartialHomeomorph #align stereographic' stereographic' @[simp] @@ -430,12 +430,12 @@ instance smoothMfldWithCorners {n : ℕ} [Fact (finrank ℝ E = n + 1)] : (ℝ ∙ (v : E))ᗮ.subtypeL.contDiff).comp U.symm.contDiff convert H₁.comp' (H₂.contDiffOn : ContDiffOn ℝ ⊤ _ Set.univ) using 1 -- -- squeezed from `ext, simp [sphere_ext_iff, stereographic'_symm_apply, real_inner_comm]` - simp only [LocalHomeomorph.trans_toLocalEquiv, LocalHomeomorph.symm_toLocalEquiv, + simp only [PartialHomeomorph.trans_toLocalEquiv, PartialHomeomorph.symm_toLocalEquiv, LocalEquiv.trans_source, LocalEquiv.symm_source, stereographic'_target, stereographic'_source] simp only [modelWithCornersSelf_coe, modelWithCornersSelf_coe_symm, Set.preimage_id, Set.range_id, Set.inter_univ, Set.univ_inter, Set.compl_singleton_eq, Set.preimage_setOf_eq] - simp only [id.def, comp_apply, Submodule.subtypeL_apply, LocalHomeomorph.coe_coe_symm, + simp only [id.def, comp_apply, Submodule.subtypeL_apply, PartialHomeomorph.coe_coe_symm, innerSL_apply, Ne.def, sphere_ext_iff, real_inner_comm (v' : E)] rfl) diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 5a71cc39c1d09..49f80e0477536 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -69,10 +69,10 @@ variable (G : StructureGroupoid H) (G' : StructureGroupoid H') to charted spaces admitting these groupoids will inherit the good behavior. -/ structure LocalInvariantProp (P : (H → H') → Set H → H → Prop) : Prop where is_local : ∀ {s x u} {f : H → H'}, IsOpen u → x ∈ u → (P f s x ↔ P f (s ∩ u) x) - right_invariance' : ∀ {s x f} {e : LocalHomeomorph H H}, + right_invariance' : ∀ {s x f} {e : PartialHomeomorph H H}, e ∈ G → x ∈ e.source → P f s x → P (f ∘ e.symm) (e.symm ⁻¹' s) (e x) congr_of_forall : ∀ {s x} {f g : H → H'}, (∀ y ∈ s, f y = g y) → f x = g x → P f s x → P g s x - left_invariance' : ∀ {s x f} {e' : LocalHomeomorph H' H'}, + left_invariance' : ∀ {s x f} {e' : PartialHomeomorph H' H'}, e' ∈ G' → s ⊆ f ⁻¹' e'.source → f x ∈ e'.source → P f s x → P (e' ∘ f) s x #align structure_groupoid.local_invariant_prop StructureGroupoid.LocalInvariantProp @@ -121,7 +121,7 @@ theorem congr' {s : Set H} {x : H} {f g : H → H'} (h : f =ᶠ[𝓝 x] g) (hP : hG.congr h.symm hP #align structure_groupoid.local_invariant_prop.congr' StructureGroupoid.LocalInvariantProp.congr' -theorem left_invariance {s : Set H} {x : H} {f : H → H'} {e' : LocalHomeomorph H' H'} +theorem left_invariance {s : Set H} {x : H} {f : H → H'} {e' : PartialHomeomorph H' H'} (he' : e' ∈ G') (hfs : ContinuousWithinAt f s x) (hxe' : f x ∈ e'.source) : P (e' ∘ f) s x ↔ P f s x := by have h2f := hfs.preimage_mem_nhdsWithin (e'.open_source.mem_nhds hxe') @@ -139,7 +139,7 @@ theorem left_invariance {s : Set H} {x : H} {f : H → H'} {e' : LocalHomeomorph exact hG.left_invariance' he' (inter_subset_right _ _) hxe' #align structure_groupoid.local_invariant_prop.left_invariance StructureGroupoid.LocalInvariantProp.left_invariance -theorem right_invariance {s : Set H} {x : H} {f : H → H'} {e : LocalHomeomorph H H} (he : e ∈ G) +theorem right_invariance {s : Set H} {x : H} {f : H → H'} {e : PartialHomeomorph H H} (he : e ∈ G) (hxe : x ∈ e.source) : P (f ∘ e.symm) (e.symm ⁻¹' s) (e x) ↔ P f s x := by refine' ⟨fun h ↦ _, hG.right_invariance' he hxe⟩ have := hG.right_invariance' (G.symm he) (e.mapsTo hxe) h @@ -208,8 +208,8 @@ open ChartedSpace namespace StructureGroupoid -variable {G : StructureGroupoid H} {G' : StructureGroupoid H'} {e e' : LocalHomeomorph M H} - {f f' : LocalHomeomorph M' H'} {P : (H → H') → Set H → H → Prop} {g g' : M → M'} {s t : Set M} +variable {G : StructureGroupoid H} {G' : StructureGroupoid H'} {e e' : PartialHomeomorph M H} + {f f' : PartialHomeomorph M' H'} {P : (H → H') → Set H → H → Prop} {g g' : M → M'} {s t : Set M} {x : M} {Q : (H → H) → Set H → H → Prop} theorem liftPropWithinAt_univ : LiftPropWithinAt P g univ x ↔ LiftPropAt P g x := Iff.rfl @@ -246,8 +246,8 @@ theorem liftPropWithinAt_iff {f : M → M'} : ((chartAt H x).target ∩ (chartAt H x).symm ⁻¹' (s ∩ f ⁻¹' (chartAt H' (f x)).source)) (chartAt H x x) := by refine' and_congr_right fun hf ↦ hG.congr_set _ - exact LocalHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter hf (mem_chart_source H x) - (chart_source_mem_nhds H' (f x)) + exact PartialHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter hf + (mem_chart_source H x) (chart_source_mem_nhds H' (f x)) #align structure_groupoid.local_invariant_prop.lift_prop_within_at_iff StructureGroupoid.LocalInvariantProp.liftPropWithinAt_iff theorem liftPropWithinAt_indep_chart_source_aux (g : M → H') (he : e ∈ G.maximalAtlas M) @@ -255,13 +255,13 @@ theorem liftPropWithinAt_indep_chart_source_aux (g : M → H') (he : e ∈ G.max P (g ∘ e.symm) (e.symm ⁻¹' s) (e x) ↔ P (g ∘ e'.symm) (e'.symm ⁻¹' s) (e' x) := by rw [← hG.right_invariance (compatible_of_mem_maximalAtlas he he')] swap; · simp only [xe, xe', mfld_simps] - simp_rw [LocalHomeomorph.trans_apply, e.left_inv xe] + simp_rw [PartialHomeomorph.trans_apply, e.left_inv xe] rw [hG.congr_iff] · refine' hG.congr_set _ refine' (eventually_of_mem _ fun y (hy : y ∈ e'.symm ⁻¹' e.source) ↦ _).set_eq · refine' (e'.symm.continuousAt <| e'.mapsTo xe').preimage_mem_nhds (e.open_source.mem_nhds _) simp_rw [e'.left_inv xe', xe] - simp_rw [mem_preimage, LocalHomeomorph.coe_trans_symm, LocalHomeomorph.symm_symm, + simp_rw [mem_preimage, PartialHomeomorph.coe_trans_symm, PartialHomeomorph.symm_symm, Function.comp_apply, e.left_inv hy] · refine' ((e'.eventually_nhds' _ xe').mpr <| e.eventually_left_inverse xe).mono fun y hy ↦ _ simp only [mfld_simps] @@ -278,7 +278,7 @@ theorem liftPropWithinAt_indep_chart_target_aux2 (g : H → M') {x : H} {s : Set exact (hgs.eventually <| f.eventually_left_inverse xf).mono fun y ↦ congr_arg f' #align structure_groupoid.local_invariant_prop.lift_prop_within_at_indep_chart_target_aux2 StructureGroupoid.LocalInvariantProp.liftPropWithinAt_indep_chart_target_aux2 -theorem liftPropWithinAt_indep_chart_target_aux {g : X → M'} {e : LocalHomeomorph X H} {x : X} +theorem liftPropWithinAt_indep_chart_target_aux {g : X → M'} {e : PartialHomeomorph X H} {x : X} {s : Set X} (xe : x ∈ e.source) (hf : f ∈ G'.maximalAtlas M') (xf : g x ∈ f.source) (hf' : f' ∈ G'.maximalAtlas M') (xf' : g x ∈ f'.source) (hgs : ContinuousWithinAt g s x) : P (f ∘ g ∘ e.symm) (e.symm ⁻¹' s) (e x) ↔ P (f' ∘ g ∘ e.symm) (e.symm ⁻¹' s) (e x) := by @@ -527,12 +527,12 @@ theorem liftPropOn_chart_symm [HasGroupoid M G] (hG : G.LocalInvariantProp G Q) #align structure_groupoid.local_invariant_prop.lift_prop_on_chart_symm StructureGroupoid.LocalInvariantProp.liftPropOn_chart_symm theorem liftPropAt_of_mem_groupoid (hG : G.LocalInvariantProp G Q) (hQ : ∀ y, Q id univ y) - {f : LocalHomeomorph H H} (hf : f ∈ G) {x : H} (hx : x ∈ f.source) : LiftPropAt Q f x := + {f : PartialHomeomorph H H} (hf : f ∈ G) {x : H} (hx : x ∈ f.source) : LiftPropAt Q f x := liftPropAt_of_mem_maximalAtlas hG hQ (G.mem_maximalAtlas_of_mem_groupoid hf) hx #align structure_groupoid.local_invariant_prop.lift_prop_at_of_mem_groupoid StructureGroupoid.LocalInvariantProp.liftPropAt_of_mem_groupoid theorem liftPropOn_of_mem_groupoid (hG : G.LocalInvariantProp G Q) (hQ : ∀ y, Q id univ y) - {f : LocalHomeomorph H H} (hf : f ∈ G) : LiftPropOn Q f f.source := + {f : PartialHomeomorph H H} (hf : f ∈ G) : LiftPropOn Q f f.source := liftPropOn_of_mem_maximalAtlas hG hQ (G.mem_maximalAtlas_of_mem_groupoid hf) #align structure_groupoid.local_invariant_prop.lift_prop_on_of_mem_groupoid StructureGroupoid.LocalInvariantProp.liftPropOn_of_mem_groupoid @@ -584,13 +584,13 @@ section LocalStructomorph variable (G) -open LocalHomeomorph +open PartialHomeomorph /-- A function from a model space `H` to itself is a local structomorphism, with respect to a structure groupoid `G` for `H`, relative to a set `s` in `H`, if for all points `x` in the set, the function agrees with a `G`-structomorphism on `s` in a neighbourhood of `x`. -/ def IsLocalStructomorphWithinAt (f : H → H) (s : Set H) (x : H) : Prop := - x ∈ s → ∃ e : LocalHomeomorph H H, e ∈ G ∧ EqOn f e.toFun (s ∩ e.source) ∧ x ∈ e.source + x ∈ s → ∃ e : PartialHomeomorph H H, e ∈ G ∧ EqOn f e.toFun (s ∩ e.source) ∧ x ∈ e.source #align structure_groupoid.is_local_structomorph_within_at StructureGroupoid.IsLocalStructomorphWithinAt /-- For a groupoid `G` which is `ClosedUnderRestriction`, being a local structomorphism is a local @@ -638,17 +638,17 @@ theorem isLocalStructomorphWithinAt_localInvariantProp [ClosedUnderRestriction G /-- A slight reformulation of `IsLocalStructomorphWithinAt` when `f` is a local homeomorph. This gives us an `e` that is defined on a subset of `f.source`. -/ -theorem _root_.LocalHomeomorph.isLocalStructomorphWithinAt_iff {G : StructureGroupoid H} - [ClosedUnderRestriction G] (f : LocalHomeomorph H H) {s : Set H} {x : H} +theorem _root_.PartialHomeomorph.isLocalStructomorphWithinAt_iff {G : StructureGroupoid H} + [ClosedUnderRestriction G] (f : PartialHomeomorph H H) {s : Set H} {x : H} (hx : x ∈ f.source ∪ sᶜ) : G.IsLocalStructomorphWithinAt (⇑f) s x ↔ - x ∈ s → ∃ e : LocalHomeomorph H H, + x ∈ s → ∃ e : PartialHomeomorph H H, e ∈ G ∧ e.source ⊆ f.source ∧ EqOn f (⇑e) (s ∩ e.source) ∧ x ∈ e.source := by constructor · intro hf h2x obtain ⟨e, he, hfe, hxe⟩ := hf h2x refine' ⟨e.restr f.source, closedUnderRestriction' he f.open_source, _, _, hxe, _⟩ - · simp_rw [LocalHomeomorph.restr_source] + · simp_rw [PartialHomeomorph.restr_source] refine' (inter_subset_right _ _).trans interior_subset · intro x' hx' exact hfe ⟨hx'.1, hx'.2.1⟩ @@ -657,32 +657,32 @@ theorem _root_.LocalHomeomorph.isLocalStructomorphWithinAt_iff {G : StructureGro · intro hf hx obtain ⟨e, he, _, hfe, hxe⟩ := hf hx exact ⟨e, he, hfe, hxe⟩ -#align local_homeomorph.is_local_structomorph_within_at_iff LocalHomeomorph.isLocalStructomorphWithinAt_iff +#align local_homeomorph.is_local_structomorph_within_at_iff PartialHomeomorph.isLocalStructomorphWithinAt_iff /-- A slight reformulation of `IsLocalStructomorphWithinAt` when `f` is a local homeomorph and the set we're considering is a superset of `f.source`. -/ -theorem _root_.LocalHomeomorph.isLocalStructomorphWithinAt_iff' {G : StructureGroupoid H} - [ClosedUnderRestriction G] (f : LocalHomeomorph H H) {s : Set H} {x : H} (hs : f.source ⊆ s) +theorem _root_.PartialHomeomorph.isLocalStructomorphWithinAt_iff' {G : StructureGroupoid H} + [ClosedUnderRestriction G] (f : PartialHomeomorph H H) {s : Set H} {x : H} (hs : f.source ⊆ s) (hx : x ∈ f.source ∪ sᶜ) : G.IsLocalStructomorphWithinAt (⇑f) s x ↔ - x ∈ s → ∃ e : LocalHomeomorph H H, + x ∈ s → ∃ e : PartialHomeomorph H H, e ∈ G ∧ e.source ⊆ f.source ∧ EqOn f (⇑e) e.source ∧ x ∈ e.source := by rw [f.isLocalStructomorphWithinAt_iff hx] refine' imp_congr_right fun _ ↦ exists_congr fun e ↦ and_congr_right fun _ ↦ _ refine' and_congr_right fun h2e ↦ _ rw [inter_eq_right.mpr (h2e.trans hs)] -#align local_homeomorph.is_local_structomorph_within_at_iff' LocalHomeomorph.isLocalStructomorphWithinAt_iff' +#align local_homeomorph.is_local_structomorph_within_at_iff' PartialHomeomorph.isLocalStructomorphWithinAt_iff' /-- A slight reformulation of `IsLocalStructomorphWithinAt` when `f` is a local homeomorph and the set we're considering is `f.source`. -/ -theorem _root_.LocalHomeomorph.isLocalStructomorphWithinAt_source_iff {G : StructureGroupoid H} - [ClosedUnderRestriction G] (f : LocalHomeomorph H H) {x : H} : +theorem _root_.PartialHomeomorph.isLocalStructomorphWithinAt_source_iff {G : StructureGroupoid H} + [ClosedUnderRestriction G] (f : PartialHomeomorph H H) {x : H} : G.IsLocalStructomorphWithinAt (⇑f) f.source x ↔ - x ∈ f.source → ∃ e : LocalHomeomorph H H, + x ∈ f.source → ∃ e : PartialHomeomorph H H, e ∈ G ∧ e.source ⊆ f.source ∧ EqOn f (⇑e) e.source ∧ x ∈ e.source := haveI : x ∈ f.source ∪ f.sourceᶜ := by simp_rw [union_compl_self, mem_univ] f.isLocalStructomorphWithinAt_iff' Subset.rfl this -#align local_homeomorph.is_local_structomorph_within_at_source_iff LocalHomeomorph.isLocalStructomorphWithinAt_source_iff +#align local_homeomorph.is_local_structomorph_within_at_source_iff PartialHomeomorph.isLocalStructomorphWithinAt_source_iff variable {H₁ : Type*} [TopologicalSpace H₁] {H₂ : Type*} [TopologicalSpace H₂] {H₃ : Type*} [TopologicalSpace H₃] [ChartedSpace H₁ H₂] [ChartedSpace H₂ H₃] {G₁ : StructureGroupoid H₁} @@ -703,15 +703,15 @@ theorem HasGroupoid.comp obtain ⟨φ, hφG₁, hφ, hφ_dom⟩ := LocalInvariantProp.liftPropOn_indep_chart (isLocalStructomorphWithinAt_localInvariantProp G₁) (G₁.subset_maximalAtlas hf) (G₁.subset_maximalAtlas hf') (H _ (G₂.compatible he he')) hxs' hxs - simp_rw [← LocalHomeomorph.coe_trans, LocalHomeomorph.trans_assoc] at hφ - simp_rw [LocalHomeomorph.trans_symm_eq_symm_trans_symm, LocalHomeomorph.trans_assoc] + simp_rw [← PartialHomeomorph.coe_trans, PartialHomeomorph.trans_assoc] at hφ + simp_rw [PartialHomeomorph.trans_symm_eq_symm_trans_symm, PartialHomeomorph.trans_assoc] have hs : IsOpen (f.symm ≫ₕ e.symm ≫ₕ e' ≫ₕ f').source := (f.symm ≫ₕ e.symm ≫ₕ e' ≫ₕ f').open_source refine' ⟨_, hs.inter φ.open_source, _, _⟩ · simp only [hx, hφ_dom, mfld_simps] · refine' G₁.eq_on_source (closedUnderRestriction' hφG₁ hs) _ - rw [LocalHomeomorph.restr_source_inter] - refine' LocalHomeomorph.Set.EqOn.restr_eqOn_source (hφ.mono _) + rw [PartialHomeomorph.restr_source_inter] + refine' PartialHomeomorph.Set.EqOn.restr_eqOn_source (hφ.mono _) mfld_set_tac } #align structure_groupoid.has_groupoid.comp StructureGroupoid.HasGroupoid.comp diff --git a/Mathlib/Geometry/Manifold/MFDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv.lean index 8ea9f73b508e3..049672a6d9bd2 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv.lean @@ -243,9 +243,9 @@ def MDifferentiable (f : M → M') := #align mdifferentiable MDifferentiable /-- Prop registering if a local homeomorphism is a local diffeomorphism on its source -/ -def LocalHomeomorph.MDifferentiable (f : LocalHomeomorph M M') := +def PartialHomeomorph.MDifferentiable (f : PartialHomeomorph M M') := MDifferentiableOn I I' f f.source ∧ MDifferentiableOn I' I f.symm f.target -#align local_homeomorph.mdifferentiable LocalHomeomorph.MDifferentiable +#align local_homeomorph.mdifferentiable PartialHomeomorph.MDifferentiable variable [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] @@ -1804,7 +1804,7 @@ end ModelWithCorners section Charts -variable {e : LocalHomeomorph M H} +variable {e : PartialHomeomorph M H} theorem mdifferentiableAt_atlas (h : e ∈ atlas H M) {x : M} (hx : x ∈ e.source) : MDifferentiableAt I I e x := by @@ -1822,7 +1822,7 @@ theorem mdifferentiableAt_atlas (h : e ∈ atlas H M) {x : M} (hx : x ∈ e.sour simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] - · apply IsOpen.mem_nhds ((LocalHomeomorph.open_source _).preimage I.continuous_symm) mem.1 + · apply IsOpen.mem_nhds ((PartialHomeomorph.open_source _).preimage I.continuous_symm) mem.1 #align mdifferentiable_at_atlas mdifferentiableAt_atlas theorem mdifferentiableOn_atlas (h : e ∈ atlas H M) : MDifferentiableOn I I e e.source := @@ -1844,7 +1844,7 @@ theorem mdifferentiableAt_atlas_symm (h : e ∈ atlas H M) {x : H} (hx : x ∈ e simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] - · apply IsOpen.mem_nhds ((LocalHomeomorph.open_source _).preimage I.continuous_symm) mem.1 + · apply IsOpen.mem_nhds ((PartialHomeomorph.open_source _).preimage I.continuous_symm) mem.1 #align mdifferentiable_at_atlas_symm mdifferentiableAt_atlas_symm theorem mdifferentiableOn_atlas_symm (h : e ∈ atlas H M) : MDifferentiableOn I I e.symm e.target := @@ -1893,7 +1893,7 @@ end SpecificFunctions /-! ### Differentiable local homeomorphisms -/ -namespace LocalHomeomorph.MDifferentiable +namespace PartialHomeomorph.MDifferentiable variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} @@ -1901,19 +1901,19 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} - {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] {e : LocalHomeomorph M M'} - (he : e.MDifferentiable I I') {e' : LocalHomeomorph M' M''} + {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] {e : PartialHomeomorph M M'} + (he : e.MDifferentiable I I') {e' : PartialHomeomorph M' M''} nonrec theorem symm : e.symm.MDifferentiable I' I := he.symm -#align local_homeomorph.mdifferentiable.symm LocalHomeomorph.MDifferentiable.symm +#align local_homeomorph.mdifferentiable.symm PartialHomeomorph.MDifferentiable.symm protected theorem mdifferentiableAt {x : M} (hx : x ∈ e.source) : MDifferentiableAt I I' e x := (he.1 x hx).mdifferentiableAt (IsOpen.mem_nhds e.open_source hx) -#align local_homeomorph.mdifferentiable.mdifferentiable_at LocalHomeomorph.MDifferentiable.mdifferentiableAt +#align local_homeomorph.mdifferentiable.mdifferentiable_at PartialHomeomorph.MDifferentiable.mdifferentiableAt theorem mdifferentiableAt_symm {x : M'} (hx : x ∈ e.target) : MDifferentiableAt I' I e.symm x := (he.2 x hx).mdifferentiableAt (IsOpen.mem_nhds e.open_target hx) -#align local_homeomorph.mdifferentiable.mdifferentiable_at_symm LocalHomeomorph.MDifferentiable.mdifferentiableAt_symm +#align local_homeomorph.mdifferentiable.mdifferentiable_at_symm PartialHomeomorph.MDifferentiable.mdifferentiableAt_symm variable [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] [SmoothManifoldWithCorners I'' M''] @@ -1929,13 +1929,13 @@ theorem symm_comp_deriv {x : M} (hx : x ∈ e.source) : apply Filter.EventuallyEq.mfderiv_eq have : e.source ∈ 𝓝 x := IsOpen.mem_nhds e.open_source hx exact Filter.mem_of_superset this (by mfld_set_tac) -#align local_homeomorph.mdifferentiable.symm_comp_deriv LocalHomeomorph.MDifferentiable.symm_comp_deriv +#align local_homeomorph.mdifferentiable.symm_comp_deriv PartialHomeomorph.MDifferentiable.symm_comp_deriv theorem comp_symm_deriv {x : M'} (hx : x ∈ e.target) : (mfderiv I I' e (e.symm x)).comp (mfderiv I' I e.symm x) = ContinuousLinearMap.id 𝕜 (TangentSpace I' x) := he.symm.symm_comp_deriv hx -#align local_homeomorph.mdifferentiable.comp_symm_deriv LocalHomeomorph.MDifferentiable.comp_symm_deriv +#align local_homeomorph.mdifferentiable.comp_symm_deriv PartialHomeomorph.MDifferentiable.comp_symm_deriv /-- The derivative of a differentiable local homeomorphism, as a continuous linear equivalence between the tangent spaces at `x` and `e x`. -/ @@ -1954,31 +1954,31 @@ protected def mfderiv {x : M} (hx : x ∈ e.source) : TangentSpace I x ≃L[𝕜 conv_rhs => rw [← this, ← he.comp_symm_deriv (e.map_source hx)] rw [e.left_inv hx] rfl } -#align local_homeomorph.mdifferentiable.mfderiv LocalHomeomorph.MDifferentiable.mfderiv +#align local_homeomorph.mdifferentiable.mfderiv PartialHomeomorph.MDifferentiable.mfderiv theorem mfderiv_bijective {x : M} (hx : x ∈ e.source) : Function.Bijective (mfderiv I I' e x) := (he.mfderiv hx).bijective -#align local_homeomorph.mdifferentiable.mfderiv_bijective LocalHomeomorph.MDifferentiable.mfderiv_bijective +#align local_homeomorph.mdifferentiable.mfderiv_bijective PartialHomeomorph.MDifferentiable.mfderiv_bijective theorem mfderiv_injective {x : M} (hx : x ∈ e.source) : Function.Injective (mfderiv I I' e x) := (he.mfderiv hx).injective -#align local_homeomorph.mdifferentiable.mfderiv_injective LocalHomeomorph.MDifferentiable.mfderiv_injective +#align local_homeomorph.mdifferentiable.mfderiv_injective PartialHomeomorph.MDifferentiable.mfderiv_injective theorem mfderiv_surjective {x : M} (hx : x ∈ e.source) : Function.Surjective (mfderiv I I' e x) := (he.mfderiv hx).surjective -#align local_homeomorph.mdifferentiable.mfderiv_surjective LocalHomeomorph.MDifferentiable.mfderiv_surjective +#align local_homeomorph.mdifferentiable.mfderiv_surjective PartialHomeomorph.MDifferentiable.mfderiv_surjective theorem ker_mfderiv_eq_bot {x : M} (hx : x ∈ e.source) : LinearMap.ker (mfderiv I I' e x) = ⊥ := (he.mfderiv hx).toLinearEquiv.ker -#align local_homeomorph.mdifferentiable.ker_mfderiv_eq_bot LocalHomeomorph.MDifferentiable.ker_mfderiv_eq_bot +#align local_homeomorph.mdifferentiable.ker_mfderiv_eq_bot PartialHomeomorph.MDifferentiable.ker_mfderiv_eq_bot theorem range_mfderiv_eq_top {x : M} (hx : x ∈ e.source) : LinearMap.range (mfderiv I I' e x) = ⊤ := (he.mfderiv hx).toLinearEquiv.range -#align local_homeomorph.mdifferentiable.range_mfderiv_eq_top LocalHomeomorph.MDifferentiable.range_mfderiv_eq_top +#align local_homeomorph.mdifferentiable.range_mfderiv_eq_top PartialHomeomorph.MDifferentiable.range_mfderiv_eq_top theorem range_mfderiv_eq_univ {x : M} (hx : x ∈ e.source) : range (mfderiv I I' e x) = univ := (he.mfderiv_surjective hx).range_eq -#align local_homeomorph.mdifferentiable.range_mfderiv_eq_univ LocalHomeomorph.MDifferentiable.range_mfderiv_eq_univ +#align local_homeomorph.mdifferentiable.range_mfderiv_eq_univ PartialHomeomorph.MDifferentiable.range_mfderiv_eq_univ theorem trans (he' : e'.MDifferentiable I' I'') : (e.trans e').MDifferentiable I I'' := by constructor @@ -1991,9 +1991,9 @@ theorem trans (he' : e'.MDifferentiable I' I'') : (e.trans e').MDifferentiable I exact ((he.symm.mdifferentiableAt hx.2).comp _ (he'.symm.mdifferentiableAt hx.1)).mdifferentiableWithinAt -#align local_homeomorph.mdifferentiable.trans LocalHomeomorph.MDifferentiable.trans +#align local_homeomorph.mdifferentiable.trans PartialHomeomorph.MDifferentiable.trans -end LocalHomeomorph.MDifferentiable +end PartialHomeomorph.MDifferentiable /-! ### Differentiability of `extChartAt` -/ @@ -2067,7 +2067,7 @@ theorem UniqueMDiffOn.image_denseRange (hs : UniqueMDiffOn I s) {f : M → M'} hs.image_denseRange' (fun x hx ↦ (hf x hx).hasMFDerivWithinAt) hd protected theorem UniqueMDiffWithinAt.preimage_localHomeomorph (hs : UniqueMDiffWithinAt I s x) - {e : LocalHomeomorph M M'} (he : e.MDifferentiable I I') (hx : x ∈ e.source) : + {e : PartialHomeomorph M M'} (he : e.MDifferentiable I I') (hx : x ∈ e.source) : UniqueMDiffWithinAt I' (e.target ∩ e.symm ⁻¹' s) (e x) := by rw [← e.image_source_inter_eq', inter_comm] exact (hs.inter (e.open_source.mem_nhds hx)).image_denseRange @@ -2076,7 +2076,7 @@ protected theorem UniqueMDiffWithinAt.preimage_localHomeomorph (hs : UniqueMDiff /-- If a set has the unique differential property, then its image under a local diffeomorphism also has the unique differential property. -/ -theorem UniqueMDiffOn.uniqueMDiffOn_preimage (hs : UniqueMDiffOn I s) {e : LocalHomeomorph M M'} +theorem UniqueMDiffOn.uniqueMDiffOn_preimage (hs : UniqueMDiffOn I s) {e : PartialHomeomorph M M'} (he : e.MDifferentiable I I') : UniqueMDiffOn I' (e.target ∩ e.symm ⁻¹' s) := fun _x hx ↦ e.right_inv hx.1 ▸ (hs _ hx.2).preimage_localHomeomorph he (e.map_target hx.1) #align unique_mdiff_on.unique_mdiff_on_preimage UniqueMDiffOn.uniqueMDiffOn_preimage @@ -2116,7 +2116,7 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type [∀ b, Module 𝕜 (Z b)] [FiberBundle F Z] [VectorBundle 𝕜 F Z] [SmoothVectorBundle F Z I] theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : - e.toLocalHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := + e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := ⟨(e.smoothOn I).mdifferentiableOn, (e.smoothOn_symm I).mdifferentiableOn⟩ theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} @@ -2130,7 +2130,7 @@ theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} rw [← e.left_inv hp] refine (this.preimage_localHomeomorph e.mdifferentiable.symm (e.map_source hp)).mono ?_ rintro y ⟨hy, hys, -⟩ - rwa [LocalHomeomorph.symm_symm, e.coe_coe, e.coe_fst hy] at hys + rwa [PartialHomeomorph.symm_symm, e.coe_coe, e.coe_fst hy] at hys variable (Z) diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 059b502b06fe5..de3cdb468ed3b 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -642,7 +642,7 @@ theorem IsOpen.exists_msmooth_support_eq (hs : IsOpen s) : Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 := by intro i apply IsOpen.exists_msmooth_support_eq_aux - exact LocalHomeomorph.isOpen_inter_preimage_symm _ hs + exact PartialHomeomorph.isOpen_inter_preimage_symm _ hs choose g g_supp g_diff hg using A have h'g : ∀ c x, 0 ≤ g c x := fun c x ↦ (hg c (mem_range_self (f := g c) x)).1 have h''g : ∀ c x, 0 ≤ f c x * g c (chartAt H c x) := diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 40c7bb08fe3ac..3ec595e594c4c 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -330,7 +330,7 @@ theorem unique_diff_preimage {s : Set H} (hs : IsOpen s) : exact I.unique_diff.inter (hs.preimage I.continuous_invFun) #align model_with_corners.unique_diff_preimage ModelWithCorners.unique_diff_preimage -theorem unique_diff_preimage_source {β : Type*} [TopologicalSpace β] {e : LocalHomeomorph H β} : +theorem unique_diff_preimage_source {β : Type*} [TopologicalSpace β] {e : PartialHomeomorph H β} : UniqueDiffOn 𝕜 (I.symm ⁻¹' e.source ∩ range I) := I.unique_diff_preimage e.open_source #align model_with_corners.unique_diff_preimage_source ModelWithCorners.unique_diff_preimage_source @@ -603,7 +603,7 @@ variable (n) /-- An identity local homeomorphism belongs to the `C^n` groupoid. -/ theorem ofSet_mem_contDiffGroupoid {s : Set H} (hs : IsOpen s) : - LocalHomeomorph.ofSet s hs ∈ contDiffGroupoid n I := by + PartialHomeomorph.ofSet s hs ∈ contDiffGroupoid n I := by rw [contDiffGroupoid, mem_groupoid_of_pregroupoid] suffices h : ContDiffOn 𝕜 n (I ∘ I.symm) (I.symm ⁻¹' s ∩ range I) · simp [h] @@ -613,10 +613,10 @@ theorem ofSet_mem_contDiffGroupoid {s : Set H} (hs : IsOpen s) : /-- The composition of a local homeomorphism from `H` to `M` and its inverse belongs to the `C^n` groupoid. -/ -theorem symm_trans_mem_contDiffGroupoid (e : LocalHomeomorph M H) : +theorem symm_trans_mem_contDiffGroupoid (e : PartialHomeomorph M H) : e.symm.trans e ∈ contDiffGroupoid n I := - haveI : e.symm.trans e ≈ LocalHomeomorph.ofSet e.target e.open_target := - LocalHomeomorph.trans_symm_self _ + haveI : e.symm.trans e ≈ PartialHomeomorph.ofSet e.target e.open_target := + PartialHomeomorph.trans_symm_self _ StructureGroupoid.eq_on_source _ (ofSet_mem_contDiffGroupoid n I e.open_target) this #align symm_trans_mem_cont_diff_groupoid symm_trans_mem_contDiffGroupoid @@ -624,12 +624,12 @@ variable {E' H' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [Topologi /-- The product of two smooth local homeomorphisms is smooth. -/ theorem contDiffGroupoid_prod {I : ModelWithCorners 𝕜 E H} {I' : ModelWithCorners 𝕜 E' H'} - {e : LocalHomeomorph H H} {e' : LocalHomeomorph H' H'} (he : e ∈ contDiffGroupoid ⊤ I) + {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid ⊤ I) (he' : e' ∈ contDiffGroupoid ⊤ I') : e.prod e' ∈ contDiffGroupoid ⊤ (I.prod I') := by cases' he with he he_symm cases' he' with he' he'_symm simp only at he he_symm he' he'_symm - constructor <;> simp only [LocalEquiv.prod_source, LocalHomeomorph.prod_toLocalEquiv] + constructor <;> simp only [LocalEquiv.prod_source, PartialHomeomorph.prod_toLocalEquiv] · have h3 := ContDiffOn.prod_map he he' rw [← I.image_eq, ← I'.image_eq, prod_image_image_eq] at h3 rw [← (I.prod I').image_eq] @@ -746,15 +746,15 @@ def analyticGroupoid : StructureGroupoid H := /-- An identity local homeomorphism belongs to the analytic groupoid. -/ theorem ofSet_mem_analyticGroupoid {s : Set H} (hs : IsOpen s) : - LocalHomeomorph.ofSet s hs ∈ analyticGroupoid I := by + PartialHomeomorph.ofSet s hs ∈ analyticGroupoid I := by rw [analyticGroupoid] refine And.intro (ofSet_mem_contDiffGroupoid ∞ I hs) ?_ apply mem_groupoid_of_pregroupoid.mpr suffices h : AnalyticOn 𝕜 (I ∘ I.symm) (I.symm ⁻¹' s ∩ interior (range I)) ∧ (I.symm ⁻¹' s ∩ interior (range I)).image (I ∘ I.symm) ⊆ interior (range I) - · simp only [LocalHomeomorph.ofSet_apply, left_id, LocalHomeomorph.ofSet_toLocalEquiv, + · simp only [PartialHomeomorph.ofSet_apply, left_id, PartialHomeomorph.ofSet_toLocalEquiv, LocalEquiv.ofSet_source, h, comp_apply, mem_range, image_subset_iff, true_and, - LocalHomeomorph.ofSet_symm, LocalEquiv.ofSet_target, and_self] + PartialHomeomorph.ofSet_symm, LocalEquiv.ofSet_target, and_self] intro x hx refine mem_preimage.mpr ?_ rw [← I.right_inv (interior_subset hx.right)] at hx @@ -772,10 +772,10 @@ theorem ofSet_mem_analyticGroupoid {s : Set H} (hs : IsOpen s) : /-- The composition of a local homeomorphism from `H` to `M` and its inverse belongs to the analytic groupoid. -/ -theorem symm_trans_mem_analyticGroupoid (e : LocalHomeomorph M H) : +theorem symm_trans_mem_analyticGroupoid (e : PartialHomeomorph M H) : e.symm.trans e ∈ analyticGroupoid I := - haveI : e.symm.trans e ≈ LocalHomeomorph.ofSet e.target e.open_target := - LocalHomeomorph.trans_symm_self _ + haveI : e.symm.trans e ≈ PartialHomeomorph.ofSet e.target e.open_target := + PartialHomeomorph.trans_symm_self _ StructureGroupoid.eq_on_source _ (ofSet_mem_analyticGroupoid I e.open_target) this /-- The analytic groupoid is closed under restriction. -/ @@ -790,7 +790,7 @@ instance : ClosedUnderRestriction (analyticGroupoid I) := /-- The analytic groupoid on a boundaryless charted space modeled on a complete vector space consists of the local homeomorphisms which are analytic and have analytic inverse. -/ theorem mem_analyticGroupoid_of_boundaryless [CompleteSpace E] [I.Boundaryless] - (e : LocalHomeomorph H H) : + (e : PartialHomeomorph H H) : e ∈ analyticGroupoid I ↔ AnalyticOn 𝕜 (I ∘ e ∘ I.symm) (I '' e.source) ∧ AnalyticOn 𝕜 (I ∘ e.symm ∘ I.symm) (I '' e.target) := by apply Iff.intro @@ -830,7 +830,7 @@ theorem SmoothManifoldWithCorners.mk' {𝕜 : Type*} [NontriviallyNormedField theorem smoothManifoldWithCorners_of_contDiffOn {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] - (h : ∀ e e' : LocalHomeomorph M H, e ∈ atlas H M → e' ∈ atlas H M → + (h : ∀ e e' : PartialHomeomorph M H, e ∈ atlas H M → e' ∈ atlas H M → ContDiffOn 𝕜 ⊤ (I ∘ e.symm ≫ₕ e' ∘ I.symm) (I.symm ⁻¹' (e.symm ≫ₕ e').source ∩ range I)) : SmoothManifoldWithCorners I M where compatible := by @@ -875,7 +875,7 @@ theorem chart_mem_maximalAtlas [SmoothManifoldWithCorners I M] (x : M) : variable {I} -theorem compatible_of_mem_maximalAtlas {e e' : LocalHomeomorph M H} (he : e ∈ maximalAtlas I M) +theorem compatible_of_mem_maximalAtlas {e e' : PartialHomeomorph M H} (he : e ∈ maximalAtlas I M) (he' : e' ∈ maximalAtlas I M) : e.symm.trans e' ∈ contDiffGroupoid ∞ I := StructureGroupoid.compatible_of_mem_maximalAtlas he he' #align smooth_manifold_with_corners.compatible_of_mem_maximal_atlas SmoothManifoldWithCorners.compatible_of_mem_maximalAtlas @@ -889,7 +889,7 @@ instance prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedA [SmoothManifoldWithCorners I' M'] : SmoothManifoldWithCorners (I.prod I') (M × M') where compatible := by rintro f g ⟨f1, f2, hf1, hf2, rfl⟩ ⟨g1, g2, hg1, hg2, rfl⟩ - rw [LocalHomeomorph.prod_symm, LocalHomeomorph.prod_trans] + rw [PartialHomeomorph.prod_symm, PartialHomeomorph.prod_trans] have h1 := (contDiffGroupoid ⊤ I).compatible hf1 hg1 have h2 := (contDiffGroupoid ⊤ I').compatible hf2 hg2 exact contDiffGroupoid_prod h1 h2 @@ -897,21 +897,21 @@ instance prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedA end SmoothManifoldWithCorners -theorem LocalHomeomorph.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] - {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] (e : LocalHomeomorph M H) - (h : e.source = Set.univ) : +theorem PartialHomeomorph.singleton_smoothManifoldWithCorners + {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) + {M : Type*} [TopologicalSpace M] (e : PartialHomeomorph M H) (h : e.source = Set.univ) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ (e.singletonChartedSpace h) := @SmoothManifoldWithCorners.mk' _ _ _ _ _ _ _ _ _ _ (id _) <| e.singleton_hasGroupoid h (contDiffGroupoid ∞ I) -#align local_homeomorph.singleton_smooth_manifold_with_corners LocalHomeomorph.singleton_smoothManifoldWithCorners +#align local_homeomorph.singleton_smooth_manifold_with_corners PartialHomeomorph.singleton_smoothManifoldWithCorners theorem OpenEmbedding.singleton_smoothManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [Nonempty M] {f : M → H} (h : OpenEmbedding f) : @SmoothManifoldWithCorners 𝕜 _ E _ _ H _ I M _ h.singletonChartedSpace := - (h.toLocalHomeomorph f).singleton_smoothManifoldWithCorners I (by simp) + (h.toPartialHomeomorph f).singleton_smoothManifoldWithCorners I (by simp) #align open_embedding.singleton_smooth_manifold_with_corners OpenEmbedding.singleton_smoothManifoldWithCorners namespace TopologicalSpace.Opens @@ -932,7 +932,7 @@ section ExtendedCharts open scoped Topology variable {𝕜 E M H E' M' H' : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] - [NormedSpace 𝕜 E] [TopologicalSpace H] [TopologicalSpace M] (f f' : LocalHomeomorph M H) + [NormedSpace 𝕜 E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) (I : ModelWithCorners 𝕜 E H) [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] [TopologicalSpace M'] (I' : ModelWithCorners 𝕜 E' H') (x : M) {s t : Set M} @@ -941,51 +941,51 @@ variable {𝕜 E M H E' M' H' : Type*} [NontriviallyNormedField 𝕜] [NormedAdd In a smooth manifold with corners, the model space is the space `H`. However, we will also need to use extended charts taking values in the model vector space `E`. These extended charts are -not `LocalHomeomorph` as the target is not open in `E` in general, but we can still register them +not `PartialHomeomorph` as the target is not open in `E` in general, but we can still register them as `LocalEquiv`. -/ -namespace LocalHomeomorph +namespace PartialHomeomorph /-- Given a chart `f` on a manifold with corners, `f.extend I` is the extended chart to the model vector space. -/ @[simp, mfld_simps] def extend : LocalEquiv M E := f.toLocalEquiv ≫ I.toLocalEquiv -#align local_homeomorph.extend LocalHomeomorph.extend +#align local_homeomorph.extend PartialHomeomorph.extend theorem extend_coe : ⇑(f.extend I) = I ∘ f := rfl -#align local_homeomorph.extend_coe LocalHomeomorph.extend_coe +#align local_homeomorph.extend_coe PartialHomeomorph.extend_coe theorem extend_coe_symm : ⇑(f.extend I).symm = f.symm ∘ I.symm := rfl -#align local_homeomorph.extend_coe_symm LocalHomeomorph.extend_coe_symm +#align local_homeomorph.extend_coe_symm PartialHomeomorph.extend_coe_symm theorem extend_source : (f.extend I).source = f.source := by rw [extend, LocalEquiv.trans_source, I.source_eq, preimage_univ, inter_univ] -#align local_homeomorph.extend_source LocalHomeomorph.extend_source +#align local_homeomorph.extend_source PartialHomeomorph.extend_source theorem isOpen_extend_source : IsOpen (f.extend I).source := by rw [extend_source] exact f.open_source -#align local_homeomorph.is_open_extend_source LocalHomeomorph.isOpen_extend_source +#align local_homeomorph.is_open_extend_source PartialHomeomorph.isOpen_extend_source theorem extend_target : (f.extend I).target = I.symm ⁻¹' f.target ∩ range I := by simp_rw [extend, LocalEquiv.trans_target, I.target_eq, I.toLocalEquiv_coe_symm, inter_comm] -#align local_homeomorph.extend_target LocalHomeomorph.extend_target +#align local_homeomorph.extend_target PartialHomeomorph.extend_target theorem mapsTo_extend (hs : s ⊆ f.source) : MapsTo (f.extend I) s ((f.extend I).symm ⁻¹' s ∩ range I) := by rw [mapsTo', extend_coe, extend_coe_symm, preimage_comp, ← I.image_eq, image_comp, f.image_eq_target_inter_inv_preimage hs] exact image_subset _ (inter_subset_right _ _) -#align local_homeomorph.maps_to_extend LocalHomeomorph.mapsTo_extend +#align local_homeomorph.maps_to_extend PartialHomeomorph.mapsTo_extend theorem extend_left_inv {x : M} (hxf : x ∈ f.source) : (f.extend I).symm (f.extend I x) = x := (f.extend I).left_inv <| by rwa [f.extend_source] -#align local_homeomorph.extend_left_inv LocalHomeomorph.extend_left_inv +#align local_homeomorph.extend_left_inv PartialHomeomorph.extend_left_inv /-- Variant of `f.extend_left_inv I`, stated in terms of images. -/ lemma extend_left_inv' (ht: t ⊆ f.source) : ((f.extend I).symm ∘ (f.extend I)) '' t = t := @@ -993,71 +993,71 @@ lemma extend_left_inv' (ht: t ⊆ f.source) : ((f.extend I).symm ∘ (f.extend I theorem extend_source_mem_nhds {x : M} (h : x ∈ f.source) : (f.extend I).source ∈ 𝓝 x := (isOpen_extend_source f I).mem_nhds <| by rwa [f.extend_source I] -#align local_homeomorph.extend_source_mem_nhds LocalHomeomorph.extend_source_mem_nhds +#align local_homeomorph.extend_source_mem_nhds PartialHomeomorph.extend_source_mem_nhds theorem extend_source_mem_nhdsWithin {x : M} (h : x ∈ f.source) : (f.extend I).source ∈ 𝓝[s] x := mem_nhdsWithin_of_mem_nhds <| extend_source_mem_nhds f I h -#align local_homeomorph.extend_source_mem_nhds_within LocalHomeomorph.extend_source_mem_nhdsWithin +#align local_homeomorph.extend_source_mem_nhds_within PartialHomeomorph.extend_source_mem_nhdsWithin theorem continuousOn_extend : ContinuousOn (f.extend I) (f.extend I).source := by refine' I.continuous.comp_continuousOn _ rw [extend_source] exact f.continuousOn -#align local_homeomorph.continuous_on_extend LocalHomeomorph.continuousOn_extend +#align local_homeomorph.continuous_on_extend PartialHomeomorph.continuousOn_extend theorem continuousAt_extend {x : M} (h : x ∈ f.source) : ContinuousAt (f.extend I) x := (continuousOn_extend f I).continuousAt <| extend_source_mem_nhds f I h -#align local_homeomorph.continuous_at_extend LocalHomeomorph.continuousAt_extend +#align local_homeomorph.continuous_at_extend PartialHomeomorph.continuousAt_extend theorem map_extend_nhds {x : M} (hy : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝[range I] f.extend I x := by rwa [extend_coe, comp_apply, ← I.map_nhds_eq, ← f.map_nhds_eq, map_map] -#align local_homeomorph.map_extend_nhds LocalHomeomorph.map_extend_nhds +#align local_homeomorph.map_extend_nhds PartialHomeomorph.map_extend_nhds theorem extend_target_mem_nhdsWithin {y : M} (hy : y ∈ f.source) : (f.extend I).target ∈ 𝓝[range I] f.extend I y := by rw [← LocalEquiv.image_source_eq_target, ← map_extend_nhds f I hy] exact image_mem_map (extend_source_mem_nhds _ _ hy) -#align local_homeomorph.extend_target_mem_nhds_within LocalHomeomorph.extend_target_mem_nhdsWithin +#align local_homeomorph.extend_target_mem_nhds_within PartialHomeomorph.extend_target_mem_nhdsWithin theorem extend_target_subset_range : (f.extend I).target ⊆ range I := by simp only [mfld_simps] -#align local_homeomorph.extend_target_subset_range LocalHomeomorph.extend_target_subset_range +#align local_homeomorph.extend_target_subset_range PartialHomeomorph.extend_target_subset_range theorem nhdsWithin_extend_target_eq {y : M} (hy : y ∈ f.source) : 𝓝[(f.extend I).target] f.extend I y = 𝓝[range I] f.extend I y := (nhdsWithin_mono _ (extend_target_subset_range _ _)).antisymm <| nhdsWithin_le_of_mem (extend_target_mem_nhdsWithin _ _ hy) -#align local_homeomorph.nhds_within_extend_target_eq LocalHomeomorph.nhdsWithin_extend_target_eq +#align local_homeomorph.nhds_within_extend_target_eq PartialHomeomorph.nhdsWithin_extend_target_eq theorem continuousAt_extend_symm' {x : E} (h : x ∈ (f.extend I).target) : ContinuousAt (f.extend I).symm x := (f.continuousAt_symm h.2).comp I.continuous_symm.continuousAt -#align local_homeomorph.continuous_at_extend_symm' LocalHomeomorph.continuousAt_extend_symm' +#align local_homeomorph.continuous_at_extend_symm' PartialHomeomorph.continuousAt_extend_symm' theorem continuousAt_extend_symm {x : M} (h : x ∈ f.source) : ContinuousAt (f.extend I).symm (f.extend I x) := continuousAt_extend_symm' f I <| (f.extend I).map_source <| by rwa [f.extend_source] -#align local_homeomorph.continuous_at_extend_symm LocalHomeomorph.continuousAt_extend_symm +#align local_homeomorph.continuous_at_extend_symm PartialHomeomorph.continuousAt_extend_symm theorem continuousOn_extend_symm : ContinuousOn (f.extend I).symm (f.extend I).target := fun _ h => (continuousAt_extend_symm' _ _ h).continuousWithinAt -#align local_homeomorph.continuous_on_extend_symm LocalHomeomorph.continuousOn_extend_symm +#align local_homeomorph.continuous_on_extend_symm PartialHomeomorph.continuousOn_extend_symm theorem extend_symm_continuousWithinAt_comp_right_iff {X} [TopologicalSpace X] {g : M → X} {s : Set M} {x : M} : ContinuousWithinAt (g ∘ (f.extend I).symm) ((f.extend I).symm ⁻¹' s ∩ range I) (f.extend I x) ↔ ContinuousWithinAt (g ∘ f.symm) (f.symm ⁻¹' s) (f x) := by rw [← I.symm_continuousWithinAt_comp_right_iff]; rfl -#align local_homeomorph.extend_symm_continuous_within_at_comp_right_iff LocalHomeomorph.extend_symm_continuousWithinAt_comp_right_iff +#align local_homeomorph.extend_symm_continuous_within_at_comp_right_iff PartialHomeomorph.extend_symm_continuousWithinAt_comp_right_iff theorem isOpen_extend_preimage' {s : Set E} (hs : IsOpen s) : IsOpen ((f.extend I).source ∩ f.extend I ⁻¹' s) := (continuousOn_extend f I).isOpen_inter_preimage (isOpen_extend_source _ _) hs -#align local_homeomorph.is_open_extend_preimage' LocalHomeomorph.isOpen_extend_preimage' +#align local_homeomorph.is_open_extend_preimage' PartialHomeomorph.isOpen_extend_preimage' theorem isOpen_extend_preimage {s : Set E} (hs : IsOpen s) : IsOpen (f.source ∩ f.extend I ⁻¹' s) := by rw [← extend_source f I]; exact isOpen_extend_preimage' f I hs -#align local_homeomorph.is_open_extend_preimage LocalHomeomorph.isOpen_extend_preimage +#align local_homeomorph.is_open_extend_preimage PartialHomeomorph.isOpen_extend_preimage theorem map_extend_nhdsWithin_eq_image {y : M} (hy : y ∈ f.source) : map (f.extend I) (𝓝[s] y) = 𝓝[f.extend I '' ((f.extend I).source ∩ s)] f.extend I y := by @@ -1070,7 +1070,7 @@ theorem map_extend_nhdsWithin_eq_image {y : M} (hy : y ∈ f.source) : ((f.extend I).left_inv <| by rwa [f.extend_source]) (continuousAt_extend_symm f I hy).continuousWithinAt (continuousAt_extend f I hy).continuousWithinAt -#align local_homeomorph.map_extend_nhds_within_eq_image LocalHomeomorph.map_extend_nhdsWithin_eq_image +#align local_homeomorph.map_extend_nhds_within_eq_image PartialHomeomorph.map_extend_nhdsWithin_eq_image theorem map_extend_nhdsWithin_eq_image_of_subset {y : M} (hy : y ∈ f.source) (hs : s ⊆ f.source) : map (f.extend I) (𝓝[s] y) = 𝓝[f.extend I '' s] f.extend I y := by @@ -1082,18 +1082,18 @@ theorem map_extend_nhdsWithin {y : M} (hy : y ∈ f.source) : rw [map_extend_nhdsWithin_eq_image f I hy, nhdsWithin_inter, ← nhdsWithin_extend_target_eq _ _ hy, ← nhdsWithin_inter, (f.extend I).image_source_inter_eq', inter_comm] -#align local_homeomorph.map_extend_nhds_within LocalHomeomorph.map_extend_nhdsWithin +#align local_homeomorph.map_extend_nhds_within PartialHomeomorph.map_extend_nhdsWithin theorem map_extend_symm_nhdsWithin {y : M} (hy : y ∈ f.source) : map (f.extend I).symm (𝓝[(f.extend I).symm ⁻¹' s ∩ range I] f.extend I y) = 𝓝[s] y := by rw [← map_extend_nhdsWithin f I hy, map_map, Filter.map_congr, map_id] exact (f.extend I).leftInvOn.eqOn.eventuallyEq_of_mem (extend_source_mem_nhdsWithin _ _ hy) -#align local_homeomorph.map_extend_symm_nhds_within LocalHomeomorph.map_extend_symm_nhdsWithin +#align local_homeomorph.map_extend_symm_nhds_within PartialHomeomorph.map_extend_symm_nhdsWithin theorem map_extend_symm_nhdsWithin_range {y : M} (hy : y ∈ f.source) : map (f.extend I).symm (𝓝[range I] f.extend I y) = 𝓝 y := by rw [← nhdsWithin_univ, ← map_extend_symm_nhdsWithin f I hy, preimage_univ, univ_inter] -#align local_homeomorph.map_extend_symm_nhds_within_range LocalHomeomorph.map_extend_symm_nhdsWithin_range +#align local_homeomorph.map_extend_symm_nhds_within_range PartialHomeomorph.map_extend_symm_nhdsWithin_range theorem tendsto_extend_comp_iff {l : Filter α} {g : α → M} (hg : ∀ᶠ z in l, g z ∈ f.source) (hy : y ∈ f.source) : @@ -1105,7 +1105,7 @@ theorem tendsto_extend_comp_iff {l : Filter α} {g : α → M} (hg : ∀ᶠ z in simpa only [(· ∘ ·), extend_left_inv _ _ hz, mem_preimage] using hzu -- there is no definition `writtenInExtend` but we already use some made-up names in this file -theorem continuousWithinAt_writtenInExtend_iff {f' : LocalHomeomorph M' H'} {g : M → M'} +theorem continuousWithinAt_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g : M → M'} (hy : y ∈ f.source) (hgy : g y ∈ f'.source) (hmaps : MapsTo g s f'.source) : ContinuousWithinAt (f'.extend I' ∘ g ∘ (f.extend I).symm) ((f.extend I).symm ⁻¹' s ∩ range I) (f.extend I y) ↔ ContinuousWithinAt g s y := by @@ -1122,7 +1122,7 @@ theorem continuousWithinAt_writtenInExtend_iff {f' : LocalHomeomorph M' H'} {g : /-- If `s ⊆ f.source` and `g x ∈ f'.source` whenever `x ∈ s`, then `g` is continuous on `s` if and only if `g` written in charts `f.extend I` and `f'.extend I'` is continuous on `f.extend I '' s`. -/ -theorem continuousOn_writtenInExtend_iff {f' : LocalHomeomorph M' H'} {g : M → M'} +theorem continuousOn_writtenInExtend_iff {f' : PartialHomeomorph M' H'} {g : M → M'} (hs : s ⊆ f.source) (hmaps : MapsTo g s f'.source) : ContinuousOn (f'.extend I' ∘ g ∘ (f.extend I).symm) (f.extend I '' s) ↔ ContinuousOn g s := by refine ball_image_iff.trans <| forall₂_congr fun x hx ↦ ?_ @@ -1136,14 +1136,14 @@ in the source is a neighborhood of the preimage, within a set. -/ theorem extend_preimage_mem_nhdsWithin {x : M} (h : x ∈ f.source) (ht : t ∈ 𝓝[s] x) : (f.extend I).symm ⁻¹' t ∈ 𝓝[(f.extend I).symm ⁻¹' s ∩ range I] f.extend I x := by rwa [← map_extend_symm_nhdsWithin f I h, mem_map] at ht -#align local_homeomorph.extend_preimage_mem_nhds_within LocalHomeomorph.extend_preimage_mem_nhdsWithin +#align local_homeomorph.extend_preimage_mem_nhds_within PartialHomeomorph.extend_preimage_mem_nhdsWithin theorem extend_preimage_mem_nhds {x : M} (h : x ∈ f.source) (ht : t ∈ 𝓝 x) : (f.extend I).symm ⁻¹' t ∈ 𝓝 (f.extend I x) := by apply (continuousAt_extend_symm f I h).preimage_mem_nhds rwa [(f.extend I).left_inv] rwa [f.extend_source] -#align local_homeomorph.extend_preimage_mem_nhds LocalHomeomorph.extend_preimage_mem_nhds +#align local_homeomorph.extend_preimage_mem_nhds PartialHomeomorph.extend_preimage_mem_nhds /-- Technical lemma to rewrite suitably the preimage of an intersection under an extended chart, to bring it into a convenient form to apply derivative lemmas. -/ @@ -1151,7 +1151,7 @@ theorem extend_preimage_inter_eq : (f.extend I).symm ⁻¹' (s ∩ t) ∩ range I = (f.extend I).symm ⁻¹' s ∩ range I ∩ (f.extend I).symm ⁻¹' t := by mfld_set_tac -#align local_homeomorph.extend_preimage_inter_eq LocalHomeomorph.extend_preimage_inter_eq +#align local_homeomorph.extend_preimage_inter_eq PartialHomeomorph.extend_preimage_inter_eq -- porting note: an `aux` lemma that is no longer needed. Delete? theorem extend_symm_preimage_inter_range_eventuallyEq_aux {s : Set M} {x : M} (hx : x ∈ f.source) : @@ -1165,14 +1165,14 @@ theorem extend_symm_preimage_inter_range_eventuallyEq_aux {s : Set M} {x : M} (h refine' (eventuallyEq_univ.mpr _).symm.inter EventuallyEq.rfl refine' I.continuousAt_symm.preimage_mem_nhds (f.open_target.mem_nhds _) simp_rw [f.extend_coe, Function.comp_apply, I.left_inv, f.mapsTo hx] -#align local_homeomorph.extend_symm_preimage_inter_range_eventually_eq_aux LocalHomeomorph.extend_symm_preimage_inter_range_eventuallyEq_aux +#align local_homeomorph.extend_symm_preimage_inter_range_eventually_eq_aux PartialHomeomorph.extend_symm_preimage_inter_range_eventuallyEq_aux theorem extend_symm_preimage_inter_range_eventuallyEq {s : Set M} {x : M} (hs : s ⊆ f.source) (hx : x ∈ f.source) : ((f.extend I).symm ⁻¹' s ∩ range I : Set _) =ᶠ[𝓝 (f.extend I x)] f.extend I '' s := by rw [← nhdsWithin_eq_iff_eventuallyEq, ← map_extend_nhdsWithin _ _ hx, map_extend_nhdsWithin_eq_image_of_subset _ _ hx hs] -#align local_homeomorph.extend_symm_preimage_inter_range_eventually_eq LocalHomeomorph.extend_symm_preimage_inter_range_eventuallyEq +#align local_homeomorph.extend_symm_preimage_inter_range_eventually_eq PartialHomeomorph.extend_symm_preimage_inter_range_eventuallyEq /-! We use the name `extend_coord_change` for `(f'.extend I).symm ≫ f.extend I`. -/ @@ -1181,13 +1181,13 @@ theorem extend_coord_change_source : simp_rw [LocalEquiv.trans_source, I.image_eq, extend_source, LocalEquiv.symm_source, extend_target, inter_right_comm _ (range I)] rfl -#align local_homeomorph.extend_coord_change_source LocalHomeomorph.extend_coord_change_source +#align local_homeomorph.extend_coord_change_source PartialHomeomorph.extend_coord_change_source theorem extend_image_source_inter : f.extend I '' (f.source ∩ f'.source) = ((f.extend I).symm ≫ f'.extend I).source := by simp_rw [f.extend_coord_change_source, f.extend_coe, image_comp I f, trans_source'', symm_symm, symm_target] -#align local_homeomorph.extend_image_source_inter LocalHomeomorph.extend_image_source_inter +#align local_homeomorph.extend_image_source_inter PartialHomeomorph.extend_image_source_inter theorem extend_coord_change_source_mem_nhdsWithin {x : E} (hx : x ∈ ((f.extend I).symm ≫ f'.extend I).source) : @@ -1195,8 +1195,8 @@ theorem extend_coord_change_source_mem_nhdsWithin {x : E} rw [f.extend_coord_change_source] at hx ⊢ obtain ⟨x, hx, rfl⟩ := hx refine' I.image_mem_nhdsWithin _ - refine' (LocalHomeomorph.open_source _).mem_nhds hx -#align local_homeomorph.extend_coord_change_source_mem_nhds_within LocalHomeomorph.extend_coord_change_source_mem_nhdsWithin + refine' (PartialHomeomorph.open_source _).mem_nhds hx +#align local_homeomorph.extend_coord_change_source_mem_nhds_within PartialHomeomorph.extend_coord_change_source_mem_nhdsWithin theorem extend_coord_change_source_mem_nhdsWithin' {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) : @@ -1204,7 +1204,7 @@ theorem extend_coord_change_source_mem_nhdsWithin' {x : M} (hxf : x ∈ f.source apply extend_coord_change_source_mem_nhdsWithin rw [← extend_image_source_inter] exact mem_image_of_mem _ ⟨hxf, hxf'⟩ -#align local_homeomorph.extend_coord_change_source_mem_nhds_within' LocalHomeomorph.extend_coord_change_source_mem_nhdsWithin' +#align local_homeomorph.extend_coord_change_source_mem_nhds_within' PartialHomeomorph.extend_coord_change_source_mem_nhdsWithin' variable {f f'} @@ -1215,7 +1215,7 @@ theorem contDiffOn_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtl ContDiffOn 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) ((f'.extend I).symm ≫ f.extend I).source := by rw [extend_coord_change_source, I.image_eq] exact (StructureGroupoid.compatible_of_mem_maximalAtlas hf' hf).1 -#align local_homeomorph.cont_diff_on_extend_coord_change LocalHomeomorph.contDiffOn_extend_coord_change +#align local_homeomorph.cont_diff_on_extend_coord_change PartialHomeomorph.contDiffOn_extend_coord_change theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : E} (hx : x ∈ ((f'.extend I).symm ≫ f.extend I).source) : @@ -1223,8 +1223,8 @@ theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maxi apply (contDiffOn_extend_coord_change I hf hf' x hx).mono_of_mem rw [extend_coord_change_source] at hx ⊢ obtain ⟨z, hz, rfl⟩ := hx - exact I.image_mem_nhdsWithin ((LocalHomeomorph.open_source _).mem_nhds hz) -#align local_homeomorph.cont_diff_within_at_extend_coord_change LocalHomeomorph.contDiffWithinAt_extend_coord_change + exact I.image_mem_nhdsWithin ((PartialHomeomorph.open_source _).mem_nhds hz) +#align local_homeomorph.cont_diff_within_at_extend_coord_change PartialHomeomorph.contDiffWithinAt_extend_coord_change theorem contDiffWithinAt_extend_coord_change' [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) : @@ -1232,11 +1232,11 @@ theorem contDiffWithinAt_extend_coord_change' [ChartedSpace H M] (hf : f ∈ max refine' contDiffWithinAt_extend_coord_change I hf hf' _ rw [← extend_image_source_inter] exact mem_image_of_mem _ ⟨hxf', hxf⟩ -#align local_homeomorph.cont_diff_within_at_extend_coord_change' LocalHomeomorph.contDiffWithinAt_extend_coord_change' +#align local_homeomorph.cont_diff_within_at_extend_coord_change' PartialHomeomorph.contDiffWithinAt_extend_coord_change' -end LocalHomeomorph +end PartialHomeomorph -open LocalHomeomorph +open PartialHomeomorph variable [ChartedSpace H M] [ChartedSpace H' M'] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index 2a00233017f6b..2edf4b8f90bb7 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -64,7 +64,7 @@ set_option autoImplicit true assert_not_exists mfderiv -open Bundle Set LocalHomeomorph +open Bundle Set PartialHomeomorph open Function (id_def) @@ -85,15 +85,15 @@ variable [TopologicalSpace F] [TopologicalSpace (TotalSpace F E)] [∀ x, Topolo /-- A fiber bundle `E` over a base `B` with model fiber `F` is naturally a charted space modelled on `B × F`. -/ instance FiberBundle.chartedSpace' : ChartedSpace (B × F) (TotalSpace F E) where - atlas := (fun e : Trivialization F (π F E) => e.toLocalHomeomorph) '' trivializationAtlas F E - chartAt x := (trivializationAt F E x.proj).toLocalHomeomorph + atlas := (fun e : Trivialization F (π F E) => e.toPartialHomeomorph) '' trivializationAtlas F E + chartAt x := (trivializationAt F E x.proj).toPartialHomeomorph mem_chart_source x := (trivializationAt F E x.proj).mem_source.mpr (mem_baseSet_trivializationAt F E x.proj) chart_mem_atlas _ := mem_image_of_mem _ (trivialization_mem_atlas F E _) #align fiber_bundle.charted_space FiberBundle.chartedSpace' theorem FiberBundle.chartedSpace'_chartAt (x : TotalSpace F E) : - chartAt (B × F) x = (trivializationAt F E x.proj).toLocalHomeomorph := + chartAt (B × F) x = (trivializationAt F E x.proj).toPartialHomeomorph := rfl /- Porting note: In Lean 3, the next instance was inside a section with locally reducible @@ -110,8 +110,8 @@ instance FiberBundle.chartedSpace : ChartedSpace (ModelProd HB F) (TotalSpace F theorem FiberBundle.chartedSpace_chartAt (x : TotalSpace F E) : chartAt (ModelProd HB F) x = - (trivializationAt F E x.proj).toLocalHomeomorph ≫ₕ - (chartAt HB x.proj).prod (LocalHomeomorph.refl F) := by + (trivializationAt F E x.proj).toPartialHomeomorph ≫ₕ + (chartAt HB x.proj).prod (PartialHomeomorph.refl F) := by dsimp only [chartAt_comp, prodChartedSpace_chartAt, FiberBundle.chartedSpace'_chartAt, chartAt_self_eq] rw [Trivialization.coe_coe, Trivialization.coe_fst' _ (mem_baseSet_trivializationAt F E x.proj)] @@ -164,7 +164,7 @@ theorem FiberBundle.writtenInExtChartAt_trivializationAt {x : TotalSpace F E} {y theorem FiberBundle.writtenInExtChartAt_trivializationAt_symm {x : TotalSpace F E} {y} (hy : y ∈ (extChartAt (IB.prod 𝓘(𝕜, F)) x).target) : writtenInExtChartAt (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) (trivializationAt F E x.proj x) - (trivializationAt F E x.proj).toLocalHomeomorph.symm y = y := + (trivializationAt F E x.proj).toPartialHomeomorph.symm y = y := writtenInExtChartAt_chartAt_symm_comp _ _ hy /-! ### Smoothness of maps in/out fiber bundles @@ -193,7 +193,7 @@ theorem contMDiffWithinAt_totalSpace (f : M → TotalSpace F E) {s : Set M} {x ((FiberBundle.continuous_proj F E).continuousWithinAt.comp hf (mapsTo_image f s)) ((Trivialization.open_baseSet _).mem_nhds (mem_baseSet_trivializationAt F E _)) refine EventuallyEq.contMDiffWithinAt_iff (eventually_of_mem h1 fun x hx => ?_) ?_ - · simp_rw [Function.comp, LocalHomeomorph.coe_coe, Trivialization.coe_coe] + · simp_rw [Function.comp, PartialHomeomorph.coe_coe, Trivialization.coe_coe] rw [Trivialization.coe_fst'] exact hx · simp only [mfld_simps] @@ -431,7 +431,7 @@ variable (IB e e') theorem Trivialization.contMDiffOn_symm_trans : ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) n - (e.toLocalHomeomorph.symm ≫ₕ e'.toLocalHomeomorph) (e.target ∩ e'.target) := by + (e.toPartialHomeomorph.symm ≫ₕ e'.toPartialHomeomorph) (e.target ∩ e'.target) := by have Hmaps : MapsTo Prod.fst (e.target ∩ e'.target) (e.baseSet ∩ e'.baseSet) := fun x hx ↦ ⟨e.mem_target.1 hx.1, e'.mem_target.1 hx.2⟩ rw [mapsTo_inter] at Hmaps @@ -440,7 +440,7 @@ theorem Trivialization.contMDiffOn_symm_trans : (contMDiffOn_fst.coordChange contMDiffOn_snd Hmaps.1 Hmaps.2)).congr ?_ rintro ⟨b, x⟩ hb refine Prod.ext ?_ rfl - · have : (e.toLocalHomeomorph.symm (b, x)).1 ∈ e'.baseSet + · have : (e.toPartialHomeomorph.symm (b, x)).1 ∈ e'.baseSet · simp_all only [Trivialization.mem_target, mfld_simps] exact (e'.coe_fst' this).trans (e.proj_symm_apply hb.1) @@ -478,7 +478,7 @@ instance SmoothFiberwiseLinear.hasGroupoid : rw [mem_smoothFiberwiseLinear_iff] refine' ⟨_, _, e.open_baseSet.inter e'.open_baseSet, smoothOn_coordChangeL IB e e', smoothOn_symm_coordChangeL IB e e', _⟩ - refine LocalHomeomorph.eqOnSourceSetoid.symm ⟨?_, ?_⟩ + refine PartialHomeomorph.eqOnSourceSetoid.symm ⟨?_, ?_⟩ · simp only [e.symm_trans_source_eq e', FiberwiseLinear.localHomeomorph, trans_toLocalEquiv, symm_toLocalEquiv] · rintro ⟨b, v⟩ hb @@ -566,8 +566,8 @@ theorem Trivialization.smoothOn (e : Trivialization F (π F E)) [MemTrivializati exact (this.1.prod_mk this.2).congr fun x hx ↦ (e.mk_proj_snd hx).symm theorem Trivialization.smoothOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : - SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e.toLocalHomeomorph.symm e.target := by - rw [e.smoothOn_iff IB e.toLocalHomeomorph.symm_mapsTo] + SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e.toPartialHomeomorph.symm e.target := by + rw [e.smoothOn_iff IB e.toPartialHomeomorph.symm_mapsTo] refine ⟨smoothOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, smoothOn_snd.congr fun x hx ↦ ?_⟩ rw [e.apply_symm_apply hx] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean index fb1125210b8b3..d2a2a3978cc57 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean @@ -35,7 +35,7 @@ determines a local homeomorphism from `B × F` to itself by its action fiberwise def localHomeomorph (φ : B → F ≃L[𝕜] F) (hU : IsOpen U) (hφ : ContinuousOn (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : ContinuousOn (fun x => (φ x).symm : B → F →L[𝕜] F) U) : - LocalHomeomorph (B × F) (B × F) where + PartialHomeomorph (B × F) (B × F) where toFun x := (x.1, φ x.1 x.2) invFun x := (x.1, (φ x.1).symm x.2) source := U ×ˢ univ @@ -106,7 +106,7 @@ local homeomorphism. Then the source of `e` is of the form `U ×ˢ univ`, for some set `U` in `B`, and, at any point `x` in `U`, admits a neighbourhood `u` of `x` such that `e` is equal on `u ×ˢ univ` to some bi-smooth fiberwise linear local homeomorphism. -/ -theorem SmoothFiberwiseLinear.locality_aux₁ (e : LocalHomeomorph (B × F) (B × F)) +theorem SmoothFiberwiseLinear.locality_aux₁ (e : PartialHomeomorph (B × F) (B × F)) (h : ∀ p ∈ e.source, ∃ s : Set (B × F), IsOpen s ∧ p ∈ s ∧ ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) @@ -158,7 +158,7 @@ together the various bi-smooth fiberwise linear local homeomorphism which exist The `U` in the conclusion is the same `U` as in the hypothesis. We state it like this, because this is exactly what we need for `smoothFiberwiseLinear`. -/ -theorem SmoothFiberwiseLinear.locality_aux₂ (e : LocalHomeomorph (B × F) (B × F)) (U : Set B) +theorem SmoothFiberwiseLinear.locality_aux₂ (e : PartialHomeomorph (B × F) (B × F)) (U : Set B) (hU : e.source = U ×ˢ univ) (h : ∀ x ∈ U, ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (_hUu : u ⊆ U) (_hux : x ∈ u) @@ -223,7 +223,7 @@ theorem SmoothFiberwiseLinear.locality_aux₂ (e : LocalHomeomorph (B × F) (B /- Porting note: `simp only [mem_iUnion]` fails in the next definition. This aux lemma is a workaround. -/ -private theorem mem_aux {e : LocalHomeomorph (B × F) (B × F)} : +private theorem mem_aux {e : PartialHomeomorph (B × F) (B × F)} : (e ∈ ⋃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), @@ -251,7 +251,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where simp only [mem_aux] rintro e e' ⟨φ, U, hU, hφ, h2φ, heφ⟩ ⟨φ', U', hU', hφ', h2φ', heφ'⟩ refine' ⟨fun b => (φ b).trans (φ' b), _, hU.inter hU', _, _, - Setoid.trans (LocalHomeomorph.EqOnSource.trans' heφ heφ') ⟨_, _⟩⟩ + Setoid.trans (PartialHomeomorph.EqOnSource.trans' heφ heφ') ⟨_, _⟩⟩ · show SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x : B => (φ' x).toContinuousLinearMap ∘L (φ x).toContinuousLinearMap) (U ∩ U') @@ -267,7 +267,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where symm' := fun e ↦ by simp only [mem_iUnion] rintro ⟨φ, U, hU, hφ, h2φ, heφ⟩ - refine' ⟨fun b => (φ b).symm, U, hU, h2φ, _, LocalHomeomorph.EqOnSource.symm' heφ⟩ + refine' ⟨fun b => (φ b).symm, U, hU, h2φ, _, PartialHomeomorph.EqOnSource.symm' heφ⟩ simp_rw [ContinuousLinearEquiv.symm_symm] exact hφ id_mem' := by @@ -277,7 +277,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where -/ refine mem_iUnion.2 ⟨fun _ ↦ .refl 𝕜 F, mem_iUnion.2 ⟨univ, mem_iUnion.2 ⟨isOpen_univ, ?_⟩⟩⟩ refine mem_iUnion.2 ⟨contMDiffOn_const, mem_iUnion.2 ⟨contMDiffOn_const, ?_, ?_⟩⟩ - · simp only [FiberwiseLinear.localHomeomorph, LocalHomeomorph.refl_localEquiv, + · simp only [FiberwiseLinear.localHomeomorph, PartialHomeomorph.refl_localEquiv, LocalEquiv.refl_source, univ_prod_univ] · exact eqOn_refl id _ locality' := by @@ -293,7 +293,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where #align smooth_fiberwise_linear smoothFiberwiseLinear @[simp] -theorem mem_smoothFiberwiseLinear_iff (e : LocalHomeomorph (B × F) (B × F)) : +theorem mem_smoothFiberwiseLinear_iff (e : PartialHomeomorph (B × F) (B × F)) : e ∈ smoothFiberwiseLinear B F IB ↔ ∃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index bdd48637bd1d1..2e07c956dd408 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -20,7 +20,7 @@ To do it for semilinear maps, we would need to generalize `ContinuousLinearMap.c noncomputable section -open Bundle Set LocalHomeomorph ContinuousLinearMap Pretrivialization +open Bundle Set PartialHomeomorph ContinuousLinearMap Pretrivialization open scoped Manifold Bundle @@ -55,8 +55,9 @@ theorem smoothOn_continuousLinearMapCoordChange theorem hom_chart (y₀ y : LE₁E₂) : chartAt (ModelProd HB (F₁ →L[𝕜] F₂)) y₀ y = (chartAt HB y₀.1 y.1, inCoordinates F₁ E₁ F₂ E₂ y₀.1 y.1 y₀.1 y.1 y.2) := by - rw [FiberBundle.chartedSpace_chartAt, trans_apply, LocalHomeomorph.prod_apply, - Trivialization.coe_coe, LocalHomeomorph.refl_apply, Function.id_def, hom_trivializationAt_apply] + rw [FiberBundle.chartedSpace_chartAt, trans_apply, PartialHomeomorph.prod_apply, + Trivialization.coe_coe, PartialHomeomorph.refl_apply, Function.id_def, + hom_trivializationAt_apply] #align hom_chart hom_chart variable {IB} diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index aeb31731a39da..ba0ff1b0533b5 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -14,8 +14,8 @@ This file defines the tangent bundle as a smooth vector bundle. Let `M` be a smooth manifold with corners with model `I` on `(E, H)`. We define the tangent bundle of `M` using the `VectorBundleCore` construction indexed by the charts of `M` with fibers `E`. -Given two charts `i, j : LocalHomeomorph M H`, the coordinate change between `i` and `j` at a point -`x : M` is the derivative of the composite +Given two charts `i, j : PartialHomeomorph M H`, the coordinate change between `i` and `j` +at a point `x : M` is the derivative of the composite ``` I.symm i.symm j I E -----> H -----> M --> H --> E @@ -32,7 +32,7 @@ This defines a smooth vector bundle `TangentBundle` with fibers `TangentSpace`. -/ -open Bundle Set SmoothManifoldWithCorners LocalHomeomorph ContinuousLinearMap +open Bundle Set SmoothManifoldWithCorners PartialHomeomorph ContinuousLinearMap open scoped Manifold Topology Bundle @@ -58,7 +58,7 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : rw [i.1.extend_coord_change_source]; apply image_subset_range intro x hx refine' (ContDiffWithinAt.fderivWithin_right _ I.unique_diff le_top <| h hx).mono h - refine' (LocalHomeomorph.contDiffOn_extend_coord_change I (subset_maximalAtlas I j.2) + refine' (PartialHomeomorph.contDiffOn_extend_coord_change I (subset_maximalAtlas I j.2) (subset_maximalAtlas I i.2) x hx).mono_of_mem _ exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 I hx #align cont_diff_on_fderiv_coord_change contDiffOn_fderiv_coord_change @@ -181,8 +181,8 @@ namespace TangentBundle protected theorem chartAt (p : TM) : chartAt (ModelProd H E) p = - ((tangentBundleCore I M).toFiberBundleCore.localTriv (achart H p.1)).toLocalHomeomorph ≫ₕ - (chartAt H p.1).prod (LocalHomeomorph.refl E) := + ((tangentBundleCore I M).toFiberBundleCore.localTriv (achart H p.1)).toPartialHomeomorph ≫ₕ + (chartAt H p.1).prod (PartialHomeomorph.refl E) := rfl #align tangent_bundle.chart_at TangentBundle.chartAt @@ -342,14 +342,14 @@ theorem tangentBundle_model_space_chartAt (p : TangentBundle I H) : @[simp, mfld_simps] theorem tangentBundle_model_space_coe_chartAt (p : TangentBundle I H) : ⇑(chartAt (ModelProd H E) p) = TotalSpace.toProd H E := by - rw [← LocalHomeomorph.coe_coe, tangentBundle_model_space_chartAt]; rfl + rw [← PartialHomeomorph.coe_coe, tangentBundle_model_space_chartAt]; rfl #align tangent_bundle_model_space_coe_chart_at tangentBundle_model_space_coe_chartAt @[simp, mfld_simps] theorem tangentBundle_model_space_coe_chartAt_symm (p : TangentBundle I H) : ((chartAt (ModelProd H E) p).symm : ModelProd H E → TangentBundle I H) = (TotalSpace.toProd H E).symm := by - rw [← LocalHomeomorph.coe_coe, LocalHomeomorph.symm_toLocalEquiv, + rw [← PartialHomeomorph.coe_coe, PartialHomeomorph.symm_toLocalEquiv, tangentBundle_model_space_chartAt]; rfl #align tangent_bundle_model_space_coe_chart_at_symm tangentBundle_model_space_coe_chartAt_symm diff --git a/Mathlib/Logic/Equiv/LocalEquiv.lean b/Mathlib/Logic/Equiv/LocalEquiv.lean index 814b8d495151e..3326f433af59c 100644 --- a/Mathlib/Logic/Equiv/LocalEquiv.lean +++ b/Mathlib/Logic/Equiv/LocalEquiv.lean @@ -71,7 +71,7 @@ then it should use `e.source ∩ s` or `e.target ∩ t`, not `s ∩ e.source` or open Lean Meta Elab Tactic /-! Implementation of the `mfld_set_tac` tactic for working with the domains of partially-defined -functions (`LocalEquiv`, `LocalHomeomorph`, etc). +functions (`LocalEquiv`, `PartialHomeomorph`, etc). This is in a separate file from `Mathlib.Logic.Equiv.MfldSimpsAttr` because attributes need a new file to become functional. diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 995cdca6dea03..c5a1504e8473d 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -1252,7 +1252,7 @@ theorem integral_image_eq_integral_abs_deriv_smul {s : Set ℝ} {f : ℝ → ℝ (fun x hx => (hf' x hx).hasFDerivWithinAt) hf g #align measure_theory.integral_image_eq_integral_abs_deriv_smul MeasureTheory.integral_image_eq_integral_abs_deriv_smul -theorem integral_target_eq_integral_abs_det_fderiv_smul {f : LocalHomeomorph E E} +theorem integral_target_eq_integral_abs_det_fderiv_smul {f : PartialHomeomorph E E} (hf' : ∀ x ∈ f.source, HasFDerivAt f (f' x) x) (g : E → F) : ∫ x in f.target, g x ∂μ = ∫ x in f.source, |(f' x).det| • g (f x) ∂μ := by have : f '' f.source = f.target := LocalEquiv.image_source_eq_target f.toLocalEquiv diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index 88839645af3e0..575db2d4e5814 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -108,7 +108,7 @@ protected theorem isLocalHomeomorphOn (hf : IsCoveringMapOn f s) : have h := (hf (f x) hx).mem_toTrivialization_baseSet let he := e.mem_source.2 h refine' - ⟨e.toLocalHomeomorph.trans + ⟨e.toPartialHomeomorph.trans { toFun := fun p => p.1 invFun := fun p => ⟨p, x, rfl⟩ source := e.baseSet ×ˢ ({⟨x, rfl⟩} : Set (f ⁻¹' {f x})) @@ -122,7 +122,7 @@ protected theorem isLocalHomeomorphOn (hf : IsCoveringMapOn f s) : right_inv' := fun p _ => rfl continuousOn_toFun := continuous_fst.continuousOn continuousOn_invFun := (continuous_id'.prod_mk continuous_const).continuousOn }, - ⟨he, by rwa [e.toLocalHomeomorph.symm_symm, e.proj_toFun x he], + ⟨he, by rwa [e.toPartialHomeomorph.symm_symm, e.proj_toFun x he], (hf (f x) hx).toTrivialization_apply⟩, fun p h => (e.proj_toFun p h.1).symm⟩ #align is_covering_map_on.is_locally_homeomorph_on IsCoveringMapOn.isLocalHomeomorphOn diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 64394bf49d6ee..1859d8f6865e8 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -444,7 +444,7 @@ def proj : Z.TotalSpace → B := #align fiber_bundle_core.proj FiberBundleCore.proj /-- Local homeomorphism version of the trivialization change. -/ -def trivChange (i j : ι) : LocalHomeomorph (B × F) (B × F) where +def trivChange (i j : ι) : PartialHomeomorph (B × F) (B × F) where source := (Z.baseSet i ∩ Z.baseSet j) ×ˢ univ target := (Z.baseSet i ∩ Z.baseSet j) ×ˢ univ toFun p := ⟨p.1, Z.coordChange i j p.1 p.2⟩ @@ -623,7 +623,7 @@ theorem continuous_const_section (v : F) refine continuous_iff_continuousAt.2 fun x => ?_ have A : Z.baseSet (Z.indexAt x) ∈ 𝓝 x := IsOpen.mem_nhds (Z.isOpen_baseSet (Z.indexAt x)) (Z.mem_baseSet_at x) - refine ((Z.localTrivAt x).toLocalHomeomorph.continuousAt_iff_continuousAt_comp_left ?_).2 ?_ + refine ((Z.localTrivAt x).toPartialHomeomorph.continuousAt_iff_continuousAt_comp_left ?_).2 ?_ · exact A · apply continuousAt_id.prod simp only [(· ∘ ·), mfld_simps, localTrivAt_snd] @@ -703,7 +703,7 @@ theorem mem_localTrivAt_target (p : B × F) (b : B) : @[simp, mfld_simps] theorem localTriv_symm_apply (p : B × F) : - (Z.localTriv i).toLocalHomeomorph.symm p = ⟨p.1, Z.coordChange i (Z.indexAt p.1) p.1 p.2⟩ := + (Z.localTriv i).toPartialHomeomorph.symm p = ⟨p.1, Z.coordChange i (Z.indexAt p.1) p.1 p.2⟩ := rfl #align fiber_bundle_core.local_triv_symm_apply FiberBundleCore.localTriv_symm_apply @@ -848,7 +848,7 @@ theorem continuous_totalSpaceMk (b : B) : Continuous[_, a.totalSpaceTopology] (TotalSpace.mk b) := by letI := a.totalSpaceTopology let e := a.trivializationOfMemPretrivializationAtlas (a.pretrivialization_mem_atlas b) - rw [e.toLocalHomeomorph.continuous_iff_continuous_comp_left + rw [e.toPartialHomeomorph.continuous_iff_continuous_comp_left (a.totalSpaceMk_preimage_source b)] exact continuous_iff_le_induced.mpr (le_antisymm_iff.mp (a.totalSpaceMk_inducing b).induced).1 #align fiber_prebundle.continuous_total_space_mk FiberPrebundle.continuous_totalSpaceMk diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index b279d4cde31c6..598e7c52ec3a8 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -60,7 +60,7 @@ def homeomorphProd : TotalSpace F (Trivial B F) ≃ₜ B × F := /-- Local trivialization for trivial bundle. -/ def trivialization : Trivialization F (π F (Bundle.Trivial B F)) where -- porting note: golfed - toLocalHomeomorph := (homeomorphProd B F).toLocalHomeomorph + toPartialHomeomorph := (homeomorphProd B F).toPartialHomeomorph baseSet := univ open_baseSet := isOpen_univ source_eq := rfl @@ -153,7 +153,7 @@ theorem Prod.continuous_to_fun : ContinuousOn (Prod.toFun' e₁ e₂) let f₃ : (B × F₁) × B × F₂ → B × F₁ × F₂ := fun p ↦ ⟨p.1.1, p.1.2, p.2.2⟩ have hf₁ : Continuous f₁ := (Prod.inducing_diag F₁ E₁ F₂ E₂).continuous have hf₂ : ContinuousOn f₂ (e₁.source ×ˢ e₂.source) := - e₁.toLocalHomeomorph.continuousOn.prod_map e₂.toLocalHomeomorph.continuousOn + e₁.toPartialHomeomorph.continuousOn.prod_map e₂.toPartialHomeomorph.continuousOn have hf₃ : Continuous f₃ := (continuous_fst.comp continuous_fst).prod_mk (continuous_snd.prod_map continuous_snd) refine' ((hf₃.comp_continuousOn hf₂).comp hf₁.continuousOn _).congr _ diff --git a/Mathlib/Topology/FiberBundle/Trivialization.lean b/Mathlib/Topology/FiberBundle/Trivialization.lean index ae8f2f643d730..7a3a063f875dd 100644 --- a/Mathlib/Topology/FiberBundle/Trivialization.lean +++ b/Mathlib/Topology/FiberBundle/Trivialization.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Data.Bundle import Mathlib.Topology.Algebra.Order.Field -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph #align_import topology.fiber_bundle.trivialization from "leanprover-community/mathlib"@"e473c3198bb41f68560cab68a0529c854b618833" @@ -304,12 +304,12 @@ variable [TopologicalSpace Z] [TopologicalSpace (TotalSpace F E)] sets of the form `proj ⁻¹' baseSet` and `baseSet × F`, acting trivially on the first coordinate. -/ -- porting note: todo: was @[nolint has_nonempty_instance] -structure Trivialization (proj : Z → B) extends LocalHomeomorph Z (B × F) where +structure Trivialization (proj : Z → B) extends PartialHomeomorph Z (B × F) where baseSet : Set B open_baseSet : IsOpen baseSet source_eq : source = proj ⁻¹' baseSet target_eq : target = baseSet ×ˢ univ - proj_toFun : ∀ p ∈ source, (toLocalHomeomorph p).1 = proj p + proj_toFun : ∀ p ∈ source, (toPartialHomeomorph p).1 = proj p #align trivialization Trivialization namespace Trivialization @@ -318,7 +318,7 @@ variable {F} variable (e : Trivialization F proj) {x : Z} @[ext] -lemma ext' (e e' : Trivialization F proj) (h₁ : e.toLocalHomeomorph = e'.toLocalHomeomorph) +lemma ext' (e e' : Trivialization F proj) (h₁ : e.toPartialHomeomorph = e'.toPartialHomeomorph) (h₂ : e.baseSet = e'.baseSet) : e = e' := by cases e; cases e'; congr #align trivialization.ext Trivialization.ext' @@ -342,12 +342,12 @@ instance : Coe (Trivialization F proj) (Pretrivialization F proj) := theorem toPretrivialization_injective : Function.Injective fun e : Trivialization F proj => e.toPretrivialization := fun e e' h => by ext1 - exacts [LocalHomeomorph.toLocalEquiv_injective (congr_arg Pretrivialization.toLocalEquiv h), + exacts [PartialHomeomorph.toLocalEquiv_injective (congr_arg Pretrivialization.toLocalEquiv h), congr_arg Pretrivialization.baseSet h] #align trivialization.to_pretrivialization_injective Trivialization.toPretrivialization_injective @[simp, mfld_simps] -theorem coe_coe : ⇑e.toLocalHomeomorph = e := +theorem coe_coe : ⇑e.toPartialHomeomorph = e := rfl #align trivialization.coe_coe Trivialization.coe_coe @@ -376,11 +376,11 @@ theorem mk_proj_snd' (ex : proj x ∈ e.baseSet) : (proj x, (e x).2) = e x := theorem source_inter_preimage_target_inter (s : Set (B × F)) : e.source ∩ e ⁻¹' (e.target ∩ s) = e.source ∩ e ⁻¹' s := - e.toLocalHomeomorph.source_inter_preimage_target_inter s + e.toPartialHomeomorph.source_inter_preimage_target_inter s #align trivialization.source_inter_preimage_target_inter Trivialization.source_inter_preimage_target_inter @[simp, mfld_simps] -theorem coe_mk (e : LocalHomeomorph Z (B × F)) (i j k l m) (x : Z) : +theorem coe_mk (e : PartialHomeomorph Z (B × F)) (i j k l m) (x : Z) : (Trivialization.mk e i j k l m : Trivialization F proj) x = e x := rfl #align trivialization.coe_mk Trivialization.coe_mk @@ -389,16 +389,17 @@ theorem mem_target {x : B × F} : x ∈ e.target ↔ x.1 ∈ e.baseSet := e.toPretrivialization.mem_target #align trivialization.mem_target Trivialization.mem_target -theorem map_target {x : B × F} (hx : x ∈ e.target) : e.toLocalHomeomorph.symm x ∈ e.source := - e.toLocalHomeomorph.map_target hx +theorem map_target {x : B × F} (hx : x ∈ e.target) : e.toPartialHomeomorph.symm x ∈ e.source := + e.toPartialHomeomorph.map_target hx #align trivialization.map_target Trivialization.map_target -theorem proj_symm_apply {x : B × F} (hx : x ∈ e.target) : proj (e.toLocalHomeomorph.symm x) = x.1 := +theorem proj_symm_apply {x : B × F} (hx : x ∈ e.target) : + proj (e.toPartialHomeomorph.symm x) = x.1 := e.toPretrivialization.proj_symm_apply hx #align trivialization.proj_symm_apply Trivialization.proj_symm_apply theorem proj_symm_apply' {b : B} {x : F} (hx : b ∈ e.baseSet) : - proj (e.toLocalHomeomorph.symm (b, x)) = b := + proj (e.toPartialHomeomorph.symm (b, x)) = b := e.toPretrivialization.proj_symm_apply' hx #align trivialization.proj_symm_apply' Trivialization.proj_symm_apply' @@ -406,17 +407,17 @@ theorem proj_surjOn_baseSet [Nonempty F] : Set.SurjOn proj e.source e.baseSet := e.toPretrivialization.proj_surjOn_baseSet #align trivialization.proj_surj_on_base_set Trivialization.proj_surjOn_baseSet -theorem apply_symm_apply {x : B × F} (hx : x ∈ e.target) : e (e.toLocalHomeomorph.symm x) = x := - e.toLocalHomeomorph.right_inv hx +theorem apply_symm_apply {x : B × F} (hx : x ∈ e.target) : e (e.toPartialHomeomorph.symm x) = x := + e.toPartialHomeomorph.right_inv hx #align trivialization.apply_symm_apply Trivialization.apply_symm_apply theorem apply_symm_apply' {b : B} {x : F} (hx : b ∈ e.baseSet) : - e (e.toLocalHomeomorph.symm (b, x)) = (b, x) := + e (e.toPartialHomeomorph.symm (b, x)) = (b, x) := e.toPretrivialization.apply_symm_apply' hx #align trivialization.apply_symm_apply' Trivialization.apply_symm_apply' @[simp, mfld_simps] -theorem symm_apply_mk_proj (ex : x ∈ e.source) : e.toLocalHomeomorph.symm (proj x, (e x).2) = x := +theorem symm_apply_mk_proj (ex : x ∈ e.source) : e.toPartialHomeomorph.symm (proj x, (e x).2) = x := e.toPretrivialization.symm_apply_mk_proj ex #align trivialization.symm_apply_mk_proj Trivialization.symm_apply_mk_proj @@ -478,7 +479,7 @@ theorem nhds_eq_inf_comap {z : Z} (hz : z ∈ e.source) : /-- The preimage of a subset of the base set is homeomorphic to the product with the fiber. -/ def preimageHomeomorph {s : Set B} (hb : s ⊆ e.baseSet) : proj ⁻¹' s ≃ₜ s × F := - (e.toLocalHomeomorph.homeomorphOfImageSubsetSource (e.preimage_subset_source hb) + (e.toPartialHomeomorph.homeomorphOfImageSubsetSource (e.preimage_subset_source hb) (e.image_preimage_eq_prod_univ hb)).trans ((Homeomorph.Set.prod s univ).trans ((Homeomorph.refl s).prodCongr (Homeomorph.Set.univ F))) #align trivialization.preimage_homeomorph Trivialization.preimageHomeomorph @@ -541,7 +542,7 @@ theorem continuousAt_proj (ex : x ∈ e.source) : ContinuousAt proj x := /-- Composition of a `Trivialization` and a `Homeomorph`. -/ protected def compHomeomorph {Z' : Type*} [TopologicalSpace Z'] (h : Z' ≃ₜ Z) : Trivialization F (proj ∘ h) where - toLocalHomeomorph := h.toLocalHomeomorph.trans e.toLocalHomeomorph + toPartialHomeomorph := h.toPartialHomeomorph.trans e.toPartialHomeomorph baseSet := e.baseSet open_baseSet := e.open_baseSet source_eq := by simp [source_eq, preimage_preimage, (· ∘ ·)] @@ -559,8 +560,8 @@ theorem continuousAt_of_comp_right {X : Type*} [TopologicalSpace X] {f : Z → X have hez : z ∈ e.toLocalEquiv.symm.target := by rw [LocalEquiv.symm_target, e.mem_source] exact he - rwa [e.toLocalHomeomorph.symm.continuousAt_iff_continuousAt_comp_right hez, - LocalHomeomorph.symm_symm] + rwa [e.toPartialHomeomorph.symm.continuousAt_iff_continuousAt_comp_right hez, + PartialHomeomorph.symm_symm] #align trivialization.continuous_at_of_comp_right Trivialization.continuousAt_of_comp_right /-- Read off the continuity of a function `f : X → Z` at `x : X` by transferring via a @@ -584,7 +585,7 @@ theorem coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.baseSet := e'.mem_source #align trivialization.coe_mem_source Trivialization.coe_mem_source -@[deprecated LocalHomeomorph.open_target] +@[deprecated PartialHomeomorph.open_target] theorem open_target' : IsOpen e'.target := e'.open_target #align trivialization.open_target Trivialization.open_target' @@ -598,13 +599,13 @@ theorem mk_mem_target {y : F} : (b, y) ∈ e'.target ↔ b ∈ e'.baseSet := #align trivialization.mk_mem_target Trivialization.mk_mem_target theorem symm_apply_apply {x : TotalSpace F E} (hx : x ∈ e'.source) : - e'.toLocalHomeomorph.symm (e' x) = x := + e'.toPartialHomeomorph.symm (e' x) = x := e'.toLocalEquiv.left_inv hx #align trivialization.symm_apply_apply Trivialization.symm_apply_apply @[simp, mfld_simps] theorem symm_coe_proj {x : B} {y : F} (e : Trivialization F (π F E)) (h : x ∈ e.baseSet) : - (e.toLocalHomeomorph.symm (x, y)).1 = x := + (e.toPartialHomeomorph.symm (x, y)).1 = x := e.proj_symm_apply' h #align trivialization.symm_coe_proj Trivialization.symm_coe_proj @@ -619,7 +620,7 @@ protected noncomputable def symm (e : Trivialization F (π F E)) (b : B) (y : F) #align trivialization.symm Trivialization.symm theorem symm_apply (e : Trivialization F (π F E)) {b : B} (hb : b ∈ e.baseSet) (y : F) : - e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.toLocalHomeomorph.symm (b, y)).2 := + e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.toPartialHomeomorph.symm (b, y)).2 := dif_pos hb #align trivialization.symm_apply Trivialization.symm_apply @@ -629,7 +630,7 @@ theorem symm_apply_of_not_mem (e : Trivialization F (π F E)) {b : B} (hb : b #align trivialization.symm_apply_of_not_mem Trivialization.symm_apply_of_not_mem theorem mk_symm (e : Trivialization F (π F E)) {b : B} (hb : b ∈ e.baseSet) (y : F) : - TotalSpace.mk b (e.symm b y) = e.toLocalHomeomorph.symm (b, y) := + TotalSpace.mk b (e.symm b y) = e.toPartialHomeomorph.symm (b, y) := e.toPretrivialization.mk_symm hb y #align trivialization.mk_symm Trivialization.mk_symm @@ -651,12 +652,12 @@ theorem apply_mk_symm (e : Trivialization F (π F E)) {b : B} (hb : b ∈ e.base theorem continuousOn_symm (e : Trivialization F (π F E)) : ContinuousOn (fun z : B × F => TotalSpace.mk' F z.1 (e.symm z.1 z.2)) (e.baseSet ×ˢ univ) := by have : ∀ z ∈ e.baseSet ×ˢ (univ : Set F), - TotalSpace.mk z.1 (e.symm z.1 z.2) = e.toLocalHomeomorph.symm z := by + TotalSpace.mk z.1 (e.symm z.1 z.2) = e.toPartialHomeomorph.symm z := by rintro x ⟨hx : x.1 ∈ e.baseSet, _⟩ rw [e.mk_symm hx] refine' ContinuousOn.congr _ this rw [← e.target_eq] - exact e.toLocalHomeomorph.continuousOn_symm + exact e.toPartialHomeomorph.continuousOn_symm #align trivialization.continuous_on_symm Trivialization.continuousOn_symm end Zero @@ -666,7 +667,7 @@ end Zero that sends `p : Z` to `((e p).1, h (e p).2)`. -/ def transFiberHomeomorph {F' : Type*} [TopologicalSpace F'] (e : Trivialization F proj) (h : F ≃ₜ F') : Trivialization F' proj where - toLocalHomeomorph := e.toLocalHomeomorph.transHomeomorph <| (Homeomorph.refl _).prodCongr h + toPartialHomeomorph := e.toPartialHomeomorph.transHomeomorph <| (Homeomorph.refl _).prodCongr h baseSet := e.baseSet open_baseSet := e.open_baseSet source_eq := e.source_eq @@ -683,12 +684,12 @@ theorem transFiberHomeomorph_apply {F' : Type*} [TopologicalSpace F'] (e : Trivi /-- Coordinate transformation in the fiber induced by a pair of bundle trivializations. See also `Trivialization.coordChangeHomeomorph` for a version bundled as `F ≃ₜ F`. -/ def coordChange (e₁ e₂ : Trivialization F proj) (b : B) (x : F) : F := - (e₂ <| e₁.toLocalHomeomorph.symm (b, x)).2 + (e₂ <| e₁.toPartialHomeomorph.symm (b, x)).2 #align trivialization.coord_change Trivialization.coordChange theorem mk_coordChange (e₁ e₂ : Trivialization F proj) {b : B} (h₁ : b ∈ e₁.baseSet) (h₂ : b ∈ e₂.baseSet) (x : F) : - (b, e₁.coordChange e₂ b x) = e₂ (e₁.toLocalHomeomorph.symm (b, x)) := by + (b, e₁.coordChange e₂ b x) = e₂ (e₁.toPartialHomeomorph.symm (b, x)) := by refine' Prod.ext _ rfl rw [e₂.coe_fst', ← e₁.coe_fst', e₁.apply_symm_apply' h₁] · rwa [e₁.proj_symm_apply' h₁] @@ -718,8 +719,8 @@ theorem coordChange_coordChange (e₁ e₂ e₃ : Trivialization F proj) {b : B} theorem continuous_coordChange (e₁ e₂ : Trivialization F proj) {b : B} (h₁ : b ∈ e₁.baseSet) (h₂ : b ∈ e₂.baseSet) : Continuous (e₁.coordChange e₂ b) := by - refine' continuous_snd.comp (e₂.toLocalHomeomorph.continuousOn.comp_continuous - (e₁.toLocalHomeomorph.continuousOn_symm.comp_continuous _ _) _) + refine' continuous_snd.comp (e₂.toPartialHomeomorph.continuousOn.comp_continuous + (e₁.toPartialHomeomorph.continuousOn_symm.comp_continuous _ _) _) · exact continuous_const.prod_mk continuous_id · exact fun x => e₁.mem_target.2 h₁ · intro x @@ -747,13 +748,13 @@ theorem coordChangeHomeomorph_coe (e₁ e₂ : Trivialization F proj) {b : B} (h variable {B' : Type*} [TopologicalSpace B'] theorem isImage_preimage_prod (e : Trivialization F proj) (s : Set B) : - e.toLocalHomeomorph.IsImage (proj ⁻¹' s) (s ×ˢ univ) := fun x hx => by simp [e.coe_fst', hx] + e.toPartialHomeomorph.IsImage (proj ⁻¹' s) (s ×ˢ univ) := fun x hx => by simp [e.coe_fst', hx] #align trivialization.is_image_preimage_prod Trivialization.isImage_preimage_prod /-- Restrict a `Trivialization` to an open set in the base. -/ protected def restrOpen (e : Trivialization F proj) (s : Set B) (hs : IsOpen s) : Trivialization F proj where - toLocalHomeomorph := + toPartialHomeomorph := ((e.isImage_preimage_prod s).symm.restr (IsOpen.inter e.open_target (hs.prod isOpen_univ))).symm baseSet := e.baseSet ∩ s open_baseSet := IsOpen.inter e.open_baseSet hs @@ -778,8 +779,8 @@ otherwise. -/ noncomputable def piecewise (e e' : Trivialization F proj) (s : Set B) (Hs : e.baseSet ∩ frontier s = e'.baseSet ∩ frontier s) (Heq : EqOn e e' <| proj ⁻¹' (e.baseSet ∩ frontier s)) : Trivialization F proj where - toLocalHomeomorph := - e.toLocalHomeomorph.piecewise e'.toLocalHomeomorph (proj ⁻¹' s) (s ×ˢ univ) + toPartialHomeomorph := + e.toPartialHomeomorph.piecewise e'.toPartialHomeomorph (proj ⁻¹' s) (s ×ˢ univ) (e.isImage_preimage_prod s) (e'.isImage_preimage_prod s) (by rw [e.frontier_preimage, e'.frontier_preimage, Hs]) (by rwa [e.frontier_preimage]) baseSet := s.ite e.baseSet e'.baseSet @@ -828,8 +829,8 @@ bundle trivialization over the union of the base sets that agrees with `e` and ` base sets. -/ noncomputable def disjointUnion (e e' : Trivialization F proj) (H : Disjoint e.baseSet e'.baseSet) : Trivialization F proj where - toLocalHomeomorph := - e.toLocalHomeomorph.disjointUnion e'.toLocalHomeomorph + toPartialHomeomorph := + e.toPartialHomeomorph.disjointUnion e'.toPartialHomeomorph (by rw [e.source_eq, e'.source_eq] exact H.preimage _) diff --git a/Mathlib/Topology/IsLocalHomeomorph.lean b/Mathlib/Topology/IsLocalHomeomorph.lean index 43d9f4bcb3c93..c1f1a82d1bfc4 100644 --- a/Mathlib/Topology/IsLocalHomeomorph.lean +++ b/Mathlib/Topology/IsLocalHomeomorph.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph import Mathlib.Topology.SeparatedMap #align_import topology.is_locally_homeomorph from "leanprover-community/mathlib"@"e97cf15cd1aec9bd5c193b2ffac5a6dc9118912b" @@ -20,7 +20,7 @@ This file defines local homeomorphisms. between `U` and an open subset of `Y`. Note that `IsLocalHomeomorph` is a global condition. This is in contrast to - `LocalHomeomorph`, which is a homeomorphism between specific open subsets. + `PartialHomeomorph`, which is a homeomorphism between specific open subsets. -/ @@ -30,9 +30,9 @@ variable {X Y Z : Type*} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalS (f : X → Y) (s : Set X) (t : Set Y) /-- A function `f : X → Y` satisfies `IsLocalHomeomorphOn f s` if each `x ∈ s` is contained in -the source of some `e : LocalHomeomorph X Y` with `f = e`. -/ +the source of some `e : PartialHomeomorph X Y` with `f = e`. -/ def IsLocalHomeomorphOn := - ∀ x ∈ s, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ f = e + ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ f = e #align is_locally_homeomorph_on IsLocalHomeomorphOn theorem isLocalHomeomorphOn_iff_openEmbedding_restrict {f : X → Y} : @@ -46,16 +46,16 @@ theorem isLocalHomeomorphOn_iff_openEmbedding_restrict {f : X → Y} : rw [Set.range_inclusion]; exact isOpen_induced isOpen_interior obtain ⟨cont, inj, openMap⟩ := openEmbedding_iff_continuous_injective_open.mp this haveI : Nonempty X := ⟨x⟩ - exact ⟨LocalHomeomorph.ofContinuousOpenRestrict (Set.injOn_iff_injective.mpr inj).toLocalEquiv + exact ⟨PartialHomeomorph.ofContinuousOpenRestrict (Set.injOn_iff_injective.mpr inj).toLocalEquiv (continuousOn_iff_continuous_restrict.mpr cont) openMap isOpen_interior, mem_interior_iff_mem_nhds.mpr hU, rfl⟩ namespace IsLocalHomeomorphOn /-- Proves that `f` satisfies `IsLocalHomeomorphOn f s`. The condition `h` is weaker than the -definition of `IsLocalHomeomorphOn f s`, since it only requires `e : LocalHomeomorph X Y` to +definition of `IsLocalHomeomorphOn f s`, since it only requires `e : PartialHomeomorph X Y` to agree with `f` on its source `e.source`, as opposed to on the whole space `X`. -/ -theorem mk (h : ∀ x ∈ s, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : +theorem mk (h : ∀ x ∈ s, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : IsLocalHomeomorphOn f s := by intro x hx obtain ⟨e, hx, he⟩ := h x hx @@ -121,13 +121,13 @@ protected theorem comp (hg : IsLocalHomeomorphOn g t) (hf : IsLocalHomeomorphOn end IsLocalHomeomorphOn /-- A function `f : X → Y` satisfies `IsLocalHomeomorph f` if each `x : x` is contained in - the source of some `e : LocalHomeomorph X Y` with `f = e`. -/ + the source of some `e : PartialHomeomorph X Y` with `f = e`. -/ def IsLocalHomeomorph := - ∀ x : X, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ f = e + ∀ x : X, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ f = e #align is_locally_homeomorph IsLocalHomeomorph theorem Homeomorph.isLocalHomeomorph (f : X ≃ₜ Y) : IsLocalHomeomorph f := - fun _ ↦ ⟨f.toLocalHomeomorph, trivial, rfl⟩ + fun _ ↦ ⟨f.toPartialHomeomorph, trivial, rfl⟩ variable {f s} @@ -154,9 +154,9 @@ variable (f) namespace IsLocalHomeomorph /-- Proves that `f` satisfies `IsLocalHomeomorph f`. The condition `h` is weaker than the -definition of `IsLocalHomeomorph f`, since it only requires `e : LocalHomeomorph X Y` to +definition of `IsLocalHomeomorph f`, since it only requires `e : PartialHomeomorph X Y` to agree with `f` on its source `e.source`, as opposed to on the whole space `X`. -/ -theorem mk (h : ∀ x : X, ∃ e : LocalHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : +theorem mk (h : ∀ x : X, ∃ e : PartialHomeomorph X Y, x ∈ e.source ∧ ∀ y ∈ e.source, f y = e y) : IsLocalHomeomorph f := isLocalHomeomorph_iff_isLocalHomeomorphOn_univ.mpr (IsLocalHomeomorphOn.mk f Set.univ fun x _hx => h x) diff --git a/Mathlib/Topology/LocalHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean similarity index 57% rename from Mathlib/Topology/LocalHomeomorph.lean rename to Mathlib/Topology/PartialHomeomorph.lean index a39c69ce16469..9a92d543ee735 100644 --- a/Mathlib/Topology/LocalHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -9,10 +9,10 @@ import Mathlib.Topology.Sets.Opens #align_import topology.local_homeomorph from "leanprover-community/mathlib"@"431589bce478b2229eba14b14a283250428217db" /-! -# Local homeomorphisms +# Partial homeomorphisms This file defines homeomorphisms between open subsets of topological spaces. An element `e` of -`LocalHomeomorph α β` is an extension of `LocalEquiv α β`, i.e., it is a pair of functions +`PartialHomeomorph α β` is an extension of `LocalEquiv α β`, i.e., it is a pair of functions `e.toFun` and `e.invFun`, inverse of each other on the sets `e.source` and `e.target`. Additionally, we require that these sets are open, and that the functions are continuous on them. Equivalently, they are homeomorphisms there. @@ -22,14 +22,14 @@ instead of `e.toFun x` and `e.invFun x`. ## Main definitions -* `Homeomorph.toLocalHomeomorph`: associating a local homeomorphism to a homeomorphism, with +* `Homeomorph.toPartialHomeomorph`: associating a partial homeomorphism to a homeomorphism, with `source = target = Set.univ`; -* `LocalHomeomorph.symm`: the inverse of a local homeomorphism -* `LocalHomeomorph.trans`: the composition of two local homeomorphisms -* `LocalHomeomorph.refl`: the identity local homeomorphism -* `LocalHomeomorph.ofSet`: the identity on a set `s` -* `LocalHomeomorph.EqOnSource`: equivalence relation describing the "right" notion of equality - for local homeomorphisms +* `PartialHomeomorph.symm`: the inverse of a partial homeomorphism +* `PartialHomeomorph.trans`: the composition of two partial homeomorphisms +* `PartialHomeomorph.refl`: the identity partial homeomorphism +* `PartialHomeomorph.ofSet`: the identity on a set `s` +* `PartialHomeomorph.EqOnSource`: equivalence relation describing the "right" notion of equality + for partial homeomorphisms ## Implementation notes @@ -50,98 +50,99 @@ open Function Set Filter Topology variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] -/-- local homeomorphisms, defined on open subsets of the space -/ +/-- partial homeomorphisms, defined on open subsets of the space -/ -- porting note: commented @[nolint has_nonempty_instance] -structure LocalHomeomorph (α : Type*) (β : Type*) [TopologicalSpace α] +structure PartialHomeomorph (α : Type*) (β : Type*) [TopologicalSpace α] [TopologicalSpace β] extends LocalEquiv α β where open_source : IsOpen source open_target : IsOpen target continuousOn_toFun : ContinuousOn toFun source continuousOn_invFun : ContinuousOn invFun target -#align local_homeomorph LocalHomeomorph +#align local_homeomorph PartialHomeomorph -namespace LocalHomeomorph +namespace PartialHomeomorph -variable (e : LocalHomeomorph α β) (e' : LocalHomeomorph β γ) +variable (e : PartialHomeomorph α β) (e' : PartialHomeomorph β γ) -/-- Coercion of a local homeomorphisms to a function. We don't use `e.toFun` because it is actually -`e.toLocalEquiv.toFun`, so `simp` will apply lemmas about `toLocalEquiv`. While we may want to -switch to this behavior later, doing it mid-port will break a lot of proofs. -/ +/-- Coercion of a partial homeomorphisms to a function. We don't use `e.toFun` because it is +actually `e.toLocalEquiv.toFun`, so `simp` will apply lemmas about `toLocalEquiv`. +While we may want to switch to this behavior later, doing it mid-port will break a lot of proofs. -/ @[coe] def toFun' : α → β := e.toFun -/-- Coercion of a `LocalHomeomorph` to function. Note that a `LocalHomeomorph` is not `FunLike`. -/ -instance : CoeFun (LocalHomeomorph α β) fun _ => α → β := +/-- Coercion of a `PartialHomeomorph` to function. +Note that a `PartialHomeomorph` is not `FunLike`. -/ +instance : CoeFun (PartialHomeomorph α β) fun _ => α → β := ⟨fun e => e.toFun'⟩ -/-- The inverse of a local homeomorphism -/ -protected def symm : LocalHomeomorph β α where +/-- The inverse of a partial homeomorphism -/ +protected def symm : PartialHomeomorph β α where toLocalEquiv := e.toLocalEquiv.symm open_source := e.open_target open_target := e.open_source continuousOn_toFun := e.continuousOn_invFun continuousOn_invFun := e.continuousOn_toFun -#align local_homeomorph.symm LocalHomeomorph.symm +#align local_homeomorph.symm PartialHomeomorph.symm /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ -def Simps.apply (e : LocalHomeomorph α β) : α → β := e -#align local_homeomorph.simps.apply LocalHomeomorph.Simps.apply +def Simps.apply (e : PartialHomeomorph α β) : α → β := e +#align local_homeomorph.simps.apply PartialHomeomorph.Simps.apply /-- See Note [custom simps projection] -/ -def Simps.symm_apply (e : LocalHomeomorph α β) : β → α := e.symm -#align local_homeomorph.simps.symm_apply LocalHomeomorph.Simps.symm_apply +def Simps.symm_apply (e : PartialHomeomorph α β) : β → α := e.symm +#align local_homeomorph.simps.symm_apply PartialHomeomorph.Simps.symm_apply -initialize_simps_projections LocalHomeomorph (toFun → apply, invFun → symm_apply) +initialize_simps_projections PartialHomeomorph (toFun → apply, invFun → symm_apply) protected theorem continuousOn : ContinuousOn e e.source := e.continuousOn_toFun -#align local_homeomorph.continuous_on LocalHomeomorph.continuousOn +#align local_homeomorph.continuous_on PartialHomeomorph.continuousOn theorem continuousOn_symm : ContinuousOn e.symm e.target := e.continuousOn_invFun -#align local_homeomorph.continuous_on_symm LocalHomeomorph.continuousOn_symm +#align local_homeomorph.continuous_on_symm PartialHomeomorph.continuousOn_symm @[simp, mfld_simps] -theorem mk_coe (e : LocalEquiv α β) (a b c d) : (LocalHomeomorph.mk e a b c d : α → β) = e := +theorem mk_coe (e : LocalEquiv α β) (a b c d) : (PartialHomeomorph.mk e a b c d : α → β) = e := rfl -#align local_homeomorph.mk_coe LocalHomeomorph.mk_coe +#align local_homeomorph.mk_coe PartialHomeomorph.mk_coe @[simp, mfld_simps] theorem mk_coe_symm (e : LocalEquiv α β) (a b c d) : - ((LocalHomeomorph.mk e a b c d).symm : β → α) = e.symm := + ((PartialHomeomorph.mk e a b c d).symm : β → α) = e.symm := rfl -#align local_homeomorph.mk_coe_symm LocalHomeomorph.mk_coe_symm +#align local_homeomorph.mk_coe_symm PartialHomeomorph.mk_coe_symm -theorem toLocalEquiv_injective : Injective (toLocalEquiv : LocalHomeomorph α β → LocalEquiv α β) +theorem toLocalEquiv_injective : Injective (toLocalEquiv : PartialHomeomorph α β → LocalEquiv α β) | ⟨_, _, _, _, _⟩, ⟨_, _, _, _, _⟩, rfl => rfl -#align local_homeomorph.to_local_equiv_injective LocalHomeomorph.toLocalEquiv_injective +#align local_homeomorph.to_local_equiv_injective PartialHomeomorph.toLocalEquiv_injective /- Register a few simp lemmas to make sure that `simp` puts the application of a local homeomorphism in its normal form, i.e., in terms of its coercion to a function. -/ @[simp, mfld_simps] -theorem toFun_eq_coe (e : LocalHomeomorph α β) : e.toFun = e := +theorem toFun_eq_coe (e : PartialHomeomorph α β) : e.toFun = e := rfl -#align local_homeomorph.to_fun_eq_coe LocalHomeomorph.toFun_eq_coe +#align local_homeomorph.to_fun_eq_coe PartialHomeomorph.toFun_eq_coe @[simp, mfld_simps] -theorem invFun_eq_coe (e : LocalHomeomorph α β) : e.invFun = e.symm := +theorem invFun_eq_coe (e : PartialHomeomorph α β) : e.invFun = e.symm := rfl -#align local_homeomorph.inv_fun_eq_coe LocalHomeomorph.invFun_eq_coe +#align local_homeomorph.inv_fun_eq_coe PartialHomeomorph.invFun_eq_coe @[simp, mfld_simps] theorem coe_coe : (e.toLocalEquiv : α → β) = e := rfl -#align local_homeomorph.coe_coe LocalHomeomorph.coe_coe +#align local_homeomorph.coe_coe PartialHomeomorph.coe_coe @[simp, mfld_simps] theorem coe_coe_symm : (e.toLocalEquiv.symm : β → α) = e.symm := rfl -#align local_homeomorph.coe_coe_symm LocalHomeomorph.coe_coe_symm +#align local_homeomorph.coe_coe_symm PartialHomeomorph.coe_coe_symm @[simp, mfld_simps] theorem map_source {x : α} (h : x ∈ e.source) : e x ∈ e.target := e.map_source' h -#align local_homeomorph.map_source LocalHomeomorph.map_source +#align local_homeomorph.map_source PartialHomeomorph.map_source /-- Variant of `map_source`, stated for images of subsets of `source`. -/ lemma map_source'' : e '' e.source ⊆ e.target := @@ -150,244 +151,244 @@ lemma map_source'' : e '' e.source ⊆ e.target := @[simp, mfld_simps] theorem map_target {x : β} (h : x ∈ e.target) : e.symm x ∈ e.source := e.map_target' h -#align local_homeomorph.map_target LocalHomeomorph.map_target +#align local_homeomorph.map_target PartialHomeomorph.map_target @[simp, mfld_simps] theorem left_inv {x : α} (h : x ∈ e.source) : e.symm (e x) = x := e.left_inv' h -#align local_homeomorph.left_inv LocalHomeomorph.left_inv +#align local_homeomorph.left_inv PartialHomeomorph.left_inv @[simp, mfld_simps] theorem right_inv {x : β} (h : x ∈ e.target) : e (e.symm x) = x := e.right_inv' h -#align local_homeomorph.right_inv LocalHomeomorph.right_inv +#align local_homeomorph.right_inv PartialHomeomorph.right_inv theorem eq_symm_apply {x : α} {y : β} (hx : x ∈ e.source) (hy : y ∈ e.target) : x = e.symm y ↔ e x = y := e.toLocalEquiv.eq_symm_apply hx hy -#align local_homeomorph.eq_symm_apply LocalHomeomorph.eq_symm_apply +#align local_homeomorph.eq_symm_apply PartialHomeomorph.eq_symm_apply protected theorem mapsTo : MapsTo e e.source e.target := fun _ => e.map_source -#align local_homeomorph.maps_to LocalHomeomorph.mapsTo +#align local_homeomorph.maps_to PartialHomeomorph.mapsTo protected theorem symm_mapsTo : MapsTo e.symm e.target e.source := e.symm.mapsTo -#align local_homeomorph.symm_maps_to LocalHomeomorph.symm_mapsTo +#align local_homeomorph.symm_maps_to PartialHomeomorph.symm_mapsTo protected theorem leftInvOn : LeftInvOn e.symm e e.source := fun _ => e.left_inv -#align local_homeomorph.left_inv_on LocalHomeomorph.leftInvOn +#align local_homeomorph.left_inv_on PartialHomeomorph.leftInvOn protected theorem rightInvOn : RightInvOn e.symm e e.target := fun _ => e.right_inv -#align local_homeomorph.right_inv_on LocalHomeomorph.rightInvOn +#align local_homeomorph.right_inv_on PartialHomeomorph.rightInvOn protected theorem invOn : InvOn e.symm e e.source e.target := ⟨e.leftInvOn, e.rightInvOn⟩ -#align local_homeomorph.inv_on LocalHomeomorph.invOn +#align local_homeomorph.inv_on PartialHomeomorph.invOn protected theorem injOn : InjOn e e.source := e.leftInvOn.injOn -#align local_homeomorph.inj_on LocalHomeomorph.injOn +#align local_homeomorph.inj_on PartialHomeomorph.injOn protected theorem bijOn : BijOn e e.source e.target := e.invOn.bijOn e.mapsTo e.symm_mapsTo -#align local_homeomorph.bij_on LocalHomeomorph.bijOn +#align local_homeomorph.bij_on PartialHomeomorph.bijOn protected theorem surjOn : SurjOn e e.source e.target := e.bijOn.surjOn -#align local_homeomorph.surj_on LocalHomeomorph.surjOn +#align local_homeomorph.surj_on PartialHomeomorph.surjOn -/-- Interpret a `Homeomorph` as a `LocalHomeomorph` by restricting it +/-- Interpret a `Homeomorph` as a `PartialHomeomorph` by restricting it to an open set `s` in the domain and to `t` in the codomain. -/ @[simps! (config := .asFn) apply symm_apply toLocalEquiv, simps! (config := .lemmasOnly) source target] -def _root_.Homeomorph.toLocalHomeomorphOfImageEq (e : α ≃ₜ β) (s : Set α) (hs : IsOpen s) - (t : Set β) (h : e '' s = t) : LocalHomeomorph α β where +def _root_.Homeomorph.toPartialHomeomorphOfImageEq (e : α ≃ₜ β) (s : Set α) (hs : IsOpen s) + (t : Set β) (h : e '' s = t) : PartialHomeomorph α β where toLocalEquiv := e.toLocalEquivOfImageEq s t h open_source := hs open_target := by simpa [← h] continuousOn_toFun := e.continuous.continuousOn continuousOn_invFun := e.symm.continuous.continuousOn -/-- A homeomorphism induces a local homeomorphism on the whole space -/ +/-- A homeomorphism induces a partial homeomorphism on the whole space -/ @[simps! (config := mfld_cfg)] -def _root_.Homeomorph.toLocalHomeomorph (e : α ≃ₜ β) : LocalHomeomorph α β := - e.toLocalHomeomorphOfImageEq univ isOpen_univ univ <| by rw [image_univ, e.surjective.range_eq] -#align homeomorph.to_local_homeomorph Homeomorph.toLocalHomeomorph +def _root_.Homeomorph.toPartialHomeomorph (e : α ≃ₜ β) : PartialHomeomorph α β := + e.toPartialHomeomorphOfImageEq univ isOpen_univ univ <| by rw [image_univ, e.surjective.range_eq] +#align homeomorph.to_local_homeomorph Homeomorph.toPartialHomeomorph /-- Replace `toLocalEquiv` field to provide better definitional equalities. -/ -def replaceEquiv (e : LocalHomeomorph α β) (e' : LocalEquiv α β) (h : e.toLocalEquiv = e') : - LocalHomeomorph α β where +def replaceEquiv (e : PartialHomeomorph α β) (e' : LocalEquiv α β) (h : e.toLocalEquiv = e') : + PartialHomeomorph α β where toLocalEquiv := e' open_source := h ▸ e.open_source open_target := h ▸ e.open_target continuousOn_toFun := h ▸ e.continuousOn_toFun continuousOn_invFun := h ▸ e.continuousOn_invFun -#align local_homeomorph.replace_equiv LocalHomeomorph.replaceEquiv +#align local_homeomorph.replace_equiv PartialHomeomorph.replaceEquiv -theorem replaceEquiv_eq_self (e : LocalHomeomorph α β) (e' : LocalEquiv α β) +theorem replaceEquiv_eq_self (e : PartialHomeomorph α β) (e' : LocalEquiv α β) (h : e.toLocalEquiv = e') : e.replaceEquiv e' h = e := by cases e subst e' rfl -#align local_homeomorph.replace_equiv_eq_self LocalHomeomorph.replaceEquiv_eq_self +#align local_homeomorph.replace_equiv_eq_self PartialHomeomorph.replaceEquiv_eq_self theorem source_preimage_target : e.source ⊆ e ⁻¹' e.target := e.mapsTo -#align local_homeomorph.source_preimage_target LocalHomeomorph.source_preimage_target +#align local_homeomorph.source_preimage_target PartialHomeomorph.source_preimage_target @[deprecated toLocalEquiv_injective] -theorem eq_of_localEquiv_eq {e e' : LocalHomeomorph α β} (h : e.toLocalEquiv = e'.toLocalEquiv) : +theorem eq_of_localEquiv_eq {e e' : PartialHomeomorph α β} (h : e.toLocalEquiv = e'.toLocalEquiv) : e = e' := toLocalEquiv_injective h -#align local_homeomorph.eq_of_local_equiv_eq LocalHomeomorph.eq_of_localEquiv_eq +#align local_homeomorph.eq_of_local_equiv_eq PartialHomeomorph.eq_of_localEquiv_eq -theorem eventually_left_inverse (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) : +theorem eventually_left_inverse (e : PartialHomeomorph α β) {x} (hx : x ∈ e.source) : ∀ᶠ y in 𝓝 x, e.symm (e y) = y := (e.open_source.eventually_mem hx).mono e.left_inv' -#align local_homeomorph.eventually_left_inverse LocalHomeomorph.eventually_left_inverse +#align local_homeomorph.eventually_left_inverse PartialHomeomorph.eventually_left_inverse -theorem eventually_left_inverse' (e : LocalHomeomorph α β) {x} (hx : x ∈ e.target) : +theorem eventually_left_inverse' (e : PartialHomeomorph α β) {x} (hx : x ∈ e.target) : ∀ᶠ y in 𝓝 (e.symm x), e.symm (e y) = y := e.eventually_left_inverse (e.map_target hx) -#align local_homeomorph.eventually_left_inverse' LocalHomeomorph.eventually_left_inverse' +#align local_homeomorph.eventually_left_inverse' PartialHomeomorph.eventually_left_inverse' -theorem eventually_right_inverse (e : LocalHomeomorph α β) {x} (hx : x ∈ e.target) : +theorem eventually_right_inverse (e : PartialHomeomorph α β) {x} (hx : x ∈ e.target) : ∀ᶠ y in 𝓝 x, e (e.symm y) = y := (e.open_target.eventually_mem hx).mono e.right_inv' -#align local_homeomorph.eventually_right_inverse LocalHomeomorph.eventually_right_inverse +#align local_homeomorph.eventually_right_inverse PartialHomeomorph.eventually_right_inverse -theorem eventually_right_inverse' (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) : +theorem eventually_right_inverse' (e : PartialHomeomorph α β) {x} (hx : x ∈ e.source) : ∀ᶠ y in 𝓝 (e x), e (e.symm y) = y := e.eventually_right_inverse (e.map_source hx) -#align local_homeomorph.eventually_right_inverse' LocalHomeomorph.eventually_right_inverse' +#align local_homeomorph.eventually_right_inverse' PartialHomeomorph.eventually_right_inverse' -theorem eventually_ne_nhdsWithin (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) : +theorem eventually_ne_nhdsWithin (e : PartialHomeomorph α β) {x} (hx : x ∈ e.source) : ∀ᶠ x' in 𝓝[≠] x, e x' ≠ e x := eventually_nhdsWithin_iff.2 <| (e.eventually_left_inverse hx).mono fun x' hx' => mt fun h => by rw [mem_singleton_iff, ← e.left_inv hx, ← h, hx'] -#align local_homeomorph.eventually_ne_nhds_within LocalHomeomorph.eventually_ne_nhdsWithin +#align local_homeomorph.eventually_ne_nhds_within PartialHomeomorph.eventually_ne_nhdsWithin theorem nhdsWithin_source_inter {x} (hx : x ∈ e.source) (s : Set α) : 𝓝[e.source ∩ s] x = 𝓝[s] x := nhdsWithin_inter_of_mem (mem_nhdsWithin_of_mem_nhds <| IsOpen.mem_nhds e.open_source hx) -#align local_homeomorph.nhds_within_source_inter LocalHomeomorph.nhdsWithin_source_inter +#align local_homeomorph.nhds_within_source_inter PartialHomeomorph.nhdsWithin_source_inter theorem nhdsWithin_target_inter {x} (hx : x ∈ e.target) (s : Set β) : 𝓝[e.target ∩ s] x = 𝓝[s] x := e.symm.nhdsWithin_source_inter hx s -#align local_homeomorph.nhds_within_target_inter LocalHomeomorph.nhdsWithin_target_inter +#align local_homeomorph.nhds_within_target_inter PartialHomeomorph.nhdsWithin_target_inter theorem image_eq_target_inter_inv_preimage {s : Set α} (h : s ⊆ e.source) : e '' s = e.target ∩ e.symm ⁻¹' s := e.toLocalEquiv.image_eq_target_inter_inv_preimage h -#align local_homeomorph.image_eq_target_inter_inv_preimage LocalHomeomorph.image_eq_target_inter_inv_preimage +#align local_homeomorph.image_eq_target_inter_inv_preimage PartialHomeomorph.image_eq_target_inter_inv_preimage theorem image_source_inter_eq' (s : Set α) : e '' (e.source ∩ s) = e.target ∩ e.symm ⁻¹' s := e.toLocalEquiv.image_source_inter_eq' s -#align local_homeomorph.image_source_inter_eq' LocalHomeomorph.image_source_inter_eq' +#align local_homeomorph.image_source_inter_eq' PartialHomeomorph.image_source_inter_eq' theorem image_source_inter_eq (s : Set α) : e '' (e.source ∩ s) = e.target ∩ e.symm ⁻¹' (e.source ∩ s) := e.toLocalEquiv.image_source_inter_eq s -#align local_homeomorph.image_source_inter_eq LocalHomeomorph.image_source_inter_eq +#align local_homeomorph.image_source_inter_eq PartialHomeomorph.image_source_inter_eq theorem symm_image_eq_source_inter_preimage {s : Set β} (h : s ⊆ e.target) : e.symm '' s = e.source ∩ e ⁻¹' s := e.symm.image_eq_target_inter_inv_preimage h -#align local_homeomorph.symm_image_eq_source_inter_preimage LocalHomeomorph.symm_image_eq_source_inter_preimage +#align local_homeomorph.symm_image_eq_source_inter_preimage PartialHomeomorph.symm_image_eq_source_inter_preimage theorem symm_image_target_inter_eq (s : Set β) : e.symm '' (e.target ∩ s) = e.source ∩ e ⁻¹' (e.target ∩ s) := e.symm.image_source_inter_eq _ -#align local_homeomorph.symm_image_target_inter_eq LocalHomeomorph.symm_image_target_inter_eq +#align local_homeomorph.symm_image_target_inter_eq PartialHomeomorph.symm_image_target_inter_eq theorem source_inter_preimage_inv_preimage (s : Set α) : e.source ∩ e ⁻¹' (e.symm ⁻¹' s) = e.source ∩ s := e.toLocalEquiv.source_inter_preimage_inv_preimage s -#align local_homeomorph.source_inter_preimage_inv_preimage LocalHomeomorph.source_inter_preimage_inv_preimage +#align local_homeomorph.source_inter_preimage_inv_preimage PartialHomeomorph.source_inter_preimage_inv_preimage theorem target_inter_inv_preimage_preimage (s : Set β) : e.target ∩ e.symm ⁻¹' (e ⁻¹' s) = e.target ∩ s := e.symm.source_inter_preimage_inv_preimage _ -#align local_homeomorph.target_inter_inv_preimage_preimage LocalHomeomorph.target_inter_inv_preimage_preimage +#align local_homeomorph.target_inter_inv_preimage_preimage PartialHomeomorph.target_inter_inv_preimage_preimage theorem source_inter_preimage_target_inter (s : Set β) : e.source ∩ e ⁻¹' (e.target ∩ s) = e.source ∩ e ⁻¹' s := e.toLocalEquiv.source_inter_preimage_target_inter s -#align local_homeomorph.source_inter_preimage_target_inter LocalHomeomorph.source_inter_preimage_target_inter +#align local_homeomorph.source_inter_preimage_target_inter PartialHomeomorph.source_inter_preimage_target_inter -theorem image_source_eq_target (e : LocalHomeomorph α β) : e '' e.source = e.target := +theorem image_source_eq_target (e : PartialHomeomorph α β) : e '' e.source = e.target := e.toLocalEquiv.image_source_eq_target -#align local_homeomorph.image_source_eq_target LocalHomeomorph.image_source_eq_target +#align local_homeomorph.image_source_eq_target PartialHomeomorph.image_source_eq_target -theorem symm_image_target_eq_source (e : LocalHomeomorph α β) : e.symm '' e.target = e.source := +theorem symm_image_target_eq_source (e : PartialHomeomorph α β) : e.symm '' e.target = e.source := e.symm.image_source_eq_target -#align local_homeomorph.symm_image_target_eq_source LocalHomeomorph.symm_image_target_eq_source +#align local_homeomorph.symm_image_target_eq_source PartialHomeomorph.symm_image_target_eq_source -/-- Two local homeomorphisms are equal when they have equal `toFun`, `invFun` and `source`. +/-- Two partial homeomorphisms are equal when they have equal `toFun`, `invFun` and `source`. It is not sufficient to have equal `toFun` and `source`, as this only determines `invFun` on the target. This would only be true for a weaker notion of equality, arguably the right one, called `EqOnSource`. -/ @[ext] -protected theorem ext (e' : LocalHomeomorph α β) (h : ∀ x, e x = e' x) +protected theorem ext (e' : PartialHomeomorph α β) (h : ∀ x, e x = e' x) (hinv : ∀ x, e.symm x = e'.symm x) (hs : e.source = e'.source) : e = e' := toLocalEquiv_injective (LocalEquiv.ext h hinv hs) -#align local_homeomorph.ext LocalHomeomorph.ext +#align local_homeomorph.ext PartialHomeomorph.ext -protected theorem ext_iff {e e' : LocalHomeomorph α β} : +protected theorem ext_iff {e e' : PartialHomeomorph α β} : e = e' ↔ (∀ x, e x = e' x) ∧ (∀ x, e.symm x = e'.symm x) ∧ e.source = e'.source := ⟨by rintro rfl exact ⟨fun x => rfl, fun x => rfl, rfl⟩, fun h => e.ext e' h.1 h.2.1 h.2.2⟩ -#align local_homeomorph.ext_iff LocalHomeomorph.ext_iff +#align local_homeomorph.ext_iff PartialHomeomorph.ext_iff @[simp, mfld_simps] theorem symm_toLocalEquiv : e.symm.toLocalEquiv = e.toLocalEquiv.symm := rfl -#align local_homeomorph.symm_to_local_equiv LocalHomeomorph.symm_toLocalEquiv +#align local_homeomorph.symm_to_local_equiv PartialHomeomorph.symm_toLocalEquiv -- The following lemmas are already simp via `LocalEquiv` theorem symm_source : e.symm.source = e.target := rfl -#align local_homeomorph.symm_source LocalHomeomorph.symm_source +#align local_homeomorph.symm_source PartialHomeomorph.symm_source theorem symm_target : e.symm.target = e.source := rfl -#align local_homeomorph.symm_target LocalHomeomorph.symm_target +#align local_homeomorph.symm_target PartialHomeomorph.symm_target @[simp, mfld_simps] theorem symm_symm : e.symm.symm = e := rfl -#align local_homeomorph.symm_symm LocalHomeomorph.symm_symm +#align local_homeomorph.symm_symm PartialHomeomorph.symm_symm theorem symm_bijective : Function.Bijective - (LocalHomeomorph.symm : LocalHomeomorph α β → LocalHomeomorph β α) := + (PartialHomeomorph.symm : PartialHomeomorph α β → PartialHomeomorph β α) := Function.bijective_iff_has_inverse.mpr ⟨_, symm_symm, symm_symm⟩ -/-- A local homeomorphism is continuous at any point of its source -/ +/-- A partial homeomorphism is continuous at any point of its source -/ protected theorem continuousAt {x : α} (h : x ∈ e.source) : ContinuousAt e x := (e.continuousOn x h).continuousAt (e.open_source.mem_nhds h) -#align local_homeomorph.continuous_at LocalHomeomorph.continuousAt +#align local_homeomorph.continuous_at PartialHomeomorph.continuousAt -/-- A local homeomorphism inverse is continuous at any point of its target -/ +/-- A partial homeomorphism inverse is continuous at any point of its target -/ theorem continuousAt_symm {x : β} (h : x ∈ e.target) : ContinuousAt e.symm x := e.symm.continuousAt h -#align local_homeomorph.continuous_at_symm LocalHomeomorph.continuousAt_symm +#align local_homeomorph.continuous_at_symm PartialHomeomorph.continuousAt_symm theorem tendsto_symm {x} (hx : x ∈ e.source) : Tendsto e.symm (𝓝 (e x)) (𝓝 x) := by simpa only [ContinuousAt, e.left_inv hx] using e.continuousAt_symm (e.map_source hx) -#align local_homeomorph.tendsto_symm LocalHomeomorph.tendsto_symm +#align local_homeomorph.tendsto_symm PartialHomeomorph.tendsto_symm theorem map_nhds_eq {x} (hx : x ∈ e.source) : map e (𝓝 x) = 𝓝 (e x) := le_antisymm (e.continuousAt hx) <| le_map_of_right_inverse (e.eventually_right_inverse' hx) (e.tendsto_symm hx) -#align local_homeomorph.map_nhds_eq LocalHomeomorph.map_nhds_eq +#align local_homeomorph.map_nhds_eq PartialHomeomorph.map_nhds_eq theorem symm_map_nhds_eq {x} (hx : x ∈ e.source) : map e.symm (𝓝 (e x)) = 𝓝 x := (e.symm.map_nhds_eq <| e.map_source hx).trans <| by rw [e.left_inv hx] -#align local_homeomorph.symm_map_nhds_eq LocalHomeomorph.symm_map_nhds_eq +#align local_homeomorph.symm_map_nhds_eq PartialHomeomorph.symm_map_nhds_eq theorem image_mem_nhds {x} (hx : x ∈ e.source) {s : Set α} (hs : s ∈ 𝓝 x) : e '' s ∈ 𝓝 (e x) := e.map_nhds_eq hx ▸ Filter.image_mem_map hs -#align local_homeomorph.image_mem_nhds LocalHomeomorph.image_mem_nhds +#align local_homeomorph.image_mem_nhds PartialHomeomorph.image_mem_nhds -theorem map_nhdsWithin_eq (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) (s : Set α) : +theorem map_nhdsWithin_eq (e : PartialHomeomorph α β) {x} (hx : x ∈ e.source) (s : Set α) : map e (𝓝[s] x) = 𝓝[e '' (e.source ∩ s)] e x := calc map e (𝓝[s] x) = map e (𝓝[e.source ∩ s] x) := @@ -396,44 +397,44 @@ theorem map_nhdsWithin_eq (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) (e.leftInvOn.mono <| inter_subset_left _ _).map_nhdsWithin_eq (e.left_inv hx) (e.continuousAt_symm (e.map_source hx)).continuousWithinAt (e.continuousAt hx).continuousWithinAt -#align local_homeomorph.map_nhds_within_eq LocalHomeomorph.map_nhdsWithin_eq +#align local_homeomorph.map_nhds_within_eq PartialHomeomorph.map_nhdsWithin_eq -theorem map_nhdsWithin_preimage_eq (e : LocalHomeomorph α β) {x} (hx : x ∈ e.source) (s : Set β) : +theorem map_nhdsWithin_preimage_eq (e : PartialHomeomorph α β) {x} (hx : x ∈ e.source) (s : Set β) : map e (𝓝[e ⁻¹' s] x) = 𝓝[s] e x := by rw [e.map_nhdsWithin_eq hx, e.image_source_inter_eq', e.target_inter_inv_preimage_preimage, e.nhdsWithin_target_inter (e.map_source hx)] -#align local_homeomorph.map_nhds_within_preimage_eq LocalHomeomorph.map_nhdsWithin_preimage_eq +#align local_homeomorph.map_nhds_within_preimage_eq PartialHomeomorph.map_nhdsWithin_preimage_eq -theorem eventually_nhds (e : LocalHomeomorph α β) {x : α} (p : β → Prop) (hx : x ∈ e.source) : +theorem eventually_nhds (e : PartialHomeomorph α β) {x : α} (p : β → Prop) (hx : x ∈ e.source) : (∀ᶠ y in 𝓝 (e x), p y) ↔ ∀ᶠ x in 𝓝 x, p (e x) := Iff.trans (by rw [e.map_nhds_eq hx]) eventually_map -#align local_homeomorph.eventually_nhds LocalHomeomorph.eventually_nhds +#align local_homeomorph.eventually_nhds PartialHomeomorph.eventually_nhds -theorem eventually_nhds' (e : LocalHomeomorph α β) {x : α} (p : α → Prop) (hx : x ∈ e.source) : +theorem eventually_nhds' (e : PartialHomeomorph α β) {x : α} (p : α → Prop) (hx : x ∈ e.source) : (∀ᶠ y in 𝓝 (e x), p (e.symm y)) ↔ ∀ᶠ x in 𝓝 x, p x := by rw [e.eventually_nhds _ hx] refine' eventually_congr ((e.eventually_left_inverse hx).mono fun y hy => _) rw [hy] -#align local_homeomorph.eventually_nhds' LocalHomeomorph.eventually_nhds' +#align local_homeomorph.eventually_nhds' PartialHomeomorph.eventually_nhds' -theorem eventually_nhdsWithin (e : LocalHomeomorph α β) {x : α} (p : β → Prop) {s : Set α} +theorem eventually_nhdsWithin (e : PartialHomeomorph α β) {x : α} (p : β → Prop) {s : Set α} (hx : x ∈ e.source) : (∀ᶠ y in 𝓝[e.symm ⁻¹' s] e x, p y) ↔ ∀ᶠ x in 𝓝[s] x, p (e x) := by refine' Iff.trans _ eventually_map rw [e.map_nhdsWithin_eq hx, e.image_source_inter_eq', e.nhdsWithin_target_inter (e.mapsTo hx)] -#align local_homeomorph.eventually_nhds_within LocalHomeomorph.eventually_nhdsWithin +#align local_homeomorph.eventually_nhds_within PartialHomeomorph.eventually_nhdsWithin -theorem eventually_nhdsWithin' (e : LocalHomeomorph α β) {x : α} (p : α → Prop) {s : Set α} +theorem eventually_nhdsWithin' (e : PartialHomeomorph α β) {x : α} (p : α → Prop) {s : Set α} (hx : x ∈ e.source) : (∀ᶠ y in 𝓝[e.symm ⁻¹' s] e x, p (e.symm y)) ↔ ∀ᶠ x in 𝓝[s] x, p x := by rw [e.eventually_nhdsWithin _ hx] refine eventually_congr <| (eventually_nhdsWithin_of_eventually_nhds <| e.eventually_left_inverse hx).mono fun y hy => ?_ rw [hy] -#align local_homeomorph.eventually_nhds_within' LocalHomeomorph.eventually_nhdsWithin' +#align local_homeomorph.eventually_nhds_within' PartialHomeomorph.eventually_nhdsWithin' /-- This lemma is useful in the manifold library in the case that `e` is a chart. It states that locally around `e x` the set `e.symm ⁻¹' s` is the same as the set intersected with the target of `e` and some other neighborhood of `f x` (which will be the source of a chart on `γ`). -/ -theorem preimage_eventuallyEq_target_inter_preimage_inter {e : LocalHomeomorph α β} {s : Set α} +theorem preimage_eventuallyEq_target_inter_preimage_inter {e : PartialHomeomorph α β} {s : Set α} {t : Set γ} {x : α} {f : α → γ} (hf : ContinuousWithinAt f s x) (hxe : x ∈ e.source) (ht : t ∈ 𝓝 (f x)) : e.symm ⁻¹' s =ᶠ[𝓝 (e x)] (e.target ∩ e.symm ⁻¹' (s ∩ f ⁻¹' t) : Set β) := by @@ -443,40 +444,40 @@ theorem preimage_eventuallyEq_target_inter_preimage_inter {e : LocalHomeomorph intro y hy hyu simp_rw [mem_inter_iff, mem_preimage, mem_inter_iff, e.mapsTo hy, true_and_iff, iff_self_and, e.left_inv hy, iff_true_intro hyu] -#align local_homeomorph.preimage_eventually_eq_target_inter_preimage_inter LocalHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter +#align local_homeomorph.preimage_eventually_eq_target_inter_preimage_inter PartialHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter theorem isOpen_inter_preimage {s : Set β} (hs : IsOpen s) : IsOpen (e.source ∩ e ⁻¹' s) := e.continuousOn.isOpen_inter_preimage e.open_source hs -#align local_homeomorph.preimage_open_of_open LocalHomeomorph.isOpen_inter_preimage +#align local_homeomorph.preimage_open_of_open PartialHomeomorph.isOpen_inter_preimage -/-- A local homeomorphism is an open map on its source. -/ +/-- A partial homeomorphism is an open map on its source. -/ lemma isOpen_image_of_subset_source {s : Set α} (hs : IsOpen s) (hse : s ⊆ e.source) : IsOpen (e '' s) := by rw [(image_eq_target_inter_inv_preimage (e := e) hse)] exact e.continuousOn_invFun.isOpen_inter_preimage e.open_target hs -/-- The inverse of a local homeomorphism `e` is an open map on `e.target`. -/ +/-- The inverse of a partial homeomorphism `e` is an open map on `e.target`. -/ lemma isOpen_image_symm_of_subset_target {t : Set β} (ht : IsOpen t) (hte : t ⊆ e.target) : IsOpen (e.symm '' t) := isOpen_image_of_subset_source e.symm ht (e.symm_source ▸ hte) /-! -### `LocalHomeomorph.IsImage` relation +### `PartialHomeomorph.IsImage` relation -We say that `t : Set β` is an image of `s : Set α` under a local homeomorphism `e` if any of the +We say that `t : Set β` is an image of `s : Set α` under a partial homeomorphism `e` if any of the following equivalent conditions hold: * `e '' (e.source ∩ s) = e.target ∩ t`; * `e.source ∩ e ⁻¹ t = e.source ∩ s`; * `∀ x ∈ e.source, e x ∈ t ↔ x ∈ s` (this one is used in the definition). -This definition is a restatement of `LocalEquiv.IsImage` for local homeomorphisms. In this section -we transfer API about `LocalEquiv.IsImage` to local homeomorphisms and add a few -`LocalHomeomorph`-specific lemmas like `LocalHomeomorph.IsImage.closure`. +This definition is a restatement of `LocalEquiv.IsImage` for partial homeomorphisms. In this section +we transfer API about `LocalEquiv.IsImage` to partial homeomorphisms and add a few +`PartialHomeomorph`-specific lemmas like `PartialHomeomorph.IsImage.closure`. -/ -/-- We say that `t : Set β` is an image of `s : Set α` under a local homeomorphism `e` if any of the -following equivalent conditions hold: +/-- We say that `t : Set β` is an image of `s : Set α` under a partial homeomorphism `e` +if any of the following equivalent conditions hold: * `e '' (e.source ∩ s) = e.target ∩ t`; * `e.source ∩ e ⁻¹ t = e.source ∩ s`; @@ -484,7 +485,7 @@ following equivalent conditions hold: -/ def IsImage (s : Set α) (t : Set β) : Prop := ∀ ⦃x⦄, x ∈ e.source → (e x ∈ t ↔ x ∈ s) -#align local_homeomorph.is_image LocalHomeomorph.IsImage +#align local_homeomorph.is_image PartialHomeomorph.IsImage namespace IsImage @@ -492,178 +493,178 @@ variable {e} {s : Set α} {t : Set β} {x : α} {y : β} theorem toLocalEquiv (h : e.IsImage s t) : e.toLocalEquiv.IsImage s t := h -#align local_homeomorph.is_image.to_local_equiv LocalHomeomorph.IsImage.toLocalEquiv +#align local_homeomorph.is_image.to_local_equiv PartialHomeomorph.IsImage.toLocalEquiv theorem apply_mem_iff (h : e.IsImage s t) (hx : x ∈ e.source) : e x ∈ t ↔ x ∈ s := h hx -#align local_homeomorph.is_image.apply_mem_iff LocalHomeomorph.IsImage.apply_mem_iff +#align local_homeomorph.is_image.apply_mem_iff PartialHomeomorph.IsImage.apply_mem_iff protected theorem symm (h : e.IsImage s t) : e.symm.IsImage t s := h.toLocalEquiv.symm -#align local_homeomorph.is_image.symm LocalHomeomorph.IsImage.symm +#align local_homeomorph.is_image.symm PartialHomeomorph.IsImage.symm theorem symm_apply_mem_iff (h : e.IsImage s t) (hy : y ∈ e.target) : e.symm y ∈ s ↔ y ∈ t := h.symm hy -#align local_homeomorph.is_image.symm_apply_mem_iff LocalHomeomorph.IsImage.symm_apply_mem_iff +#align local_homeomorph.is_image.symm_apply_mem_iff PartialHomeomorph.IsImage.symm_apply_mem_iff @[simp] theorem symm_iff : e.symm.IsImage t s ↔ e.IsImage s t := ⟨fun h => h.symm, fun h => h.symm⟩ -#align local_homeomorph.is_image.symm_iff LocalHomeomorph.IsImage.symm_iff +#align local_homeomorph.is_image.symm_iff PartialHomeomorph.IsImage.symm_iff protected theorem mapsTo (h : e.IsImage s t) : MapsTo e (e.source ∩ s) (e.target ∩ t) := h.toLocalEquiv.mapsTo -#align local_homeomorph.is_image.maps_to LocalHomeomorph.IsImage.mapsTo +#align local_homeomorph.is_image.maps_to PartialHomeomorph.IsImage.mapsTo theorem symm_mapsTo (h : e.IsImage s t) : MapsTo e.symm (e.target ∩ t) (e.source ∩ s) := h.symm.mapsTo -#align local_homeomorph.is_image.symm_maps_to LocalHomeomorph.IsImage.symm_mapsTo +#align local_homeomorph.is_image.symm_maps_to PartialHomeomorph.IsImage.symm_mapsTo theorem image_eq (h : e.IsImage s t) : e '' (e.source ∩ s) = e.target ∩ t := h.toLocalEquiv.image_eq -#align local_homeomorph.is_image.image_eq LocalHomeomorph.IsImage.image_eq +#align local_homeomorph.is_image.image_eq PartialHomeomorph.IsImage.image_eq theorem symm_image_eq (h : e.IsImage s t) : e.symm '' (e.target ∩ t) = e.source ∩ s := h.symm.image_eq -#align local_homeomorph.is_image.symm_image_eq LocalHomeomorph.IsImage.symm_image_eq +#align local_homeomorph.is_image.symm_image_eq PartialHomeomorph.IsImage.symm_image_eq theorem iff_preimage_eq : e.IsImage s t ↔ e.source ∩ e ⁻¹' t = e.source ∩ s := LocalEquiv.IsImage.iff_preimage_eq -#align local_homeomorph.is_image.iff_preimage_eq LocalHomeomorph.IsImage.iff_preimage_eq +#align local_homeomorph.is_image.iff_preimage_eq PartialHomeomorph.IsImage.iff_preimage_eq alias ⟨preimage_eq, of_preimage_eq⟩ := iff_preimage_eq -#align local_homeomorph.is_image.preimage_eq LocalHomeomorph.IsImage.preimage_eq -#align local_homeomorph.is_image.of_preimage_eq LocalHomeomorph.IsImage.of_preimage_eq +#align local_homeomorph.is_image.preimage_eq PartialHomeomorph.IsImage.preimage_eq +#align local_homeomorph.is_image.of_preimage_eq PartialHomeomorph.IsImage.of_preimage_eq theorem iff_symm_preimage_eq : e.IsImage s t ↔ e.target ∩ e.symm ⁻¹' s = e.target ∩ t := symm_iff.symm.trans iff_preimage_eq -#align local_homeomorph.is_image.iff_symm_preimage_eq LocalHomeomorph.IsImage.iff_symm_preimage_eq +#align local_homeomorph.is_image.iff_symm_preimage_eq PartialHomeomorph.IsImage.iff_symm_preimage_eq alias ⟨symm_preimage_eq, of_symm_preimage_eq⟩ := iff_symm_preimage_eq -#align local_homeomorph.is_image.symm_preimage_eq LocalHomeomorph.IsImage.symm_preimage_eq -#align local_homeomorph.is_image.of_symm_preimage_eq LocalHomeomorph.IsImage.of_symm_preimage_eq +#align local_homeomorph.is_image.symm_preimage_eq PartialHomeomorph.IsImage.symm_preimage_eq +#align local_homeomorph.is_image.of_symm_preimage_eq PartialHomeomorph.IsImage.of_symm_preimage_eq theorem iff_symm_preimage_eq' : e.IsImage s t ↔ e.target ∩ e.symm ⁻¹' (e.source ∩ s) = e.target ∩ t := by rw [iff_symm_preimage_eq, ← image_source_inter_eq, ← image_source_inter_eq'] -#align local_homeomorph.is_image.iff_symm_preimage_eq' LocalHomeomorph.IsImage.iff_symm_preimage_eq' +#align local_homeomorph.is_image.iff_symm_preimage_eq' PartialHomeomorph.IsImage.iff_symm_preimage_eq' alias ⟨symm_preimage_eq', of_symm_preimage_eq'⟩ := iff_symm_preimage_eq' -#align local_homeomorph.is_image.symm_preimage_eq' LocalHomeomorph.IsImage.symm_preimage_eq' -#align local_homeomorph.is_image.of_symm_preimage_eq' LocalHomeomorph.IsImage.of_symm_preimage_eq' +#align local_homeomorph.is_image.symm_preimage_eq' PartialHomeomorph.IsImage.symm_preimage_eq' +#align local_homeomorph.is_image.of_symm_preimage_eq' PartialHomeomorph.IsImage.of_symm_preimage_eq' theorem iff_preimage_eq' : e.IsImage s t ↔ e.source ∩ e ⁻¹' (e.target ∩ t) = e.source ∩ s := symm_iff.symm.trans iff_symm_preimage_eq' -#align local_homeomorph.is_image.iff_preimage_eq' LocalHomeomorph.IsImage.iff_preimage_eq' +#align local_homeomorph.is_image.iff_preimage_eq' PartialHomeomorph.IsImage.iff_preimage_eq' alias ⟨preimage_eq', of_preimage_eq'⟩ := iff_preimage_eq' -#align local_homeomorph.is_image.preimage_eq' LocalHomeomorph.IsImage.preimage_eq' -#align local_homeomorph.is_image.of_preimage_eq' LocalHomeomorph.IsImage.of_preimage_eq' +#align local_homeomorph.is_image.preimage_eq' PartialHomeomorph.IsImage.preimage_eq' +#align local_homeomorph.is_image.of_preimage_eq' PartialHomeomorph.IsImage.of_preimage_eq' theorem of_image_eq (h : e '' (e.source ∩ s) = e.target ∩ t) : e.IsImage s t := LocalEquiv.IsImage.of_image_eq h -#align local_homeomorph.is_image.of_image_eq LocalHomeomorph.IsImage.of_image_eq +#align local_homeomorph.is_image.of_image_eq PartialHomeomorph.IsImage.of_image_eq theorem of_symm_image_eq (h : e.symm '' (e.target ∩ t) = e.source ∩ s) : e.IsImage s t := LocalEquiv.IsImage.of_symm_image_eq h -#align local_homeomorph.is_image.of_symm_image_eq LocalHomeomorph.IsImage.of_symm_image_eq +#align local_homeomorph.is_image.of_symm_image_eq PartialHomeomorph.IsImage.of_symm_image_eq protected theorem compl (h : e.IsImage s t) : e.IsImage sᶜ tᶜ := fun _ hx => (h hx).not -#align local_homeomorph.is_image.compl LocalHomeomorph.IsImage.compl +#align local_homeomorph.is_image.compl PartialHomeomorph.IsImage.compl protected theorem inter {s' t'} (h : e.IsImage s t) (h' : e.IsImage s' t') : e.IsImage (s ∩ s') (t ∩ t') := fun _ hx => (h hx).and (h' hx) -#align local_homeomorph.is_image.inter LocalHomeomorph.IsImage.inter +#align local_homeomorph.is_image.inter PartialHomeomorph.IsImage.inter protected theorem union {s' t'} (h : e.IsImage s t) (h' : e.IsImage s' t') : e.IsImage (s ∪ s') (t ∪ t') := fun _ hx => (h hx).or (h' hx) -#align local_homeomorph.is_image.union LocalHomeomorph.IsImage.union +#align local_homeomorph.is_image.union PartialHomeomorph.IsImage.union protected theorem diff {s' t'} (h : e.IsImage s t) (h' : e.IsImage s' t') : e.IsImage (s \ s') (t \ t') := h.inter h'.compl -#align local_homeomorph.is_image.diff LocalHomeomorph.IsImage.diff +#align local_homeomorph.is_image.diff PartialHomeomorph.IsImage.diff -theorem leftInvOn_piecewise {e' : LocalHomeomorph α β} [∀ i, Decidable (i ∈ s)] +theorem leftInvOn_piecewise {e' : PartialHomeomorph α β} [∀ i, Decidable (i ∈ s)] [∀ i, Decidable (i ∈ t)] (h : e.IsImage s t) (h' : e'.IsImage s t) : LeftInvOn (t.piecewise e.symm e'.symm) (s.piecewise e e') (s.ite e.source e'.source) := h.toLocalEquiv.leftInvOn_piecewise h' -#align local_homeomorph.is_image.left_inv_on_piecewise LocalHomeomorph.IsImage.leftInvOn_piecewise +#align local_homeomorph.is_image.left_inv_on_piecewise PartialHomeomorph.IsImage.leftInvOn_piecewise -theorem inter_eq_of_inter_eq_of_eqOn {e' : LocalHomeomorph α β} (h : e.IsImage s t) +theorem inter_eq_of_inter_eq_of_eqOn {e' : PartialHomeomorph α β} (h : e.IsImage s t) (h' : e'.IsImage s t) (hs : e.source ∩ s = e'.source ∩ s) (Heq : EqOn e e' (e.source ∩ s)) : e.target ∩ t = e'.target ∩ t := h.toLocalEquiv.inter_eq_of_inter_eq_of_eqOn h' hs Heq -#align local_homeomorph.is_image.inter_eq_of_inter_eq_of_eq_on LocalHomeomorph.IsImage.inter_eq_of_inter_eq_of_eqOn +#align local_homeomorph.is_image.inter_eq_of_inter_eq_of_eq_on PartialHomeomorph.IsImage.inter_eq_of_inter_eq_of_eqOn -theorem symm_eqOn_of_inter_eq_of_eqOn {e' : LocalHomeomorph α β} (h : e.IsImage s t) +theorem symm_eqOn_of_inter_eq_of_eqOn {e' : PartialHomeomorph α β} (h : e.IsImage s t) (hs : e.source ∩ s = e'.source ∩ s) (Heq : EqOn e e' (e.source ∩ s)) : EqOn e.symm e'.symm (e.target ∩ t) := h.toLocalEquiv.symm_eq_on_of_inter_eq_of_eqOn hs Heq -#align local_homeomorph.is_image.symm_eq_on_of_inter_eq_of_eq_on LocalHomeomorph.IsImage.symm_eqOn_of_inter_eq_of_eqOn +#align local_homeomorph.is_image.symm_eq_on_of_inter_eq_of_eq_on PartialHomeomorph.IsImage.symm_eqOn_of_inter_eq_of_eqOn theorem map_nhdsWithin_eq (h : e.IsImage s t) (hx : x ∈ e.source) : map e (𝓝[s] x) = 𝓝[t] e x := by rw [e.map_nhdsWithin_eq hx, h.image_eq, e.nhdsWithin_target_inter (e.map_source hx)] -#align local_homeomorph.is_image.map_nhds_within_eq LocalHomeomorph.IsImage.map_nhdsWithin_eq +#align local_homeomorph.is_image.map_nhds_within_eq PartialHomeomorph.IsImage.map_nhdsWithin_eq protected theorem closure (h : e.IsImage s t) : e.IsImage (closure s) (closure t) := fun x hx => by simp only [mem_closure_iff_nhdsWithin_neBot, ← h.map_nhdsWithin_eq hx, map_neBot_iff] -#align local_homeomorph.is_image.closure LocalHomeomorph.IsImage.closure +#align local_homeomorph.is_image.closure PartialHomeomorph.IsImage.closure protected theorem interior (h : e.IsImage s t) : e.IsImage (interior s) (interior t) := by simpa only [closure_compl, compl_compl] using h.compl.closure.compl -#align local_homeomorph.is_image.interior LocalHomeomorph.IsImage.interior +#align local_homeomorph.is_image.interior PartialHomeomorph.IsImage.interior protected theorem frontier (h : e.IsImage s t) : e.IsImage (frontier s) (frontier t) := h.closure.diff h.interior -#align local_homeomorph.is_image.frontier LocalHomeomorph.IsImage.frontier +#align local_homeomorph.is_image.frontier PartialHomeomorph.IsImage.frontier theorem isOpen_iff (h : e.IsImage s t) : IsOpen (e.source ∩ s) ↔ IsOpen (e.target ∩ t) := ⟨fun hs => h.symm_preimage_eq' ▸ e.symm.isOpen_inter_preimage hs, fun hs => h.preimage_eq' ▸ e.isOpen_inter_preimage hs⟩ -#align local_homeomorph.is_image.is_open_iff LocalHomeomorph.IsImage.isOpen_iff +#align local_homeomorph.is_image.is_open_iff PartialHomeomorph.IsImage.isOpen_iff -/-- Restrict a `LocalHomeomorph` to a pair of corresponding open sets. -/ +/-- Restrict a `PartialHomeomorph` to a pair of corresponding open sets. -/ @[simps toLocalEquiv] -def restr (h : e.IsImage s t) (hs : IsOpen (e.source ∩ s)) : LocalHomeomorph α β where +def restr (h : e.IsImage s t) (hs : IsOpen (e.source ∩ s)) : PartialHomeomorph α β where toLocalEquiv := h.toLocalEquiv.restr open_source := hs open_target := h.isOpen_iff.1 hs continuousOn_toFun := e.continuousOn.mono (inter_subset_left _ _) continuousOn_invFun := e.symm.continuousOn.mono (inter_subset_left _ _) -#align local_homeomorph.is_image.restr LocalHomeomorph.IsImage.restr +#align local_homeomorph.is_image.restr PartialHomeomorph.IsImage.restr end IsImage theorem isImage_source_target : e.IsImage e.source e.target := e.toLocalEquiv.isImage_source_target -#align local_homeomorph.is_image_source_target LocalHomeomorph.isImage_source_target +#align local_homeomorph.is_image_source_target PartialHomeomorph.isImage_source_target -theorem isImage_source_target_of_disjoint (e' : LocalHomeomorph α β) +theorem isImage_source_target_of_disjoint (e' : PartialHomeomorph α β) (hs : Disjoint e.source e'.source) (ht : Disjoint e.target e'.target) : e.IsImage e'.source e'.target := e.toLocalEquiv.isImage_source_target_of_disjoint e'.toLocalEquiv hs ht -#align local_homeomorph.is_image_source_target_of_disjoint LocalHomeomorph.isImage_source_target_of_disjoint +#align local_homeomorph.is_image_source_target_of_disjoint PartialHomeomorph.isImage_source_target_of_disjoint -/-- Preimage of interior or interior of preimage coincide for local homeomorphisms, when restricted -to the source. -/ +/-- Preimage of interior or interior of preimage coincide for partial homeomorphisms, +when restricted to the source. -/ theorem preimage_interior (s : Set β) : e.source ∩ e ⁻¹' interior s = e.source ∩ interior (e ⁻¹' s) := (IsImage.of_preimage_eq rfl).interior.preimage_eq -#align local_homeomorph.preimage_interior LocalHomeomorph.preimage_interior +#align local_homeomorph.preimage_interior PartialHomeomorph.preimage_interior theorem preimage_closure (s : Set β) : e.source ∩ e ⁻¹' closure s = e.source ∩ closure (e ⁻¹' s) := (IsImage.of_preimage_eq rfl).closure.preimage_eq -#align local_homeomorph.preimage_closure LocalHomeomorph.preimage_closure +#align local_homeomorph.preimage_closure PartialHomeomorph.preimage_closure theorem preimage_frontier (s : Set β) : e.source ∩ e ⁻¹' frontier s = e.source ∩ frontier (e ⁻¹' s) := (IsImage.of_preimage_eq rfl).frontier.preimage_eq -#align local_homeomorph.preimage_frontier LocalHomeomorph.preimage_frontier +#align local_homeomorph.preimage_frontier PartialHomeomorph.preimage_frontier theorem isOpen_inter_preimage_symm {s : Set α} (hs : IsOpen s) : IsOpen (e.target ∩ e.symm ⁻¹' s) := e.symm.continuousOn.isOpen_inter_preimage e.open_target hs -#align local_homeomorph.preimage_open_of_open_symm LocalHomeomorph.isOpen_inter_preimage_symm +#align local_homeomorph.preimage_open_of_open_symm PartialHomeomorph.isOpen_inter_preimage_symm /-- The image of an open set in the source is open. -/ theorem image_isOpen_of_isOpen {s : Set α} (hs : IsOpen s) (h : s ⊆ e.source) : @@ -671,102 +672,102 @@ theorem image_isOpen_of_isOpen {s : Set α} (hs : IsOpen s) (h : s ⊆ e.source) have : e '' s = e.target ∩ e.symm ⁻¹' s := e.toLocalEquiv.image_eq_target_inter_inv_preimage h rw [this] exact e.continuousOn_symm.isOpen_inter_preimage e.open_target hs -#align local_homeomorph.image_open_of_open LocalHomeomorph.image_isOpen_of_isOpen +#align local_homeomorph.image_open_of_open PartialHomeomorph.image_isOpen_of_isOpen /-- The image of the restriction of an open set to the source is open. -/ theorem image_isOpen_of_isOpen' {s : Set α} (hs : IsOpen s) : IsOpen (e '' (e.source ∩ s)) := image_isOpen_of_isOpen _ (IsOpen.inter e.open_source hs) (inter_subset_left _ _) -#align local_homeomorph.image_open_of_open' LocalHomeomorph.image_isOpen_of_isOpen' +#align local_homeomorph.image_open_of_open' PartialHomeomorph.image_isOpen_of_isOpen' -/-- A `LocalEquiv` with continuous open forward map and an open source is a `LocalHomeomorph`. -/ +/-- A `LocalEquiv` with continuous open forward map and an open source is a `PartialHomeomorph`. -/ def ofContinuousOpenRestrict (e : LocalEquiv α β) (hc : ContinuousOn e e.source) - (ho : IsOpenMap (e.source.restrict e)) (hs : IsOpen e.source) : LocalHomeomorph α β where + (ho : IsOpenMap (e.source.restrict e)) (hs : IsOpen e.source) : PartialHomeomorph α β where toLocalEquiv := e open_source := hs open_target := by simpa only [range_restrict, e.image_source_eq_target] using ho.isOpen_range continuousOn_toFun := hc continuousOn_invFun := e.image_source_eq_target ▸ ho.continuousOn_image_of_leftInvOn e.leftInvOn -#align local_homeomorph.of_continuous_open_restrict LocalHomeomorph.ofContinuousOpenRestrict +#align local_homeomorph.of_continuous_open_restrict PartialHomeomorph.ofContinuousOpenRestrict -/-- A `LocalEquiv` with continuous open forward map and an open source is a `LocalHomeomorph`. -/ +/-- A `LocalEquiv` with continuous open forward map and an open source is a `PartialHomeomorph`. -/ def ofContinuousOpen (e : LocalEquiv α β) (hc : ContinuousOn e e.source) (ho : IsOpenMap e) - (hs : IsOpen e.source) : LocalHomeomorph α β := + (hs : IsOpen e.source) : PartialHomeomorph α β := ofContinuousOpenRestrict e hc (ho.restrict hs) hs -#align local_homeomorph.of_continuous_open LocalHomeomorph.ofContinuousOpen +#align local_homeomorph.of_continuous_open PartialHomeomorph.ofContinuousOpen -/-- Restricting a local homeomorphism `e` to `e.source ∩ s` when `s` is open. This is sometimes hard -to use because of the openness assumption, but it has the advantage that when it can -be used then its local_equiv is defeq to local_equiv.restr -/ -protected def restrOpen (s : Set α) (hs : IsOpen s) : LocalHomeomorph α β := +/-- Restricting a partial homeomorphism `e` to `e.source ∩ s` when `s` is open. +This is sometimes hard to use because of the openness assumption, but it has the advantage that +when it can be used then its `LocalEquiv` is defeq to `LocalEquiv.restr`. -/ +protected def restrOpen (s : Set α) (hs : IsOpen s) : PartialHomeomorph α β := (@IsImage.of_symm_preimage_eq α β _ _ e s (e.symm ⁻¹' s) rfl).restr (IsOpen.inter e.open_source hs) -#align local_homeomorph.restr_open LocalHomeomorph.restrOpen +#align local_homeomorph.restr_open PartialHomeomorph.restrOpen @[simp, mfld_simps] theorem restrOpen_toLocalEquiv (s : Set α) (hs : IsOpen s) : (e.restrOpen s hs).toLocalEquiv = e.toLocalEquiv.restr s := rfl -#align local_homeomorph.restr_open_to_local_equiv LocalHomeomorph.restrOpen_toLocalEquiv +#align local_homeomorph.restr_open_to_local_equiv PartialHomeomorph.restrOpen_toLocalEquiv -- Already simp via `LocalEquiv` theorem restrOpen_source (s : Set α) (hs : IsOpen s) : (e.restrOpen s hs).source = e.source ∩ s := rfl -#align local_homeomorph.restr_open_source LocalHomeomorph.restrOpen_source +#align local_homeomorph.restr_open_source PartialHomeomorph.restrOpen_source -/-- Restricting a local homeomorphism `e` to `e.source ∩ interior s`. We use the interior to make -sure that the restriction is well defined whatever the set s, since local homeomorphisms are by +/-- Restricting a partial homeomorphism `e` to `e.source ∩ interior s`. We use the interior to make +sure that the restriction is well defined whatever the set s, since partial homeomorphisms are by definition defined on open sets. In applications where `s` is open, this coincides with the restriction of local equivalences -/ @[simps! (config := mfld_cfg) apply symm_apply, simps! (config := .lemmasOnly) source target] -protected def restr (s : Set α) : LocalHomeomorph α β := +protected def restr (s : Set α) : PartialHomeomorph α β := e.restrOpen (interior s) isOpen_interior -#align local_homeomorph.restr LocalHomeomorph.restr +#align local_homeomorph.restr PartialHomeomorph.restr @[simp, mfld_simps] theorem restr_toLocalEquiv (s : Set α) : (e.restr s).toLocalEquiv = e.toLocalEquiv.restr (interior s) := rfl -#align local_homeomorph.restr_to_local_equiv LocalHomeomorph.restr_toLocalEquiv +#align local_homeomorph.restr_to_local_equiv PartialHomeomorph.restr_toLocalEquiv theorem restr_source' (s : Set α) (hs : IsOpen s) : (e.restr s).source = e.source ∩ s := by rw [e.restr_source, hs.interior_eq] -#align local_homeomorph.restr_source' LocalHomeomorph.restr_source' +#align local_homeomorph.restr_source' PartialHomeomorph.restr_source' theorem restr_toLocalEquiv' (s : Set α) (hs : IsOpen s) : (e.restr s).toLocalEquiv = e.toLocalEquiv.restr s := by rw [e.restr_toLocalEquiv, hs.interior_eq] -#align local_homeomorph.restr_to_local_equiv' LocalHomeomorph.restr_toLocalEquiv' +#align local_homeomorph.restr_to_local_equiv' PartialHomeomorph.restr_toLocalEquiv' -theorem restr_eq_of_source_subset {e : LocalHomeomorph α β} {s : Set α} (h : e.source ⊆ s) : +theorem restr_eq_of_source_subset {e : PartialHomeomorph α β} {s : Set α} (h : e.source ⊆ s) : e.restr s = e := toLocalEquiv_injective <| LocalEquiv.restr_eq_of_source_subset <| interior_maximal h e.open_source -#align local_homeomorph.restr_eq_of_source_subset LocalHomeomorph.restr_eq_of_source_subset +#align local_homeomorph.restr_eq_of_source_subset PartialHomeomorph.restr_eq_of_source_subset @[simp, mfld_simps] -theorem restr_univ {e : LocalHomeomorph α β} : e.restr univ = e := +theorem restr_univ {e : PartialHomeomorph α β} : e.restr univ = e := restr_eq_of_source_subset (subset_univ _) -#align local_homeomorph.restr_univ LocalHomeomorph.restr_univ +#align local_homeomorph.restr_univ PartialHomeomorph.restr_univ theorem restr_source_inter (s : Set α) : e.restr (e.source ∩ s) = e.restr s := by - refine' LocalHomeomorph.ext _ _ (fun x => rfl) (fun x => rfl) _ + refine' PartialHomeomorph.ext _ _ (fun x => rfl) (fun x => rfl) _ simp [e.open_source.interior_eq, ← inter_assoc] -#align local_homeomorph.restr_source_inter LocalHomeomorph.restr_source_inter +#align local_homeomorph.restr_source_inter PartialHomeomorph.restr_source_inter -/-- The identity on the whole space as a local homeomorphism. -/ +/-- The identity on the whole space as a partial homeomorphism. -/ @[simps! (config := mfld_cfg) apply, simps! (config := .lemmasOnly) source target] -protected def refl (α : Type*) [TopologicalSpace α] : LocalHomeomorph α α := - (Homeomorph.refl α).toLocalHomeomorph -#align local_homeomorph.refl LocalHomeomorph.refl +protected def refl (α : Type*) [TopologicalSpace α] : PartialHomeomorph α α := + (Homeomorph.refl α).toPartialHomeomorph +#align local_homeomorph.refl PartialHomeomorph.refl @[simp, mfld_simps] -theorem refl_localEquiv : (LocalHomeomorph.refl α).toLocalEquiv = LocalEquiv.refl α := +theorem refl_localEquiv : (PartialHomeomorph.refl α).toLocalEquiv = LocalEquiv.refl α := rfl -#align local_homeomorph.refl_local_equiv LocalHomeomorph.refl_localEquiv +#align local_homeomorph.refl_local_equiv PartialHomeomorph.refl_localEquiv @[simp, mfld_simps] -theorem refl_symm : (LocalHomeomorph.refl α).symm = LocalHomeomorph.refl α := +theorem refl_symm : (PartialHomeomorph.refl α).symm = PartialHomeomorph.refl α := rfl -#align local_homeomorph.refl_symm LocalHomeomorph.refl_symm +#align local_homeomorph.refl_symm PartialHomeomorph.refl_symm section @@ -774,238 +775,238 @@ variable {s : Set α} (hs : IsOpen s) /-- The identity local equiv on a set `s` -/ @[simps! (config := mfld_cfg) apply, simps! (config := .lemmasOnly) source target] -def ofSet (s : Set α) (hs : IsOpen s) : LocalHomeomorph α α where +def ofSet (s : Set α) (hs : IsOpen s) : PartialHomeomorph α α where toLocalEquiv := LocalEquiv.ofSet s open_source := hs open_target := hs continuousOn_toFun := continuous_id.continuousOn continuousOn_invFun := continuous_id.continuousOn -#align local_homeomorph.of_set LocalHomeomorph.ofSet +#align local_homeomorph.of_set PartialHomeomorph.ofSet @[simp, mfld_simps] theorem ofSet_toLocalEquiv : (ofSet s hs).toLocalEquiv = LocalEquiv.ofSet s := rfl -#align local_homeomorph.of_set_to_local_equiv LocalHomeomorph.ofSet_toLocalEquiv +#align local_homeomorph.of_set_to_local_equiv PartialHomeomorph.ofSet_toLocalEquiv @[simp, mfld_simps] theorem ofSet_symm : (ofSet s hs).symm = ofSet s hs := rfl -#align local_homeomorph.of_set_symm LocalHomeomorph.ofSet_symm +#align local_homeomorph.of_set_symm PartialHomeomorph.ofSet_symm @[simp, mfld_simps] -theorem ofSet_univ_eq_refl : ofSet univ isOpen_univ = LocalHomeomorph.refl α := by ext <;> simp -#align local_homeomorph.of_set_univ_eq_refl LocalHomeomorph.ofSet_univ_eq_refl +theorem ofSet_univ_eq_refl : ofSet univ isOpen_univ = PartialHomeomorph.refl α := by ext <;> simp +#align local_homeomorph.of_set_univ_eq_refl PartialHomeomorph.ofSet_univ_eq_refl end -/-- Composition of two local homeomorphisms when the target of the first and the source of +/-- Composition of two partial homeomorphisms when the target of the first and the source of the second coincide. -/ @[simps! apply symm_apply toLocalEquiv, simps! (config := .lemmasOnly) source target] -protected def trans' (h : e.target = e'.source) : LocalHomeomorph α γ where +protected def trans' (h : e.target = e'.source) : PartialHomeomorph α γ where toLocalEquiv := LocalEquiv.trans' e.toLocalEquiv e'.toLocalEquiv h open_source := e.open_source open_target := e'.open_target continuousOn_toFun := e'.continuousOn.comp e.continuousOn <| h ▸ e.mapsTo continuousOn_invFun := e.continuousOn_symm.comp e'.continuousOn_symm <| h.symm ▸ e'.symm_mapsTo -#align local_homeomorph.trans' LocalHomeomorph.trans' +#align local_homeomorph.trans' PartialHomeomorph.trans' -/-- Composing two local homeomorphisms, by restricting to the maximal domain where their +/-- Composing two partial homeomorphisms, by restricting to the maximal domain where their composition is well defined. -/ -protected def trans : LocalHomeomorph α γ := - LocalHomeomorph.trans' (e.symm.restrOpen e'.source e'.open_source).symm +protected def trans : PartialHomeomorph α γ := + PartialHomeomorph.trans' (e.symm.restrOpen e'.source e'.open_source).symm (e'.restrOpen e.target e.open_target) (by simp [inter_comm]) -#align local_homeomorph.trans LocalHomeomorph.trans +#align local_homeomorph.trans PartialHomeomorph.trans @[simp, mfld_simps] theorem trans_toLocalEquiv : (e.trans e').toLocalEquiv = e.toLocalEquiv.trans e'.toLocalEquiv := rfl -#align local_homeomorph.trans_to_local_equiv LocalHomeomorph.trans_toLocalEquiv +#align local_homeomorph.trans_to_local_equiv PartialHomeomorph.trans_toLocalEquiv @[simp, mfld_simps] theorem coe_trans : (e.trans e' : α → γ) = e' ∘ e := rfl -#align local_homeomorph.coe_trans LocalHomeomorph.coe_trans +#align local_homeomorph.coe_trans PartialHomeomorph.coe_trans @[simp, mfld_simps] theorem coe_trans_symm : ((e.trans e').symm : γ → α) = e.symm ∘ e'.symm := rfl -#align local_homeomorph.coe_trans_symm LocalHomeomorph.coe_trans_symm +#align local_homeomorph.coe_trans_symm PartialHomeomorph.coe_trans_symm theorem trans_apply {x : α} : (e.trans e') x = e' (e x) := rfl -#align local_homeomorph.trans_apply LocalHomeomorph.trans_apply +#align local_homeomorph.trans_apply PartialHomeomorph.trans_apply theorem trans_symm_eq_symm_trans_symm : (e.trans e').symm = e'.symm.trans e.symm := rfl -#align local_homeomorph.trans_symm_eq_symm_trans_symm LocalHomeomorph.trans_symm_eq_symm_trans_symm +#align local_homeomorph.trans_symm_eq_symm_trans_symm PartialHomeomorph.trans_symm_eq_symm_trans_symm /- This could be considered as a simp lemma, but there are many situations where it makes something simple into something more complicated. -/ theorem trans_source : (e.trans e').source = e.source ∩ e ⁻¹' e'.source := LocalEquiv.trans_source e.toLocalEquiv e'.toLocalEquiv -#align local_homeomorph.trans_source LocalHomeomorph.trans_source +#align local_homeomorph.trans_source PartialHomeomorph.trans_source theorem trans_source' : (e.trans e').source = e.source ∩ e ⁻¹' (e.target ∩ e'.source) := LocalEquiv.trans_source' e.toLocalEquiv e'.toLocalEquiv -#align local_homeomorph.trans_source' LocalHomeomorph.trans_source' +#align local_homeomorph.trans_source' PartialHomeomorph.trans_source' theorem trans_source'' : (e.trans e').source = e.symm '' (e.target ∩ e'.source) := LocalEquiv.trans_source'' e.toLocalEquiv e'.toLocalEquiv -#align local_homeomorph.trans_source'' LocalHomeomorph.trans_source'' +#align local_homeomorph.trans_source'' PartialHomeomorph.trans_source'' theorem image_trans_source : e '' (e.trans e').source = e.target ∩ e'.source := LocalEquiv.image_trans_source e.toLocalEquiv e'.toLocalEquiv -#align local_homeomorph.image_trans_source LocalHomeomorph.image_trans_source +#align local_homeomorph.image_trans_source PartialHomeomorph.image_trans_source theorem trans_target : (e.trans e').target = e'.target ∩ e'.symm ⁻¹' e.target := rfl -#align local_homeomorph.trans_target LocalHomeomorph.trans_target +#align local_homeomorph.trans_target PartialHomeomorph.trans_target theorem trans_target' : (e.trans e').target = e'.target ∩ e'.symm ⁻¹' (e'.source ∩ e.target) := trans_source' e'.symm e.symm -#align local_homeomorph.trans_target' LocalHomeomorph.trans_target' +#align local_homeomorph.trans_target' PartialHomeomorph.trans_target' theorem trans_target'' : (e.trans e').target = e' '' (e'.source ∩ e.target) := trans_source'' e'.symm e.symm -#align local_homeomorph.trans_target'' LocalHomeomorph.trans_target'' +#align local_homeomorph.trans_target'' PartialHomeomorph.trans_target'' theorem inv_image_trans_target : e'.symm '' (e.trans e').target = e'.source ∩ e.target := image_trans_source e'.symm e.symm -#align local_homeomorph.inv_image_trans_target LocalHomeomorph.inv_image_trans_target +#align local_homeomorph.inv_image_trans_target PartialHomeomorph.inv_image_trans_target -theorem trans_assoc (e'' : LocalHomeomorph γ δ) : +theorem trans_assoc (e'' : PartialHomeomorph γ δ) : (e.trans e').trans e'' = e.trans (e'.trans e'') := toLocalEquiv_injective <| e.1.trans_assoc _ _ -#align local_homeomorph.trans_assoc LocalHomeomorph.trans_assoc +#align local_homeomorph.trans_assoc PartialHomeomorph.trans_assoc @[simp, mfld_simps] -theorem trans_refl : e.trans (LocalHomeomorph.refl β) = e := +theorem trans_refl : e.trans (PartialHomeomorph.refl β) = e := toLocalEquiv_injective e.1.trans_refl -#align local_homeomorph.trans_refl LocalHomeomorph.trans_refl +#align local_homeomorph.trans_refl PartialHomeomorph.trans_refl @[simp, mfld_simps] -theorem refl_trans : (LocalHomeomorph.refl α).trans e = e := +theorem refl_trans : (PartialHomeomorph.refl α).trans e = e := toLocalEquiv_injective e.1.refl_trans -#align local_homeomorph.refl_trans LocalHomeomorph.refl_trans +#align local_homeomorph.refl_trans PartialHomeomorph.refl_trans theorem trans_ofSet {s : Set β} (hs : IsOpen s) : e.trans (ofSet s hs) = e.restr (e ⁻¹' s) := - LocalHomeomorph.ext _ _ (fun _ => rfl) (fun _ => rfl) <| by + PartialHomeomorph.ext _ _ (fun _ => rfl) (fun _ => rfl) <| by rw [trans_source, restr_source, ofSet_source, ← preimage_interior, hs.interior_eq] -#align local_homeomorph.trans_of_set LocalHomeomorph.trans_ofSet +#align local_homeomorph.trans_of_set PartialHomeomorph.trans_ofSet theorem trans_of_set' {s : Set β} (hs : IsOpen s) : e.trans (ofSet s hs) = e.restr (e.source ∩ e ⁻¹' s) := by rw [trans_ofSet, restr_source_inter] -#align local_homeomorph.trans_of_set' LocalHomeomorph.trans_of_set' +#align local_homeomorph.trans_of_set' PartialHomeomorph.trans_of_set' theorem ofSet_trans {s : Set α} (hs : IsOpen s) : (ofSet s hs).trans e = e.restr s := - LocalHomeomorph.ext _ _ (fun x => rfl) (fun x => rfl) <| by simp [hs.interior_eq, inter_comm] -#align local_homeomorph.of_set_trans LocalHomeomorph.ofSet_trans + PartialHomeomorph.ext _ _ (fun x => rfl) (fun x => rfl) <| by simp [hs.interior_eq, inter_comm] +#align local_homeomorph.of_set_trans PartialHomeomorph.ofSet_trans theorem ofSet_trans' {s : Set α} (hs : IsOpen s) : (ofSet s hs).trans e = e.restr (e.source ∩ s) := by rw [ofSet_trans, restr_source_inter] -#align local_homeomorph.of_set_trans' LocalHomeomorph.ofSet_trans' +#align local_homeomorph.of_set_trans' PartialHomeomorph.ofSet_trans' @[simp, mfld_simps] theorem ofSet_trans_ofSet {s : Set α} (hs : IsOpen s) {s' : Set α} (hs' : IsOpen s') : (ofSet s hs).trans (ofSet s' hs') = ofSet (s ∩ s') (IsOpen.inter hs hs') := by rw [(ofSet s hs).trans_ofSet hs'] ext <;> simp [hs'.interior_eq] -#align local_homeomorph.of_set_trans_of_set LocalHomeomorph.ofSet_trans_ofSet +#align local_homeomorph.of_set_trans_of_set PartialHomeomorph.ofSet_trans_ofSet theorem restr_trans (s : Set α) : (e.restr s).trans e' = (e.trans e').restr s := toLocalEquiv_injective <| LocalEquiv.restr_trans e.toLocalEquiv e'.toLocalEquiv (interior s) -#align local_homeomorph.restr_trans LocalHomeomorph.restr_trans +#align local_homeomorph.restr_trans PartialHomeomorph.restr_trans -/-- Postcompose a local homeomorphism with a homeomorphism. +/-- Postcompose a partial homeomorphism with a homeomorphism. We modify the source and target to have better definitional behavior. -/ @[simps! (config := .asFn)] -def transHomeomorph (e' : β ≃ₜ γ) : LocalHomeomorph α γ where +def transHomeomorph (e' : β ≃ₜ γ) : PartialHomeomorph α γ where toLocalEquiv := e.toLocalEquiv.transEquiv e'.toEquiv open_source := e.open_source open_target := e.open_target.preimage e'.symm.continuous continuousOn_toFun := e'.continuous.comp_continuousOn e.continuousOn continuousOn_invFun := e.symm.continuousOn.comp e'.symm.continuous.continuousOn fun _ => id -#align local_homeomorph.trans_homeomorph LocalHomeomorph.transHomeomorph +#align local_homeomorph.trans_homeomorph PartialHomeomorph.transHomeomorph theorem transHomeomorph_eq_trans (e' : β ≃ₜ γ) : - e.transHomeomorph e' = e.trans e'.toLocalHomeomorph := + e.transHomeomorph e' = e.trans e'.toPartialHomeomorph := toLocalEquiv_injective <| LocalEquiv.transEquiv_eq_trans _ _ -#align local_homeomorph.trans_equiv_eq_trans LocalHomeomorph.transHomeomorph_eq_trans +#align local_homeomorph.trans_equiv_eq_trans PartialHomeomorph.transHomeomorph_eq_trans -/-- Precompose a local homeomorphism with a homeomorphism. +/-- Precompose a partial homeomorphism with a homeomorphism. We modify the source and target to have better definitional behavior. -/ @[simps! (config := .asFn)] -def _root_.Homeomorph.transLocalHomeomorph (e : α ≃ₜ β) : LocalHomeomorph α γ where +def _root_.Homeomorph.transPartialHomeomorph (e : α ≃ₜ β) : PartialHomeomorph α γ where toLocalEquiv := e.toEquiv.transLocalEquiv e'.toLocalEquiv open_source := e'.open_source.preimage e.continuous open_target := e'.open_target continuousOn_toFun := e'.continuousOn.comp e.continuous.continuousOn fun _ => id continuousOn_invFun := e.symm.continuous.comp_continuousOn e'.symm.continuousOn -#align homeomorph.trans_local_homeomorph Homeomorph.transLocalHomeomorph +#align homeomorph.trans_local_homeomorph Homeomorph.transPartialHomeomorph -theorem _root_.Homeomorph.transLocalHomeomorph_eq_trans (e : α ≃ₜ β) : - e.transLocalHomeomorph e' = e.toLocalHomeomorph.trans e' := +theorem _root_.Homeomorph.transPartialHomeomorph_eq_trans (e : α ≃ₜ β) : + e.transPartialHomeomorph e' = e.toPartialHomeomorph.trans e' := toLocalEquiv_injective <| Equiv.transLocalEquiv_eq_trans _ _ -#align homeomorph.trans_local_homeomorph_eq_trans Homeomorph.transLocalHomeomorph_eq_trans +#align homeomorph.trans_local_homeomorph_eq_trans Homeomorph.transPartialHomeomorph_eq_trans /-- `EqOnSource e e'` means that `e` and `e'` have the same source, and coincide there. They -should really be considered the same local equiv. -/ -def EqOnSource (e e' : LocalHomeomorph α β) : Prop := +should really be considered the same local equivalence. -/ +def EqOnSource (e e' : PartialHomeomorph α β) : Prop := e.source = e'.source ∧ EqOn e e' e.source -#align local_homeomorph.eq_on_source LocalHomeomorph.EqOnSource +#align local_homeomorph.eq_on_source PartialHomeomorph.EqOnSource -theorem eqOnSource_iff (e e' : LocalHomeomorph α β) : +theorem eqOnSource_iff (e e' : PartialHomeomorph α β) : EqOnSource e e' ↔ LocalEquiv.EqOnSource e.toLocalEquiv e'.toLocalEquiv := Iff.rfl -#align local_homeomorph.eq_on_source_iff LocalHomeomorph.eqOnSource_iff +#align local_homeomorph.eq_on_source_iff PartialHomeomorph.eqOnSource_iff /-- `EqOnSource` is an equivalence relation. -/ -instance eqOnSourceSetoid : Setoid (LocalHomeomorph α β) := +instance eqOnSourceSetoid : Setoid (PartialHomeomorph α β) := { LocalEquiv.eqOnSourceSetoid.comap toLocalEquiv with r := EqOnSource } theorem eqOnSource_refl : e ≈ e := Setoid.refl _ -#align local_homeomorph.eq_on_source_refl LocalHomeomorph.eqOnSource_refl +#align local_homeomorph.eq_on_source_refl PartialHomeomorph.eqOnSource_refl -/-- If two local homeomorphisms are equivalent, so are their inverses. -/ -theorem EqOnSource.symm' {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.symm ≈ e'.symm := +/-- If two partial homeomorphisms are equivalent, so are their inverses. -/ +theorem EqOnSource.symm' {e e' : PartialHomeomorph α β} (h : e ≈ e') : e.symm ≈ e'.symm := LocalEquiv.EqOnSource.symm' h -#align local_homeomorph.eq_on_source.symm' LocalHomeomorph.EqOnSource.symm' +#align local_homeomorph.eq_on_source.symm' PartialHomeomorph.EqOnSource.symm' -/-- Two equivalent local homeomorphisms have the same source. -/ -theorem EqOnSource.source_eq {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.source = e'.source := +/-- Two equivalent partial homeomorphisms have the same source. -/ +theorem EqOnSource.source_eq {e e' : PartialHomeomorph α β} (h : e ≈ e') : e.source = e'.source := h.1 -#align local_homeomorph.eq_on_source.source_eq LocalHomeomorph.EqOnSource.source_eq +#align local_homeomorph.eq_on_source.source_eq PartialHomeomorph.EqOnSource.source_eq -/-- Two equivalent local homeomorphisms have the same target. -/ -theorem EqOnSource.target_eq {e e' : LocalHomeomorph α β} (h : e ≈ e') : e.target = e'.target := +/-- Two equivalent partial homeomorphisms have the same target. -/ +theorem EqOnSource.target_eq {e e' : PartialHomeomorph α β} (h : e ≈ e') : e.target = e'.target := h.symm'.1 -#align local_homeomorph.eq_on_source.target_eq LocalHomeomorph.EqOnSource.target_eq +#align local_homeomorph.eq_on_source.target_eq PartialHomeomorph.EqOnSource.target_eq -/-- Two equivalent local homeomorphisms have coinciding `toFun` on the source -/ -theorem EqOnSource.eqOn {e e' : LocalHomeomorph α β} (h : e ≈ e') : EqOn e e' e.source := +/-- Two equivalent partial homeomorphisms have coinciding `toFun` on the source -/ +theorem EqOnSource.eqOn {e e' : PartialHomeomorph α β} (h : e ≈ e') : EqOn e e' e.source := h.2 -#align local_homeomorph.eq_on_source.eq_on LocalHomeomorph.EqOnSource.eqOn +#align local_homeomorph.eq_on_source.eq_on PartialHomeomorph.EqOnSource.eqOn -/-- Two equivalent local homeomorphisms have coinciding `invFun` on the target -/ -theorem EqOnSource.symm_eqOn_target {e e' : LocalHomeomorph α β} (h : e ≈ e') : +/-- Two equivalent partial homeomorphisms have coinciding `invFun` on the target -/ +theorem EqOnSource.symm_eqOn_target {e e' : PartialHomeomorph α β} (h : e ≈ e') : EqOn e.symm e'.symm e.target := h.symm'.2 -#align local_homeomorph.eq_on_source.symm_eq_on_target LocalHomeomorph.EqOnSource.symm_eqOn_target +#align local_homeomorph.eq_on_source.symm_eq_on_target PartialHomeomorph.EqOnSource.symm_eqOn_target -/-- Composition of local homeomorphisms respects equivalence -/ -theorem EqOnSource.trans' {e e' : LocalHomeomorph α β} {f f' : LocalHomeomorph β γ} (he : e ≈ e') - (hf : f ≈ f') : e.trans f ≈ e'.trans f' := +/-- Composition of partial homeomorphisms respects equivalence. -/ +theorem EqOnSource.trans' {e e' : PartialHomeomorph α β} {f f' : PartialHomeomorph β γ} + (he : e ≈ e') (hf : f ≈ f') : e.trans f ≈ e'.trans f' := LocalEquiv.EqOnSource.trans' he hf -#align local_homeomorph.eq_on_source.trans' LocalHomeomorph.EqOnSource.trans' +#align local_homeomorph.eq_on_source.trans' PartialHomeomorph.EqOnSource.trans' -/-- Restriction of local homeomorphisms respects equivalence -/ -theorem EqOnSource.restr {e e' : LocalHomeomorph α β} (he : e ≈ e') (s : Set α) : +/-- Restriction of partial homeomorphisms respects equivalence -/ +theorem EqOnSource.restr {e e' : PartialHomeomorph α β} (he : e ≈ e') (s : Set α) : e.restr s ≈ e'.restr s := LocalEquiv.EqOnSource.restr he _ -#align local_homeomorph.eq_on_source.restr LocalHomeomorph.EqOnSource.restr +#align local_homeomorph.eq_on_source.restr PartialHomeomorph.EqOnSource.restr -/- Two equivalent local homeomorphisms are equal when the source and target are `univ`. -/ -theorem Set.EqOn.restr_eqOn_source {e e' : LocalHomeomorph α β} +/- Two equivalent partial homeomorphisms are equal when the source and target are `univ`. -/ +theorem Set.EqOn.restr_eqOn_source {e e' : PartialHomeomorph α β} (h : EqOn e e' (e.source ∩ e'.source)) : e.restr e'.source ≈ e'.restr e.source := by constructor · rw [e'.restr_source' _ e.open_source] @@ -1013,90 +1014,90 @@ theorem Set.EqOn.restr_eqOn_source {e e' : LocalHomeomorph α β} exact Set.inter_comm _ _ · rw [e.restr_source' _ e'.open_source] refine' (EqOn.trans _ h).trans _ <;> simp only [mfld_simps, eqOn_refl] -#align local_homeomorph.set.eq_on.restr_eq_on_source LocalHomeomorph.Set.EqOn.restr_eqOn_source +#align local_homeomorph.set.eq_on.restr_eq_on_source PartialHomeomorph.Set.EqOn.restr_eqOn_source -/-- Composition of a local homeomorphism and its inverse is equivalent to the restriction of the +/-- Composition of a partial homeomorphism and its inverse is equivalent to the restriction of the identity to the source -/ -theorem trans_self_symm : e.trans e.symm ≈ LocalHomeomorph.ofSet e.source e.open_source := +theorem trans_self_symm : e.trans e.symm ≈ PartialHomeomorph.ofSet e.source e.open_source := LocalEquiv.trans_self_symm _ -#align local_homeomorph.trans_self_symm LocalHomeomorph.trans_self_symm +#align local_homeomorph.trans_self_symm PartialHomeomorph.trans_self_symm -theorem trans_symm_self : e.symm.trans e ≈ LocalHomeomorph.ofSet e.target e.open_target := +theorem trans_symm_self : e.symm.trans e ≈ PartialHomeomorph.ofSet e.target e.open_target := e.symm.trans_self_symm -#align local_homeomorph.trans_symm_self LocalHomeomorph.trans_symm_self +#align local_homeomorph.trans_symm_self PartialHomeomorph.trans_symm_self -theorem eq_of_eqOnSource_univ {e e' : LocalHomeomorph α β} (h : e ≈ e') (s : e.source = univ) +theorem eq_of_eqOnSource_univ {e e' : PartialHomeomorph α β} (h : e ≈ e') (s : e.source = univ) (t : e.target = univ) : e = e' := toLocalEquiv_injective <| LocalEquiv.eq_of_eqOnSource_univ _ _ h s t -#align local_homeomorph.eq_of_eq_on_source_univ LocalHomeomorph.eq_of_eqOnSource_univ +#align local_homeomorph.eq_of_eq_on_source_univ PartialHomeomorph.eq_of_eqOnSource_univ section Prod -/-- The product of two local homeomorphisms, as a local homeomorphism on the product space. -/ +/-- The product of two partial homeomorphisms, as a partial homeomorphism on the product space. -/ @[simps! (config := mfld_cfg) toLocalEquiv apply, simps! (config := .lemmasOnly) source target symm_apply] -def prod (e : LocalHomeomorph α β) (e' : LocalHomeomorph γ δ) : - LocalHomeomorph (α × γ) (β × δ) where +def prod (e : PartialHomeomorph α β) (e' : PartialHomeomorph γ δ) : + PartialHomeomorph (α × γ) (β × δ) where open_source := e.open_source.prod e'.open_source open_target := e.open_target.prod e'.open_target continuousOn_toFun := e.continuousOn.prod_map e'.continuousOn continuousOn_invFun := e.continuousOn_symm.prod_map e'.continuousOn_symm toLocalEquiv := e.toLocalEquiv.prod e'.toLocalEquiv -#align local_homeomorph.prod LocalHomeomorph.prod +#align local_homeomorph.prod PartialHomeomorph.prod @[simp, mfld_simps] -theorem prod_symm (e : LocalHomeomorph α β) (e' : LocalHomeomorph γ δ) : +theorem prod_symm (e : PartialHomeomorph α β) (e' : PartialHomeomorph γ δ) : (e.prod e').symm = e.symm.prod e'.symm := rfl -#align local_homeomorph.prod_symm LocalHomeomorph.prod_symm +#align local_homeomorph.prod_symm PartialHomeomorph.prod_symm @[simp] theorem refl_prod_refl {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] : - (LocalHomeomorph.refl α).prod (LocalHomeomorph.refl β) = LocalHomeomorph.refl (α × β) := - LocalHomeomorph.ext _ _ (fun _ => rfl) (fun _ => rfl) univ_prod_univ -#align local_homeomorph.refl_prod_refl LocalHomeomorph.refl_prod_refl + (PartialHomeomorph.refl α).prod (PartialHomeomorph.refl β) = PartialHomeomorph.refl (α × β) := + PartialHomeomorph.ext _ _ (fun _ => rfl) (fun _ => rfl) univ_prod_univ +#align local_homeomorph.refl_prod_refl PartialHomeomorph.refl_prod_refl @[simp, mfld_simps] theorem prod_trans {η : Type*} {ε : Type*} [TopologicalSpace η] [TopologicalSpace ε] - (e : LocalHomeomorph α β) (f : LocalHomeomorph β γ) (e' : LocalHomeomorph δ η) - (f' : LocalHomeomorph η ε) : (e.prod e').trans (f.prod f') = (e.trans f).prod (e'.trans f') := + (e : PartialHomeomorph α β) (f : PartialHomeomorph β γ) (e' : PartialHomeomorph δ η) + (f' : PartialHomeomorph η ε) : (e.prod e').trans (f.prod f') = (e.trans f).prod (e'.trans f') := toLocalEquiv_injective <| e.1.prod_trans .. -#align local_homeomorph.prod_trans LocalHomeomorph.prod_trans +#align local_homeomorph.prod_trans PartialHomeomorph.prod_trans -theorem prod_eq_prod_of_nonempty {e₁ e₁' : LocalHomeomorph α β} {e₂ e₂' : LocalHomeomorph γ δ} +theorem prod_eq_prod_of_nonempty {e₁ e₁' : PartialHomeomorph α β} {e₂ e₂' : PartialHomeomorph γ δ} (h : (e₁.prod e₂).source.Nonempty) : e₁.prod e₂ = e₁'.prod e₂' ↔ e₁ = e₁' ∧ e₂ = e₂' := by obtain ⟨⟨x, y⟩, -⟩ := id h haveI : Nonempty α := ⟨x⟩ haveI : Nonempty β := ⟨e₁ x⟩ haveI : Nonempty γ := ⟨y⟩ haveI : Nonempty δ := ⟨e₂ y⟩ - simp_rw [LocalHomeomorph.ext_iff, prod_apply, prod_symm_apply, prod_source, Prod.ext_iff, + simp_rw [PartialHomeomorph.ext_iff, prod_apply, prod_symm_apply, prod_source, Prod.ext_iff, Set.prod_eq_prod_iff_of_nonempty h, forall_and, Prod.forall, forall_const, and_assoc, and_left_comm] -#align local_homeomorph.prod_eq_prod_of_nonempty LocalHomeomorph.prod_eq_prod_of_nonempty +#align local_homeomorph.prod_eq_prod_of_nonempty PartialHomeomorph.prod_eq_prod_of_nonempty -theorem prod_eq_prod_of_nonempty' {e₁ e₁' : LocalHomeomorph α β} {e₂ e₂' : LocalHomeomorph γ δ} +theorem prod_eq_prod_of_nonempty' {e₁ e₁' : PartialHomeomorph α β} {e₂ e₂' : PartialHomeomorph γ δ} (h : (e₁'.prod e₂').source.Nonempty) : e₁.prod e₂ = e₁'.prod e₂' ↔ e₁ = e₁' ∧ e₂ = e₂' := by rw [eq_comm, prod_eq_prod_of_nonempty h, eq_comm, @eq_comm _ e₂'] -#align local_homeomorph.prod_eq_prod_of_nonempty' LocalHomeomorph.prod_eq_prod_of_nonempty' +#align local_homeomorph.prod_eq_prod_of_nonempty' PartialHomeomorph.prod_eq_prod_of_nonempty' end Prod section Piecewise -/-- Combine two `LocalHomeomorph`s using `Set.piecewise`. The source of the new `LocalHomeomorph` -is `s.ite e.source e'.source = e.source ∩ s ∪ e'.source \ s`, and similarly for target. The -function sends `e.source ∩ s` to `e.target ∩ t` using `e` and `e'.source \ s` to `e'.target \ t` -using `e'`, and similarly for the inverse function. To ensure that the maps `toFun` and `invFun` -are inverse of each other on the new `source` and `target`, the definition assumes that the sets `s` -and `t` are related both by `e.is_image` and `e'.is_image`. To ensure that the new maps are -continuous on `source`/`target`, it also assumes that `e.source` and `e'.source` meet `frontier s` -on the same set and `e x = e' x` on this intersection. -/ +/-- Combine two `PartialHomeomorph`s using `Set.piecewise`. The source of the new +`PartialHomeomorph` is `s.ite e.source e'.source = e.source ∩ s ∪ e'.source \ s`, and similarly for +target. The function sends `e.source ∩ s` to `e.target ∩ t` using `e` and +`e'.source \ s` to `e'.target \ t` using `e'`, and similarly for the inverse function. +To ensure the maps `toFun` and `invFun` are inverse of each other on the new `source` and `target`, +the definition assumes that the sets `s` and `t` are related both by `e.is_image` and `e'.is_image`. +To ensure that the new maps are continuous on `source`/`target`, it also assumes that `e.source` and +`e'.source` meet `frontier s` on the same set and `e x = e' x` on this intersection. -/ @[simps! (config := .asFn) toLocalEquiv apply] -def piecewise (e e' : LocalHomeomorph α β) (s : Set α) (t : Set β) [∀ x, Decidable (x ∈ s)] +def piecewise (e e' : PartialHomeomorph α β) (s : Set α) (t : Set β) [∀ x, Decidable (x ∈ s)] [∀ y, Decidable (y ∈ t)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : e.source ∩ frontier s = e'.source ∩ frontier s) - (Heq : EqOn e e' (e.source ∩ frontier s)) : LocalHomeomorph α β where + (Heq : EqOn e e' (e.source ∩ frontier s)) : PartialHomeomorph α β where toLocalEquiv := e.toLocalEquiv.piecewise e'.toLocalEquiv s t H H' open_source := e.open_source.ite e'.open_source Hs open_target := @@ -1106,11 +1107,11 @@ def piecewise (e e' : LocalHomeomorph α β) (s : Set α) (t : Set β) [∀ x, D continuousOn_piecewise_ite e.continuousOn_symm e'.continuousOn_symm (H.frontier.inter_eq_of_inter_eq_of_eqOn H'.frontier Hs Heq) (H.frontier.symm_eqOn_of_inter_eq_of_eqOn Hs Heq) -#align local_homeomorph.piecewise LocalHomeomorph.piecewise +#align local_homeomorph.piecewise PartialHomeomorph.piecewise @[simp] -theorem symm_piecewise (e e' : LocalHomeomorph α β) {s : Set α} {t : Set β} [∀ x, Decidable (x ∈ s)] - [∀ y, Decidable (y ∈ t)] (H : e.IsImage s t) (H' : e'.IsImage s t) +theorem symm_piecewise (e e' : PartialHomeomorph α β) {s : Set α} {t : Set β} + [∀ x, Decidable (x ∈ s)] [∀ y, Decidable (y ∈ t)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : e.source ∩ frontier s = e'.source ∩ frontier s) (Heq : EqOn e e' (e.source ∩ frontier s)) : (e.piecewise e' s t H H' Hs Heq).symm = @@ -1118,14 +1119,14 @@ theorem symm_piecewise (e e' : LocalHomeomorph α β) {s : Set α} {t : Set β} (H.frontier.inter_eq_of_inter_eq_of_eqOn H'.frontier Hs Heq) (H.frontier.symm_eqOn_of_inter_eq_of_eqOn Hs Heq) := rfl -#align local_homeomorph.symm_piecewise LocalHomeomorph.symm_piecewise +#align local_homeomorph.symm_piecewise PartialHomeomorph.symm_piecewise -/-- Combine two `LocalHomeomorph`s with disjoint sources and disjoint targets. We reuse -`LocalHomeomorph.piecewise` then override `toLocalEquiv` to `LocalEquiv.disjointUnion`. +/-- Combine two `PartialHomeomorph`s with disjoint sources and disjoint targets. We reuse +`PartialHomeomorph.piecewise` then override `toLocalEquiv` to `LocalEquiv.disjointUnion`. This way we have better definitional equalities for `source` and `target`. -/ -def disjointUnion (e e' : LocalHomeomorph α β) [∀ x, Decidable (x ∈ e.source)] +def disjointUnion (e e' : PartialHomeomorph α β) [∀ x, Decidable (x ∈ e.source)] [∀ y, Decidable (y ∈ e.target)] (Hs : Disjoint e.source e'.source) - (Ht : Disjoint e.target e'.target) : LocalHomeomorph α β := + (Ht : Disjoint e.target e'.target) : PartialHomeomorph α β := (e.piecewise e' e.source e.target e.isImage_source_target (e'.isImage_source_target_of_disjoint e Hs.symm Ht.symm) (by rw [e.open_source.inter_frontier_eq, (Hs.symm.frontier_right e'.open_source).inter_eq]) @@ -1134,18 +1135,18 @@ def disjointUnion (e e' : LocalHomeomorph α β) [∀ x, Decidable (x ∈ e.sour exact eqOn_empty _ _)).replaceEquiv (e.toLocalEquiv.disjointUnion e'.toLocalEquiv Hs Ht) (LocalEquiv.disjointUnion_eq_piecewise _ _ _ _).symm -#align local_homeomorph.disjoint_union LocalHomeomorph.disjointUnion +#align local_homeomorph.disjoint_union PartialHomeomorph.disjointUnion end Piecewise section Pi variable {ι : Type*} [Fintype ι] {Xi Yi : ι → Type*} [∀ i, TopologicalSpace (Xi i)] - [∀ i, TopologicalSpace (Yi i)] (ei : ∀ i, LocalHomeomorph (Xi i) (Yi i)) + [∀ i, TopologicalSpace (Yi i)] (ei : ∀ i, PartialHomeomorph (Xi i) (Yi i)) -/-- The product of a finite family of `LocalHomeomorph`s. -/ +/-- The product of a finite family of `PartialHomeomorph`s. -/ @[simps toLocalEquiv] -def pi : LocalHomeomorph (∀ i, Xi i) (∀ i, Yi i) where +def pi : PartialHomeomorph (∀ i, Xi i) (∀ i, Yi i) where toLocalEquiv := LocalEquiv.pi fun i => (ei i).toLocalEquiv open_source := isOpen_set_pi finite_univ fun i _ => (ei i).open_source open_target := isOpen_set_pi finite_univ fun i _ => (ei i).open_target @@ -1153,7 +1154,7 @@ def pi : LocalHomeomorph (∀ i, Xi i) (∀ i, Yi i) where (ei i).continuousOn.comp (continuous_apply _).continuousOn fun _f hf => hf i trivial continuousOn_invFun := continuousOn_pi.2 fun i => (ei i).continuousOn_symm.comp (continuous_apply _).continuousOn fun _f hf => hf i trivial -#align local_homeomorph.pi LocalHomeomorph.pi +#align local_homeomorph.pi PartialHomeomorph.pi end Pi @@ -1166,17 +1167,17 @@ theorem continuousWithinAt_iff_continuousWithinAt_comp_right {f : β → γ} {s ContinuousWithinAt f s x ↔ ContinuousWithinAt (f ∘ e) (e ⁻¹' s) (e.symm x) := by simp_rw [ContinuousWithinAt, ← @tendsto_map'_iff _ _ _ _ e, e.map_nhdsWithin_preimage_eq (e.map_target h), (· ∘ ·), e.right_inv h] -#align local_homeomorph.continuous_within_at_iff_continuous_within_at_comp_right LocalHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_right +#align local_homeomorph.continuous_within_at_iff_continuous_within_at_comp_right PartialHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_right -/-- Continuity at a point can be read under right composition with a local homeomorphism, if the +/-- Continuity at a point can be read under right composition with a partial homeomorphism, if the point is in its target -/ theorem continuousAt_iff_continuousAt_comp_right {f : β → γ} {x : β} (h : x ∈ e.target) : ContinuousAt f x ↔ ContinuousAt (f ∘ e) (e.symm x) := by rw [← continuousWithinAt_univ, e.continuousWithinAt_iff_continuousWithinAt_comp_right h, preimage_univ, continuousWithinAt_univ] -#align local_homeomorph.continuous_at_iff_continuous_at_comp_right LocalHomeomorph.continuousAt_iff_continuousAt_comp_right +#align local_homeomorph.continuous_at_iff_continuous_at_comp_right PartialHomeomorph.continuousAt_iff_continuousAt_comp_right -/-- A function is continuous on a set if and only if its composition with a local homeomorphism +/-- A function is continuous on a set if and only if its composition with a partial homeomorphism on the right is continuous on the corresponding set. -/ theorem continuousOn_iff_continuousOn_comp_right {f : β → γ} {s : Set β} (h : s ⊆ e.target) : ContinuousOn f s ↔ ContinuousOn (f ∘ e) (e.source ∩ e ⁻¹' s) := by @@ -1185,7 +1186,7 @@ theorem continuousOn_iff_continuousOn_comp_right {f : β → γ} {s : Set β} (h rw [e.continuousWithinAt_iff_continuousWithinAt_comp_right (h hx), e.symm_image_eq_source_inter_preimage h, inter_comm, continuousWithinAt_inter] exact IsOpen.mem_nhds e.open_source (e.map_target (h hx)) -#align local_homeomorph.continuous_on_iff_continuous_on_comp_right LocalHomeomorph.continuousOn_iff_continuousOn_comp_right +#align local_homeomorph.continuous_on_iff_continuous_on_comp_right PartialHomeomorph.continuousOn_iff_continuousOn_comp_right /-- Continuity within a set at a point can be read under left composition with a local homeomorphism if a neighborhood of the initial point is sent to the source of the local @@ -1200,38 +1201,38 @@ theorem continuousWithinAt_iff_continuousWithinAt_comp_left {f : γ → α} {s : (e.continuousAt_symm (e.map_source hx)).continuousWithinAt ContinuousWithinAt.comp this fe_cont (subset_univ _) exact this.congr (fun y hy => by simp [e.left_inv hy.2]) (by simp [e.left_inv hx]) -#align local_homeomorph.continuous_within_at_iff_continuous_within_at_comp_left LocalHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_left +#align local_homeomorph.continuous_within_at_iff_continuous_within_at_comp_left PartialHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_left -/-- Continuity at a point can be read under left composition with a local homeomorphism if a -neighborhood of the initial point is sent to the source of the local homeomorphism-/ +/-- Continuity at a point can be read under left composition with a partial homeomorphism if a +neighborhood of the initial point is sent to the source of the partial homeomorphism-/ theorem continuousAt_iff_continuousAt_comp_left {f : γ → α} {x : γ} (h : f ⁻¹' e.source ∈ 𝓝 x) : ContinuousAt f x ↔ ContinuousAt (e ∘ f) x := by have hx : f x ∈ e.source := (mem_of_mem_nhds h : _) have h' : f ⁻¹' e.source ∈ 𝓝[univ] x := by rwa [nhdsWithin_univ] rw [← continuousWithinAt_univ, ← continuousWithinAt_univ, e.continuousWithinAt_iff_continuousWithinAt_comp_left hx h'] -#align local_homeomorph.continuous_at_iff_continuous_at_comp_left LocalHomeomorph.continuousAt_iff_continuousAt_comp_left +#align local_homeomorph.continuous_at_iff_continuous_at_comp_left PartialHomeomorph.continuousAt_iff_continuousAt_comp_left -/-- A function is continuous on a set if and only if its composition with a local homeomorphism +/-- A function is continuous on a set if and only if its composition with a partial homeomorphism on the left is continuous on the corresponding set. -/ theorem continuousOn_iff_continuousOn_comp_left {f : γ → α} {s : Set γ} (h : s ⊆ f ⁻¹' e.source) : ContinuousOn f s ↔ ContinuousOn (e ∘ f) s := forall₂_congr fun _x hx => e.continuousWithinAt_iff_continuousWithinAt_comp_left (h hx) (mem_of_superset self_mem_nhdsWithin h) -#align local_homeomorph.continuous_on_iff_continuous_on_comp_left LocalHomeomorph.continuousOn_iff_continuousOn_comp_left +#align local_homeomorph.continuous_on_iff_continuous_on_comp_left PartialHomeomorph.continuousOn_iff_continuousOn_comp_left -/-- A function is continuous if and only if its composition with a local homeomorphism +/-- A function is continuous if and only if its composition with a partial homeomorphism on the left is continuous and its image is contained in the source. -/ theorem continuous_iff_continuous_comp_left {f : γ → α} (h : f ⁻¹' e.source = univ) : Continuous f ↔ Continuous (e ∘ f) := by simp only [continuous_iff_continuousOn_univ] exact e.continuousOn_iff_continuousOn_comp_left (Eq.symm h).subset -#align local_homeomorph.continuous_iff_continuous_comp_left LocalHomeomorph.continuous_iff_continuous_comp_left +#align local_homeomorph.continuous_iff_continuous_comp_left PartialHomeomorph.continuous_iff_continuous_comp_left end Continuity -/-- The homeomorphism obtained by restricting a `LocalHomeomorph` to a subset of the source. -/ +/-- The homeomorphism obtained by restricting a `PartialHomeomorph` to a subset of the source. -/ @[simps] def homeomorphOfImageSubsetSource {s : Set α} {t : Set β} (hs : s ⊆ e.source) (ht : e '' s = t) : s ≃ₜ t := @@ -1245,18 +1246,18 @@ def homeomorphOfImageSubsetSource {s : Set α} {t : Set β} (hs : s ⊆ e.source right_inv := fun b => Subtype.eq <| e.right_inv (h₂ b.2) continuous_toFun := (e.continuousOn.mono hs).restrict_mapsTo h₁ continuous_invFun := (e.continuousOn_symm.mono h₂).restrict_mapsTo h₃ } -#align local_homeomorph.homeomorph_of_image_subset_source LocalHomeomorph.homeomorphOfImageSubsetSource +#align local_homeomorph.homeomorph_of_image_subset_source PartialHomeomorph.homeomorphOfImageSubsetSource -/-- A local homeomorphism defines a homeomorphism between its source and target. -/ +/-- A partial homeomorphism defines a homeomorphism between its source and target. -/ @[simps!] -- porting note: new `simps` def toHomeomorphSourceTarget : e.source ≃ₜ e.target := e.homeomorphOfImageSubsetSource subset_rfl e.image_source_eq_target -#align local_homeomorph.to_homeomorph_source_target LocalHomeomorph.toHomeomorphSourceTarget +#align local_homeomorph.to_homeomorph_source_target PartialHomeomorph.toHomeomorphSourceTarget -theorem secondCountableTopology_source [SecondCountableTopology β] (e : LocalHomeomorph α β) : +theorem secondCountableTopology_source [SecondCountableTopology β] (e : PartialHomeomorph α β) : SecondCountableTopology e.source := e.toHomeomorphSourceTarget.secondCountableTopology -#align local_homeomorph.second_countable_topology_source LocalHomeomorph.secondCountableTopology_source +#align local_homeomorph.second_countable_topology_source PartialHomeomorph.secondCountableTopology_source theorem nhds_eq_comap_inf_principal {x} (hx : x ∈ e.source) : 𝓝 x = comap e (𝓝 (e x)) ⊓ 𝓟 e.source := by @@ -1265,7 +1266,7 @@ theorem nhds_eq_comap_inf_principal {x} (hx : x ∈ e.source) : e.toHomeomorphSourceTarget.nhds_eq_comap, nhds_subtype_eq_comap] simp only [(· ∘ ·), toHomeomorphSourceTarget_apply_coe, comap_comap] -/-- If a local homeomorphism has source and target equal to univ, then it induces a homeomorphism +/-- If a partial homeomorphism has source and target equal to univ, then it induces a homeomorphism between the whole spaces, expressed in this definition. -/ @[simps (config := mfld_cfg) apply symm_apply] -- porting note: todo: add a `LocalEquiv` version def toHomeomorphOfSourceEqUnivTargetEqUniv (h : e.source = (univ : Set α)) (h' : e.target = univ) : @@ -1284,7 +1285,7 @@ def toHomeomorphOfSourceEqUnivTargetEqUniv (h : e.source = (univ : Set α)) (h' simpa only [continuous_iff_continuousOn_univ, h] using e.continuousOn continuous_invFun := by simpa only [continuous_iff_continuousOn_univ, h'] using e.continuousOn_symm -#align local_homeomorph.to_homeomorph_of_source_eq_univ_target_eq_univ LocalHomeomorph.toHomeomorphOfSourceEqUnivTargetEqUniv +#align local_homeomorph.to_homeomorph_of_source_eq_univ_target_eq_univ PartialHomeomorph.toHomeomorphOfSourceEqUnivTargetEqUniv theorem openEmbedding_restrict : OpenEmbedding (e.source.restrict e) := by refine openEmbedding_of_continuous_injective_open (e.continuousOn.comp_continuous @@ -1293,37 +1294,38 @@ theorem openEmbedding_restrict : OpenEmbedding (e.source.restrict e) := by exact e.image_isOpen_of_isOpen (e.open_source.isOpenMap_subtype_val V hV) fun _ ⟨x, _, h⟩ ↦ h ▸ x.2 -/-- A local homeomorphism whose source is all of `α` defines an open embedding of `α` into `β`. The -converse is also true; see `OpenEmbedding.toLocalHomeomorph`. -/ +/-- A partial homeomorphism whose source is all of `α` defines an open embedding of `α` into `β`. +The converse is also true; see `OpenEmbedding.toPartialHomeomorph`. -/ theorem to_openEmbedding (h : e.source = Set.univ) : OpenEmbedding e := e.openEmbedding_restrict.comp ((Homeomorph.setCongr h).trans <| Homeomorph.Set.univ α).symm.openEmbedding -#align local_homeomorph.to_open_embedding LocalHomeomorph.to_openEmbedding +#align local_homeomorph.to_open_embedding PartialHomeomorph.to_openEmbedding -end LocalHomeomorph +end PartialHomeomorph namespace Homeomorph variable (e : α ≃ₜ β) (e' : β ≃ₜ γ) -/- Register as simp lemmas that the fields of a local homeomorphism built from a homeomorphism +/- Register as simp lemmas that the fields of a partial homeomorphism built from a homeomorphism correspond to the fields of the original homeomorphism. -/ @[simp, mfld_simps] -theorem refl_toLocalHomeomorph : (Homeomorph.refl α).toLocalHomeomorph = LocalHomeomorph.refl α := +theorem refl_toPartialHomeomorph : + (Homeomorph.refl α).toPartialHomeomorph = PartialHomeomorph.refl α := rfl -#align homeomorph.refl_to_local_homeomorph Homeomorph.refl_toLocalHomeomorph +#align homeomorph.refl_to_local_homeomorph Homeomorph.refl_toPartialHomeomorph @[simp, mfld_simps] -theorem symm_toLocalHomeomorph : e.symm.toLocalHomeomorph = e.toLocalHomeomorph.symm := +theorem symm_toPartialHomeomorph : e.symm.toPartialHomeomorph = e.toPartialHomeomorph.symm := rfl -#align homeomorph.symm_to_local_homeomorph Homeomorph.symm_toLocalHomeomorph +#align homeomorph.symm_to_local_homeomorph Homeomorph.symm_toPartialHomeomorph @[simp, mfld_simps] -theorem trans_toLocalHomeomorph : - (e.trans e').toLocalHomeomorph = e.toLocalHomeomorph.trans e'.toLocalHomeomorph := - LocalHomeomorph.toLocalEquiv_injective <| Equiv.trans_toLocalEquiv _ _ -#align homeomorph.trans_to_local_homeomorph Homeomorph.trans_toLocalHomeomorph +theorem trans_toPartialHomeomorph : + (e.trans e').toPartialHomeomorph = e.toPartialHomeomorph.trans e'.toPartialHomeomorph := + PartialHomeomorph.toLocalEquiv_injective <| Equiv.trans_toLocalEquiv _ _ +#align homeomorph.trans_to_local_homeomorph Homeomorph.trans_toPartialHomeomorph end Homeomorph @@ -1331,13 +1333,13 @@ namespace OpenEmbedding variable (f : α → β) (h : OpenEmbedding f) -/-- An open embedding of `α` into `β`, with `α` nonempty, defines a local homeomorphism whose source -is all of `α`. The converse is also true; see `LocalHomeomorph.to_openEmbedding`. -/ +/-- An open embedding of `α` into `β`, with `α` nonempty, defines a partial homeomorphism +whose source is all of `α`. The converse is also true; see `PartialHomeomorph.to_openEmbedding`. -/ @[simps! (config := mfld_cfg) apply source target] -noncomputable def toLocalHomeomorph [Nonempty α] : LocalHomeomorph α β := - LocalHomeomorph.ofContinuousOpen ((h.toEmbedding.inj.injOn univ).toLocalEquiv _ _) +noncomputable def toPartialHomeomorph [Nonempty α] : PartialHomeomorph α β := + PartialHomeomorph.ofContinuousOpen ((h.toEmbedding.inj.injOn univ).toLocalEquiv _ _) h.continuous.continuousOn h.isOpenMap isOpen_univ -#align open_embedding.to_local_homeomorph OpenEmbedding.toLocalHomeomorph +#align open_embedding.to_local_homeomorph OpenEmbedding.toPartialHomeomorph end OpenEmbedding @@ -1347,10 +1349,10 @@ open TopologicalSpace variable (s : Opens α) [Nonempty s] -/-- The inclusion of an open subset `s` of a space `α` into `α` is a local homeomorphism from the +/-- The inclusion of an open subset `s` of a space `α` into `α` is a partial homeomorphism from the subtype `s` to `α`. -/ -noncomputable def localHomeomorphSubtypeCoe : LocalHomeomorph s α := - OpenEmbedding.toLocalHomeomorph _ s.2.openEmbedding_subtype_val +noncomputable def localHomeomorphSubtypeCoe : PartialHomeomorph s α := + OpenEmbedding.toPartialHomeomorph _ s.2.openEmbedding_subtype_val #align topological_space.opens.local_homeomorph_subtype_coe TopologicalSpace.Opens.localHomeomorphSubtypeCoe @[simp, mfld_simps] @@ -1371,34 +1373,34 @@ theorem localHomeomorphSubtypeCoe_target : s.localHomeomorphSubtypeCoe.target = end TopologicalSpace.Opens -namespace LocalHomeomorph +namespace PartialHomeomorph open TopologicalSpace -variable (e : LocalHomeomorph α β) +variable (e : PartialHomeomorph α β) variable (s : Opens α) [Nonempty s] -/-- The restriction of a local homeomorphism `e` to an open subset `s` of the domain type produces a -local homeomorphism whose domain is the subtype `s`.-/ -noncomputable def subtypeRestr : LocalHomeomorph s β := +/-- The restriction of a partial homeomorphism `e` to an open subset `s` of the domain type +produces a partial homeomorphism whose domain is the subtype `s`. -/ +noncomputable def subtypeRestr : PartialHomeomorph s β := s.localHomeomorphSubtypeCoe.trans e -#align local_homeomorph.subtype_restr LocalHomeomorph.subtypeRestr +#align local_homeomorph.subtype_restr PartialHomeomorph.subtypeRestr theorem subtypeRestr_def : e.subtypeRestr s = s.localHomeomorphSubtypeCoe.trans e := rfl -#align local_homeomorph.subtype_restr_def LocalHomeomorph.subtypeRestr_def +#align local_homeomorph.subtype_restr_def PartialHomeomorph.subtypeRestr_def @[simp, mfld_simps] theorem subtypeRestr_coe : - ((e.subtypeRestr s : LocalHomeomorph s β) : s → β) = Set.restrict ↑s (e : α → β) := + ((e.subtypeRestr s : PartialHomeomorph s β) : s → β) = Set.restrict ↑s (e : α → β) := rfl -#align local_homeomorph.subtype_restr_coe LocalHomeomorph.subtypeRestr_coe +#align local_homeomorph.subtype_restr_coe PartialHomeomorph.subtypeRestr_coe @[simp, mfld_simps] theorem subtypeRestr_source : (e.subtypeRestr s).source = (↑) ⁻¹' e.source := by simp only [subtypeRestr_def, mfld_simps] -#align local_homeomorph.subtype_restr_source LocalHomeomorph.subtypeRestr_source +#align local_homeomorph.subtype_restr_source PartialHomeomorph.subtypeRestr_source variable {s} @@ -1406,13 +1408,13 @@ theorem map_subtype_source {x : s} (hxe : (x : α) ∈ e.source): e x ∈ (e.sub refine' ⟨e.map_source hxe, _⟩ rw [s.localHomeomorphSubtypeCoe_target, mem_preimage, e.leftInvOn hxe] exact x.prop -#align local_homeomorph.map_subtype_source LocalHomeomorph.map_subtype_source +#align local_homeomorph.map_subtype_source PartialHomeomorph.map_subtype_source variable (s) /- This lemma characterizes the transition functions of an open subset in terms of the transition functions of the original space. -/ -theorem subtypeRestr_symm_trans_subtypeRestr (f f' : LocalHomeomorph α β) : +theorem subtypeRestr_symm_trans_subtypeRestr (f f' : PartialHomeomorph α β) : (f.subtypeRestr s).symm.trans (f'.subtypeRestr s) ≈ (f.symm.trans f').restr (f.target ∩ f.symm ⁻¹' s) := by simp only [subtypeRestr_def, trans_symm_eq_symm_trans_symm] @@ -1428,7 +1430,7 @@ theorem subtypeRestr_symm_trans_subtypeRestr (f f' : LocalHomeomorph α β) : -- f has been eliminated !!! refine' Setoid.trans (trans_symm_self s.localHomeomorphSubtypeCoe) _ simp only [mfld_simps, Setoid.refl] -#align local_homeomorph.subtype_restr_symm_trans_subtype_restr LocalHomeomorph.subtypeRestr_symm_trans_subtypeRestr +#align local_homeomorph.subtype_restr_symm_trans_subtype_restr PartialHomeomorph.subtypeRestr_symm_trans_subtypeRestr theorem subtypeRestr_symm_eqOn (U : Opens α) [Nonempty U] : EqOn e.symm (Subtype.val ∘ (e.subtypeRestr U).symm) (e.subtypeRestr U).target := by @@ -1443,18 +1445,18 @@ theorem subtypeRestr_symm_eqOn_of_le {U V : Opens α} [Nonempty U] [Nonempty V] (e.subtypeRestr U).target := by set i := Set.inclusion hUV intro y hy - dsimp [LocalHomeomorph.subtypeRestr_def] at hy ⊢ + dsimp [PartialHomeomorph.subtypeRestr_def] at hy ⊢ have hyV : e.symm y ∈ V.localHomeomorphSubtypeCoe.target := by rw [Opens.localHomeomorphSubtypeCoe_target] at hy ⊢ exact hUV hy.2 refine' V.localHomeomorphSubtypeCoe.injOn _ trivial _ - · rw [← LocalHomeomorph.symm_target] - apply LocalHomeomorph.map_source - rw [LocalHomeomorph.symm_source] + · rw [← PartialHomeomorph.symm_target] + apply PartialHomeomorph.map_source + rw [PartialHomeomorph.symm_source] exact hyV · rw [V.localHomeomorphSubtypeCoe.right_inv hyV] show _ = U.localHomeomorphSubtypeCoe _ rw [U.localHomeomorphSubtypeCoe.right_inv hy.2] -#align local_homeomorph.subtype_restr_symm_eq_on_of_le LocalHomeomorph.subtypeRestr_symm_eqOn_of_le +#align local_homeomorph.subtype_restr_symm_eq_on_of_le PartialHomeomorph.subtypeRestr_symm_eqOn_of_le -end LocalHomeomorph +end PartialHomeomorph diff --git a/Mathlib/Topology/Sheaves/LocalPredicate.lean b/Mathlib/Topology/Sheaves/LocalPredicate.lean index 8b1e61d4bdf67..9cbd1e3f6b1a6 100644 --- a/Mathlib/Topology/Sheaves/LocalPredicate.lean +++ b/Mathlib/Topology/Sheaves/LocalPredicate.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin, Scott Morrison, Adam Topaz -/ import Mathlib.Topology.Sheaves.SheafOfFunctions import Mathlib.Topology.Sheaves.Stalks -import Mathlib.Topology.LocalHomeomorph +import Mathlib.Topology.PartialHomeomorph import Mathlib.Topology.Sheaves.SheafCondition.UniqueGluing #align_import topology.sheaves.local_predicate from "leanprover-community/mathlib"@"5dc6092d09e5e489106865241986f7f2ad28d4c8" diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index de9ecf145514d..625d8332aaf98 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -344,7 +344,7 @@ set_option linter.uppercaseLean3 false in theorem apply_symm_apply_eq_coordChangeL (e e' : Trivialization F (π F E)) [e.IsLinear R] [e'.IsLinear R] {b : B} (hb : b ∈ e.baseSet ∩ e'.baseSet) (v : F) : - e' (e.toLocalHomeomorph.symm (b, v)) = (b, e.coordChangeL R e' b v) := by + e' (e.toPartialHomeomorph.symm (b, v)) = (b, e.coordChangeL R e' b v) := by rw [e.mk_coordChangeL e' hb, e.mk_symm hb.1] set_option linter.uppercaseLean3 false in #align trivialization.apply_symm_apply_eq_coord_changeL Trivialization.apply_symm_apply_eq_coordChangeL @@ -354,7 +354,7 @@ right-hand side is ugly, but has good definitional properties for specifically d trivializations. -/ theorem coordChangeL_apply' (e e' : Trivialization F (π F E)) [e.IsLinear R] [e'.IsLinear R] {b : B} (hb : b ∈ e.baseSet ∩ e'.baseSet) (y : F) : - coordChangeL R e e' b y = (e' (e.toLocalHomeomorph.symm (b, y))).2 := by + coordChangeL R e e' b y = (e' (e.toPartialHomeomorph.symm (b, y))).2 := by rw [e.coordChangeL_apply e' hb, e.mk_symm hb.1] set_option linter.uppercaseLean3 false in #align trivialization.coord_changeL_apply' Trivialization.coordChangeL_apply' @@ -521,7 +521,7 @@ variable {R} theorem symm_apply_eq_mk_continuousLinearEquivAt_symm (e : Trivialization F (π F E)) [e.IsLinear R] (b : B) (hb : b ∈ e.baseSet) (z : F) : - e.toLocalHomeomorph.symm ⟨b, z⟩ = ⟨b, (e.continuousLinearEquivAt R b hb).symm z⟩ := by + e.toPartialHomeomorph.symm ⟨b, z⟩ = ⟨b, (e.continuousLinearEquivAt R b hb).symm z⟩ := by have h : (b, z) ∈ e.target · rw [e.target_eq] exact ⟨hb, mem_univ _⟩ @@ -649,7 +649,7 @@ protected def TotalSpace := #align vector_bundle_core.total_space VectorBundleCore.TotalSpace /-- Local homeomorphism version of the trivialization change. -/ -def trivChange (i j : ι) : LocalHomeomorph (B × F) (B × F) := +def trivChange (i j : ι) : PartialHomeomorph (B × F) (B × F) := Z.toFiberBundleCore.trivChange i j #align vector_bundle_core.triv_change VectorBundleCore.trivChange @@ -712,7 +712,7 @@ theorem mem_localTriv_target (p : B × F) : @[simp, mfld_simps] theorem localTriv_symm_fst (p : B × F) : - (Z.localTriv i).toLocalHomeomorph.symm p = ⟨p.1, Z.coordChange i (Z.indexAt p.1) p.1 p.2⟩ := + (Z.localTriv i).toPartialHomeomorph.symm p = ⟨p.1, Z.coordChange i (Z.indexAt p.1) p.1 p.2⟩ := rfl #align vector_bundle_core.local_triv_symm_fst VectorBundleCore.localTriv_symm_fst diff --git a/scripts/nolints.json b/scripts/nolints.json index 112195fadba12..21ee2657fed9f 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -661,10 +661,10 @@ ["docBlame", "List.equiv"], ["docBlame", "List.inj_on"], ["docBlame", "List.remove"], - ["docBlame", "LocalHomeomorph.continuousOn_invFun"], - ["docBlame", "LocalHomeomorph.continuousOn_toFun"], - ["docBlame", "LocalHomeomorph.open_source"], - ["docBlame", "LocalHomeomorph.open_target"], + ["docBlame", "PartialHomeomorph.continuousOn_invFun"], + ["docBlame", "PartialHomeomorph.continuousOn_toFun"], + ["docBlame", "PartialHomeomorph.open_source"], + ["docBlame", "PartialHomeomorph.open_target"], ["docBlame", "LocallyConvexSpace.convex_basis"], ["docBlame", "LowerSetTopology.topology_eq_lowerSetTopology"], ["docBlame", "LowerTopology.topology_eq_lowerTopology"], diff --git a/test/MfldSetTac.lean b/test/MfldSetTac.lean index 1e0991e781b20..93b3665322bb7 100644 --- a/test/MfldSetTac.lean +++ b/test/MfldSetTac.lean @@ -11,7 +11,7 @@ private axiom test_sorry : ∀ {α}, α This is a test file for the tactic `mfld_set_tac`. Because this tactic applies a simp-set which mostly contains lemmas in advanced parts of mathlib, it is currently impossible to truly test it in realistic conditions. Instead, we create stub definitions and lemmas on objects such as -`LocalHomeomorph`, label them with `mfld_simps` and run tests on those. +`PartialHomeomorph`, label them with `mfld_simps` and run tests on those. -/ open Lean Meta Elab Tactic @@ -21,28 +21,28 @@ open Lean Meta Elab Tactic set_option autoImplicit true section stub_lemmas -structure LocalHomeomorph (α : Type u) (β : Type u) extends LocalEquiv α β +structure PartialHomeomorph (α : Type u) (β : Type u) extends LocalEquiv α β noncomputable -instance LocalHomeomorph.has_coe_to_fun : CoeFun (LocalHomeomorph α β) (λ _ => α → β) := test_sorry +instance PartialHomeomorph.has_coe_to_fun : CoeFun (PartialHomeomorph α β) (λ _ => α → β) := test_sorry noncomputable -def LocalHomeomorph.symm (_e : LocalHomeomorph α β) : LocalHomeomorph β α := test_sorry +def PartialHomeomorph.symm (_e : PartialHomeomorph α β) : PartialHomeomorph β α := test_sorry -@[mfld_simps] lemma LocalHomeomorph.left_inv (e : LocalHomeomorph α β) {x : α} +@[mfld_simps] lemma PartialHomeomorph.left_inv (e : PartialHomeomorph α β) {x : α} (_h : x ∈ e.toLocalEquiv.source) : e.symm (e x) = x := test_sorry -@[mfld_simps] theorem LocalHomeomorph.symm_to_LocalEquiv (e : LocalHomeomorph α β) : +@[mfld_simps] theorem PartialHomeomorph.symm_to_LocalEquiv (e : PartialHomeomorph α β) : e.symm.toLocalEquiv = e.toLocalEquiv.symm := test_sorry -@[mfld_simps] lemma LocalHomeomorph.coe_coe (e : LocalHomeomorph α β) : +@[mfld_simps] lemma PartialHomeomorph.coe_coe (e : PartialHomeomorph α β) : (e.toLocalEquiv : α → β) = e := test_sorry -@[mfld_simps] lemma LocalHomeomorph.coe_coe_symm (e : LocalHomeomorph α β) : +@[mfld_simps] lemma PartialHomeomorph.coe_coe_symm (e : PartialHomeomorph α β) : (e.toLocalEquiv.symm : β → α) = (e.symm : β → α) := test_sorry @@ -81,7 +81,7 @@ example (e : LocalEquiv α β) (e' : LocalEquiv β γ) : example (e : LocalEquiv α β) : (e.trans e.symm).source = e.source := by mfld_set_tac -example (s : Set α) (f : LocalHomeomorph α β) : +example (s : Set α) (f : PartialHomeomorph α β) : f.symm.toLocalEquiv.source ∩ (f.toLocalEquiv.target ∩ Set.preimage f.symm s) = f.symm.toLocalEquiv.source ∩ Set.preimage f.symm s := by mfld_set_tac @@ -89,9 +89,9 @@ example {I : ModelWithCorners 𝕜 E H} {I' : ModelWithCorners 𝕜 E' H'} {I'' : ModelWithCorners 𝕜 E'' H''} - (e₁ : LocalHomeomorph M H) - (e₂ : LocalHomeomorph M' H') - (e₃ : LocalHomeomorph M'' H'') + (e₁ : PartialHomeomorph M H) + (e₂ : PartialHomeomorph M' H') + (e₃ : PartialHomeomorph M'' H'') {f : M → M'} {g : M' → M''} : (Set.preimage (f ∘ ((e₁.toLocalEquiv.trans I.toLocalEquiv).symm))