From 4e53e9600604949d13355adebbcbcd280d22231c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 13 Nov 2023 15:00:36 -0800 Subject: [PATCH 001/203] doc --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 34 ++++++++++++++++++++ 1 file changed, 34 insertions(+) create mode 100644 Mathlib/Geometry/Manifold/IntegralCurve.lean diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean new file mode 100644 index 0000000000000..84308d057995f --- /dev/null +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2023 Winston Yin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Winston Yin +-/ +import Mathlib.Analysis.ODE.PicardLindelof +import Mathlib.Geometry.Manifold.ContMDiff +import Mathlib.Geometry.Manifold.MFDeriv + +/-! +# Integral curves of vector fields on a manifold + +For any continuously differentiable vector field on a manifold `M` and any chosen non-boundary point +`x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` +at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. + +As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold +without boundary. + +## Implementation details + +Since there is already an ODE solution existence theorem +`IsPicardLindelof.exists_forall_hasDerivWithinAt_Icc_eq`, the bulk of this file is to convert +statements about manifolds to statements about the model space. This comes in a few steps: +1. Express the smoothness of the vector field `v` in a single fixed chart around the starting point +`x₀`. +2. Use the ODE solution existence theorem to obtain a curve `γ : ℝ → M` whose derivative coincides +with the vector field (stated in the local chart around `x₀`). +3. Same as 2 but now stated in the local chart around `γ t`, which is how `cont_mdiff` is defined. + +## Tags + +integral curve, vector field +-/ From f14b3780672be711e46253b1059b155bddec21ed Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 12:15:57 -0800 Subject: [PATCH 002/203] compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 148 +++++++++++++++++-- 1 file changed, 135 insertions(+), 13 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 84308d057995f..c0af3c404f047 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Winston Yin -/ import Mathlib.Analysis.ODE.PicardLindelof -import Mathlib.Geometry.Manifold.ContMDiff -import Mathlib.Geometry.Manifold.MFDeriv +-- import Mathlib.Geometry.Manifold.ContMDiff +-- import Mathlib.Geometry.Manifold.MFDeriv +import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection /-! # Integral curves of vector fields on a manifold @@ -17,18 +18,139 @@ at `t` coincides with the vector field at `γ t` for all `t` within an open inte As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold without boundary. -## Implementation details - -Since there is already an ODE solution existence theorem -`IsPicardLindelof.exists_forall_hasDerivWithinAt_Icc_eq`, the bulk of this file is to convert -statements about manifolds to statements about the model space. This comes in a few steps: -1. Express the smoothness of the vector field `v` in a single fixed chart around the starting point -`x₀`. -2. Use the ODE solution existence theorem to obtain a curve `γ : ℝ → M` whose derivative coincides -with the vector field (stated in the local chart around `x₀`). -3. Same as 2 but now stated in the local chart around `γ t`, which is how `cont_mdiff` is defined. - ## Tags integral curve, vector field -/ + +open scoped Manifold +scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x +scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) + +section + +variable + {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] + {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] + {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} + {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] + +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target + +lemma ModelWithCorners.Boundaryless.isOpen_target + [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by +rw [extChartAt_target, ModelWithCorners.Boundaryless.range_eq_univ, Set.inter_univ] +exact (ModelWithCorners.continuous_symm _).isOpen_preimage _ (LocalHomeomorph.open_target _) + +lemma ModelWithCorners.Boundaryless.isInteriorPoint + [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by +rw [ModelWithCorners.IsInteriorPoint, + IsOpen.interior_eq ModelWithCorners.Boundaryless.isOpen_target] +exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) + +end + +variable + {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] + {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] + {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners ℝ E' H'} + {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] + {v : (x : M) → TangentSpace I x} {x₀ : M} + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) + +set_option trace.Meta.Tactic.simp.rewrite true + +/-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the + manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` + at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ +theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : + ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → + HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := by +rw [contMDiffAt_iff] at hv +obtain ⟨_, hv⟩ := hv +have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) +· rw [mem_nhds_iff] + refine ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), + isOpen_interior, hx⟩ +obtain ⟨ε1, hε1, f, hf1, hf2⟩ := + exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) +rw [←Real.ball_eq_Ioo] at hf2 +-- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, +-- f t ∈ interior (extChartAt I x₀).target +have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt +rw [continuousAt_def, hf1] at hcont +have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ +· apply hcont + exact IsOpen.mem_nhds isOpen_interior hx +rw [Metric.mem_nhds_iff] at hnhds +obtain ⟨ε2, hε2, hf3⟩ := hnhds +simp_rw [Set.subset_def, Set.mem_preimage] at hf3 +-- prove the theorem +refine' ⟨min ε1 ε2, lt_min hε1 hε2, (extChartAt I x₀).symm ∘ f, _, _⟩ +· apply Eq.symm + rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] +intros t ht +rw [←Real.ball_eq_Ioo] at ht +have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) +have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) +have h : HasDerivAt f + ((fderivWithin ℝ + ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (Set.range I) + (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (v ((extChartAt I x₀).symm (f t)))) + t := hf2 t ht1 +have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset +rw [HasMFDerivAt] +use ContinuousAt.comp + (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) +apply HasDerivWithinAt.hasFDerivWithinAt +rw [modelWithCornersSelf_coe, Set.range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, + writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, + LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] +-- v -> identity v +rw [←tangentBundleCore_coordChange_achart] at h +have hvsub : v ((extChartAt I x₀).symm (f t)) = (tangentBundleCore I M).coordChange + (achart H x₀) (achart H ((extChartAt I x₀).symm (f t))) ((extChartAt I x₀).symm (f t)) + ((tangentBundleCore I M).coordChange (achart H ((extChartAt I x₀).symm (f t))) (achart H x₀) + ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) +· rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] + · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] + exact mem_extChartAt_source .. + · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, + ←extChartAt_source I, ←extChartAt_source I, Set.inter_comm, ←Set.inter_assoc, Set.inter_self] + constructor + · exact mem_extChartAt_source .. + · rw [←Set.mem_preimage] + apply Set.mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) + rw [LocalEquiv.symm_source] + exact hf3' +rw [hvsub] +apply HasFDerivAt.comp_hasDerivAt _ _ h +rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWithin_of_mem_nhds] +· apply DifferentiableAt.hasFDerivAt + apply MDifferentiableAt.differentiableAt + apply MDifferentiableAt.comp (I' := I) + · exact (contMDiffAt_extChartAt (n := 1)).mdifferentiableAt (le_refl _) + · apply MDifferentiableOn.mdifferentiableAt + ((contMDiffOn_extChartAt_symm (n := 1) _).mdifferentiableOn (le_refl _)) + rw [mem_nhds_iff] + exact ⟨interior (extChartAt I x₀).target, interior_subset, isOpen_interior, hf3 _ ht2⟩ +· rw [mem_nhds_iff] + refine ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ + +/-- For any continuously differentiable vector field defined on a manifold without boundary and any + chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the + tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open + interval around `t₀`. -/ +lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : + ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → + HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := +exists_integralCurve_of_contMDiff_tangent_section hv _ ModelWithCorners.Boundaryless.isInteriorPoint From 22b4f1f2b759157837aa9e4be5242dbd3d58cc39 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 12:28:44 -0800 Subject: [PATCH 003/203] indent --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c0af3c404f047..d7a7775bb4b26 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -41,12 +41,12 @@ variable def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target lemma ModelWithCorners.Boundaryless.isOpen_target - [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by + [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by rw [extChartAt_target, ModelWithCorners.Boundaryless.range_eq_univ, Set.inter_univ] exact (ModelWithCorners.continuous_symm _).isOpen_preimage _ (LocalHomeomorph.open_target _) lemma ModelWithCorners.Boundaryless.isInteriorPoint - [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by + [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq ModelWithCorners.Boundaryless.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) @@ -63,8 +63,6 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) -set_option trace.Meta.Tactic.simp.rewrite true - /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ @@ -153,4 +151,5 @@ rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWit lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := -exists_integralCurve_of_contMDiff_tangent_section hv _ ModelWithCorners.Boundaryless.isInteriorPoint + exists_integralCurve_of_contMDiff_tangent_section hv _ + ModelWithCorners.Boundaryless.isInteriorPoint From 074717457dfdad8f5f4189dcfe175e74e127e350 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 12:32:10 -0800 Subject: [PATCH 004/203] import, indent --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index d7a7775bb4b26..cf5525de7e235 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Winston Yin -/ import Mathlib.Analysis.ODE.PicardLindelof --- import Mathlib.Geometry.Manifold.ContMDiff --- import Mathlib.Geometry.Manifold.MFDeriv -import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection +import Mathlib.Geometry.Manifold.ContMDiff +import Mathlib.Geometry.Manifold.MFDeriv /-! # Integral curves of vector fields on a manifold @@ -149,7 +148,7 @@ rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWit tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : - ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → - HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := + ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → + HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := exists_integralCurve_of_contMDiff_tangent_section hv _ ModelWithCorners.Boundaryless.isInteriorPoint From 23d691fb2ce0556efccd07df94f9077306398ae2 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 12:33:22 -0800 Subject: [PATCH 005/203] docstring --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index cf5525de7e235..33c3a1f37d989 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -37,6 +37,8 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] +/-- An interior point of a manifold is a point whose image in the model vector space is in the +interior of the chart's target. -/ def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target lemma ModelWithCorners.Boundaryless.isOpen_target From e8e7338eeb8aef7aacb6263ab48f0cd2436a9ce0 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 12:36:08 -0800 Subject: [PATCH 006/203] Mathlib.lean --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 2364513137a8d..3cee19487891b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2056,6 +2056,7 @@ import Mathlib.Geometry.Manifold.Diffeomorph import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.Instances.Sphere import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra +import Mathlib.Geometry.Manifold.IntegralCurve import Mathlib.Geometry.Manifold.LocalInvariantProperties import Mathlib.Geometry.Manifold.MFDeriv import Mathlib.Geometry.Manifold.Metrizable From ccc8f77d6f032a2014e268e07152355efbc81e71 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 17 Nov 2023 15:13:59 -0800 Subject: [PATCH 007/203] lake manifest --- lake-manifest.json | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index d5b9b3e725bac..d4d48eaca4083 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,7 +4,7 @@ [{"git": {"url": "https://github.com/leanprover/std4", "subDir?": null, - "rev": "fb07d160aff0e8bdf403a78a5167fc7acf9c8227", + "rev": "d3049643f6dded69eb7ce8124796cb1ec8df8840", "opts": {}, "name": "std", "inputRev?": "main", @@ -12,7 +12,7 @@ {"git": {"url": "https://github.com/leanprover-community/quote4", "subDir?": null, - "rev": "511eb96eca98a7474683b8ae55193a7e7c51bc39", + "rev": "d3a1d25f3eba0d93a58d5d3d027ffa78ece07755", "opts": {}, "name": "Qq", "inputRev?": "master", @@ -20,7 +20,7 @@ {"git": {"url": "https://github.com/leanprover-community/aesop", "subDir?": null, - "rev": "cb87803274405db79ec578fc07c4730c093efb90", + "rev": "bf5ab42a58e71de7ebad399ce3f90d29aae7fca9", "opts": {}, "name": "aesop", "inputRev?": "master", From 720ec35dd7b9350742a0b7727d458f922f0b92dc Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 18 Nov 2023 00:50:35 -0800 Subject: [PATCH 008/203] removed temporary notation --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 33c3a1f37d989..99ae085e17961 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -23,8 +23,6 @@ integral curve, vector field -/ open scoped Manifold -scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x -scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) section From 25f1065e058fcf75529492d81db99a3dafd313c4 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 18 Nov 2023 02:46:35 -0800 Subject: [PATCH 009/203] integralCurveAt, change PicardLindelof statement order --- Mathlib/Analysis/ODE/PicardLindelof.lean | 13 +++++++------ Mathlib/Geometry/Manifold/IntegralCurve.lean | 17 +++++++++-------- 2 files changed, 16 insertions(+), 14 deletions(-) diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index 3916ef0f0ec84..25a8d48d5984c 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -423,18 +423,19 @@ variable [CompleteSpace E] /-- A time-independent, continuously differentiable ODE admits a solution in some open interval. -/ theorem exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt (hv : ContDiffAt ℝ 1 v x₀) : - ∃ ε > (0 : ℝ), - ∃ f : ℝ → E, f t₀ = x₀ ∧ ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt f (v (f t)) t := by + ∃ f : ℝ → E, f t₀ = x₀ ∧ + ∃ ε > (0 : ℝ), ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt f (v (f t)) t := by obtain ⟨ε, hε, _, _, _, hpl⟩ := exists_isPicardLindelof_const_of_contDiffAt t₀ hv obtain ⟨f, hf1, hf2⟩ := hpl.exists_forall_hasDerivWithinAt_Icc_eq x₀ - exact ⟨ε, hε, f, hf1, fun t ht => + exact ⟨f, hf1, ε, hε, fun t ht => (hf2 t (Ioo_subset_Icc_self ht)).hasDerivAt (Icc_mem_nhds ht.1 ht.2)⟩ #align exists_forall_deriv_at_Ioo_eq_of_cont_diff_on_nhds exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt /-- A time-independent, continuously differentiable ODE admits a solution in some open interval. -/ theorem exists_forall_hasDerivAt_Ioo_eq_of_contDiff (hv : ContDiff ℝ 1 v) : - ∃ ε > (0 : ℝ), ∃ f : ℝ → E, f t₀ = x₀ ∧ ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt f (v (f t)) t := - let ⟨ε, hε, f, hf1, hf2⟩ := + ∃ f : ℝ → E, f t₀ = x₀ ∧ + ∃ ε > (0 : ℝ), ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt f (v (f t)) t := + let ⟨f, hf1, ε, hε, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ hv.contDiffAt - ⟨ε, hε, f, hf1, fun _ h => hf2 _ h⟩ + ⟨f, hf1, ε, hε, fun _ h => hf2 _ h⟩ #align exists_forall_deriv_at_Ioo_eq_of_cont_diff exists_forall_hasDerivAt_Ioo_eq_of_contDiff diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 99ae085e17961..a7bfb3e6fa13d 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -62,6 +62,10 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) +def integralCurveAt (γ : ℝ → M) := γ t₀ = x₀ ∧ + ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → + HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) + /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ @@ -75,16 +79,15 @@ have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) refine ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hx⟩ -obtain ⟨ε1, hε1, f, hf1, hf2⟩ := +obtain ⟨f, ε1, hε1, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) rw [←Real.ball_eq_Ioo] at hf2 -- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, -- f t ∈ interior (extChartAt I x₀).target have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont -have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ -· apply hcont - exact IsOpen.mem_nhds isOpen_interior hx +have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := + hcont _ (IsOpen.mem_nhds isOpen_interior hx) rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [Set.subset_def, Set.mem_preimage] at hf3 @@ -97,10 +100,8 @@ rw [←Real.ball_eq_Ioo] at ht have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) have h : HasDerivAt f - ((fderivWithin ℝ - ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (Set.range I) - (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (Set.range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) (v ((extChartAt I x₀).symm (f t)))) t := hf2 t ht1 have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset From d4139b6ad653f5c58ad14befddb008a939bec477 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 18 Nov 2023 02:59:43 -0800 Subject: [PATCH 010/203] compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 14 ++++++-------- 1 file changed, 6 insertions(+), 8 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a7bfb3e6fa13d..873dcb4f4033d 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -62,16 +62,15 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) -def integralCurveAt (γ : ℝ → M) := γ t₀ = x₀ ∧ - ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → +def HasIntegralCurveAt (v : (x : M) → TangentSpace I x) (γ : ℝ → M) (x₀ : M) (t₀ : ℝ) := + γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : - ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → - HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := by + ∃ (γ : ℝ → M), HasIntegralCurveAt v γ x₀ t₀ := by rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) @@ -79,7 +78,7 @@ have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) refine ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hx⟩ -obtain ⟨f, ε1, hε1, hf1, hf2⟩ := +obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) rw [←Real.ball_eq_Ioo] at hf2 -- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, @@ -92,7 +91,7 @@ rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [Set.subset_def, Set.mem_preimage] at hf3 -- prove the theorem -refine' ⟨min ε1 ε2, lt_min hε1 hε2, (extChartAt I x₀).symm ∘ f, _, _⟩ +refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] intros t ht @@ -149,7 +148,6 @@ rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWit tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : - ∃ ε > (0 : ℝ), ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → - HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) := + ∃ (γ : ℝ → M), HasIntegralCurveAt v γ x₀ t₀ := exists_integralCurve_of_contMDiff_tangent_section hv _ ModelWithCorners.Boundaryless.isInteriorPoint From 4b57dfe30091d387abb6270893025b2571aafc9c Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 18 Nov 2023 19:30:12 +0100 Subject: [PATCH 011/203] Sketch a section on interior and boundary of manifolds. --- Mathlib.lean | 1 + Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 +- .../Geometry/Manifold/InteriorBoundary.lean | 112 ++++++++++++++++++ 3 files changed, 114 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Geometry/Manifold/InteriorBoundary.lean diff --git a/Mathlib.lean b/Mathlib.lean index 261787fddc9d4..245083a6465b5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2070,6 +2070,7 @@ import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.Instances.Sphere import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra import Mathlib.Geometry.Manifold.IntegralCurve +import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalInvariantProperties import Mathlib.Geometry.Manifold.MFDeriv import Mathlib.Geometry.Manifold.Metrizable diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 873dcb4f4033d..abf8e43f84676 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Winston Yin -/ import Mathlib.Analysis.ODE.PicardLindelof -import Mathlib.Geometry.Manifold.ContMDiff +import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.MFDeriv /-! @@ -35,10 +35,6 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -/-- An interior point of a manifold is a point whose image in the model vector space is in the -interior of the chart's target. -/ -def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target - lemma ModelWithCorners.Boundaryless.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by rw [extChartAt_target, ModelWithCorners.Boundaryless.range_eq_univ, Set.inter_univ] diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean new file mode 100644 index 0000000000000..18d5a4e2f5755 --- /dev/null +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2023 Michael Rothgang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Rothgang, Winston Yin +-/ + +import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners + +-- FIXME: should this be its own file or go in SmoothManifoldWithCorners? +-- the latter is already huge, or its own file - move other results about boundaryless here? + +/-! +## Interior and boundary of a smooth manifold + +Define the interior and boundary of a smooth manifold. + +**Main definitions** +- **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, + `φ x` is an interior point of `φ.target`. +- **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, for `φ` being the preferred chart at `x`, +- **SmoothManifoldWithBoundary.interior I M** is the **interior** of `M`, the set of its interior points. +- **SmoothManifoldWithBoundary.boundary I M** is the **boundary** of `M`, the set of its boundary points. + +**Main results**: +- `xxx`: M is the union of its interior and boundary +- `yyy`: interior I M is open + +**TODO**: show that +- interior I M is a manifold without boundary + (need to upgrade the above; map the charts from an open ball to entire ℝ^n) +- the boundary is a submanifold of codimension 1 (once mathlib has submanifolds) + +## Tags +manifold, interior, boundary +-/ + +-- Let M be a smooth manifold with corners over the pair (I,E). +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] + +/-- `p ∈ M` is an interior point of a smooth manifold `M` iff +for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +-- Otherwise, it is a boundary point. + +def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target + +namespace SmoothManifoldWithCorners +-- FIXME(MR): can I enable dot notation, like `M.interior I` or so? + +variable (I M) in +/-- The **interior** of a manifold `M` is the set of its interior points. -/ +protected def interior : Set M := { x : M | I.IsInteriorPoint x} + +variable (I M) in +/-- The **boundary** of a manifold `M` is the set of its boundary points. -/ +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} + +-- FIXME(MR): find a better wording for the next two docstrings + +/-- Whether `x` is an interior point can equivalently be described by any chart + whose source contains `x`. -/ +lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} + (hx : x ∈ e.source) : I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by + sorry + +/-- Whether `x` is a boundary point of `M` can equivalently be described by any chart +whose source contains `x`. -/ +lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} + (hx : x ∈ e.source) : I.IsBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by + sorry + +-- underlying lemma: if e and e' are two charts, +-- the transition map maps interior points to interior points and boundary to boundary +lemma foobar {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} + (hx : x ∈ e.source ∩ e'.source) : + (e.extend I) x ∈ interior (e.extend I).target ↔ + (e'.extend I) x ∈ interior (e'.extend I).target := sorry + +lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} + (hx : x ∈ e.source ∩ e'.source) : + (e.extend I) x ∈ frontier (e.extend I).target ↔ + (e'.extend I) x ∈ frontier (e'.extend I).target := sorry + +-- more abstract result: a local homeomorphism maps interior to interior and boundary to boundary + +/-- Every point is either an interior or a boundary point. -/ +lemma bar (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by + set e := extChartAt I x + set y := extChartAt I x x + have : IsClosed I.target := I.target_eq ▸ (I.closed_range) + -- TODO: this should be obvious now! + have : IsClosed e.target := sorry + have : y ∈ interior e.target ∪ frontier e.target := by + rw [← closure_eq_interior_union_frontier (e.target), this.closure_eq] + exact mem_extChartAt_target I x + exact (Set.mem_union y _ _).mp this + +-- Decomposition of M into interior and boundary. TODO: find nice name! +lemma foo : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = M := by + -- FIXME: should follow from lemma `bar` + sorry + +/-- Ihe interior of a smooth manifold is an open subset. -/ +lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by + -- use `isInteriorPoint_iff` + sorry + +-- interior I M is a smooth manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) +end SmoothManifoldWithCorners From e9d3546c78635933c1f7a45ed4387b2541d1df49 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sun, 19 Nov 2023 10:24:19 +0100 Subject: [PATCH 012/203] More progress. Tweaks, polish and fleshing out some sorries. - use lowerCamelCase for our definitions, per the naming convention. - sketch how to prove a few more sorries. --- .../Geometry/Manifold/InteriorBoundary.lean | 89 +++++++++++++------ 1 file changed, 63 insertions(+), 26 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 18d5a4e2f5755..a0c15d0858ca4 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -10,18 +10,18 @@ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners -- the latter is already huge, or its own file - move other results about boundaryless here? /-! -## Interior and boundary of a smooth manifold +# Interior and boundary of a smooth manifold Define the interior and boundary of a smooth manifold. -**Main definitions** +## Main definitions - **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. - **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, for `φ` being the preferred chart at `x`, - **SmoothManifoldWithBoundary.interior I M** is the **interior** of `M`, the set of its interior points. - **SmoothManifoldWithBoundary.boundary I M** is the **boundary** of `M`, the set of its boundary points. -**Main results**: +## Main results - `xxx`: M is the union of its interior and boundary - `yyy`: interior I M is open @@ -34,7 +34,7 @@ Define the interior and boundary of a smooth manifold. manifold, interior, boundary -/ --- Let M be a smooth manifold with corners over the pair (I,E). +-- Let `M` be a smooth manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} @@ -42,52 +42,52 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- `p ∈ M` is an interior point of a smooth manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ -def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target --- Otherwise, it is a boundary point. +def ModelWithCorners.isInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target -def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target +/-- `p ∈ M` is a boundary point of a smooth manifold `M` iff +for `φ` being the preferred chart at `x`, `φ x` is a boundary point of `φ.target`. -/ +def ModelWithCorners.isBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target namespace SmoothManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ -protected def interior : Set M := { x : M | I.IsInteriorPoint x} +protected def interior : Set M := { x : M | I.isInteriorPoint x} variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ -protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} +protected def boundary : Set M := { x : M | I.isBoundaryPoint x} -- FIXME(MR): find a better wording for the next two docstrings - /-- Whether `x` is an interior point can equivalently be described by any chart whose source contains `x`. -/ lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by + (hx : x ∈ e.source) : I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by sorry /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.IsBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by + (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by sorry --- underlying lemma: if e and e' are two charts, --- the transition map maps interior points to interior points and boundary to boundary +/-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ lemma foobar {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} - (hx : x ∈ e.source ∩ e'.source) : - (e.extend I) x ∈ interior (e.extend I).target ↔ - (e'.extend I) x ∈ interior (e'.extend I).target := sorry + (hx : x ∈ e.source ∩ e'.source) : + (e.extend I) x ∈ interior (e.extend I).target ↔ + (e'.extend I) x ∈ interior (e'.extend I).target := sorry +/-- If `e` and `e'` are two charts, the transition map maps boundary points to boundary points. -/ lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} - (hx : x ∈ e.source ∩ e'.source) : - (e.extend I) x ∈ frontier (e.extend I).target ↔ + (hx : x ∈ e.source ∩ e'.source) : + (e.extend I) x ∈ frontier (e.extend I).target ↔ (e'.extend I) x ∈ frontier (e'.extend I).target := sorry -- more abstract result: a local homeomorphism maps interior to interior and boundary to boundary -/-- Every point is either an interior or a boundary point. -/ -lemma bar (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by +/-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isBoundaryPoint x := by set e := extChartAt I x set y := extChartAt I x x have : IsClosed I.target := I.target_eq ▸ (I.closed_range) @@ -98,14 +98,51 @@ lemma bar (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by exact mem_extChartAt_target I x exact (Set.mem_union y _ _).mp this --- Decomposition of M into interior and boundary. TODO: find nice name! -lemma foo : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = M := by - -- FIXME: should follow from lemma `bar` - sorry +/-- A smooth manifold decomposes into interior and boundary. -/ +lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ + (SmoothManifoldWithCorners.boundary I M) = (Set.univ : Set M) := by + apply le_antisymm + · exact fun x _ ↦ trivial + · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x + +-- this should be in mathlib; cannot find it now! +lemma aux {s : Set M} : IsOpen s ↔ ∀ x, x ∈ s → ∃ t : Set M, x ∈ t ∧ t ⊆ s ∧ IsOpen t := by + constructor + · intro h x hx + rcases mem_nhds_iff.mp (h.mem_nhds hx) with ⟨t, hts, htopen, hxt⟩ + use t + · sorry -- exercise for now + +lemma aux' {s : Set M} (h : ∀ x, x ∈ s → ∃ t : Set M, x ∈ t ∧ t ⊆ s ∧ IsOpen t) : IsOpen s := + aux.mpr h /-- Ihe interior of a smooth manifold is an open subset. -/ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by - -- use `isInteriorPoint_iff` + apply aux' + intro x hx + -- Consider the preferred chart at `x`. + let e := chartAt H x + -- Its extended chart has open interior (should be easy). + let U := interior (e.extend I).target + have : IsOpen U := sorry + + -- For all `y ∈ e.source`, `y` is an interior point iff its image lies in `U`. + have : ∀ y, y ∈ e.source → (I.isInteriorPoint y ↔ (e.extend I) y ∈ U) := + fun y hy ↦ isInteriorPoint_iff (chart_mem_atlas H x) hy + use (e.extend I) ⁻¹' U + refine ⟨hx, ?_, ?_⟩ + · sorry + · sorry -- this one I'm not 100% sure about + +/-- The boundary of any extended chart has empty interior. -/ +-- xxx: do I need that e is in the atlas? I think not; not double-checked. +-- xxx: is this lemma fully true, or do I need a stronger definition of boundary? +lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : + interior (frontier (e.extend I).target) = ∅ := sorry + +/-- The boundary of a smooth manifold has empty interior. -/ +lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by + -- use isBoundaryPoint_iff and the previous lemma; similar to `interior_isOpen` sorry -- interior I M is a smooth manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) From cbd22f4410bc2235941f3569cf88ce56e7b679c7 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sun, 19 Nov 2023 12:45:20 +0100 Subject: [PATCH 013/203] One more lemma. --- .../Geometry/Manifold/InteriorBoundary.lean | 54 +++++++++++-------- 1 file changed, 33 insertions(+), 21 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index a0c15d0858ca4..494ab7800f515 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -34,6 +34,8 @@ Define the interior and boundary of a smooth manifold. manifold, interior, boundary -/ +open Set + -- Let `M` be a smooth manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] @@ -59,19 +61,6 @@ variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ protected def boundary : Set M := { x : M | I.isBoundaryPoint x} --- FIXME(MR): find a better wording for the next two docstrings -/-- Whether `x` is an interior point can equivalently be described by any chart - whose source contains `x`. -/ -lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by - sorry - -/-- Whether `x` is a boundary point of `M` can equivalently be described by any chart -whose source contains `x`. -/ -lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by - sorry - /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ lemma foobar {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} (hx : x ∈ e.source ∩ e'.source) : @@ -86,6 +75,19 @@ lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ -- more abstract result: a local homeomorphism maps interior to interior and boundary to boundary +-- FIXME(MR): find a better wording for the next two docstrings +/-- Whether `x` is an interior point can equivalently be described by any chart + whose source contains `x`. -/ +lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} + (hx : x ∈ e.source) : I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by + sorry + +/-- Whether `x` is a boundary point of `M` can equivalently be described by any chart +whose source contains `x`. -/ +lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} + (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by + sorry + /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isBoundaryPoint x := by set e := extChartAt I x @@ -96,11 +98,11 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isB have : y ∈ interior e.target ∪ frontier e.target := by rw [← closure_eq_interior_union_frontier (e.target), this.closure_eq] exact mem_extChartAt_target I x - exact (Set.mem_union y _ _).mp this + exact (mem_union y _ _).mp this /-- A smooth manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ - (SmoothManifoldWithCorners.boundary I M) = (Set.univ : Set M) := by + (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by apply le_antisymm · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x @@ -124,17 +126,27 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by let e := chartAt H x -- Its extended chart has open interior (should be easy). let U := interior (e.extend I).target - have : IsOpen U := sorry - + have hU : IsOpen U := isOpen_interior -- For all `y ∈ e.source`, `y` is an interior point iff its image lies in `U`. + -- FIXME: extract this into a separate lemma? have : ∀ y, y ∈ e.source → (I.isInteriorPoint y ↔ (e.extend I) y ∈ U) := fun y hy ↦ isInteriorPoint_iff (chart_mem_atlas H x) hy - use (e.extend I) ⁻¹' U - refine ⟨hx, ?_, ?_⟩ - · sorry - · sorry -- this one I'm not 100% sure about + use (e.extend I).source ∩ (e.extend I) ⁻¹' U + refine ⟨?_, ?_, ?_⟩ + · have : x ∈ (e.extend I).source := by + rw [e.extend_source] + exact mem_chart_source H x + exact mem_inter this hx + · intro y hy + rw [e.extend_source] at hy + apply (this y (mem_of_mem_inter_left hy)).mpr + have : y ∈ (LocalHomeomorph.extend e I) ⁻¹' U := mem_of_mem_inter_right hy + exact this + · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) hU /-- The boundary of any extended chart has empty interior. -/ +-- NB: this is *false* for any set instead of (e.extend I).target: +-- for instance, $ℚ ⊆ ℝ$ has frontiert ℝ (ℚ is dense in ℝ and ℚ has empty interior). -- xxx: do I need that e is in the atlas? I think not; not double-checked. -- xxx: is this lemma fully true, or do I need a stronger definition of boundary? lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : From 0dbf3b08342b67d0e093cbc6a8046e7367933f5d Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sun, 19 Nov 2023 13:03:50 +0100 Subject: [PATCH 014/203] One lemma was in mathlib. :-) --- .../Geometry/Manifold/InteriorBoundary.lean | 21 +++++-------------- 1 file changed, 5 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 494ab7800f515..137daee7beb09 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -107,20 +107,9 @@ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x --- this should be in mathlib; cannot find it now! -lemma aux {s : Set M} : IsOpen s ↔ ∀ x, x ∈ s → ∃ t : Set M, x ∈ t ∧ t ⊆ s ∧ IsOpen t := by - constructor - · intro h x hx - rcases mem_nhds_iff.mp (h.mem_nhds hx) with ⟨t, hts, htopen, hxt⟩ - use t - · sorry -- exercise for now - -lemma aux' {s : Set M} (h : ∀ x, x ∈ s → ∃ t : Set M, x ∈ t ∧ t ⊆ s ∧ IsOpen t) : IsOpen s := - aux.mpr h - /-- Ihe interior of a smooth manifold is an open subset. -/ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by - apply aux' + apply isOpen_iff_forall_mem_open.mpr intro x hx -- Consider the preferred chart at `x`. let e := chartAt H x @@ -133,16 +122,16 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by fun y hy ↦ isInteriorPoint_iff (chart_mem_atlas H x) hy use (e.extend I).source ∩ (e.extend I) ⁻¹' U refine ⟨?_, ?_, ?_⟩ - · have : x ∈ (e.extend I).source := by - rw [e.extend_source] - exact mem_chart_source H x - exact mem_inter this hx · intro y hy rw [e.extend_source] at hy apply (this y (mem_of_mem_inter_left hy)).mpr have : y ∈ (LocalHomeomorph.extend e I) ⁻¹' U := mem_of_mem_inter_right hy exact this · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) hU + · have : x ∈ (e.extend I).source := by + rw [e.extend_source] + exact mem_chart_source H x + exact mem_inter this hx /-- The boundary of any extended chart has empty interior. -/ -- NB: this is *false* for any set instead of (e.extend I).target: From a978f02905096173a07a303a3d9a06cdec47d11a Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sun, 19 Nov 2023 19:55:03 +0100 Subject: [PATCH 015/203] Fix the build; mild clean-ups. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 18 ++++++++---------- 1 file changed, 8 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index abf8e43f84676..4cb393a7e0e75 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -35,16 +35,14 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -lemma ModelWithCorners.Boundaryless.isOpen_target - [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by -rw [extChartAt_target, ModelWithCorners.Boundaryless.range_eq_univ, Set.inter_univ] -exact (ModelWithCorners.continuous_symm _).isOpen_preimage _ (LocalHomeomorph.open_target _) +lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by + rw [extChartAt_target, I.range_eq_univ, Set.inter_univ] + exact I.continuous_symm.isOpen_preimage _ (LocalHomeomorph.open_target _) -lemma ModelWithCorners.Boundaryless.isInteriorPoint - [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by -rw [ModelWithCorners.IsInteriorPoint, - IsOpen.interior_eq ModelWithCorners.Boundaryless.isOpen_target] -exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) +lemma ModelWithCorners.Boundaryless.isInteriorPoint [I.Boundaryless] {x : M} : + I.isInteriorPoint x := by + rw [ModelWithCorners.isInteriorPoint, IsOpen.interior_eq I.isOpen_target] + exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) end @@ -65,7 +63,7 @@ def HasIntegralCurveAt (v : (x : M) → TangentSpace I x) (γ : ℝ → M) (x₀ /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : +theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.isInteriorPoint x₀) : ∃ (γ : ℝ → M), HasIntegralCurveAt v γ x₀ t₀ := by rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv From 1f6dbaad86cb96ee0c16cdda4ec30da8604aeb8b Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sun, 19 Nov 2023 20:00:56 +0100 Subject: [PATCH 016/203] Small simplifications. --- .../Geometry/Manifold/InteriorBoundary.lean | 23 ++++++++----------- 1 file changed, 9 insertions(+), 14 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 137daee7beb09..c433529f0997e 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -111,39 +111,34 @@ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by apply isOpen_iff_forall_mem_open.mpr intro x hx - -- Consider the preferred chart at `x`. + -- Consider the preferred chart at `x`. Its extended chart has open interior. let e := chartAt H x - -- Its extended chart has open interior (should be easy). let U := interior (e.extend I).target - have hU : IsOpen U := isOpen_interior -- For all `y ∈ e.source`, `y` is an interior point iff its image lies in `U`. - -- FIXME: extract this into a separate lemma? - have : ∀ y, y ∈ e.source → (I.isInteriorPoint y ↔ (e.extend I) y ∈ U) := - fun y hy ↦ isInteriorPoint_iff (chart_mem_atlas H x) hy + -- FIXME: should this be a separate lemma? use (e.extend I).source ∩ (e.extend I) ⁻¹' U refine ⟨?_, ?_, ?_⟩ · intro y hy rw [e.extend_source] at hy - apply (this y (mem_of_mem_inter_left hy)).mpr - have : y ∈ (LocalHomeomorph.extend e I) ⁻¹' U := mem_of_mem_inter_right hy - exact this - · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) hU + apply (isInteriorPoint_iff (chart_mem_atlas H x) (mem_of_mem_inter_left hy)).mpr + exact mem_of_mem_inter_right (a := e.source) hy + · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) isOpen_interior · have : x ∈ (e.extend I).source := by rw [e.extend_source] exact mem_chart_source H x exact mem_inter this hx /-- The boundary of any extended chart has empty interior. -/ --- NB: this is *false* for any set instead of (e.extend I).target: --- for instance, $ℚ ⊆ ℝ$ has frontiert ℝ (ℚ is dense in ℝ and ℚ has empty interior). +-- NB: this is *false* for any set instead of `(e.extend I).target`: +-- for instance, $ℚ ⊆ ℝ$ has frontier ℝ (ℚ is dense in ℝ and ℚ has empty interior). -- xxx: do I need that e is in the atlas? I think not; not double-checked. --- xxx: is this lemma fully true, or do I need a stronger definition of boundary? +-- xxx: is this lemma true with mathlib's current definitions? lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : interior (frontier (e.extend I).target) = ∅ := sorry /-- The boundary of a smooth manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by - -- use isBoundaryPoint_iff and the previous lemma; similar to `interior_isOpen` + -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` sorry -- interior I M is a smooth manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) From 9c586523cfea538f4a613ca3003010078ed650a5 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 12:17:58 -0800 Subject: [PATCH 017/203] has -> is, docstring, shift/stretch lemmas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 195 ++++++++++-------- .../Geometry/Manifold/InteriorBoundary.lean | 14 +- 2 files changed, 114 insertions(+), 95 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4cb393a7e0e75..5f98954c0a04b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -11,15 +11,15 @@ import Mathlib.Geometry.Manifold.MFDeriv # Integral curves of vector fields on a manifold For any continuously differentiable vector field on a manifold `M` and any chosen non-boundary point -`x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` -at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. +`x₀ : M`, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of +`γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold without boundary. ## Tags -integral curve, vector field +integral curve, vector field, local existence -/ open scoped Manifold @@ -35,13 +35,18 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := by - rw [extChartAt_target, I.range_eq_univ, Set.inter_univ] - exact I.continuous_symm.isOpen_preimage _ (LocalHomeomorph.open_target _) +lemma ModelWithCorners.isOpen_extend_target [I.Boundaryless] {e : LocalHomeomorph M H} : + IsOpen (e.extend I).target := by + rw [LocalHomeomorph.extend_target, I.range_eq_univ, Set.inter_univ] + exact I.continuous_symm.isOpen_preimage _ e.open_target -lemma ModelWithCorners.Boundaryless.isInteriorPoint [I.Boundaryless] {x : M} : - I.isInteriorPoint x := by - rw [ModelWithCorners.isInteriorPoint, IsOpen.interior_eq I.isOpen_target] +lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : + IsOpen (extChartAt I x).target := I.isOpen_extend_target + +/-- If `M` has no boundary, then every point in `M` is an interior point. -/ +lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : + I.IsInteriorPoint x := by + rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) end @@ -56,92 +61,106 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) -def HasIntegralCurveAt (v : (x : M) → TangentSpace I x) (γ : ℝ → M) (x₀ : M) (t₀ : ℝ) := +/-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means + `γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. -/ +def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) (x₀ : M) := γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) +/- +TODO: +* split the theorem below into smaller lemmas, e.g. involving HasIntegralCurveAt? +* shift and stretch theorems +-/ +example (γ : ℝ → M) (dt : ℝ) : + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ + dt) x₀ := + sorry + +example (γ : ℝ → M) (a : ℝ) : + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => a * t)) (a • v) t₀ x₀ := sorry + /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the - manifold, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` - at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.isInteriorPoint x₀) : - ∃ (γ : ℝ → M), HasIntegralCurveAt v γ x₀ t₀ := by -rw [contMDiffAt_iff] at hv -obtain ⟨_, hv⟩ := hv -have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) -· rw [mem_nhds_iff] - refine ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), - isOpen_interior, hx⟩ -obtain ⟨f, hf1, ε1, hε1, hf2⟩ := - exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) -rw [←Real.ball_eq_Ioo] at hf2 --- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, --- f t ∈ interior (extChartAt I x₀).target -have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt -rw [continuousAt_def, hf1] at hcont -have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := - hcont _ (IsOpen.mem_nhds isOpen_interior hx) -rw [Metric.mem_nhds_iff] at hnhds -obtain ⟨ε2, hε2, hf3⟩ := hnhds -simp_rw [Set.subset_def, Set.mem_preimage] at hf3 --- prove the theorem -refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ -· apply Eq.symm - rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] -intros t ht -rw [←Real.ball_eq_Ioo] at ht -have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) -have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) -have h : HasDerivAt f - ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (Set.range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) - (v ((extChartAt I x₀).symm (f t)))) - t := hf2 t ht1 -have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset -rw [HasMFDerivAt] -use ContinuousAt.comp - (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) -apply HasDerivWithinAt.hasFDerivWithinAt -rw [modelWithCornersSelf_coe, Set.range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, - writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, - LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] --- v -> identity v -rw [←tangentBundleCore_coordChange_achart] at h -have hvsub : v ((extChartAt I x₀).symm (f t)) = (tangentBundleCore I M).coordChange - (achart H x₀) (achart H ((extChartAt I x₀).symm (f t))) ((extChartAt I x₀).symm (f t)) - ((tangentBundleCore I M).coordChange (achart H ((extChartAt I x₀).symm (f t))) (achart H x₀) - ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) -· rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] - · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] - exact mem_extChartAt_source .. - · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, - ←extChartAt_source I, ←extChartAt_source I, Set.inter_comm, ←Set.inter_assoc, Set.inter_self] - constructor - · exact mem_extChartAt_source .. - · rw [←Set.mem_preimage] - apply Set.mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) - rw [LocalEquiv.symm_source] - exact hf3' -rw [hvsub] -apply HasFDerivAt.comp_hasDerivAt _ _ h -rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWithin_of_mem_nhds] -· apply DifferentiableAt.hasFDerivAt - apply MDifferentiableAt.differentiableAt - apply MDifferentiableAt.comp (I' := I) - · exact (contMDiffAt_extChartAt (n := 1)).mdifferentiableAt (le_refl _) - · apply MDifferentiableOn.mdifferentiableAt - ((contMDiffOn_extChartAt_symm (n := 1) _).mdifferentiableOn (le_refl _)) - rw [mem_nhds_iff] - exact ⟨interior (extChartAt I x₀).target, interior_subset, isOpen_interior, hf3 _ ht2⟩ -· rw [mem_nhds_iff] - refine ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ + manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector + of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around + `t₀`.-/ +theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : + ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) + · rw [mem_nhds_iff] + refine ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), + isOpen_interior, hx⟩ + obtain ⟨f, hf1, ε1, hε1, hf2⟩ := + exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) + rw [←Real.ball_eq_Ioo] at hf2 + -- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, + -- f t ∈ interior (extChartAt I x₀).target + have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt + rw [continuousAt_def, hf1] at hcont + have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := + hcont _ (IsOpen.mem_nhds isOpen_interior hx) + rw [Metric.mem_nhds_iff] at hnhds + obtain ⟨ε2, hε2, hf3⟩ := hnhds + simp_rw [Set.subset_def, Set.mem_preimage] at hf3 + -- prove the theorem + refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ + · apply Eq.symm + rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] + intros t ht + rw [←Real.ball_eq_Ioo] at ht + have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) + have h : HasDerivAt f + ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (Set.range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (v ((extChartAt I x₀).symm (f t)))) + t := hf2 t ht1 + have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset + rw [HasMFDerivAt] + use ContinuousAt.comp + (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) + apply HasDerivWithinAt.hasFDerivWithinAt + rw [modelWithCornersSelf_coe, Set.range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, + writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, + LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] + -- v -> identity v + rw [←tangentBundleCore_coordChange_achart] at h + have hvsub : v ((extChartAt I x₀).symm (f t)) = (tangentBundleCore I M).coordChange + (achart H x₀) (achart H ((extChartAt I x₀).symm (f t))) ((extChartAt I x₀).symm (f t)) + ((tangentBundleCore I M).coordChange (achart H ((extChartAt I x₀).symm (f t))) (achart H x₀) + ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) + · rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] + · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] + exact mem_extChartAt_source .. + · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, + ←extChartAt_source I, ←extChartAt_source I, Set.inter_comm, ←Set.inter_assoc, Set.inter_self] + constructor + · exact mem_extChartAt_source .. + · rw [←Set.mem_preimage] + apply Set.mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) + rw [LocalEquiv.symm_source] + exact hf3' + rw [hvsub] + apply HasFDerivAt.comp_hasDerivAt _ _ h + rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWithin_of_mem_nhds] + · apply DifferentiableAt.hasFDerivAt + apply MDifferentiableAt.differentiableAt + apply MDifferentiableAt.comp (I' := I) + · exact (contMDiffAt_extChartAt (n := 1)).mdifferentiableAt (le_refl _) + · apply MDifferentiableOn.mdifferentiableAt + ((contMDiffOn_extChartAt_symm (n := 1) _).mdifferentiableOn (le_refl _)) + rw [mem_nhds_iff] + exact ⟨interior (extChartAt I x₀).target, interior_subset, isOpen_interior, hf3 _ ht2⟩ + · rw [mem_nhds_iff] + refine ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : - ∃ (γ : ℝ → M), HasIntegralCurveAt v γ x₀ t₀ := - exists_integralCurve_of_contMDiff_tangent_section hv _ - ModelWithCorners.Boundaryless.isInteriorPoint + ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := + exists_integralCurve_of_contMDiff_tangent_section hv _ I.isInteriorPoint diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index c433529f0997e..7a7fc4a1490a0 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -44,22 +44,22 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- `p ∈ M` is an interior point of a smooth manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ -def ModelWithCorners.isInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target /-- `p ∈ M` is a boundary point of a smooth manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is a boundary point of `φ.target`. -/ -def ModelWithCorners.isBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target +def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target namespace SmoothManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ -protected def interior : Set M := { x : M | I.isInteriorPoint x} +protected def interior : Set M := { x : M | I.IsInteriorPoint x} variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ -protected def boundary : Set M := { x : M | I.isBoundaryPoint x} +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ lemma foobar {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} @@ -79,17 +79,17 @@ lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ /-- Whether `x` is an interior point can equivalently be described by any chart whose source contains `x`. -/ lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by + (hx : x ∈ e.source) : I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by sorry /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by + (hx : x ∈ e.source) : I.IsBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by sorry /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! -lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isBoundaryPoint x := by +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by set e := extChartAt I x set y := extChartAt I x x have : IsClosed I.target := I.target_eq ▸ (I.closed_range) From 6b55cbcf8baaa93fdf48125249aaace3d51ca8da Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 13:34:31 -0800 Subject: [PATCH 018/203] detailed proof steps --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 29 +++++++++++++------- 1 file changed, 19 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5f98954c0a04b..20e750bfca222 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -69,7 +69,7 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ /- TODO: -* split the theorem below into smaller lemmas, e.g. involving HasIntegralCurveAt? +* split the theorem below into smaller lemmas, e.g. involving IsIntegralCurveAt? * shift and stretch theorems -/ example (γ : ℝ → M) (dt : ℝ) : @@ -85,18 +85,20 @@ example (γ : ℝ → M) (a : ℝ) : `t₀`.-/ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by + /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv - have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) - · rw [mem_nhds_iff] + have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) := by + rw [mem_nhds_iff] refine ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hx⟩ + /- use Picard-Lindelöf theorem to extract a solution to the ODE in the chart defined by `v` -/ obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) rw [←Real.ball_eq_Ioo] at hf2 - -- use continuity of f to extract ε2 so that for t ∈ Real.ball t₀ ε2, - -- f t ∈ interior (extChartAt I x₀).target + /- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, + `f t ∈ interior (extChartAt I x₀).target` -/ have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := @@ -104,11 +106,12 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [Set.subset_def, Set.mem_preimage] at hf3 - -- prove the theorem + /- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve -/ refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] intros t ht + /- collect useful terms in convenient forms -/ rw [←Real.ball_eq_Ioo] at ht have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) @@ -118,6 +121,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin (v ((extChartAt I x₀).symm (f t)))) t := hf2 t ht1 have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset + /- express the derivative of the integral curve in the local chart -/ rw [HasMFDerivAt] use ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) @@ -125,13 +129,16 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [modelWithCornersSelf_coe, Set.range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] - -- v -> identity v - rw [←tangentBundleCore_coordChange_achart] at h + /- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of + coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use + `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires + first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of + coordinates as `D`. -/ have hvsub : v ((extChartAt I x₀).symm (f t)) = (tangentBundleCore I M).coordChange (achart H x₀) (achart H ((extChartAt I x₀).symm (f t))) ((extChartAt I x₀).symm (f t)) ((tangentBundleCore I M).coordChange (achart H ((extChartAt I x₀).symm (f t))) (achart H x₀) - ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) - · rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] + ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) := by + rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] exact mem_extChartAt_source .. · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, @@ -144,6 +151,8 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin exact hf3' rw [hvsub] apply HasFDerivAt.comp_hasDerivAt _ _ h + /- change of coordinates in the tangent bundle is exactly the derivative of composition of local + charts -/ rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWithin_of_mem_nhds] · apply DifferentiableAt.hasFDerivAt apply MDifferentiableAt.differentiableAt From 0ada08efabec14b70ea5f065eb8f0729e1a8aedb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 13:36:29 -0800 Subject: [PATCH 019/203] open Set --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 25 ++++++++++---------- 1 file changed, 13 insertions(+), 12 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 20e750bfca222..6f3924cbffa62 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -23,6 +23,7 @@ integral curve, vector field, local existence -/ open scoped Manifold +open Set section @@ -37,7 +38,7 @@ variable lemma ModelWithCorners.isOpen_extend_target [I.Boundaryless] {e : LocalHomeomorph M H} : IsOpen (e.extend I).target := by - rw [LocalHomeomorph.extend_target, I.range_eq_univ, Set.inter_univ] + rw [LocalHomeomorph.extend_target, I.range_eq_univ, inter_univ] exact I.continuous_symm.isOpen_preimage _ e.open_target lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : @@ -64,7 +65,7 @@ variable /-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. -/ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) (x₀ : M) := - γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Set.Ioo (t₀ - ε) (t₀ + ε) → + γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) /- @@ -88,7 +89,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv - have hI : Set.range I ∈ nhds (extChartAt I x₀ x₀) := by + have hI : range I ∈ nhds (extChartAt I x₀ x₀) := by rw [mem_nhds_iff] refine ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), @@ -105,7 +106,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin hcont _ (IsOpen.mem_nhds isOpen_interior hx) rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds - simp_rw [Set.subset_def, Set.mem_preimage] at hf3 + simp_rw [subset_def, mem_preimage] at hf3 /- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve -/ refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ · apply Eq.symm @@ -113,20 +114,20 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin intros t ht /- collect useful terms in convenient forms -/ rw [←Real.ball_eq_Ioo] at ht - have ht1 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) - have ht2 := Set.mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) + have ht1 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + have ht2 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) have h : HasDerivAt f ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (Set.range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) (v ((extChartAt I x₀).symm (f t)))) t := hf2 t ht1 - have hf3' := Set.mem_of_mem_of_subset (hf3 t ht2) interior_subset + have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset /- express the derivative of the integral curve in the local chart -/ rw [HasMFDerivAt] use ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) apply HasDerivWithinAt.hasFDerivWithinAt - rw [modelWithCornersSelf_coe, Set.range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, + rw [modelWithCornersSelf_coe, range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] /- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of @@ -142,11 +143,11 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] exact mem_extChartAt_source .. · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, - ←extChartAt_source I, ←extChartAt_source I, Set.inter_comm, ←Set.inter_assoc, Set.inter_self] + ←extChartAt_source I, ←extChartAt_source I, inter_comm, ←inter_assoc, inter_self] constructor · exact mem_extChartAt_source .. - · rw [←Set.mem_preimage] - apply Set.mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) + · rw [←mem_preimage] + apply mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) rw [LocalEquiv.symm_source] exact hf3' rw [hvsub] From 063822d74cd22a43fb2698447547d769905f23cb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 13:38:07 -0800 Subject: [PATCH 020/203] `IsIntegralCurveAt` in preamble --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 6f3924cbffa62..c358b268c8304 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -17,6 +17,11 @@ For any continuously differentiable vector field on a manifold `M` and any chose As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold without boundary. +## Implementation notes + +If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means +`γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. + ## Tags integral curve, vector field, local existence From fb4196f8ba22d73618101289498d1b4d7a567123 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 14:32:01 -0800 Subject: [PATCH 021/203] IsIntegralCurveAt.comp_add --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 31 +++++++++++++------ .../Geometry/Manifold/InteriorBoundary.lean | 6 ++-- 2 files changed, 26 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c358b268c8304..31f04ef64cae2 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -17,10 +17,11 @@ For any continuously differentiable vector field on a manifold `M` and any chose As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold without boundary. -## Implementation notes +## Main definition -If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means -`γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. +- **IsInteriorPoint x**: If `v : M → TM` is a vector field on `M` and `x : M`, +`IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a differentiable integral curve of `v` with +`γ x₀ = t₀`. ## Tags @@ -78,12 +79,23 @@ TODO: * split the theorem below into smaller lemmas, e.g. involving IsIntegralCurveAt? * shift and stretch theorems -/ -example (γ : ℝ → M) (dt : ℝ) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ + dt) x₀ := - sorry - -example (γ : ℝ → M) (a : ℝ) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => a * t)) (a • v) t₀ x₀ := sorry +lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : + IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by + obtain ⟨h1, ε, hε, h2⟩ := hγ + refine' ⟨by simp [h1], ε, hε, _⟩ + intros t ht + rw [sub_right_comm, sub_add_eq_add_sub, ←add_mem_Ioo_iff_left] at ht + have h2' := h2 (t + dt) ht + rw [Function.comp_apply, + ←ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] + apply HasMFDerivAt.comp t h2' (f := fun t : ℝ => t + dt) + refine' ⟨(continuous_add_right _).continuousAt, _⟩ + simp + apply HasFDerivAt.add_const + exact hasFDerivAt_id _ + +lemma IsIntegralCurveAt.comp_mul (γ : ℝ → M) (a : ℝ) : + IsIntegralCurveAt γ v t₀ x₀ → IsIntegralCurveAt (γ ∘ (fun t => a * t)) (a • v) t₀ x₀ := sorry /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector @@ -94,6 +106,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv + /- `hI` should be a separate lemma -/ have hI : range I ∈ nhds (extChartAt I x₀ x₀) := by rw [mem_nhds_iff] refine ⟨interior (extChartAt I x₀).target, diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 7a7fc4a1490a0..e7a8f2ceb41dc 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -18,8 +18,10 @@ Define the interior and boundary of a smooth manifold. - **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. - **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, for `φ` being the preferred chart at `x`, -- **SmoothManifoldWithBoundary.interior I M** is the **interior** of `M`, the set of its interior points. -- **SmoothManifoldWithBoundary.boundary I M** is the **boundary** of `M`, the set of its boundary points. +- **SmoothManifoldWithBoundary.interior I M** is the **interior** of `M`, the set of its interior +points. +- **SmoothManifoldWithBoundary.boundary I M** is the **boundary** of `M`, the set of its boundary +points. ## Main results - `xxx`: M is the union of its interior and boundary From 6f665ad1cb7ad3c3f779e6750d8b6aa27a38269c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 17:30:37 -0800 Subject: [PATCH 022/203] translation, scale, and reverse lemmas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 62 ++++++++++++++++++-- 1 file changed, 58 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 31f04ef64cae2..22a2e633777c6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -78,6 +78,7 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ TODO: * split the theorem below into smaller lemmas, e.g. involving IsIntegralCurveAt? * shift and stretch theorems +* constant curve at stationary point of v -/ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by @@ -88,14 +89,67 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v have h2' := h2 (t + dt) ht rw [Function.comp_apply, ←ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] - apply HasMFDerivAt.comp t h2' (f := fun t : ℝ => t + dt) + apply HasMFDerivAt.comp t h2' + /- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations -/ refine' ⟨(continuous_add_right _).continuousAt, _⟩ - simp + simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, + LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, + LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, + LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, + Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] apply HasFDerivAt.add_const exact hasFDerivAt_id _ -lemma IsIntegralCurveAt.comp_mul (γ : ℝ → M) (a : ℝ) : - IsIntegralCurveAt γ v t₀ x₀ → IsIntegralCurveAt (γ ∘ (fun t => a * t)) (a • v) t₀ x₀ := sorry +lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} + (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + obtain ⟨h1, ε, hε, h2⟩ := hγ + refine' ⟨by rw [Function.comp_apply, div_mul_cancel _ (ne_of_gt ha)]; exact h1, ε / a, + div_pos hε ha, _⟩ + intros t ht + have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by + rw [mem_Ioo, ←div_lt_iff ha, ←lt_div_iff ha, sub_div, add_div] + exact ht + have h2' := h2 (t * a) ht + rw [Function.comp_apply, Pi.smul_apply, ←ContinuousLinearMap.smulRight_comp] + apply HasMFDerivAt.comp t h2' + refine' ⟨(continuous_mul_right _).continuousAt, _⟩ + simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, + LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, + LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, + LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, + Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] + apply HasFDerivAt.mul_const' + exact hasFDerivAt_id _ + +lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) : + IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by + obtain ⟨h1, ε, hε, h2⟩ := hγ + refine' ⟨by simp [h1], ε, hε, _⟩ + intros t ht + rw [←neg_add', neg_add_eq_sub, ←neg_sub, ←neg_mem_Ioo_iff] at ht + have h2' := h2 (-t) ht + rw [Function.comp_apply, Pi.neg_apply, ←neg_one_smul ℝ (v (γ (-t))), + ←ContinuousLinearMap.smulRight_comp] + apply HasMFDerivAt.comp t h2' + refine' ⟨continuousAt_neg, _⟩ + simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, + LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, + LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, + LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, + Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] + apply HasDerivAt.hasFDerivAt + exact hasDerivAt_neg _ + +lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} + (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + rw [ne_iff_lt_or_gt] at ha + cases' ha with ha ha + /- this makes me think we need inverses of the above theorems -/ + · have h := (hγ.comp_mul_pos _ (neg_pos_of_neg ha)).comp_neg + have : γ ∘ (fun t ↦ t * -a) ∘ Neg.neg = γ ∘ fun t ↦ t * a := by ext; simp + rw [Function.comp.assoc, ←neg_smul, neg_neg, neg_div', neg_div_neg_eq, this] at h + exact h + · exact hγ.comp_mul_pos _ ha /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector From b59c82f3991e597debbb5e6de28508a1ac9d4cc2 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 20 Nov 2023 18:18:40 -0800 Subject: [PATCH 023/203] iff lemmas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 55 +++++++++++++++++--- 1 file changed, 48 insertions(+), 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 22a2e633777c6..2799dd103ac4b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -80,6 +80,8 @@ TODO: * shift and stretch theorems * constant curve at stationary point of v -/ +variable {t₀} + lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ @@ -100,6 +102,16 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v apply HasFDerivAt.add_const exact hasFDerivAt_id _ +lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ x₀ ↔ + IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by + refine' ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, _⟩ + intro hγ + have := hγ.comp_add (-dt) + rw [sub_neg_eq_add, sub_add_cancel] at this + convert this + ext + simp + lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ @@ -121,6 +133,17 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt apply HasFDerivAt.mul_const' exact hasFDerivAt_id _ +lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + refine' ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, _⟩ + intro hγ + have := hγ.comp_mul_pos (inv_pos_of_pos ha) + rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ←div_mul_eq_div_div_swap, + inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this + convert this + ext + simp [inv_mul_eq_div, div_self (ne_of_gt ha)] + lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) : IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ @@ -140,16 +163,34 @@ lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v apply HasDerivAt.hasFDerivAt exact hasDerivAt_neg _ +lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by + refine' ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, _⟩ + intro hγ + have := hγ.comp_neg + rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this + exact this + lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by rw [ne_iff_lt_or_gt] at ha cases' ha with ha ha - /- this makes me think we need inverses of the above theorems -/ - · have h := (hγ.comp_mul_pos _ (neg_pos_of_neg ha)).comp_neg - have : γ ∘ (fun t ↦ t * -a) ∘ Neg.neg = γ ∘ fun t ↦ t * a := by ext; simp - rw [Function.comp.assoc, ←neg_smul, neg_neg, neg_div', neg_div_neg_eq, this] at h - exact h - · exact hγ.comp_mul_pos _ ha + · apply isIntegralCurveAt_comp_neg.mpr + have : (fun t ↦ t * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp + rw [Function.comp.assoc, this, ←neg_smul, ←div_neg] + exact hγ.comp_mul_pos (neg_pos_of_neg ha) + · exact hγ.comp_mul_pos ha + +lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + refine' ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, _⟩ + intro hγ + have := hγ.comp_mul_ne_zero (inv_ne_zero ha) + rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ←div_mul_eq_div_div_swap, + inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this + convert this + ext + simp [inv_mul_eq_div, div_self ha] /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector @@ -245,4 +286,4 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := - exists_integralCurve_of_contMDiff_tangent_section hv _ I.isInteriorPoint + exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint From 3c2e4244e4afb6b7264184ea8312c3b544c1c3f2 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 12:08:18 +0100 Subject: [PATCH 024/203] Small polish. Note that smoothness is not required; all results in this file hold for topological manifolds. --- .../Geometry/Manifold/InteriorBoundary.lean | 42 ++++++++++--------- 1 file changed, 22 insertions(+), 20 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index e7a8f2ceb41dc..09537f8cf91d1 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -9,27 +9,27 @@ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners -- FIXME: should this be its own file or go in SmoothManifoldWithCorners? -- the latter is already huge, or its own file - move other results about boundaryless here? +-- NB: all results in this file hold for topological manifolds + /-! -# Interior and boundary of a smooth manifold +# Interior and boundary of a manifold -Define the interior and boundary of a smooth manifold. +Define the interior and boundary of a manifold. ## Main definitions - **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. - **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, for `φ` being the preferred chart at `x`, -- **SmoothManifoldWithBoundary.interior I M** is the **interior** of `M`, the set of its interior -points. -- **SmoothManifoldWithBoundary.boundary I M** is the **boundary** of `M`, the set of its boundary -points. +- **interior I M** is the **interior** of `M`, the set of its interior points. +- **boundary I M** is the **boundary** of `M`, the set of its boundary points. ## Main results -- `xxx`: M is the union of its interior and boundary -- `yyy`: interior I M is open +- `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary +- `interior_isOpen`: `interior I M` is open -**TODO**: show that -- interior I M is a manifold without boundary - (need to upgrade the above; map the charts from an open ball to entire ℝ^n) +**TODO** +- `interior I M` is a manifold without boundary + (need to upgrade the model used; map the charts from an open ball to entire ℝ^n) - the boundary is a submanifold of codimension 1 (once mathlib has submanifolds) ## Tags @@ -39,6 +39,7 @@ manifold, interior, boundary open Set -- Let `M` be a smooth manifold with corners over the pair `(E, H)`. +-- NB: smoothness is not required; all results in this file hold for topological manifolds. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} @@ -80,14 +81,15 @@ lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ -- FIXME(MR): find a better wording for the next two docstrings /-- Whether `x` is an interior point can equivalently be described by any chart whose source contains `x`. -/ -lemma isInteriorPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by +-- as we only need continuity properties, `e` being in the atlas is not required +lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : + I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by sorry /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ -lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} (he : e ∈ atlas H M) {x : M} - (hx : x ∈ e.source) : I.IsBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by +lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : + I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by sorry /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! @@ -102,14 +104,14 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB exact mem_extChartAt_target I x exact (mem_union y _ _).mp this -/-- A smooth manifold decomposes into interior and boundary. -/ +/-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by apply le_antisymm · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x -/-- Ihe interior of a smooth manifold is an open subset. -/ +/-- Ihe interior of a manifold is an open subset. -/ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by apply isOpen_iff_forall_mem_open.mpr intro x hx @@ -122,7 +124,7 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by refine ⟨?_, ?_, ?_⟩ · intro y hy rw [e.extend_source] at hy - apply (isInteriorPoint_iff (chart_mem_atlas H x) (mem_of_mem_inter_left hy)).mpr + apply (isInteriorPoint_iff (mem_of_mem_inter_left hy)).mpr exact mem_of_mem_inter_right (a := e.source) hy · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) isOpen_interior · have : x ∈ (e.extend I).source := by @@ -138,10 +140,10 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : interior (frontier (e.extend I).target) = ∅ := sorry -/-- The boundary of a smooth manifold has empty interior. -/ +/-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` sorry --- interior I M is a smooth manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) +-- interior I M is a manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) end SmoothManifoldWithCorners From c95243e703b2e96a8a2f7030350a7b233ce45a64 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 12:10:44 +0100 Subject: [PATCH 025/203] The boundary is closed: entails showing that interior and boundary are disjoint. --- .../Geometry/Manifold/InteriorBoundary.lean | 33 ++++++++++++++++++- 1 file changed, 32 insertions(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 09537f8cf91d1..7f822e2fd4ae9 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -26,6 +26,9 @@ Define the interior and boundary of a manifold. ## Main results - `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary - `interior_isOpen`: `interior I M` is open +- `boundary_isClosed`: `boundary I M` is closed +- `interior_boundary_eq_empty`: `boundary I M` has empty interior +(this implies it has "measure zero", see different file) **TODO** - `interior I M` is a manifold without boundary @@ -111,7 +114,27 @@ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x -/-- Ihe interior of a manifold is an open subset. -/ +-- should be in mathlib; Mathlib.Topology.Basic +lemma interior_frontier_disjoint {X : Type*} [TopologicalSpace X] {s : Set X} : + interior s ∩ frontier s = ∅ := by + rw [← closure_diff_interior s] + have : (closure s \ interior s) = closure s ∩ (interior s)ᶜ := rfl + rw [this] + rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] + +-- proper name? or _eq_emptyset? +/-- The interior and boundary of `M` are disjoint. -/ +lemma interior_boundary_disjoint : + (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by + ext x + constructor; intro h + · let e := extChartAt I x + have : e x ∈ interior e.target ∩ frontier e.target := h + rw [interior_frontier_disjoint] at this + exact this + · exfalso + +/-- The interior of a manifold is an open subset. -/ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by apply isOpen_iff_forall_mem_open.mpr intro x hx @@ -132,6 +155,14 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by exact mem_chart_source H x exact mem_inter this hx +/-- The boundary of a manifold is a closed subset. -/ +lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := by + apply isOpen_compl_iff.mp + have : (SmoothManifoldWithCorners.interior I M)ᶜ = SmoothManifoldWithCorners.boundary I M := + (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary (I := I) (M := M))) + rw [← this, compl_compl] + exact interior_isOpen + /-- The boundary of any extended chart has empty interior. -/ -- NB: this is *false* for any set instead of `(e.extend I).target`: -- for instance, $ℚ ⊆ ℝ$ has frontier ℝ (ℚ is dense in ℝ and ℚ has empty interior). From 933cc5961b26c140339d70bd1e9d808ccd33d8d8 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 13:06:20 +0100 Subject: [PATCH 026/203] Almost-proof that boundary has empty interior: need to adjust definition. The current lemma (frontier of (e.extend I).target) has contributions from boundary(I.range) (good), but also from another factor (bad). --- .../Geometry/Manifold/InteriorBoundary.lean | 31 ++++++++++++++++--- 1 file changed, 26 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 7f822e2fd4ae9..b4a507c071baf 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -163,13 +163,34 @@ lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := b rw [← this, compl_compl] exact interior_isOpen +-- XXX: this is not fully true! frontier s includes t ∖ interior t, but also closure(O) ∖ O... +-- ACTUALLY: this means the definition of boundary is not right: want to take target without interior, +-- **not** the frontier!! +lemma aux {X : Type*} [TopologicalSpace X] {s O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : + frontier s ⊆ frontier t := by + let aux := calc interior s + _ = interior O ∩ interior t := by rw [h, interior_inter] + _ = O ∩ interior t := by rw [hO.interior_eq] + + -- closure s = closure O ∩ closure t ⊆ closure t + -- frontier s = closure s ∖ interior s ⊆ closure(t) ∖ (O ∩ interior t)... not good enough!! + -- s ∖ interior s = (O ∩ t) ∖ (O ∩ interior t) = O ∩ (t ∖ interior t) ⊆ t ∖ interior t + -- ⊆ closure t ∖ interior t = frontiert t + + sorry + /-- The boundary of any extended chart has empty interior. -/ --- NB: this is *false* for any set instead of `(e.extend I).target`: --- for instance, $ℚ ⊆ ℝ$ has frontier ℝ (ℚ is dense in ℝ and ℚ has empty interior). --- xxx: do I need that e is in the atlas? I think not; not double-checked. --- xxx: is this lemma true with mathlib's current definitions? lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : - interior (frontier (e.extend I).target) = ∅ := sorry + interior (frontier (e.extend I).target) = ∅ := by + -- `e.extend_target I = (I.symm ⁻¹' e.target) ∩ range I` is the union of an open set and a + -- closed set: hence the frontier is contained in the second factor. + have h1 : frontier (e.extend I).target ⊆ frontier (range I) := by + rw [e.extend_target I] + exact aux rfl (e.open_target.preimage I.continuous_symm) + suffices interior (frontier (range I)) = ∅ by + exact subset_eq_empty (interior_mono h1) this + -- As `range I` is closed, its frontier has empty interior. + exact interior_frontier I.closed_range /-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by From 2041d3095603426c64ceb32aab3d654069d1182b Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 13:43:17 +0100 Subject: [PATCH 027/203] Correct the definition of boundary: we should ask for x not being an interior of its chart's target. (Asking for "lies in the frontier", as we did before, would also include boundary points of (extChartAt I x).source in interior (range I), which we're not interested in.) In other words, the previous definition was actually *wrong*. --- .../Geometry/Manifold/InteriorBoundary.lean | 27 +++++++++---------- 1 file changed, 12 insertions(+), 15 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index b4a507c071baf..f3a77b89bf7ca 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -52,9 +52,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target -/-- `p ∈ M` is a boundary point of a smooth manifold `M` iff -for `φ` being the preferred chart at `x`, `φ x` is a boundary point of `φ.target`. -/ -def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (extChartAt I x).target +/-- `p ∈ M` is a boundary point of a smooth manifold `M` iff it is not an interior point. +This means that, for `φ` being the preferred chart at `x`, `φ x` is not an interior point of +`φ.target`. We do not say "boundary point" as `frontier φ.target` has two components, one on the +boundary of range I and another on the boundary of e.target (which we don't want). -/ +def ModelWithCorners.isBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target namespace SmoothManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? @@ -92,20 +94,18 @@ lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.isBoundaryPoint x ↔ (e.extend I) x ∈ frontier (e.extend I).target := by + I.isBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by sorry /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by set e := extChartAt I x set y := extChartAt I x x - have : IsClosed I.target := I.target_eq ▸ (I.closed_range) - -- TODO: this should be obvious now! - have : IsClosed e.target := sorry - have : y ∈ interior e.target ∪ frontier e.target := by - rw [← closure_eq_interior_union_frontier (e.target), this.closure_eq] - exact mem_extChartAt_target I x - exact (mem_union y _ _).mp this + by_cases y ∈ interior e.target + · have : I.isInteriorPoint x := (isInteriorPoint_iff (mem_chart_source H x) (I := I)).mpr h + exact Or.inl h + · have : I.isBoundaryPoint x := (isBoundaryPoint_iff (mem_chart_source H x) (I := I)).mpr h + exact Or.inr h /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ @@ -128,10 +128,7 @@ lemma interior_boundary_disjoint : (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by ext x constructor; intro h - · let e := extChartAt I x - have : e x ∈ interior e.target ∩ frontier e.target := h - rw [interior_frontier_disjoint] at this - exact this + · exact (not_mem_of_mem_diff h) (mem_of_mem_diff h) · exfalso /-- The interior of a manifold is an open subset. -/ From f1fda329edeee953ba11d0c5d9ff1d4f95faaadb Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 13:43:33 +0100 Subject: [PATCH 028/203] Fix and complete proof that the boundary has empty interior. --- .../Geometry/Manifold/InteriorBoundary.lean | 27 +++++++++---------- 1 file changed, 13 insertions(+), 14 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index f3a77b89bf7ca..4f3f36323f233 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -160,30 +160,29 @@ lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := b rw [← this, compl_compl] exact interior_isOpen --- XXX: this is not fully true! frontier s includes t ∖ interior t, but also closure(O) ∖ O... --- ACTUALLY: this means the definition of boundary is not right: want to take target without interior, --- **not** the frontier!! +-- FIXME: good name? should go in Mathlib.Topology.Basic lemma aux {X : Type*} [TopologicalSpace X] {s O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : - frontier s ⊆ frontier t := by + s \ interior s ⊆ t \ interior t := by let aux := calc interior s _ = interior O ∩ interior t := by rw [h, interior_inter] _ = O ∩ interior t := by rw [hO.interior_eq] + calc s \ interior s + _ = (O ∩ t) \ (O ∩ interior t) := by rw [aux, h] + _ = O ∩ (t \ interior t) := by rw [inter_diff_distrib_left] + _ ⊆ t \ interior t := inter_subset_right _ _ - -- closure s = closure O ∩ closure t ⊆ closure t - -- frontier s = closure s ∖ interior s ⊆ closure(t) ∖ (O ∩ interior t)... not good enough!! - -- s ∖ interior s = (O ∩ t) ∖ (O ∩ interior t) = O ∩ (t ∖ interior t) ⊆ t ∖ interior t - -- ⊆ closure t ∖ interior t = frontiert t - - sorry +-- is this a better lemma; is `aux` useful on its own? +lemma aux2 {X : Type*} [TopologicalSpace X] {s O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) + (ht : IsClosed t) : s \ interior s ⊆ frontier t := + ht.frontier_eq ▸ aux h hO /-- The boundary of any extended chart has empty interior. -/ lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : - interior (frontier (e.extend I).target) = ∅ := by + interior ((e.extend I).target \ interior (e.extend I).target) = ∅ := by -- `e.extend_target I = (I.symm ⁻¹' e.target) ∩ range I` is the union of an open set and a -- closed set: hence the frontier is contained in the second factor. - have h1 : frontier (e.extend I).target ⊆ frontier (range I) := by - rw [e.extend_target I] - exact aux rfl (e.open_target.preimage I.continuous_symm) + have h1 : (e.extend I).target \ interior (e.extend I).target ⊆ frontier (range I) := + aux2 (e.extend_target I) (e.open_target.preimage I.continuous_symm) I.closed_range suffices interior (frontier (range I)) = ∅ by exact subset_eq_empty (interior_mono h1) this -- As `range I` is closed, its frontier has empty interior. From 6a4e39c686ca28fac025cc251a3b7bc6024dd00d Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 14:01:05 +0100 Subject: [PATCH 029/203] Small clean-ups: - make some variables explicit; fix namespacing of one result. - move all topology helper results into their own section. --- .../Geometry/Manifold/InteriorBoundary.lean | 69 +++++++++++-------- 1 file changed, 40 insertions(+), 29 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 4f3f36323f233..61251507d1712 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -41,6 +41,34 @@ manifold, interior, boundary open Set +section TopologyHelpers -- should be in mathlib; Mathlib.Topology.Basic +variable {X : Type*} [TopologicalSpace X] {s : Set X} + +-- I don't need it; still useful? +lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by + rw [← closure_diff_interior s] + have : (closure s \ interior s) = closure s ∩ (interior s)ᶜ := rfl + rw [this] + rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] + +-- FIXME: what's a good name? +lemma aux {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : + s \ interior s ⊆ t \ interior t := by + let aux := calc interior s + _ = interior O ∩ interior t := by rw [h, interior_inter] + _ = O ∩ interior t := by rw [hO.interior_eq] + calc s \ interior s + _ = (O ∩ t) \ (O ∩ interior t) := by rw [aux, h] + _ = O ∩ (t \ interior t) := by rw [inter_diff_distrib_left] + _ ⊆ t \ interior t := inter_subset_right _ _ + +-- is this a better lemma; is `aux` useful on its own? +lemma aux2 {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) + (ht : IsClosed t) : s \ interior s ⊆ frontier t := + ht.frontier_eq ▸ aux h hO + +end TopologyHelpers + -- Let `M` be a smooth manifold with corners over the pair `(E, H)`. -- NB: smoothness is not required; all results in this file hold for topological manifolds. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] @@ -84,6 +112,7 @@ lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ -- more abstract result: a local homeomorphism maps interior to interior and boundary to boundary -- FIXME(MR): find a better wording for the next two docstrings +variable (I) in /-- Whether `x` is an interior point can equivalently be described by any chart whose source contains `x`. -/ -- as we only need continuity properties, `e` being in the atlas is not required @@ -91,6 +120,7 @@ lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by sorry +variable (I) in /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : @@ -102,11 +132,12 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB set e := extChartAt I x set y := extChartAt I x x by_cases y ∈ interior e.target - · have : I.isInteriorPoint x := (isInteriorPoint_iff (mem_chart_source H x) (I := I)).mpr h + · have : I.isInteriorPoint x := (isInteriorPoint_iff I (mem_chart_source H x)).mpr h exact Or.inl h - · have : I.isBoundaryPoint x := (isBoundaryPoint_iff (mem_chart_source H x) (I := I)).mpr h + · have : I.isBoundaryPoint x := (isBoundaryPoint_iff I (mem_chart_source H x)).mpr h exact Or.inr h +variable (I M) in /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by @@ -114,14 +145,6 @@ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x --- should be in mathlib; Mathlib.Topology.Basic -lemma interior_frontier_disjoint {X : Type*} [TopologicalSpace X] {s : Set X} : - interior s ∩ frontier s = ∅ := by - rw [← closure_diff_interior s] - have : (closure s \ interior s) = closure s ∩ (interior s)ᶜ := rfl - rw [this] - rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] - -- proper name? or _eq_emptyset? /-- The interior and boundary of `M` are disjoint. -/ lemma interior_boundary_disjoint : @@ -144,7 +167,7 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by refine ⟨?_, ?_, ?_⟩ · intro y hy rw [e.extend_source] at hy - apply (isInteriorPoint_iff (mem_of_mem_inter_left hy)).mpr + apply (isInteriorPoint_iff I (mem_of_mem_inter_left hy)).mpr exact mem_of_mem_inter_right (a := e.source) hy · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) isOpen_interior · have : x ∈ (e.extend I).source := by @@ -156,28 +179,14 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := by apply isOpen_compl_iff.mp have : (SmoothManifoldWithCorners.interior I M)ᶜ = SmoothManifoldWithCorners.boundary I M := - (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary (I := I) (M := M))) + (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary I M)) rw [← this, compl_compl] exact interior_isOpen +end SmoothManifoldWithCorners --- FIXME: good name? should go in Mathlib.Topology.Basic -lemma aux {X : Type*} [TopologicalSpace X] {s O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : - s \ interior s ⊆ t \ interior t := by - let aux := calc interior s - _ = interior O ∩ interior t := by rw [h, interior_inter] - _ = O ∩ interior t := by rw [hO.interior_eq] - calc s \ interior s - _ = (O ∩ t) \ (O ∩ interior t) := by rw [aux, h] - _ = O ∩ (t \ interior t) := by rw [inter_diff_distrib_left] - _ ⊆ t \ interior t := inter_subset_right _ _ - --- is this a better lemma; is `aux` useful on its own? -lemma aux2 {X : Type*} [TopologicalSpace X] {s O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) - (ht : IsClosed t) : s \ interior s ⊆ frontier t := - ht.frontier_eq ▸ aux h hO - +variable (I) in /-- The boundary of any extended chart has empty interior. -/ -lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : +lemma LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : interior ((e.extend I).target \ interior (e.extend I).target) = ∅ := by -- `e.extend_target I = (I.symm ⁻¹' e.target) ∩ range I` is the union of an open set and a -- closed set: hence the frontier is contained in the second factor. @@ -188,6 +197,8 @@ lemma __root__.LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeo -- As `range I` is closed, its frontier has empty interior. exact interior_frontier I.closed_range +namespace SmoothManifoldWithCorners +variable (I) in /-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` From 88bea4ca167911270fc8d8f1c3f564e5409a0114 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 14:51:24 +0100 Subject: [PATCH 030/203] Clean up lemmas about boundary. --- .../Geometry/Manifold/InteriorBoundary.lean | 26 ++++++++----------- 1 file changed, 11 insertions(+), 15 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 61251507d1712..ea9bc772da559 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -98,18 +98,11 @@ variable (I M) in protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ -lemma foobar {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} - (hx : x ∈ e.source ∩ e'.source) : - (e.extend I) x ∈ interior (e.extend I).target ↔ - (e'.extend I) x ∈ interior (e'.extend I).target := sorry - -/-- If `e` and `e'` are two charts, the transition map maps boundary points to boundary points. -/ -lemma foobar' {e e' : LocalHomeomorph M H} (he : e ∈ atlas H M) (he' : e' ∈ atlas H M) {x : M} - (hx : x ∈ e.source ∩ e'.source) : - (e.extend I) x ∈ frontier (e.extend I).target ↔ - (e'.extend I) x ∈ frontier (e'.extend I).target := sorry - --- more abstract result: a local homeomorphism maps interior to interior and boundary to boundary +-- as we only need continuity property, e or e' being in the atlas is not required +lemma foobar {e e' : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source ∩ e'.source) : + (e.extend I) x ∈ interior (e.extend I).target ↔ + (e'.extend I) x ∈ interior (e'.extend I).target := sorry +-- both directions should be the same, more general lemma -- FIXME(MR): find a better wording for the next two docstrings variable (I) in @@ -117,15 +110,18 @@ variable (I) in whose source contains `x`. -/ -- as we only need continuity properties, `e` being in the atlas is not required lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := by - sorry + I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := + foobar (mem_inter (mem_chart_source H x) hx) variable (I) in /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by - sorry + -- This lemma is just the "negation" (applying not_iff_not) to isInteriorPoint_iff. + constructor <;> intro h + · exact (not_iff_not.mpr (isInteriorPoint_iff I hx)).mp h + · exact (not_iff_not.mpr (isInteriorPoint_iff I hx)).mpr h /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by From 73e203f6655f171d7b3450cdab433fb021d8d3ea Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 14:55:36 +0100 Subject: [PATCH 031/203] Small golfs. --- .../Geometry/Manifold/InteriorBoundary.lean | 26 +++++++------------ 1 file changed, 9 insertions(+), 17 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index ea9bc772da559..1e0600e98bfd3 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -119,19 +119,14 @@ whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : I.isBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by -- This lemma is just the "negation" (applying not_iff_not) to isInteriorPoint_iff. - constructor <;> intro h - · exact (not_iff_not.mpr (isInteriorPoint_iff I hx)).mp h - · exact (not_iff_not.mpr (isInteriorPoint_iff I hx)).mpr h + rw [← not_iff_not.mpr (isInteriorPoint_iff I hx)] + exact Iff.rfl /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! -lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by - set e := extChartAt I x - set y := extChartAt I x x - by_cases y ∈ interior e.target - · have : I.isInteriorPoint x := (isInteriorPoint_iff I (mem_chart_source H x)).mpr h - exact Or.inl h - · have : I.isBoundaryPoint x := (isBoundaryPoint_iff I (mem_chart_source H x)).mpr h - exact Or.inr h +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isBoundaryPoint x := by + by_cases extChartAt I x x ∈ interior (extChartAt I x).target + · exact Or.inl h + · exact Or.inr h variable (I M) in /-- A manifold decomposes into interior and boundary. -/ @@ -141,14 +136,11 @@ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x --- proper name? or _eq_emptyset? -/-- The interior and boundary of `M` are disjoint. -/ +/-- The interior and boundary of `M` are disjoint. -/ -- xxx: name `..._eq_empty` instead? lemma interior_boundary_disjoint : (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by - ext x - constructor; intro h - · exact (not_mem_of_mem_diff h) (mem_of_mem_diff h) - · exfalso + ext + exact ⟨fun h ↦ (not_mem_of_mem_diff h) (mem_of_mem_diff h), by exfalso⟩ /-- The interior of a manifold is an open subset. -/ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by From 90090b263417f79e6b17de400f182d0da0b336b9 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 15:11:34 +0100 Subject: [PATCH 032/203] Fix rebase, make it build and fix a docstring typo. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 13 +++++++------ Mathlib/Geometry/Manifold/InteriorBoundary.lean | 6 +++--- 2 files changed, 10 insertions(+), 9 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 2799dd103ac4b..52bb756a80f2e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -19,7 +19,7 @@ without boundary. ## Main definition -- **IsInteriorPoint x**: If `v : M → TM` is a vector field on `M` and `x : M`, +- **IsIntegralCurve γ v t₀ x₀**: If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. @@ -51,9 +51,10 @@ lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := I.isOpen_extend_target /-- If `M` has no boundary, then every point in `M` is an interior point. -/ -lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : - I.IsInteriorPoint x := by - rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] +-- FIXME: better name? +lemma ModelWithCorners.isInteriorPoint' [I.Boundaryless] {x : M} : + I.isInteriorPoint x := by + rw [ModelWithCorners.isInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) end @@ -196,7 +197,7 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : +theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.isInteriorPoint x₀) : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv @@ -286,4 +287,4 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := - exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint + exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint' diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 1e0600e98bfd3..e6ac9f27b6ece 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -78,7 +78,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- `p ∈ M` is an interior point of a smooth manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ -def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +def ModelWithCorners.isInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target /-- `p ∈ M` is a boundary point of a smooth manifold `M` iff it is not an interior point. This means that, for `φ` being the preferred chart at `x`, `φ x` is not an interior point of @@ -91,11 +91,11 @@ namespace SmoothManifoldWithCorners variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ -protected def interior : Set M := { x : M | I.IsInteriorPoint x} +protected def interior : Set M := { x : M | I.isInteriorPoint x} variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ -protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} +protected def boundary : Set M := { x : M | I.isBoundaryPoint x} /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ -- as we only need continuity property, e or e' being in the atlas is not required From 338ab9fca06500397c70cc17f6db5317c8b97709 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 15:16:09 +0100 Subject: [PATCH 033/203] Fix casing for real: IsInteriorPoint is correct. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +++---- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 14 +++++++------- 2 files changed, 10 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 52bb756a80f2e..b56212475b1b5 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -52,9 +52,8 @@ lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : /-- If `M` has no boundary, then every point in `M` is an interior point. -/ -- FIXME: better name? -lemma ModelWithCorners.isInteriorPoint' [I.Boundaryless] {x : M} : - I.isInteriorPoint x := by - rw [ModelWithCorners.isInteriorPoint, IsOpen.interior_eq I.isOpen_target] +lemma ModelWithCorners.isInteriorPoint' [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by + rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) end @@ -197,7 +196,7 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.isInteriorPoint x₀) : +theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index e6ac9f27b6ece..41f37fef6558b 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -78,24 +78,24 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- `p ∈ M` is an interior point of a smooth manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ -def ModelWithCorners.isInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target /-- `p ∈ M` is a boundary point of a smooth manifold `M` iff it is not an interior point. This means that, for `φ` being the preferred chart at `x`, `φ x` is not an interior point of `φ.target`. We do not say "boundary point" as `frontier φ.target` has two components, one on the boundary of range I and another on the boundary of e.target (which we don't want). -/ -def ModelWithCorners.isBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target +def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target namespace SmoothManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ -protected def interior : Set M := { x : M | I.isInteriorPoint x} +protected def interior : Set M := { x : M | I.IsInteriorPoint x} variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ -protected def boundary : Set M := { x : M | I.isBoundaryPoint x} +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ -- as we only need continuity property, e or e' being in the atlas is not required @@ -110,20 +110,20 @@ variable (I) in whose source contains `x`. -/ -- as we only need continuity properties, `e` being in the atlas is not required lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.isInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := + I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := foobar (mem_inter (mem_chart_source H x) hx) variable (I) in /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart whose source contains `x`. -/ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.isBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by + I.IsBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by -- This lemma is just the "negation" (applying not_iff_not) to isInteriorPoint_iff. rw [← not_iff_not.mpr (isInteriorPoint_iff I hx)] exact Iff.rfl /-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! -lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.isInteriorPoint x ∨ I.isBoundaryPoint x := by +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by by_cases extChartAt I x x ∈ interior (extChartAt I x).target · exact Or.inl h · exact Or.inr h From 378088c61e42b0d2d32d9ac983aa439e59e64daa Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 15:20:22 +0100 Subject: [PATCH 034/203] Name lemma in snake_case. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b56212475b1b5..628a00cb53cfc 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -50,9 +50,8 @@ lemma ModelWithCorners.isOpen_extend_target [I.Boundaryless] {e : LocalHomeomorp lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := I.isOpen_extend_target -/-- If `M` has no boundary, then every point in `M` is an interior point. -/ --- FIXME: better name? -lemma ModelWithCorners.isInteriorPoint' [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by +/-- If `M` has no boundary, every point of `M` is an interior point. -/ +lemma ModelWithCorners.is_interior_point [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) @@ -286,4 +285,4 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := - exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint' + exists_integralCurve_of_contMDiff_tangent_section hv I.is_interior_point From efe08809a1a81b373721cc81ad809aba3b056b54 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 21 Nov 2023 17:36:41 -0800 Subject: [PATCH 035/203] naming --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 628a00cb53cfc..a3b0d57cdefc6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -51,7 +51,7 @@ lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : IsOpen (extChartAt I x).target := I.isOpen_extend_target /-- If `M` has no boundary, every point of `M` is an interior point. -/ -lemma ModelWithCorners.is_interior_point [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by +lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) @@ -285,4 +285,4 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := - exists_integralCurve_of_contMDiff_tangent_section hv I.is_interior_point + exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint From d35e843e3b519db68c55f62d1f10940cfb4a4ba7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 21 Nov 2023 21:27:39 -0800 Subject: [PATCH 036/203] move lemma, style --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 11 +++-------- .../Geometry/Manifold/SmoothManifoldWithCorners.lean | 5 +++++ 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a3b0d57cdefc6..57dbfb79c6c98 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -42,13 +42,8 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -lemma ModelWithCorners.isOpen_extend_target [I.Boundaryless] {e : LocalHomeomorph M H} : - IsOpen (e.extend I).target := by - rw [LocalHomeomorph.extend_target, I.range_eq_univ, inter_univ] - exact I.continuous_symm.isOpen_preimage _ e.open_target - lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : - IsOpen (extChartAt I x).target := I.isOpen_extend_target + IsOpen (extChartAt I x).target := LocalHomeomorph.isOpen_extend_target /-- If `M` has no boundary, every point of `M` is an interior point. -/ lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by @@ -203,7 +198,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin /- `hI` should be a separate lemma -/ have hI : range I ∈ nhds (extChartAt I x₀ x₀) := by rw [mem_nhds_iff] - refine ⟨interior (extChartAt I x₀).target, + exact ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hx⟩ /- use Picard-Lindelöf theorem to extract a solution to the ODE in the chart defined by `v` -/ @@ -276,7 +271,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [mem_nhds_iff] exact ⟨interior (extChartAt I x₀).target, interior_subset, isOpen_interior, hf3 _ ht2⟩ · rw [mem_nhds_iff] - refine ⟨interior (extChartAt I x₀).target, + exact ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ /-- For any continuously differentiable vector field defined on a manifold without boundary and any diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 36c349057386d..2a7f174bceec9 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -976,6 +976,11 @@ theorem extend_target : (f.extend I).target = I.symm ⁻¹' f.target ∩ range I simp_rw [extend, LocalEquiv.trans_target, I.target_eq, I.toLocalEquiv_coe_symm, inter_comm] #align local_homeomorph.extend_target LocalHomeomorph.extend_target +lemma isOpen_extend_target [I.Boundaryless] : + IsOpen (f.extend I).target := by + rw [LocalHomeomorph.extend_target, I.range_eq_univ, inter_univ] + exact I.continuous_symm.isOpen_preimage _ f.open_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, From dce81a6f83f883f9608ac724d79708feb191e2b7 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 22:56:32 +0100 Subject: [PATCH 037/203] Define topological manifolds and use them for defining interior and boundary. TODO: add to SmoothManifoldsWithCorners' docstring - does mathlib have the implication smooth mfd => top mfd? (most is there for sure) - do I want this, or should I just "contGroupoid H" and that suffices? --- .../Geometry/Manifold/InteriorBoundary.lean | 31 +++++++++---------- .../Manifold/SmoothManifoldWithCorners.lean | 18 +++++++++++ 2 files changed, 33 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 41f37fef6558b..419db7299ac25 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -8,8 +8,8 @@ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners -- FIXME: should this be its own file or go in SmoothManifoldWithCorners? -- the latter is already huge, or its own file - move other results about boundaryless here? - --- NB: all results in this file hold for topological manifolds +-- xxx: if I can use dot notation, how set things up so they're also available for smooth manifolds? +-- manually re-declare them? /-! # Interior and boundary of a manifold @@ -69,24 +69,23 @@ lemma aux2 {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) end TopologyHelpers --- Let `M` be a smooth manifold with corners over the pair `(E, H)`. --- NB: smoothness is not required; all results in this file hold for topological manifolds. +-- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} - {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [ManifoldWithCorners I M] -/-- `p ∈ M` is an interior point of a smooth manifold `M` iff +/-- `p ∈ M` is an interior point of a manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target -/-- `p ∈ M` is a boundary point of a smooth manifold `M` iff it is not an interior point. +/-- `p ∈ M` is a boundary point of a manifold `M` iff it is not an interior point. This means that, for `φ` being the preferred chart at `x`, `φ x` is not an interior point of `φ.target`. We do not say "boundary point" as `frontier φ.target` has two components, one on the boundary of range I and another on the boundary of e.target (which we don't want). -/ def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target -namespace SmoothManifoldWithCorners +namespace ManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? variable (I M) in @@ -130,20 +129,20 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB variable (I M) in /-- A manifold decomposes into interior and boundary. -/ -lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ - (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by +lemma univ_eq_interior_union_boundary : (ManifoldWithCorners.interior I M) ∪ + (ManifoldWithCorners.boundary I M) = (univ : Set M) := by apply le_antisymm · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x /-- The interior and boundary of `M` are disjoint. -/ -- xxx: name `..._eq_empty` instead? lemma interior_boundary_disjoint : - (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by + (ManifoldWithCorners.interior I M) ∩ (ManifoldWithCorners.boundary I M) = ∅ := by ext exact ⟨fun h ↦ (not_mem_of_mem_diff h) (mem_of_mem_diff h), by exfalso⟩ /-- The interior of a manifold is an open subset. -/ -lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by +lemma interior_isOpen : IsOpen (ManifoldWithCorners.interior I M) := by apply isOpen_iff_forall_mem_open.mpr intro x hx -- Consider the preferred chart at `x`. Its extended chart has open interior. @@ -164,13 +163,13 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by exact mem_inter this hx /-- The boundary of a manifold is a closed subset. -/ -lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := by +lemma boundary_isClosed : IsClosed (ManifoldWithCorners.boundary I M) := by apply isOpen_compl_iff.mp - have : (SmoothManifoldWithCorners.interior I M)ᶜ = SmoothManifoldWithCorners.boundary I M := + have : (ManifoldWithCorners.interior I M)ᶜ = ManifoldWithCorners.boundary I M := (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary I M)) rw [← this, compl_compl] exact interior_isOpen -end SmoothManifoldWithCorners +end ManifoldWithCorners variable (I) in /-- The boundary of any extended chart has empty interior. -/ @@ -188,7 +187,7 @@ lemma LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H namespace SmoothManifoldWithCorners variable (I) in /-- The boundary of a manifold has empty interior. -/ -lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by +lemma interior_boundary_eq_empty : interior (ManifoldWithCorners.boundary I M) = ∅ := by -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` sorry diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 2a7f174bceec9..fb5cdadb5ecc1 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -807,6 +807,24 @@ theorem mem_analyticGroupoid_of_boundaryless [CompleteSpace E] [I.Boundaryless] end analyticGroupoid +/-! Topological manifolds with corners: no smoothness assumed, but boundary and/or corners +are possible. -/ +section ManifoldWithCorners +/-- Typeclass defining topological manifolds with corners with respect to a model with corners, +over a field `𝕜`. -/ +class ManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] extends + HasGroupoid M (contDiffGroupoid 0 I) : Prop + +theorem ManifoldWithCorners.mk' {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] + [gr : HasGroupoid M (contDiffGroupoid 0 I)] : ManifoldWithCorners I M := + { gr with } + +end ManifoldWithCorners + section SmoothManifoldWithCorners /-! ### Smooth manifolds with corners -/ From 79fcb13c1d0cb113cb1af8fa4266bc6868213ea3 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 23:03:45 +0100 Subject: [PATCH 038/203] Alternative, more minimal fix. Then, implication is mostly handled by the type-class system. Hmm. --- .../Geometry/Manifold/InteriorBoundary.lean | 20 +++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 419db7299ac25..c63cb1d5828fe 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -73,7 +73,7 @@ end TopologyHelpers variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} - {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [ManifoldWithCorners I M] + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [HasGroupoid M (contDiffGroupoid 0 I)] /-- `p ∈ M` is an interior point of a manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ @@ -85,7 +85,7 @@ This means that, for `φ` being the preferred chart at `x`, `φ x` is not an int boundary of range I and another on the boundary of e.target (which we don't want). -/ def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target -namespace ManifoldWithCorners +namespace SmoothManifoldWithCorners -- FIXME(MR): can I enable dot notation, like `M.interior I` or so? variable (I M) in @@ -129,20 +129,20 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB variable (I M) in /-- A manifold decomposes into interior and boundary. -/ -lemma univ_eq_interior_union_boundary : (ManifoldWithCorners.interior I M) ∪ - (ManifoldWithCorners.boundary I M) = (univ : Set M) := by +lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ + (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by apply le_antisymm · exact fun x _ ↦ trivial · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x /-- The interior and boundary of `M` are disjoint. -/ -- xxx: name `..._eq_empty` instead? lemma interior_boundary_disjoint : - (ManifoldWithCorners.interior I M) ∩ (ManifoldWithCorners.boundary I M) = ∅ := by + (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by ext exact ⟨fun h ↦ (not_mem_of_mem_diff h) (mem_of_mem_diff h), by exfalso⟩ /-- The interior of a manifold is an open subset. -/ -lemma interior_isOpen : IsOpen (ManifoldWithCorners.interior I M) := by +lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by apply isOpen_iff_forall_mem_open.mpr intro x hx -- Consider the preferred chart at `x`. Its extended chart has open interior. @@ -163,13 +163,13 @@ lemma interior_isOpen : IsOpen (ManifoldWithCorners.interior I M) := by exact mem_inter this hx /-- The boundary of a manifold is a closed subset. -/ -lemma boundary_isClosed : IsClosed (ManifoldWithCorners.boundary I M) := by +lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := by apply isOpen_compl_iff.mp - have : (ManifoldWithCorners.interior I M)ᶜ = ManifoldWithCorners.boundary I M := + have : (SmoothManifoldWithCorners.interior I M)ᶜ = SmoothManifoldWithCorners.boundary I M := (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary I M)) rw [← this, compl_compl] exact interior_isOpen -end ManifoldWithCorners +end SmoothManifoldWithCorners variable (I) in /-- The boundary of any extended chart has empty interior. -/ @@ -187,7 +187,7 @@ lemma LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H namespace SmoothManifoldWithCorners variable (I) in /-- The boundary of a manifold has empty interior. -/ -lemma interior_boundary_eq_empty : interior (ManifoldWithCorners.boundary I M) = ∅ := by +lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` sorry From e1a93636ef3516ffd510a9d412a058c42311847e Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Tue, 21 Nov 2023 23:56:58 +0100 Subject: [PATCH 039/203] sketch boundary argument --- .../Geometry/Manifold/InteriorBoundary.lean | 58 ++++++++++++++++++- 1 file changed, 56 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index c63cb1d5828fe..0d156a4ba80f9 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -188,8 +188,62 @@ namespace SmoothManifoldWithCorners variable (I) in /-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by - -- use `isBoundaryPoint_iff` and the previous lemma; similar to `interior_isOpen` - sorry + -- Easy argument. Postponed for now. + have aux1 : ⋃ x : M, (extChartAt I x).source = univ := sorry + + let bd := SmoothManifoldWithCorners.boundary I M + have := calc interior (SmoothManifoldWithCorners.boundary I M) + _ = interior (bd) := rfl + _ = interior (bd) ∩ univ := by rw [inter_univ] + _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [aux1] + -- perhaps permute the next steps a bit + --_ = interior (bd ∩ ⋃ (x : M), (extChartAt I x).source) := by sorry + --_ = ⋃ (x : M), interior bd ∩ (extChartAt I x).source := sorry + _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := sorry + _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] + + -- Apply my characterisation of boundary points. + have : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by + intro x + have r' : (chartAt H x).extend I = extChartAt I x := rfl + rw [← r'] + ext y + constructor + · rintro ⟨hbd, hsource⟩ + rw [(chartAt H x).extend_source] at hsource + let s := (isBoundaryPoint_iff I hsource).mp hbd + apply mem_inter + · sorry -- is y ∈ target? true for local homeos... + · apply s + · intro hy + have : y ∈ (chartAt H x).source := sorry -- TODO! + apply mem_inter + · apply (isBoundaryPoint_iff I this).mpr (not_mem_of_mem_diff hy) + · rw [(chartAt H x).extend_source] + exact this + + have := calc interior (bd) + _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := sorry -- computation 1 + _ = ⋃ (x : M), interior ((extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [this] + + -- this step is SCIFI: very happy if true!! need a rigorous argument, though + -- extChart is continuous on its source, so this might hold? + -- lemma: f is continuous, then interior f⁻¹'B = f⁻¹ (interior B) + -- next up: f continuous on A, then A ∩ interior f⁻¹'(B) = f⁻¹(interior B) assuming f⁻¹B ⊆ A somehow + _ = ⋃ (x : M), (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := sorry + + _ = ⋃ (x : M), ∅ := by --sorry + have aux : ∀ x : M, (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by + intro x + set e := extChartAt I x + have : interior ((e.target \ interior e.target)) = ∅ := by + have : (chartAt H x).extend I = e := rfl + apply this ▸ ((chartAt H x).extend_interior_boundary_eq_empty (I := I)) + rw [this, preimage_empty, inter_empty] + simp_rw [aux] + _ = ∅ := iUnion_empty + -- use `isBoundaryPoint_iff`; similar to `interior_isOpen` + exact this--sorry -- interior I M is a manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) end SmoothManifoldWithCorners From f38f823c3a7fcc5ee981cf64a8a6fb25fdf8d900 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 00:06:31 +0100 Subject: [PATCH 040/203] MAYBE this tweak makes the proof easier. It shouldn't make it harder, though. One proof broken, not fixed yet. Should not be bad in principle. --- .../Geometry/Manifold/InteriorBoundary.lean | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 0d156a4ba80f9..10d97241955db 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -203,28 +203,28 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] -- Apply my characterisation of boundary points. - have : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by + have : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by intro x have r' : (chartAt H x).extend I = extChartAt I x := rfl rw [← r'] ext y constructor · rintro ⟨hbd, hsource⟩ + apply mem_inter hsource ?_ -- discharge first condition, easy rw [(chartAt H x).extend_source] at hsource let s := (isBoundaryPoint_iff I hsource).mp hbd - apply mem_inter - · sorry -- is y ∈ target? true for local homeos... - · apply s - · intro hy - have : y ∈ (chartAt H x).source := sorry -- TODO! - apply mem_inter - · apply (isBoundaryPoint_iff I this).mpr (not_mem_of_mem_diff hy) - · rw [(chartAt H x).extend_source] - exact this + sorry --apply (mem_diff y).mp ?_ s + --apply s--mem_inter + --· sorry -- is y ∈ target? true for local homeos... + --· apply s + · rintro ⟨hsource, hbd⟩ + apply mem_inter ?_ hsource + rw [(chartAt H x).extend_source] at hsource + apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) have := calc interior (bd) _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := sorry -- computation 1 - _ = ⋃ (x : M), interior ((extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [this] + _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [this] -- this step is SCIFI: very happy if true!! need a rigorous argument, though -- extChart is continuous on its source, so this might hold? From 07d1dbca44a9f0264dfc866d3b8a9df8ca2a45ec Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 00:11:31 +0100 Subject: [PATCH 041/203] Cleanup, one easy sorry. --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 17 +++++++++++------ 1 file changed, 11 insertions(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 10d97241955db..0ad637675fef6 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -188,8 +188,14 @@ namespace SmoothManifoldWithCorners variable (I) in /-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by - -- Easy argument. Postponed for now. - have aux1 : ⋃ x : M, (extChartAt I x).source = univ := sorry + -- The chart domains of M cover M. Should be an easy lemma; extract if not already exists. + -- Or can I avoid this? + have aux1 : ⋃ x : M, (extChartAt I x).source = univ := by + apply subset_antisymm <;> intro y hy + · trivial + · rw [mem_iUnion] + use y + exact mem_extChartAt_source I y let bd := SmoothManifoldWithCorners.boundary I M have := calc interior (SmoothManifoldWithCorners.boundary I M) @@ -223,7 +229,7 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) have := calc interior (bd) - _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := sorry -- computation 1 + _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := sorry -- previous computation _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [this] -- this step is SCIFI: very happy if true!! need a rigorous argument, though @@ -232,7 +238,7 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary -- next up: f continuous on A, then A ∩ interior f⁻¹'(B) = f⁻¹(interior B) assuming f⁻¹B ⊆ A somehow _ = ⋃ (x : M), (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := sorry - _ = ⋃ (x : M), ∅ := by --sorry + _ = ⋃ (x : M), ∅ := by have aux : ∀ x : M, (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by intro x set e := extChartAt I x @@ -242,8 +248,7 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary rw [this, preimage_empty, inter_empty] simp_rw [aux] _ = ∅ := iUnion_empty - -- use `isBoundaryPoint_iff`; similar to `interior_isOpen` - exact this--sorry + exact this -- interior I M is a manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) end SmoothManifoldWithCorners From c331ee276f549f59c955478d25485abd01d9eda4 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 00:20:51 +0100 Subject: [PATCH 042/203] Another easy sorry. --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 0ad637675fef6..7bdbf770c26bc 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -202,10 +202,14 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary _ = interior (bd) := rfl _ = interior (bd) ∩ univ := by rw [inter_univ] _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [aux1] - -- perhaps permute the next steps a bit - --_ = interior (bd ∩ ⋃ (x : M), (extChartAt I x).source) := by sorry - --_ = ⋃ (x : M), interior bd ∩ (extChartAt I x).source := sorry - _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := sorry + _ = interior (bd) ∩ ⋃ (x : M), interior ((extChartAt I x).source) := by + have : ∀ x : M, interior ((extChartAt I x).source) = (extChartAt I x).source := by + intro x + have : extChartAt I x = (chartAt H x).extend I := rfl + rw [this] + exact (chartAt H x).isOpen_extend_source I (M := M).interior_eq + simp_rw [this] + _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := inter_iUnion _ _ _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] -- Apply my characterisation of boundary points. From 191a5c14e9bd6a6ebcbde2e18de5b9f4f48f015a Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 00:22:47 +0100 Subject: [PATCH 043/203] Move a helper upfront; eliminating another sorry. --- .../Geometry/Manifold/InteriorBoundary.lean | 34 +++++++++---------- 1 file changed, 16 insertions(+), 18 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 7bdbf770c26bc..5f5627afc3277 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -198,22 +198,8 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary exact mem_extChartAt_source I y let bd := SmoothManifoldWithCorners.boundary I M - have := calc interior (SmoothManifoldWithCorners.boundary I M) - _ = interior (bd) := rfl - _ = interior (bd) ∩ univ := by rw [inter_univ] - _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [aux1] - _ = interior (bd) ∩ ⋃ (x : M), interior ((extChartAt I x).source) := by - have : ∀ x : M, interior ((extChartAt I x).source) = (extChartAt I x).source := by - intro x - have : extChartAt I x = (chartAt H x).extend I := rfl - rw [this] - exact (chartAt H x).isOpen_extend_source I (M := M).interior_eq - simp_rw [this] - _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := inter_iUnion _ _ - _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] - -- Apply my characterisation of boundary points. - have : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by + have aux2 : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by intro x have r' : (chartAt H x).extend I = extChartAt I x := rfl rw [← r'] @@ -232,9 +218,21 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary rw [(chartAt H x).extend_source] at hsource apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) - have := calc interior (bd) - _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := sorry -- previous computation - _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [this] + -- Now, compute. + have := calc interior (SmoothManifoldWithCorners.boundary I M) + _ = interior (bd) := rfl + _ = interior (bd) ∩ univ := by rw [inter_univ] + _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [aux1] + _ = interior (bd) ∩ ⋃ (x : M), interior ((extChartAt I x).source) := by + have : ∀ x : M, interior ((extChartAt I x).source) = (extChartAt I x).source := by + intro x + have : extChartAt I x = (chartAt H x).extend I := rfl + rw [this] + exact (chartAt H x).isOpen_extend_source I (M := M).interior_eq + simp_rw [this] + _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := inter_iUnion _ _ + _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] + _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [aux2] -- this step is SCIFI: very happy if true!! need a rigorous argument, though -- extChart is continuous on its source, so this might hold? From 66b1943bfee426b8a121b792fc52e78fdfac9c86 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 00:38:06 +0100 Subject: [PATCH 044/203] Starting at scifi lemma: might be science *fiction* actually... --- .../Geometry/Manifold/InteriorBoundary.lean | 35 ++++++++++++++++++- 1 file changed, 34 insertions(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 5f5627afc3277..941efb53dbf0d 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -5,6 +5,7 @@ Authors: Michael Rothgang, Winston Yin -/ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners +import Mathlib.Topology.Maps -- FIXME: should this be its own file or go in SmoothManifoldWithCorners? -- the latter is already huge, or its own file - move other results about boundaryless here? @@ -184,6 +185,24 @@ lemma LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H -- As `range I` is closed, its frontier has empty interior. exact interior_frontier I.closed_range + +lemma auxaux {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : X → Y} + {O : Set X} {t : Set Y} (hO : IsOpen O) (hf : ContinuousOn f O) : + interior (O ∩ f ⁻¹' t) = O ∩ f ⁻¹' (interior t) := by + -- well, ⊇ always holds! + have : interior (O ∩ f ⁻¹' t) ⊇ O ∩ f ⁻¹' (interior t) := by + rw [interior_inter, hO.interior_eq] + exact hf.preimage_interior_subset_interior_preimage hO + + -- **IF** f were an open map, I'd be happy... that only holds without boundary, though... + have pretend : IsOpenMap f := sorry + let r := pretend.interior_preimage_subset_preimage_interior (s := t) + have : interior (O ∩ f ⁻¹' t) ⊆ O ∩ f ⁻¹' (interior t) := by + rw [interior_inter, hO.interior_eq] + exact inter_subset_inter_right O r + sorry -- sCIFI! + + namespace SmoothManifoldWithCorners variable (I) in /-- The boundary of a manifold has empty interior. -/ @@ -238,7 +257,21 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary -- extChart is continuous on its source, so this might hold? -- lemma: f is continuous, then interior f⁻¹'B = f⁻¹ (interior B) -- next up: f continuous on A, then A ∩ interior f⁻¹'(B) = f⁻¹(interior B) assuming f⁻¹B ⊆ A somehow - _ = ⋃ (x : M), (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := sorry + _ = ⋃ (x : M), (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by + have goal : ∀ x : M, interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) = (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by + intro x + set e := extChartAt I x + let r := (chartAt H x).continuousOn_extend I + have : (chartAt H x).extend I = e := rfl + rw [this] at r + -- interior (e.source ∩ ↑e ⁻¹' (e.target \ interior e.target)) = e.source ∩ ↑e ⁻¹' interior (e.target \ interior e.target) + have : IsOpen e.source := sorry -- easy + -- abstracted this into a lemma. now, let's see if that is actually true!!! + -- well, one direction holds - but it's the wrong one... + apply auxaux (O := e.source) this r (t := e.target \ interior e.target) + simp_rw [goal] + --let r := (chartAt H x).continuousOn_extend I + --sorry _ = ⋃ (x : M), ∅ := by have aux : ∀ x : M, (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by From 7840b91d927177d4cf8d937a41616878ab9f09c0 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 13:53:44 +0100 Subject: [PATCH 045/203] Small golfs. --- .../Geometry/Manifold/InteriorBoundary.lean | 22 +++++++------------ 1 file changed, 8 insertions(+), 14 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 941efb53dbf0d..085acdf72097e 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -5,7 +5,6 @@ Authors: Michael Rothgang, Winston Yin -/ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners -import Mathlib.Topology.Maps -- FIXME: should this be its own file or go in SmoothManifoldWithCorners? -- the latter is already huge, or its own file - move other results about boundaryless here? @@ -45,28 +44,23 @@ open Set section TopologyHelpers -- should be in mathlib; Mathlib.Topology.Basic variable {X : Type*} [TopologicalSpace X] {s : Set X} --- I don't need it; still useful? +-- I don't need this lemma; is is useful independently itself? lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by - rw [← closure_diff_interior s] - have : (closure s \ interior s) = closure s ∩ (interior s)ᶜ := rfl - rw [this] + rw [← closure_diff_interior s, diff_eq] rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] -- FIXME: what's a good name? lemma aux {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : s \ interior s ⊆ t \ interior t := by - let aux := calc interior s - _ = interior O ∩ interior t := by rw [h, interior_inter] - _ = O ∩ interior t := by rw [hO.interior_eq] - calc s \ interior s - _ = (O ∩ t) \ (O ∩ interior t) := by rw [aux, h] - _ = O ∩ (t \ interior t) := by rw [inter_diff_distrib_left] - _ ⊆ t \ interior t := inter_subset_right _ _ + rw [h, interior_inter, hO.interior_eq, ← inter_diff_distrib_left] + exact inter_subset_right O (t \ interior t) -- is this a better lemma; is `aux` useful on its own? lemma aux2 {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) - (ht : IsClosed t) : s \ interior s ⊆ frontier t := - ht.frontier_eq ▸ aux h hO + (ht : IsClosed t) : s \ interior s ⊆ frontier t := by + rw [ht.frontier_eq, h, interior_inter, hO.interior_eq, ← inter_diff_distrib_left] + exact inter_subset_right _ _ + -- alternative proof, if `aux` is useful: ht.frontier_eq ▸ aux h hO end TopologyHelpers From 61fd70ef9c38b56aadd52f6de0cf5cd501aa0be7 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 14:07:18 +0100 Subject: [PATCH 046/203] Extract lemmas used in the proof so far. --- .../Geometry/Manifold/InteriorBoundary.lean | 74 ++++++++++--------- 1 file changed, 40 insertions(+), 34 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 085acdf72097e..730892f218465 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -196,46 +196,50 @@ lemma auxaux {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : X → exact inter_subset_inter_right O r sorry -- sCIFI! +/-- The charts of a charted space cover its domain. -/ +-- {H M : Type*} [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] +lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by + apply subset_antisymm <;> intro y _ + · trivial + · rw [mem_iUnion] + use y + exact mem_chart_source H y namespace SmoothManifoldWithCorners +-- can I avoid using this lemma? +lemma covering : ⋃ x : M, (extChartAt I x).source = univ := by + simp_rw [extChartAt_source] + exact ChartedSpace.covering + +lemma isBoundaryPoint_iff' {x : M} : + SmoothManifoldWithCorners.boundary I M ∩ (extChartAt I x).source = + (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' + ((extChartAt I x).target \ interior (extChartAt I x).target) := by + have r' : (chartAt H x).extend I = extChartAt I x := rfl + rw [← r'] + ext y + constructor + · rintro ⟨hbd, hsource⟩ + apply mem_inter hsource ?_ -- discharge first condition, easy + rw [(chartAt H x).extend_source] at hsource + let s := (isBoundaryPoint_iff I hsource).mp hbd + sorry --apply (mem_diff y).mp ?_ s + --apply s--mem_inter + --· sorry -- is y ∈ target? true for local homeos... + --· apply s + · rintro ⟨hsource, hbd⟩ + apply mem_inter ?_ hsource + rw [(chartAt H x).extend_source] at hsource + apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) + variable (I) in /-- The boundary of a manifold has empty interior. -/ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by - -- The chart domains of M cover M. Should be an easy lemma; extract if not already exists. - -- Or can I avoid this? - have aux1 : ⋃ x : M, (extChartAt I x).source = univ := by - apply subset_antisymm <;> intro y hy - · trivial - · rw [mem_iUnion] - use y - exact mem_extChartAt_source I y - - let bd := SmoothManifoldWithCorners.boundary I M - -- Apply my characterisation of boundary points. - have aux2 : ∀ x : M, bd ∩ (extChartAt I x).source = (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by - intro x - have r' : (chartAt H x).extend I = extChartAt I x := rfl - rw [← r'] - ext y - constructor - · rintro ⟨hbd, hsource⟩ - apply mem_inter hsource ?_ -- discharge first condition, easy - rw [(chartAt H x).extend_source] at hsource - let s := (isBoundaryPoint_iff I hsource).mp hbd - sorry --apply (mem_diff y).mp ?_ s - --apply s--mem_inter - --· sorry -- is y ∈ target? true for local homeos... - --· apply s - · rintro ⟨hsource, hbd⟩ - apply mem_inter ?_ hsource - rw [(chartAt H x).extend_source] at hsource - apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) - + set bd := SmoothManifoldWithCorners.boundary I M -- Now, compute. - have := calc interior (SmoothManifoldWithCorners.boundary I M) - _ = interior (bd) := rfl + have := calc interior (bd) _ = interior (bd) ∩ univ := by rw [inter_univ] - _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [aux1] + _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [covering] _ = interior (bd) ∩ ⋃ (x : M), interior ((extChartAt I x).source) := by have : ∀ x : M, interior ((extChartAt I x).source) = (extChartAt I x).source := by intro x @@ -245,7 +249,9 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary simp_rw [this] _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := inter_iUnion _ _ _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] - _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [aux2] + _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ + (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by + simp_rw [isBoundaryPoint_iff'] -- this step is SCIFI: very happy if true!! need a rigorous argument, though -- extChart is continuous on its source, so this might hold? From 5869bd54a39e0df77c3d44bd85ff38524cf48816 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 14:23:08 +0100 Subject: [PATCH 047/203] Fill sorry in description of boundary. --- .../Geometry/Manifold/InteriorBoundary.lean | 18 ++++++++++++++---- 1 file changed, 14 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 730892f218465..05eed129a9640 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -218,15 +218,25 @@ lemma isBoundaryPoint_iff' {x : M} : have r' : (chartAt H x).extend I = extChartAt I x := rfl rw [← r'] ext y + -- This can surely be golfed: first three lines on both cases are the same. + -- First steps: reorder target conditions; then try and_congr or so... constructor · rintro ⟨hbd, hsource⟩ apply mem_inter hsource ?_ -- discharge first condition, easy rw [(chartAt H x).extend_source] at hsource let s := (isBoundaryPoint_iff I hsource).mp hbd - sorry --apply (mem_diff y).mp ?_ s - --apply s--mem_inter - --· sorry -- is y ∈ target? true for local homeos... - --· apply s + -- This part can surely also be golfed! + set e := chartAt H x + set e' := e.extend I -- readability + rw [mem_preimage] + apply (mem_diff (e' y)).mpr + constructor + · rw [r', extChartAt_target] + apply mem_inter ?_ (mem_range_self _) + show I (e y) ∈ I.symm ⁻¹' e.target + rw [mem_preimage, I.left_inv] + exact e.map_source hsource + · exact s · rintro ⟨hsource, hbd⟩ apply mem_inter ?_ hsource rw [(chartAt H x).extend_source] at hsource From 058bb709c9b22811c78a27a1cd4beee8fbf96fb6 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 14:25:16 +0100 Subject: [PATCH 048/203] Extract one more useful lemma; tiny golf. --- .../Geometry/Manifold/InteriorBoundary.lean | 33 ++++++++++++------- 1 file changed, 21 insertions(+), 12 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 05eed129a9640..80acafbe56a49 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -211,12 +211,30 @@ lemma covering : ⋃ x : M, (extChartAt I x).source = univ := by simp_rw [extChartAt_source] exact ChartedSpace.covering +-- XXX: fix name; move to LocalHomeomorph +lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : + (e.extend I) x ∈ (e.extend I).target := by + rw [e.extend_target] + apply mem_inter ?_ (mem_range_self _) + show I (e x) ∈ I.symm ⁻¹' e.target + rw [mem_preimage, I.left_inv] + exact e.map_source hx + +-- XXX: fix name; move to LocalHomeomorph +lemma extend_source_map_target' {e : LocalHomeomorph M H} : + (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := by + --rw [e.extend_source] + --intro y hy + --rw [mem_image.mp (e.extend I) (e.extend I.source)] + --apply (extend_source_map_target hx) + sorry -- easy + lemma isBoundaryPoint_iff' {x : M} : SmoothManifoldWithCorners.boundary I M ∩ (extChartAt I x).source = (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by have r' : (chartAt H x).extend I = extChartAt I x := rfl - rw [← r'] + rw [← r'] -- XXX: does this mean extChart needs more lemmas? or a named rewrite equation? ext y -- This can surely be golfed: first three lines on both cases are the same. -- First steps: reorder target conditions; then try and_congr or so... @@ -224,19 +242,10 @@ lemma isBoundaryPoint_iff' {x : M} : · rintro ⟨hbd, hsource⟩ apply mem_inter hsource ?_ -- discharge first condition, easy rw [(chartAt H x).extend_source] at hsource - let s := (isBoundaryPoint_iff I hsource).mp hbd -- This part can surely also be golfed! - set e := chartAt H x - set e' := e.extend I -- readability rw [mem_preimage] - apply (mem_diff (e' y)).mpr - constructor - · rw [r', extChartAt_target] - apply mem_inter ?_ (mem_range_self _) - show I (e y) ∈ I.symm ⁻¹' e.target - rw [mem_preimage, I.left_inv] - exact e.map_source hsource - · exact s + apply (mem_diff ((chartAt H x).extend I y)).mpr + exact ⟨extend_source_map_target hsource, (isBoundaryPoint_iff I hsource).mp hbd⟩ · rintro ⟨hsource, hbd⟩ apply mem_inter ?_ hsource rw [(chartAt H x).extend_source] at hsource From 7e1542d83dc903dd2c9078c27fd8dd3ec05ea03e Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 14:50:53 +0100 Subject: [PATCH 049/203] Simplify: use chart source instead of extChart.source. That's the normal form (and conceptually simpler); jumping between these different forms means I need to rewrite more often. Most of the time, I don't care which one I use. Delete a helper lemma about this bad normal form: shouldn't use that. --- .../Geometry/Manifold/InteriorBoundary.lean | 48 +++++++------------ 1 file changed, 16 insertions(+), 32 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 80acafbe56a49..a025c95021b66 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -206,11 +206,6 @@ lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by exact mem_chart_source H y namespace SmoothManifoldWithCorners --- can I avoid using this lemma? -lemma covering : ⋃ x : M, (extChartAt I x).source = univ := by - simp_rw [extChartAt_source] - exact ChartedSpace.covering - -- XXX: fix name; move to LocalHomeomorph lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : (e.extend I) x ∈ (e.extend I).target := by @@ -230,25 +225,20 @@ lemma extend_source_map_target' {e : LocalHomeomorph M H} : sorry -- easy lemma isBoundaryPoint_iff' {x : M} : - SmoothManifoldWithCorners.boundary I M ∩ (extChartAt I x).source = - (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' + SmoothManifoldWithCorners.boundary I M ∩ (chartAt H x).source = + (chartAt H x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target) := by - have r' : (chartAt H x).extend I = extChartAt I x := rfl - rw [← r'] -- XXX: does this mean extChart needs more lemmas? or a named rewrite equation? ext y -- This can surely be golfed: first three lines on both cases are the same. -- First steps: reorder target conditions; then try and_congr or so... constructor · rintro ⟨hbd, hsource⟩ - apply mem_inter hsource ?_ -- discharge first condition, easy - rw [(chartAt H x).extend_source] at hsource - -- This part can surely also be golfed! + apply mem_inter hsource ?_ rw [mem_preimage] apply (mem_diff ((chartAt H x).extend I y)).mpr exact ⟨extend_source_map_target hsource, (isBoundaryPoint_iff I hsource).mp hbd⟩ · rintro ⟨hsource, hbd⟩ apply mem_inter ?_ hsource - rw [(chartAt H x).extend_source] at hsource apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) variable (I) in @@ -258,42 +248,36 @@ lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary -- Now, compute. have := calc interior (bd) _ = interior (bd) ∩ univ := by rw [inter_univ] - _ = interior (bd) ∩ ⋃ (x : M), (extChartAt I x).source := by simp_rw [covering] - _ = interior (bd) ∩ ⋃ (x : M), interior ((extChartAt I x).source) := by - have : ∀ x : M, interior ((extChartAt I x).source) = (extChartAt I x).source := by + _ = interior (bd) ∩ ⋃ (x : M), (chartAt H x).source := by simp_rw [ChartedSpace.covering] + _ = interior (bd) ∩ ⋃ (x : M), interior ((chartAt H x).source) := by + have : ∀ x : M, interior ((chartAt H x).source) = (chartAt H x).source := by intro x - have : extChartAt I x = (chartAt H x).extend I := rfl - rw [this] - exact (chartAt H x).isOpen_extend_source I (M := M).interior_eq + exact (chartAt H x).open_source.interior_eq simp_rw [this] - _ = ⋃ (x : M), interior bd ∩ interior ((extChartAt I x).source) := inter_iUnion _ _ - _ = ⋃ (x : M), interior (bd ∩ (extChartAt I x).source) := by simp_rw [interior_inter] - _ = ⋃ (x : M), interior ((extChartAt I x).source ∩ + _ = ⋃ (x : M), interior bd ∩ interior ((chartAt H x).source) := inter_iUnion _ _ + _ = ⋃ (x : M), interior (bd ∩ (chartAt H x).source) := by simp_rw [interior_inter] + _ = ⋃ (x : M), interior ((chartAt H x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by simp_rw [isBoundaryPoint_iff'] - -- this step is SCIFI: very happy if true!! need a rigorous argument, though -- extChart is continuous on its source, so this might hold? -- lemma: f is continuous, then interior f⁻¹'B = f⁻¹ (interior B) -- next up: f continuous on A, then A ∩ interior f⁻¹'(B) = f⁻¹(interior B) assuming f⁻¹B ⊆ A somehow - _ = ⋃ (x : M), (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by - have goal : ∀ x : M, interior ((extChartAt I x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) = (extChartAt I x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by + _ = ⋃ (x : M), (chartAt H x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by + have goal : ∀ x : M, interior ((chartAt H x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) = (chartAt H x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by intro x set e := extChartAt I x let r := (chartAt H x).continuousOn_extend I have : (chartAt H x).extend I = e := rfl rw [this] at r -- interior (e.source ∩ ↑e ⁻¹' (e.target \ interior e.target)) = e.source ∩ ↑e ⁻¹' interior (e.target \ interior e.target) - have : IsOpen e.source := sorry -- easy -- abstracted this into a lemma. now, let's see if that is actually true!!! -- well, one direction holds - but it's the wrong one... - apply auxaux (O := e.source) this r (t := e.target \ interior e.target) + rw [← (chartAt H x).extend_source I] + exact auxaux (O := e.source) (isOpen_extChartAt_source I x) r (t := e.target \ interior e.target) simp_rw [goal] - --let r := (chartAt H x).continuousOn_extend I - --sorry - - _ = ⋃ (x : M), ∅ := by - have aux : ∀ x : M, (extChartAt I x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by + _ = ⋃ (_ : M), ∅ := by + have aux : ∀ x : M, (chartAt H x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by intro x set e := extChartAt I x have : interior ((e.target \ interior e.target)) = ∅ := by From 219f85cf41971e3b6dd8148cdf9e9f78dd39ffa1 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 14:59:56 +0100 Subject: [PATCH 050/203] Fill in one more helper sorry. --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index a025c95021b66..27af15c52bad4 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -216,13 +216,14 @@ lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.s exact e.map_source hx -- XXX: fix name; move to LocalHomeomorph +-- xxx: can this be golfed? lemma extend_source_map_target' {e : LocalHomeomorph M H} : (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := by - --rw [e.extend_source] - --intro y hy - --rw [mem_image.mp (e.extend I) (e.extend I.source)] - --apply (extend_source_map_target hx) - sorry -- easy + rw [e.extend_source] + intro y hy + choose x hx hxy using ((mem_image (e.extend I) e.source) y).mp hy + rw [← hxy] + exact extend_source_map_target hx lemma isBoundaryPoint_iff' {x : M} : SmoothManifoldWithCorners.boundary I M ∩ (chartAt H x).source = From b97611ea20b82d6515a822b4b072a24c067e4d53 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 15:13:21 +0100 Subject: [PATCH 051/203] Move helpers to better namespace; show a MapsTo version; mathlib has a related version, with a slightly different formula... --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 17 ++++++++++++----- 1 file changed, 12 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 27af15c52bad4..65e20a84d34ea 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -205,8 +205,8 @@ lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by use y exact mem_chart_source H y -namespace SmoothManifoldWithCorners --- XXX: fix name; move to LocalHomeomorph +namespace LocalHomeomorph -- FIXME: move to SmoothManifoldWithCorners +-- XXX: fix name lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : (e.extend I) x ∈ (e.extend I).target := by rw [e.extend_target] @@ -215,8 +215,13 @@ lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.s rw [mem_preimage, I.left_inv] exact e.map_source hx --- XXX: fix name; move to LocalHomeomorph --- xxx: can this be golfed? +-- XXX: mapsTo_extend has a different formula from extend_target... +-- should unify these/relate these better! +lemma LocalHomeomorph.mapsTo_source_target {e : LocalHomeomorph M H} : + MapsTo (e.extend I) (e.extend I).source (e.extend I).target := + fun _ hx ↦(e.extend I).map_source hx + +-- XXX: fix name; can this be golfed? lemma extend_source_map_target' {e : LocalHomeomorph M H} : (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := by rw [e.extend_source] @@ -224,7 +229,9 @@ lemma extend_source_map_target' {e : LocalHomeomorph M H} : choose x hx hxy using ((mem_image (e.extend I) e.source) y).mp hy rw [← hxy] exact extend_source_map_target hx +end LocalHomeomorph +namespace SmoothManifoldWithCorners lemma isBoundaryPoint_iff' {x : M} : SmoothManifoldWithCorners.boundary I M ∩ (chartAt H x).source = (chartAt H x).source ∩ (extChartAt I x) ⁻¹' @@ -237,7 +244,7 @@ lemma isBoundaryPoint_iff' {x : M} : apply mem_inter hsource ?_ rw [mem_preimage] apply (mem_diff ((chartAt H x).extend I y)).mpr - exact ⟨extend_source_map_target hsource, (isBoundaryPoint_iff I hsource).mp hbd⟩ + exact ⟨(chartAt H x).extend_source_map_target hsource, (isBoundaryPoint_iff I hsource).mp hbd⟩ · rintro ⟨hsource, hbd⟩ apply mem_inter ?_ hsource apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) From ef4c2324d32c245009bcf37ecc6508a472369480 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 16:35:53 +0100 Subject: [PATCH 052/203] Reduce interior independence statement to a local lemma. Extracting all those tiny lemmas still feels slightly weird, but I guess they are still useful, somewhen. apply? cannot find them, at least. --- .../Geometry/Manifold/InteriorBoundary.lean | 47 +++++++++++++++++-- 1 file changed, 44 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 65e20a84d34ea..e7f400855c350 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -90,14 +90,55 @@ protected def interior : Set M := { x : M | I.IsInteriorPoint x} variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} +end SmoothManifoldWithCorners + +namespace LocalHomeomorph -- move to SmoothManifoldsWithCorners! +-- more general lemma underlying foobar. xxx: find a better name! +lemma foobar_abstract {f : LocalHomeomorph H H} {y : H} (hy : y ∈ f.source) + (h : I y ∈ interior (range I)) : I (f y) ∈ interior (range I) := by + sorry + +-- xxx: needs better name! +-- the interior of the target of an extended local homeo is contained in the interior of it's model's range +lemma extend_interior_target_subset {e : LocalHomeomorph M H} : + interior (e.extend I).target ⊆ interior (range I) := by + rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq] + exact inter_subset_right _ _ + +-- xxx: find a good name!! +lemma foobaz {e : LocalHomeomorph M H} {y : H} (hy : y ∈ e.target) + (hy' : I y ∈ interior (range ↑I)) : I y ∈ interior (e.extend I).target := by + rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq, + mem_inter_iff, mem_preimage] + exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ -- as we only need continuity property, e or e' being in the atlas is not required lemma foobar {e e' : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source ∩ e'.source) : (e.extend I) x ∈ interior (e.extend I).target ↔ - (e'.extend I) x ∈ interior (e'.extend I).target := sorry --- both directions should be the same, more general lemma + (e'.extend I) x ∈ interior (e'.extend I).target := by + rcases ((mem_inter_iff x _ _).mp hx) with ⟨hxe, hxe'⟩ + -- reduction, step 1: simplify what the interior means + have : (e.extend I) x ∈ interior (e.extend I).target ↔ I (e x) ∈ interior (range I) := + ⟨fun hx ↦ extend_interior_target_subset hx, fun hx ↦ foobaz (e.map_source hxe) hx⟩ + rw [this] + have : (e'.extend I) x ∈ interior (e'.extend I).target ↔ I (e' x) ∈ interior (range I) := + ⟨fun hx ↦ extend_interior_target_subset hx, fun hx ↦ foobaz (e'.map_source hxe') hx⟩ + rw [this] + -- step 2: rewrite in terms of coordinate changes + constructor + · intro h + let f := e.symm.trans e' + have h2 : e x ∈ f.source := by + have : e.symm (e x) = x := e.left_inv' hxe + rw [LocalHomeomorph.trans_source, mem_inter_iff (e x), e.symm_source, mem_preimage, this] + exact ⟨e.map_source hxe, hxe'⟩ + rw [← (e.left_inv' hxe)] + exact foobar_abstract h2 h + · sorry -- exactly the same... what's the best way to deduplicate? +end LocalHomeomorph +namespace SmoothManifoldWithCorners -- FIXME(MR): find a better wording for the next two docstrings variable (I) in /-- Whether `x` is an interior point can equivalently be described by any chart @@ -105,7 +146,7 @@ variable (I) in -- as we only need continuity properties, `e` being in the atlas is not required lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := - foobar (mem_inter (mem_chart_source H x) hx) + (chartAt H x).foobar (mem_inter (mem_chart_source H x) hx) variable (I) in /-- Whether `x` is a boundary point of `M` can equivalently be described by any chart From 184a0625d5a2e000f16be252a353c308d1df7cd2 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 18:08:38 +0100 Subject: [PATCH 053/203] Give up: remove WIP claim that boundary as empty interior. I don't see how to show this, given the current hypotheses. --- .../Geometry/Manifold/InteriorBoundary.lean | 118 +----------------- 1 file changed, 5 insertions(+), 113 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index e7f400855c350..5c347ffdb1081 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -27,13 +27,14 @@ Define the interior and boundary of a manifold. - `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary - `interior_isOpen`: `interior I M` is open - `boundary_isClosed`: `boundary I M` is closed -- `interior_boundary_eq_empty`: `boundary I M` has empty interior -(this implies it has "measure zero", see different file) **TODO** +- under suitable assumptions, `boundary I M` has empty interior +(if `M` is finite-dimensional, `boundary I M` should have measure 0, which implies this) - `interior I M` is a manifold without boundary (need to upgrade the model used; map the charts from an open ball to entire ℝ^n) -- the boundary is a submanifold of codimension 1 (once mathlib has submanifolds) +- the boundary is a submanifold of codimension 1, perhaps with boundary and corners +(this requires a definition of submanifolds) ## Tags manifold, interior, boundary @@ -49,19 +50,6 @@ lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by rw [← closure_diff_interior s, diff_eq] rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] --- FIXME: what's a good name? -lemma aux {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) : - s \ interior s ⊆ t \ interior t := by - rw [h, interior_inter, hO.interior_eq, ← inter_diff_distrib_left] - exact inter_subset_right O (t \ interior t) - --- is this a better lemma; is `aux` useful on its own? -lemma aux2 {O t : Set X} (h : s = O ∩ t) (hO : IsOpen O) - (ht : IsClosed t) : s \ interior s ⊆ frontier t := by - rw [ht.frontier_eq, h, interior_inter, hO.interior_eq, ← inter_diff_distrib_left] - exact inter_subset_right _ _ - -- alternative proof, if `aux` is useful: ht.frontier_eq ▸ aux h hO - end TopologyHelpers -- Let `M` be a manifold with corners over the pair `(E, H)`. @@ -207,36 +195,6 @@ lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := b exact interior_isOpen end SmoothManifoldWithCorners -variable (I) in -/-- The boundary of any extended chart has empty interior. -/ -lemma LocalHomeomorph.extend_interior_boundary_eq_empty {e : LocalHomeomorph M H} : - interior ((e.extend I).target \ interior (e.extend I).target) = ∅ := by - -- `e.extend_target I = (I.symm ⁻¹' e.target) ∩ range I` is the union of an open set and a - -- closed set: hence the frontier is contained in the second factor. - have h1 : (e.extend I).target \ interior (e.extend I).target ⊆ frontier (range I) := - aux2 (e.extend_target I) (e.open_target.preimage I.continuous_symm) I.closed_range - suffices interior (frontier (range I)) = ∅ by - exact subset_eq_empty (interior_mono h1) this - -- As `range I` is closed, its frontier has empty interior. - exact interior_frontier I.closed_range - - -lemma auxaux {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : X → Y} - {O : Set X} {t : Set Y} (hO : IsOpen O) (hf : ContinuousOn f O) : - interior (O ∩ f ⁻¹' t) = O ∩ f ⁻¹' (interior t) := by - -- well, ⊇ always holds! - have : interior (O ∩ f ⁻¹' t) ⊇ O ∩ f ⁻¹' (interior t) := by - rw [interior_inter, hO.interior_eq] - exact hf.preimage_interior_subset_interior_preimage hO - - -- **IF** f were an open map, I'd be happy... that only holds without boundary, though... - have pretend : IsOpenMap f := sorry - let r := pretend.interior_preimage_subset_preimage_interior (s := t) - have : interior (O ∩ f ⁻¹' t) ⊆ O ∩ f ⁻¹' (interior t) := by - rw [interior_inter, hO.interior_eq] - exact inter_subset_inter_right O r - sorry -- sCIFI! - /-- The charts of a charted space cover its domain. -/ -- {H M : Type*} [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by @@ -272,70 +230,4 @@ lemma extend_source_map_target' {e : LocalHomeomorph M H} : exact extend_source_map_target hx end LocalHomeomorph -namespace SmoothManifoldWithCorners -lemma isBoundaryPoint_iff' {x : M} : - SmoothManifoldWithCorners.boundary I M ∩ (chartAt H x).source = - (chartAt H x).source ∩ (extChartAt I x) ⁻¹' - ((extChartAt I x).target \ interior (extChartAt I x).target) := by - ext y - -- This can surely be golfed: first three lines on both cases are the same. - -- First steps: reorder target conditions; then try and_congr or so... - constructor - · rintro ⟨hbd, hsource⟩ - apply mem_inter hsource ?_ - rw [mem_preimage] - apply (mem_diff ((chartAt H x).extend I y)).mpr - exact ⟨(chartAt H x).extend_source_map_target hsource, (isBoundaryPoint_iff I hsource).mp hbd⟩ - · rintro ⟨hsource, hbd⟩ - apply mem_inter ?_ hsource - apply (isBoundaryPoint_iff I hsource).mpr (not_mem_of_mem_diff hbd) - -variable (I) in -/-- The boundary of a manifold has empty interior. -/ -lemma interior_boundary_eq_empty : interior (SmoothManifoldWithCorners.boundary I M) = ∅ := by - set bd := SmoothManifoldWithCorners.boundary I M - -- Now, compute. - have := calc interior (bd) - _ = interior (bd) ∩ univ := by rw [inter_univ] - _ = interior (bd) ∩ ⋃ (x : M), (chartAt H x).source := by simp_rw [ChartedSpace.covering] - _ = interior (bd) ∩ ⋃ (x : M), interior ((chartAt H x).source) := by - have : ∀ x : M, interior ((chartAt H x).source) = (chartAt H x).source := by - intro x - exact (chartAt H x).open_source.interior_eq - simp_rw [this] - _ = ⋃ (x : M), interior bd ∩ interior ((chartAt H x).source) := inter_iUnion _ _ - _ = ⋃ (x : M), interior (bd ∩ (chartAt H x).source) := by simp_rw [interior_inter] - _ = ⋃ (x : M), interior ((chartAt H x).source ∩ - (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) := by - simp_rw [isBoundaryPoint_iff'] - -- this step is SCIFI: very happy if true!! need a rigorous argument, though - -- extChart is continuous on its source, so this might hold? - -- lemma: f is continuous, then interior f⁻¹'B = f⁻¹ (interior B) - -- next up: f continuous on A, then A ∩ interior f⁻¹'(B) = f⁻¹(interior B) assuming f⁻¹B ⊆ A somehow - _ = ⋃ (x : M), (chartAt H x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by - have goal : ∀ x : M, interior ((chartAt H x).source ∩ (extChartAt I x) ⁻¹' ((extChartAt I x).target \ interior (extChartAt I x).target)) = (chartAt H x).source ∩ ((extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target))) := by - intro x - set e := extChartAt I x - let r := (chartAt H x).continuousOn_extend I - have : (chartAt H x).extend I = e := rfl - rw [this] at r - -- interior (e.source ∩ ↑e ⁻¹' (e.target \ interior e.target)) = e.source ∩ ↑e ⁻¹' interior (e.target \ interior e.target) - -- abstracted this into a lemma. now, let's see if that is actually true!!! - -- well, one direction holds - but it's the wrong one... - rw [← (chartAt H x).extend_source I] - exact auxaux (O := e.source) (isOpen_extChartAt_source I x) r (t := e.target \ interior e.target) - simp_rw [goal] - _ = ⋃ (_ : M), ∅ := by - have aux : ∀ x : M, (chartAt H x).source ∩ (extChartAt I x) ⁻¹' (interior ((extChartAt I x).target \ interior (extChartAt I x).target)) = ∅ := by - intro x - set e := extChartAt I x - have : interior ((e.target \ interior e.target)) = ∅ := by - have : (chartAt H x).extend I = e := rfl - apply this ▸ ((chartAt H x).extend_interior_boundary_eq_empty (I := I)) - rw [this, preimage_empty, inter_empty] - simp_rw [aux] - _ = ∅ := iUnion_empty - exact this - --- interior I M is a manifold (use TopologicalSpace.Opens.instSmoothManifoldWithCornersSubtypeMemOpensInstMembershipInstSetLikeOpensInstTopologicalSpaceSubtypeInstChartedSpace) -end SmoothManifoldWithCorners +-- TODO: interior I M is a manifold From c634f1a7f87451bc382e94a92a13d0549bdde152 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 18:28:32 +0100 Subject: [PATCH 054/203] Rename and golf mapsTo results for extended local homeos. --- .../Geometry/Manifold/InteriorBoundary.lean | 25 +++++++------------ 1 file changed, 9 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 5c347ffdb1081..8875cf3816ac6 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -205,29 +205,22 @@ lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by exact mem_chart_source H y namespace LocalHomeomorph -- FIXME: move to SmoothManifoldWithCorners --- XXX: fix name -lemma extend_source_map_target {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : +lemma extend_map_source {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : (e.extend I) x ∈ (e.extend I).target := by - rw [e.extend_target] - apply mem_inter ?_ (mem_range_self _) - show I (e x) ∈ I.symm ⁻¹' e.target - rw [mem_preimage, I.left_inv] - exact e.map_source hx + rw [← e.extend_source I] at hx + exact (e.extend I).map_source hx -- XXX: mapsTo_extend has a different formula from extend_target... -- should unify these/relate these better! -lemma LocalHomeomorph.mapsTo_source_target {e : LocalHomeomorph M H} : +lemma mapsTo_extend' {e : LocalHomeomorph M H} : MapsTo (e.extend I) (e.extend I).source (e.extend I).target := fun _ hx ↦(e.extend I).map_source hx --- XXX: fix name; can this be golfed? -lemma extend_source_map_target' {e : LocalHomeomorph M H} : - (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := by - rw [e.extend_source] - intro y hy - choose x hx hxy using ((mem_image (e.extend I) e.source) y).mp hy - rw [← hxy] - exact extend_source_map_target hx +/-- Variant of `extend_map_source`, stated for images of subsets. -/ +lemma extend_map_source' {e : LocalHomeomorph M H} : + (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := + Set.mapsTo'.mp e.mapsTo_extend' + end LocalHomeomorph -- TODO: interior I M is a manifold From dd9e73f9f118a6b12e9a9d156af3c592fc8c5168 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 18:41:50 +0100 Subject: [PATCH 055/203] Small golfs and clean-ups. --- .../Geometry/Manifold/InteriorBoundary.lean | 31 +++++++++---------- 1 file changed, 15 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 8875cf3816ac6..a258eec189baf 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -45,10 +45,10 @@ open Set section TopologyHelpers -- should be in mathlib; Mathlib.Topology.Basic variable {X : Type*} [TopologicalSpace X] {s : Set X} --- I don't need this lemma; is is useful independently itself? +/-- Interior and frontier are disjoint. -/ lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by - rw [← closure_diff_interior s, diff_eq] - rw [← inter_assoc, inter_comm, ← inter_assoc, compl_inter_self, empty_inter] + rw [← closure_diff_interior s, diff_eq, ← inter_assoc, inter_comm, ← inter_assoc, + compl_inter_self, empty_inter] end TopologyHelpers @@ -81,28 +81,30 @@ protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} end SmoothManifoldWithCorners namespace LocalHomeomorph -- move to SmoothManifoldsWithCorners! +variable {e e' : LocalHomeomorph M H} + -- more general lemma underlying foobar. xxx: find a better name! lemma foobar_abstract {f : LocalHomeomorph H H} {y : H} (hy : y ∈ f.source) (h : I y ∈ interior (range I)) : I (f y) ∈ interior (range I) := by sorry -- xxx: needs better name! --- the interior of the target of an extended local homeo is contained in the interior of it's model's range -lemma extend_interior_target_subset {e : LocalHomeomorph M H} : - interior (e.extend I).target ⊆ interior (range I) := by +-- the interior of the target of an extended local homeo is contained in the interior of its +-- model's range +lemma extend_interior_target_subset : interior (e.extend I).target ⊆ interior (range I) := by rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq] exact inter_subset_right _ _ -- xxx: find a good name!! -lemma foobaz {e : LocalHomeomorph M H} {y : H} (hy : y ∈ e.target) - (hy' : I y ∈ interior (range ↑I)) : I y ∈ interior (e.extend I).target := by +lemma foobaz {y : H} (hy : y ∈ e.target) (hy' : I y ∈ interior (range ↑I)) : + I y ∈ interior (e.extend I).target := by rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq, mem_inter_iff, mem_preimage] exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ /-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ -- as we only need continuity property, e or e' being in the atlas is not required -lemma foobar {e e' : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source ∩ e'.source) : +lemma foobar {x : M} (hx : x ∈ e.source ∩ e'.source) : (e.extend I) x ∈ interior (e.extend I).target ↔ (e'.extend I) x ∈ interior (e'.extend I).target := by rcases ((mem_inter_iff x _ _).mp hx) with ⟨hxe, hxe'⟩ @@ -145,7 +147,7 @@ lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source rw [← not_iff_not.mpr (isInteriorPoint_iff I hx)] exact Iff.rfl -/-- Every point is either an interior or a boundary point. -/ -- FIXME: better name?! +/-- Every point is either an interior or a boundary point. -/ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by by_cases extChartAt I x x ∈ interior (extChartAt I x).target · exact Or.inl h @@ -154,10 +156,8 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB variable (I M) in /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ - (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := by - apply le_antisymm - · exact fun x _ ↦ trivial - · exact fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x + (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := + le_antisymm (fun _ _ ↦ trivial) (fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x) /-- The interior and boundary of `M` are disjoint. -/ -- xxx: name `..._eq_empty` instead? lemma interior_boundary_disjoint : @@ -174,8 +174,7 @@ lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by let U := interior (e.extend I).target -- For all `y ∈ e.source`, `y` is an interior point iff its image lies in `U`. -- FIXME: should this be a separate lemma? - use (e.extend I).source ∩ (e.extend I) ⁻¹' U - refine ⟨?_, ?_, ?_⟩ + refine ⟨(e.extend I).source ∩ (e.extend I) ⁻¹' U, ?_, ?_, ?_⟩ · intro y hy rw [e.extend_source] at hy apply (isInteriorPoint_iff I (mem_of_mem_inter_left hy)).mpr From 3d820bebc5531535db5d7f827511175724fb4593 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 22 Nov 2023 18:42:36 +0100 Subject: [PATCH 056/203] Move complete (and now unused) helper results to their files: I decided two MapsTo helpers are obvious enough to be left out, but disjointness of interior and frontier is interesting enough. --- Mathlib/Geometry/Manifold/ChartedSpace.lean | 7 ++++ .../Geometry/Manifold/InteriorBoundary.lean | 38 ------------------- .../Manifold/SmoothManifoldWithCorners.lean | 3 ++ Mathlib/Topology/Basic.lean | 5 +++ 4 files changed, 15 insertions(+), 38 deletions(-) diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 5db698d70957e..98bb16723d077 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -584,6 +584,13 @@ lemma chart_mem_atlas (H : Type*) {M : Type*} [TopologicalSpace H] [TopologicalS section ChartedSpace +/-- The charts of a charted space cover its domain. -/ +lemma ChartedSpace.covering {H M : Type*} [TopologicalSpace H] [TopologicalSpace M] + [ChartedSpace H M] : ⋃ x : M, (chartAt H x).source = univ := by + refine subset_antisymm (fun y _ ↦ trivial) (fun y _ ↦ ?_) + rw [mem_iUnion] + exact ⟨y, mem_chart_source H y⟩ + /-- 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} diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index a258eec189baf..8fedf1d2b60bd 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -42,16 +42,6 @@ manifold, interior, boundary open Set -section TopologyHelpers -- should be in mathlib; Mathlib.Topology.Basic -variable {X : Type*} [TopologicalSpace X] {s : Set X} - -/-- Interior and frontier are disjoint. -/ -lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by - rw [← closure_diff_interior s, diff_eq, ← inter_assoc, inter_comm, ← inter_assoc, - compl_inter_self, empty_inter] - -end TopologyHelpers - -- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] @@ -194,32 +184,4 @@ lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := b exact interior_isOpen end SmoothManifoldWithCorners -/-- The charts of a charted space cover its domain. -/ --- {H M : Type*} [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] -lemma ChartedSpace.covering : ⋃ x : M, (chartAt H x).source = univ := by - apply subset_antisymm <;> intro y _ - · trivial - · rw [mem_iUnion] - use y - exact mem_chart_source H y - -namespace LocalHomeomorph -- FIXME: move to SmoothManifoldWithCorners -lemma extend_map_source {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - (e.extend I) x ∈ (e.extend I).target := by - rw [← e.extend_source I] at hx - exact (e.extend I).map_source hx - --- XXX: mapsTo_extend has a different formula from extend_target... --- should unify these/relate these better! -lemma mapsTo_extend' {e : LocalHomeomorph M H} : - MapsTo (e.extend I) (e.extend I).source (e.extend I).target := - fun _ hx ↦(e.extend I).map_source hx - -/-- Variant of `extend_map_source`, stated for images of subsets. -/ -lemma extend_map_source' {e : LocalHomeomorph M H} : - (e.extend I) '' (e.extend I).source ⊆ (e.extend I).target := - Set.mapsTo'.mp e.mapsTo_extend' - -end LocalHomeomorph - -- TODO: interior I M is a manifold diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index fb5cdadb5ecc1..000befebde963 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -1006,6 +1006,9 @@ theorem mapsTo_extend (hs : s ⊆ f.source) : exact image_subset _ (inter_subset_right _ _) #align local_homeomorph.maps_to_extend LocalHomeomorph.mapsTo_extend +lemma mapsTo_extend' : MapsTo (f.extend I) (f.extend I).source (f.extend I).target := + fun _ hx ↦(f.extend I).map_source hx + 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 diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index d2b36293f56bc..61f9fdcb86f2e 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -709,6 +709,11 @@ theorem closure_diff_interior (s : Set α) : closure s \ interior s = frontier s rfl #align closure_diff_interior closure_diff_interior +/-- Interior and frontier are disjoint. -/ +lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by + rw [← closure_diff_interior s, diff_eq, ← inter_assoc, inter_comm, ← inter_assoc, + compl_inter_self, empty_inter] + @[simp] theorem closure_diff_frontier (s : Set α) : closure s \ frontier s = interior s := by rw [frontier, diff_diff_right_self, inter_eq_self_of_subset_right interior_subset_closure] From 2c6955356b280e3bcc0716b07c39dc5c5a7db9fe Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 22 Nov 2023 14:16:53 -0800 Subject: [PATCH 057/203] compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 57dbfb79c6c98..b4cf8a5d83b95 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -10,7 +10,7 @@ import Mathlib.Geometry.Manifold.MFDeriv /-! # Integral curves of vector fields on a manifold -For any continuously differentiable vector field on a manifold `M` and any chosen non-boundary point +For any continuously differentiable vector field on a manifold `M` and any chosen interior point `x₀ : M`, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. @@ -43,7 +43,7 @@ variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : - IsOpen (extChartAt I x).target := LocalHomeomorph.isOpen_extend_target + IsOpen (extChartAt I x).target := LocalHomeomorph.isOpen_extend_target .. /-- If `M` has no boundary, every point of `M` is an interior point. -/ lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by From 7c6bccff129106b0205b417c7bb81e1945e48527 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 23 Nov 2023 03:39:13 -0800 Subject: [PATCH 058/203] new lemmas and slight golfing --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 70 ++++++++++++-------- 1 file changed, 43 insertions(+), 27 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b4cf8a5d83b95..4ea4ffb50239a 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -50,6 +50,34 @@ lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPo rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) +variable (I) +def tangentCoordChange (x y : M) := (tangentBundleCore I M).coordChange (achart H x) (achart H y) + +variable {I} +lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = + fderivWithin 𝕜 (extChartAt I y ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl + +lemma tangentCoordChange_self {x z : M} {v : E} (h : z ∈ (extChartAt I x).source) : + tangentCoordChange I x x z v = v := by + apply (tangentBundleCore I M).coordChange_self + rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + exact h + +-- continuousOn? + +lemma tangentCoordChange_comp {w x y z : M} {v : E} + (h : z ∈ (extChartAt I w).source ∩ (extChartAt I x).source ∩ (extChartAt I y).source) : + tangentCoordChange I x y z (tangentCoordChange I w x z v) = tangentCoordChange I w y z v := by + apply (tangentBundleCore I M).coordChange_comp + simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + exact h + +lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} + (h : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source) : + HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) + (range I) (extChartAt I x z) := + ((contDiffWithinAt_ext_coord_change I y x h).differentiableWithinAt (by simp)).hasFDerivWithinAt + end variable @@ -228,6 +256,9 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) (v ((extChartAt I x₀).symm (f t)))) t := hf2 t ht1 + rw [←tangentCoordChange_def] at h + + have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset /- express the derivative of the integral curve in the local chart -/ rw [HasMFDerivAt] @@ -242,37 +273,22 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of coordinates as `D`. -/ - have hvsub : v ((extChartAt I x₀).symm (f t)) = (tangentBundleCore I M).coordChange - (achart H x₀) (achart H ((extChartAt I x₀).symm (f t))) ((extChartAt I x₀).symm (f t)) - ((tangentBundleCore I M).coordChange (achart H ((extChartAt I x₀).symm (f t))) (achart H x₀) - ((extChartAt I x₀).symm (f t)) (v ((extChartAt I x₀).symm (f t)))) := by - rw [(tangentBundleCore I M).coordChange_comp, (tangentBundleCore I M).coordChange_self] - · rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I, ←LocalEquiv.symm_target] - exact mem_extChartAt_source .. - · rw [tangentBundleCore_baseSet, tangentBundleCore_baseSet, coe_achart, coe_achart, - ←extChartAt_source I, ←extChartAt_source I, inter_comm, ←inter_assoc, inter_self] - constructor - · exact mem_extChartAt_source .. - · rw [←mem_preimage] - apply mem_of_mem_of_subset _ (LocalEquiv.source_subset_preimage_target _) - rw [LocalEquiv.symm_source] - exact hf3' - rw [hvsub] + rw [←tangentCoordChange_self (I := I) (M := M) (x := (extChartAt I x₀).symm (f t)) + (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t)))] + rw [←tangentCoordChange_comp (x := x₀)] apply HasFDerivAt.comp_hasDerivAt _ _ h - /- change of coordinates in the tangent bundle is exactly the derivative of composition of local - charts -/ - rw [tangentBundleCore_coordChange_achart, LocalEquiv.right_inv _ hf3', fderivWithin_of_mem_nhds] - · apply DifferentiableAt.hasFDerivAt - apply MDifferentiableAt.differentiableAt - apply MDifferentiableAt.comp (I' := I) - · exact (contMDiffAt_extChartAt (n := 1)).mdifferentiableAt (le_refl _) - · apply MDifferentiableOn.mdifferentiableAt - ((contMDiffOn_extChartAt_symm (n := 1) _).mdifferentiableOn (le_refl _)) - rw [mem_nhds_iff] - exact ⟨interior (extChartAt I x₀).target, interior_subset, isOpen_interior, hf3 _ ht2⟩ + apply HasFDerivWithinAt.hasFDerivAt (s := range I) + nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] + apply hasFDerivWithinAt_tangentCoordChange + · sorry · rw [mem_nhds_iff] exact ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ + · rw [inter_right_comm, inter_self, mem_inter_iff] + use mem_extChartAt_source .. + rw [←mem_preimage] + exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + · sorry /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the From e74b8603983c5b03aa3351425815db7cbdaf9bb1 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 25 Nov 2023 12:09:58 +0100 Subject: [PATCH 059/203] Essentials are complete. --- .../Geometry/Manifold/InteriorBoundary.lean | 117 ++++++++++++++++++ 1 file changed, 117 insertions(+) create mode 100644 Mathlib/Geometry/Manifold/InteriorBoundary.lean diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean new file mode 100644 index 0000000000000..e17c1e26ead80 --- /dev/null +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -0,0 +1,117 @@ +/- +Copyright (c) 2023 Michael Rothgang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Rothgang +-/ + +import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners + +/-! +# Interior and boundary of a manifold +Define the interior and boundary of a manifold. + +## Main definitions +- **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, + `φ x` is an interior point of `φ.target`. +- **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, `(extChartAt I x) x ∈ frontier (range I)`. +- **interior I M** is the **interior** of `M`, the set of its interior points. +- **boundary I M** is the **boundary** of `M`, the set of its boundary points. + +## Main results +- `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary +- if `M` is boundary, every point is an interior point + +## Tags +manifold, interior, boundary +-/ + +open Set + +-- Let `M` be a manifold with corners over the pair `(E, H)`. +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [HasGroupoid M (contDiffGroupoid 0 I)] + +/-- `p ∈ M` is an interior point of a manifold `M` iff +for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target + +def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (range I) + +namespace SmoothManifoldWithCorners +-- FIXME(MR): can I enable dot notation as in, say, `M.interior I`? + +variable (I M) in +/-- The **interior** of a manifold `M` is the set of its interior points. -/ +protected def interior : Set M := { x : M | I.IsInteriorPoint x} + +variable (I M) in +/-- The **boundary** of a manifold `M` is the set of its boundary points. -/ +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} + +lemma _root_.ModelWithCorners.isBoundaryPoint_iff {x : M} : + I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := Iff.rfl + +-- move to LocalHomeomorph! +/-- The interior of `(e.extend I).target` is contained in the interior of its model's range. -/ +lemma _root_.LocalHomeomorph.extend_interior_subset_interior_target22 {e : LocalHomeomorph M H} : + interior (e.extend I).target ⊆ interior (range I) := by + rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq] + exact inter_subset_right _ _ + +/-- If `y ∈ e.target` and `I y ∈ interior (range I)`,, then `I y` is an interior point of `(I ∘ e).target`. -/ +lemma _root_.LocalHomeomorph.mem_interior_extend_target {e : LocalHomeomorph M H} {y : H} (hy : y ∈ e.target) + (hy' : I y ∈ interior (range I)) : I y ∈ interior (e.extend I).target := by + rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq, + mem_inter_iff, mem_preimage] + exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ + +/-- Every point is either an interior or a boundary point. -/ +lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by + by_cases h : extChartAt I x x ∈ interior (extChartAt I x).target + · exact Or.inl h + · right -- Otherwise, we have a boundary point. + rw [I.isBoundaryPoint_iff, ← closure_diff_interior, I.closed_range.closure_eq] + refine ⟨mem_range_self _, ?_⟩ + by_contra h' + exact h ((chartAt H x).mem_interior_extend_target (mem_chart_target H x) h') + +variable (I M) in +/-- A manifold decomposes into interior and boundary. -/ +lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ + (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := + le_antisymm (fun _ _ ↦ trivial) (fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x) + +/-- The interior and boundary of `M` are disjoint. -/ +lemma interior_boundary_disjoint : + (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by + by_contra h + -- Choose some x in the intersection of interior and boundary. + choose x hx using nmem_singleton_empty.mp h + rcases hx with ⟨h1, h2⟩ + + rw [SmoothManifoldWithCorners.boundary] at h2 + have : I.IsBoundaryPoint x := sorry + rw [I.isBoundaryPoint_iff (x := x)] at this + have aux2 : frontier (range I) ∩ interior (range I) = ∅ := sorry -- topology + have : (extChartAt I x) x ∈ interior (range I) := by + sorry --apply?--sorry + have aux : (extChartAt I x) x ∈ (∅ : Set E) := by + rw [← aux2] + rw [inter_comm] --xxx + exact ⟨this, h2⟩ + exact aux + + + -- have prev : (extChartAt I x) x ∉ interior (range I) := by -- copied; deduplicate + -- by_contra h2 + -- have : I ((chartAt H x) x) ∈ interior (extChartAt I x).target := by + -- simp_rw [← Function.comp_apply] + -- exact (chartAt H x).mem_interior_extend_target (mem_chart_target H x) h2 + -- sorry--exact?-- h this + -- have : (extChartAt I x) x ∉ interior ((chartAt H x).extend I).target := by + -- by_contra h + -- exact prev ((chartAt H x).extend_interior_subset_interior_target22 (I := I) h1) + -- exact this h1 +end SmoothManifoldWithCorners From f8e334071d051b6923c0e0f780927ba323de67c5 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 25 Nov 2023 12:30:30 +0100 Subject: [PATCH 060/203] Clean up; document TODOs/next steps; move all helpers to their proper place. --- Mathlib.lean | 1 + .../Geometry/Manifold/InteriorBoundary.lean | 94 ++++++++++--------- .../Manifold/SmoothManifoldWithCorners.lean | 17 ++++ Mathlib/Topology/Basic.lean | 5 + 4 files changed, 73 insertions(+), 44 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 8c6f6e73f914a..1c6cef7db26d3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2077,6 +2077,7 @@ import Mathlib.Geometry.Manifold.Diffeomorph import Mathlib.Geometry.Manifold.Instances.Real import Mathlib.Geometry.Manifold.Instances.Sphere import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra +import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalInvariantProperties import Mathlib.Geometry.Manifold.MFDeriv import Mathlib.Geometry.Manifold.Metrizable diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index e17c1e26ead80..17e6f0108a4e4 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -19,10 +19,22 @@ Define the interior and boundary of a manifold. ## Main results - `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary -- if `M` is boundary, every point is an interior point +- `interior_boundary_disjoint`: interior and boundary of `M` are disjoint +- if `M` is boundaryless, every point is an interior point ## Tags manifold, interior, boundary + +## TODO +- if `M` is finite-dimensional, its interior is always open, + hence the boundary is closed (and nowhere dense) + This requires e.g. the homology of spheres. +- the interior of `M` is a smooth manifold without boundary +- `boundary M` is a smooth submanifold (possibly with boundary and corners): +follows from the corresponding statement for the model with corners `I`; +this requires a definition of submanifolds +- if `M` is finite-dimensional, its boundary has measure zero + -/ open Set @@ -37,35 +49,26 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +/-- `p ∈ M` is a boundary point of a manifold `M` iff its image in the extended chart +lies on the boundary of the model space. -/ def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (range I) namespace SmoothManifoldWithCorners --- FIXME(MR): can I enable dot notation as in, say, `M.interior I`? +-- TODO: can I enable dot notation as in, say, `M.interior I`? variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ protected def interior : Set M := { x : M | I.IsInteriorPoint x} +lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : + I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := Iff.rfl + variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} lemma _root_.ModelWithCorners.isBoundaryPoint_iff {x : M} : - I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := Iff.rfl - --- move to LocalHomeomorph! -/-- The interior of `(e.extend I).target` is contained in the interior of its model's range. -/ -lemma _root_.LocalHomeomorph.extend_interior_subset_interior_target22 {e : LocalHomeomorph M H} : - interior (e.extend I).target ⊆ interior (range I) := by - rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq] - exact inter_subset_right _ _ - -/-- If `y ∈ e.target` and `I y ∈ interior (range I)`,, then `I y` is an interior point of `(I ∘ e).target`. -/ -lemma _root_.LocalHomeomorph.mem_interior_extend_target {e : LocalHomeomorph M H} {y : H} (hy : y ∈ e.target) - (hy' : I y ∈ interior (range I)) : I y ∈ interior (e.extend I).target := by - rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq, - mem_inter_iff, mem_preimage] - exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ + I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := Iff.rfl /-- Every point is either an interior or a boundary point. -/ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by @@ -75,43 +78,46 @@ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsB rw [I.isBoundaryPoint_iff, ← closure_diff_interior, I.closed_range.closure_eq] refine ⟨mem_range_self _, ?_⟩ by_contra h' - exact h ((chartAt H x).mem_interior_extend_target (mem_chart_target H x) h') + exact h ((chartAt H x).mem_interior_extend_target I (mem_chart_target H x) h') -variable (I M) in /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := le_antisymm (fun _ _ ↦ trivial) (fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x) -/-- The interior and boundary of `M` are disjoint. -/ +/-- The interior and boundary of a manifold `M` are disjoint. -/ lemma interior_boundary_disjoint : (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by by_contra h -- Choose some x in the intersection of interior and boundary. choose x hx using nmem_singleton_empty.mp h rcases hx with ⟨h1, h2⟩ - - rw [SmoothManifoldWithCorners.boundary] at h2 - have : I.IsBoundaryPoint x := sorry - rw [I.isBoundaryPoint_iff (x := x)] at this - have aux2 : frontier (range I) ∩ interior (range I) = ∅ := sorry -- topology - have : (extChartAt I x) x ∈ interior (range I) := by - sorry --apply?--sorry - have aux : (extChartAt I x) x ∈ (∅ : Set E) := by - rw [← aux2] - rw [inter_comm] --xxx - exact ⟨this, h2⟩ - exact aux - - - -- have prev : (extChartAt I x) x ∉ interior (range I) := by -- copied; deduplicate - -- by_contra h2 - -- have : I ((chartAt H x) x) ∈ interior (extChartAt I x).target := by - -- simp_rw [← Function.comp_apply] - -- exact (chartAt H x).mem_interior_extend_target (mem_chart_target H x) h2 - -- sorry--exact?-- h this - -- have : (extChartAt I x) x ∉ interior ((chartAt H x).extend I).target := by - -- by_contra h - -- exact prev ((chartAt H x).extend_interior_subset_interior_target22 (I := I) h1) - -- exact this h1 + show (extChartAt I x) x ∈ (∅ : Set E) + rw [← interior_frontier_disjoint] + exact ⟨(chartAt H x).interior_extend_target_subset_interior_range I h1, h2⟩ + +/-- The boundary is the complement of the interior. -/ +lemma boundary_eq_complement_interior : + SmoothManifoldWithCorners.boundary I M = (SmoothManifoldWithCorners.interior I M)ᶜ := + (compl_unique interior_boundary_disjoint univ_eq_interior_union_boundary).symm end SmoothManifoldWithCorners + +/-- If `M` has no boundary, every point of `M` is an interior point. -/ +lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by + let r := ((chartAt H x).isOpen_extend_target I).interior_eq + have : extChartAt I x = (chartAt H x).extend I := rfl + rw [← this] at r + rw [ModelWithCorners.IsInteriorPoint, r] + exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) + +/-- If `I` is boundaryless, `M` has full interior interior. -/ +lemma ModelWithCorners.interior_eq_univ [I.Boundaryless] : + SmoothManifoldWithCorners.interior I M = univ := by + ext + refine ⟨fun _ ↦ trivial, fun _ ↦ I.isInteriorPoint⟩ + +/-- If `I` is boundaryless, `M` has empty boundary. -/ +lemma ModelWithCorners.Boundaryless.boundary_eq_empty [I.Boundaryless] : + SmoothManifoldWithCorners.boundary I M = ∅ := by + rw [SmoothManifoldWithCorners.boundary_eq_complement_interior, I.interior_eq_univ, + compl_empty_iff] diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 36c349057386d..67de61c3a60ff 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -976,6 +976,10 @@ theorem extend_target : (f.extend I).target = I.symm ⁻¹' f.target ∩ range I simp_rw [extend, LocalEquiv.trans_target, I.target_eq, I.toLocalEquiv_coe_symm, inter_comm] #align local_homeomorph.extend_target LocalHomeomorph.extend_target +lemma isOpen_extend_target [I.Boundaryless] : IsOpen (f.extend I).target := by + rw [extend_target, I.range_eq_univ, inter_univ] + exact I.continuous_symm.isOpen_preimage _ f.open_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, @@ -1023,6 +1027,19 @@ theorem extend_target_mem_nhdsWithin {y : M} (hy : y ∈ f.source) : 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 +lemma interior_extend_target_subset_interior_range : + interior (f.extend I).target ⊆ interior (range I) := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq] + exact inter_subset_right _ _ + +/-- If `y ∈ f.target` and `I y ∈ interior (range I)`, + then `I y` is an interior point of `(I ∘ f).target`. -/ +lemma mem_interior_extend_target {y : H} (hy : y ∈ f.target) + (hy' : I y ∈ interior (range I)) : I y ∈ interior (f.extend I).target := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq, + mem_inter_iff, mem_preimage] + exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ + 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 <| diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index d2b36293f56bc..61f9fdcb86f2e 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -709,6 +709,11 @@ theorem closure_diff_interior (s : Set α) : closure s \ interior s = frontier s rfl #align closure_diff_interior closure_diff_interior +/-- Interior and frontier are disjoint. -/ +lemma interior_frontier_disjoint : interior s ∩ frontier s = ∅ := by + rw [← closure_diff_interior s, diff_eq, ← inter_assoc, inter_comm, ← inter_assoc, + compl_inter_self, empty_inter] + @[simp] theorem closure_diff_frontier (s : Set α) : closure s \ frontier s = interior s := by rw [frontier, diff_diff_right_self, inter_eq_self_of_subset_right interior_subset_closure] From 4febc74c6f21809fb4840bf8ef963b7eca884795 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 25 Nov 2023 13:34:20 +0100 Subject: [PATCH 061/203] Remove helper lemmas not needed here. --- .../Manifold/SmoothManifoldWithCorners.lean | 21 ------------------- 1 file changed, 21 deletions(-) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 000befebde963..2a7f174bceec9 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -807,24 +807,6 @@ theorem mem_analyticGroupoid_of_boundaryless [CompleteSpace E] [I.Boundaryless] end analyticGroupoid -/-! Topological manifolds with corners: no smoothness assumed, but boundary and/or corners -are possible. -/ -section ManifoldWithCorners -/-- Typeclass defining topological manifolds with corners with respect to a model with corners, -over a field `𝕜`. -/ -class ManifoldWithCorners {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] extends - HasGroupoid M (contDiffGroupoid 0 I) : Prop - -theorem ManifoldWithCorners.mk' {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] - [gr : HasGroupoid M (contDiffGroupoid 0 I)] : ManifoldWithCorners I M := - { gr with } - -end ManifoldWithCorners - section SmoothManifoldWithCorners /-! ### Smooth manifolds with corners -/ @@ -1006,9 +988,6 @@ theorem mapsTo_extend (hs : s ⊆ f.source) : exact image_subset _ (inter_subset_right _ _) #align local_homeomorph.maps_to_extend LocalHomeomorph.mapsTo_extend -lemma mapsTo_extend' : MapsTo (f.extend I) (f.extend I).source (f.extend I).target := - fun _ hx ↦(f.extend I).map_source hx - 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 From c1f3395f60ee401d042cb3af7078ddf561bc9a8c Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 25 Nov 2023 13:29:20 +0100 Subject: [PATCH 062/203] Pare down interior/boundary file to the necesssary basics; sorry-free. Everything else requires advanced technology not in mathlib yet. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 - .../Geometry/Manifold/InteriorBoundary.lean | 187 ++++++------------ .../Manifold/SmoothManifoldWithCorners.lean | 13 ++ 3 files changed, 74 insertions(+), 134 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4ea4ffb50239a..c84d8da06b07b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -42,14 +42,6 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -lemma ModelWithCorners.isOpen_target [I.Boundaryless] {x : M} : - IsOpen (extChartAt I x).target := LocalHomeomorph.isOpen_extend_target .. - -/-- If `M` has no boundary, every point of `M` is an interior point. -/ -lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by - rw [ModelWithCorners.IsInteriorPoint, IsOpen.interior_eq I.isOpen_target] - exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) - variable (I) def tangentCoordChange (x y : M) := (tangentBundleCore I M).coordChange (achart H x) (achart H y) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 8fedf1d2b60bd..ac0e94c0e75a4 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -1,40 +1,36 @@ /- Copyright (c) 2023 Michael Rothgang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Michael Rothgang, Winston Yin +Authors: Michael Rothgang -/ import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners --- FIXME: should this be its own file or go in SmoothManifoldWithCorners? --- the latter is already huge, or its own file - move other results about boundaryless here? --- xxx: if I can use dot notation, how set things up so they're also available for smooth manifolds? --- manually re-declare them? - /-! # Interior and boundary of a manifold - Define the interior and boundary of a manifold. ## Main definitions - **IsInteriorPoint x**: `p ∈ M` is an interior point if, for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -- **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, for `φ` being the preferred chart at `x`, +- **IsBoundaryPoint x**: `p ∈ M` is a boundary point if, `(extChartAt I x) x ∈ frontier (range I)`. - **interior I M** is the **interior** of `M`, the set of its interior points. - **boundary I M** is the **boundary** of `M`, the set of its boundary points. ## Main results - `univ_eq_interior_union_boundary`: `M` is the union of its interior and boundary -- `interior_isOpen`: `interior I M` is open -- `boundary_isClosed`: `boundary I M` is closed - -**TODO** -- under suitable assumptions, `boundary I M` has empty interior -(if `M` is finite-dimensional, `boundary I M` should have measure 0, which implies this) -- `interior I M` is a manifold without boundary - (need to upgrade the model used; map the charts from an open ball to entire ℝ^n) -- the boundary is a submanifold of codimension 1, perhaps with boundary and corners -(this requires a definition of submanifolds) +- `interior_boundary_disjoint`: interior and boundary of `M` are disjoint +- if `M` is boundaryless, every point is an interior point + +## TODO +- if `M` is finite-dimensional, its interior is always open, + hence the boundary is closed (and nowhere dense) + This requires e.g. the homology of spheres. +- the interior of `M` is a smooth manifold without boundary +- `boundary M` is a smooth submanifold (possibly with boundary and corners): +follows from the corresponding statement for the model with corners `I`; +this requires a definition of submanifolds +- if `M` is finite-dimensional, its boundary has measure zero ## Tags manifold, interior, boundary @@ -52,136 +48,75 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target -/-- `p ∈ M` is a boundary point of a manifold `M` iff it is not an interior point. -This means that, for `φ` being the preferred chart at `x`, `φ x` is not an interior point of -`φ.target`. We do not say "boundary point" as `frontier φ.target` has two components, one on the -boundary of range I and another on the boundary of e.target (which we don't want). -/ -def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∉ interior (extChartAt I x).target +/-- `p ∈ M` is a boundary point of a manifold `M` iff its image in the extended chart +lies on the boundary of the model space. -/ +def ModelWithCorners.IsBoundaryPoint (x : M) := extChartAt I x x ∈ frontier (range I) namespace SmoothManifoldWithCorners --- FIXME(MR): can I enable dot notation, like `M.interior I` or so? +-- TODO: can I enable dot notation as in, say, `M.interior I`? variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ protected def interior : Set M := { x : M | I.IsInteriorPoint x} +lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : + I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := Iff.rfl + variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} -end SmoothManifoldWithCorners -namespace LocalHomeomorph -- move to SmoothManifoldsWithCorners! -variable {e e' : LocalHomeomorph M H} - --- more general lemma underlying foobar. xxx: find a better name! -lemma foobar_abstract {f : LocalHomeomorph H H} {y : H} (hy : y ∈ f.source) - (h : I y ∈ interior (range I)) : I (f y) ∈ interior (range I) := by - sorry - --- xxx: needs better name! --- the interior of the target of an extended local homeo is contained in the interior of its --- model's range -lemma extend_interior_target_subset : interior (e.extend I).target ⊆ interior (range I) := by - rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq] - exact inter_subset_right _ _ - --- xxx: find a good name!! -lemma foobaz {y : H} (hy : y ∈ e.target) (hy' : I y ∈ interior (range ↑I)) : - I y ∈ interior (e.extend I).target := by - rw [e.extend_target, interior_inter, (e.open_target.preimage I.continuous_symm).interior_eq, - mem_inter_iff, mem_preimage] - exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ - -/-- If `e` and `e'` are two charts, the transition map maps interior points to interior points. -/ --- as we only need continuity property, e or e' being in the atlas is not required -lemma foobar {x : M} (hx : x ∈ e.source ∩ e'.source) : - (e.extend I) x ∈ interior (e.extend I).target ↔ - (e'.extend I) x ∈ interior (e'.extend I).target := by - rcases ((mem_inter_iff x _ _).mp hx) with ⟨hxe, hxe'⟩ - -- reduction, step 1: simplify what the interior means - have : (e.extend I) x ∈ interior (e.extend I).target ↔ I (e x) ∈ interior (range I) := - ⟨fun hx ↦ extend_interior_target_subset hx, fun hx ↦ foobaz (e.map_source hxe) hx⟩ - rw [this] - have : (e'.extend I) x ∈ interior (e'.extend I).target ↔ I (e' x) ∈ interior (range I) := - ⟨fun hx ↦ extend_interior_target_subset hx, fun hx ↦ foobaz (e'.map_source hxe') hx⟩ - rw [this] - -- step 2: rewrite in terms of coordinate changes - constructor - · intro h - let f := e.symm.trans e' - have h2 : e x ∈ f.source := by - have : e.symm (e x) = x := e.left_inv' hxe - rw [LocalHomeomorph.trans_source, mem_inter_iff (e x), e.symm_source, mem_preimage, this] - exact ⟨e.map_source hxe, hxe'⟩ - rw [← (e.left_inv' hxe)] - exact foobar_abstract h2 h - · sorry -- exactly the same... what's the best way to deduplicate? -end LocalHomeomorph - -namespace SmoothManifoldWithCorners --- FIXME(MR): find a better wording for the next two docstrings -variable (I) in -/-- Whether `x` is an interior point can equivalently be described by any chart - whose source contains `x`. -/ --- as we only need continuity properties, `e` being in the atlas is not required -lemma isInteriorPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.IsInteriorPoint x ↔ (e.extend I) x ∈ interior (e.extend I).target := - (chartAt H x).foobar (mem_inter (mem_chart_source H x) hx) - -variable (I) in -/-- Whether `x` is a boundary point of `M` can equivalently be described by any chart -whose source contains `x`. -/ -lemma isBoundaryPoint_iff {e : LocalHomeomorph M H} {x : M} (hx : x ∈ e.source) : - I.IsBoundaryPoint x ↔ (e.extend I) x ∉ interior (e.extend I).target := by - -- This lemma is just the "negation" (applying not_iff_not) to isInteriorPoint_iff. - rw [← not_iff_not.mpr (isInteriorPoint_iff I hx)] - exact Iff.rfl +lemma _root_.ModelWithCorners.isBoundaryPoint_iff {x : M} : + I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := Iff.rfl /-- Every point is either an interior or a boundary point. -/ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by - by_cases extChartAt I x x ∈ interior (extChartAt I x).target + by_cases h : extChartAt I x x ∈ interior (extChartAt I x).target · exact Or.inl h - · exact Or.inr h + · right -- Otherwise, we have a boundary point. + rw [I.isBoundaryPoint_iff, ← closure_diff_interior, I.closed_range.closure_eq] + refine ⟨mem_range_self _, ?_⟩ + by_contra h' + exact h ((chartAt H x).mem_interior_extend_target I (mem_chart_target H x) h') -variable (I M) in /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ (SmoothManifoldWithCorners.boundary I M) = (univ : Set M) := le_antisymm (fun _ _ ↦ trivial) (fun x _ ↦ isInteriorPoint_or_isBoundaryPoint x) -/-- The interior and boundary of `M` are disjoint. -/ -- xxx: name `..._eq_empty` instead? +/-- The interior and boundary of a manifold `M` are disjoint. -/ lemma interior_boundary_disjoint : (SmoothManifoldWithCorners.interior I M) ∩ (SmoothManifoldWithCorners.boundary I M) = ∅ := by - ext - exact ⟨fun h ↦ (not_mem_of_mem_diff h) (mem_of_mem_diff h), by exfalso⟩ - -/-- The interior of a manifold is an open subset. -/ -lemma interior_isOpen : IsOpen (SmoothManifoldWithCorners.interior I M) := by - apply isOpen_iff_forall_mem_open.mpr - intro x hx - -- Consider the preferred chart at `x`. Its extended chart has open interior. - let e := chartAt H x - let U := interior (e.extend I).target - -- For all `y ∈ e.source`, `y` is an interior point iff its image lies in `U`. - -- FIXME: should this be a separate lemma? - refine ⟨(e.extend I).source ∩ (e.extend I) ⁻¹' U, ?_, ?_, ?_⟩ - · intro y hy - rw [e.extend_source] at hy - apply (isInteriorPoint_iff I (mem_of_mem_inter_left hy)).mpr - exact mem_of_mem_inter_right (a := e.source) hy - · exact (e.continuousOn_extend I).preimage_open_of_open (e.isOpen_extend_source I) isOpen_interior - · have : x ∈ (e.extend I).source := by - rw [e.extend_source] - exact mem_chart_source H x - exact mem_inter this hx - -/-- The boundary of a manifold is a closed subset. -/ -lemma boundary_isClosed : IsClosed (SmoothManifoldWithCorners.boundary I M) := by - apply isOpen_compl_iff.mp - have : (SmoothManifoldWithCorners.interior I M)ᶜ = SmoothManifoldWithCorners.boundary I M := - (compl_unique interior_boundary_disjoint (univ_eq_interior_union_boundary I M)) - rw [← this, compl_compl] - exact interior_isOpen + by_contra h + -- Choose some x in the intersection of interior and boundary. + choose x hx using nmem_singleton_empty.mp h + rcases hx with ⟨h1, h2⟩ + show (extChartAt I x) x ∈ (∅ : Set E) + rw [← interior_frontier_disjoint] + exact ⟨(chartAt H x).interior_extend_target_subset_interior_range I h1, h2⟩ + +/-- The boundary is the complement of the interior. -/ +lemma boundary_eq_complement_interior : + SmoothManifoldWithCorners.boundary I M = (SmoothManifoldWithCorners.interior I M)ᶜ := + (compl_unique interior_boundary_disjoint univ_eq_interior_union_boundary).symm end SmoothManifoldWithCorners --- TODO: interior I M is a manifold +/-- If `M` has no boundary, every point of `M` is an interior point. -/ +lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by + let r := ((chartAt H x).isOpen_extend_target I).interior_eq + have : extChartAt I x = (chartAt H x).extend I := rfl + rw [← this] at r + rw [ModelWithCorners.IsInteriorPoint, r] + exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) + +/-- If `I` is boundaryless, `M` has full interior interior. -/ +lemma ModelWithCorners.interior_eq_univ [I.Boundaryless] : + SmoothManifoldWithCorners.interior I M = univ := by + ext + refine ⟨fun _ ↦ trivial, fun _ ↦ I.isInteriorPoint⟩ + +/-- If `I` is boundaryless, `M` has empty boundary. -/ +lemma ModelWithCorners.Boundaryless.boundary_eq_empty [I.Boundaryless] : + SmoothManifoldWithCorners.boundary I M = ∅ := by + rw [SmoothManifoldWithCorners.boundary_eq_complement_interior, I.interior_eq_univ, + compl_empty_iff] diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 2a7f174bceec9..8d1cc6c6c805e 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -1028,6 +1028,19 @@ theorem extend_target_mem_nhdsWithin {y : M} (hy : y ∈ f.source) : 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 +lemma interior_extend_target_subset_interior_range : + interior (f.extend I).target ⊆ interior (range I) := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq] + exact inter_subset_right _ _ + +/-- If `y ∈ f.target` and `I y ∈ interior (range I)`, + then `I y` is an interior point of `(I ∘ f).target`. -/ +lemma mem_interior_extend_target {y : H} (hy : y ∈ f.target) + (hy' : I y ∈ interior (range I)) : I y ∈ interior (f.extend I).target := by + rw [f.extend_target, interior_inter, (f.open_target.preimage I.continuous_symm).interior_eq, + mem_inter_iff, mem_preimage] + exact ⟨mem_of_eq_of_mem (I.left_inv (y)) hy, hy'⟩ + 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 <| From 1459dfb6951fd462a4ddeee6548fe04b46d4a9c0 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Sat, 25 Nov 2023 15:24:42 +0100 Subject: [PATCH 063/203] Simplify. --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 17e6f0108a4e4..2d058b5506e29 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -47,7 +47,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- `p ∈ M` is an interior point of a manifold `M` iff for `φ` being the preferred chart at `x`, `φ x` is an interior point of `φ.target`. -/ -def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (extChartAt I x).target +def ModelWithCorners.IsInteriorPoint (x : M) := extChartAt I x x ∈ interior (range I) /-- `p ∈ M` is a boundary point of a manifold `M` iff its image in the extended chart lies on the boundary of the model space. -/ @@ -61,7 +61,9 @@ variable (I M) in protected def interior : Set M := { x : M | I.IsInteriorPoint x} lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : - I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := Iff.rfl + I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := + ⟨fun h ↦ (chartAt H x).mem_interior_extend_target _ (mem_chart_target H x) h, + fun h ↦ LocalHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ @@ -72,13 +74,11 @@ lemma _root_.ModelWithCorners.isBoundaryPoint_iff {x : M} : /-- Every point is either an interior or a boundary point. -/ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by - by_cases h : extChartAt I x x ∈ interior (extChartAt I x).target + by_cases h : extChartAt I x x ∈ interior (range I) · exact Or.inl h · right -- Otherwise, we have a boundary point. rw [I.isBoundaryPoint_iff, ← closure_diff_interior, I.closed_range.closure_eq] - refine ⟨mem_range_self _, ?_⟩ - by_contra h' - exact h ((chartAt H x).mem_interior_extend_target I (mem_chart_target H x) h') + exact ⟨mem_range_self _, h⟩ /-- A manifold decomposes into interior and boundary. -/ lemma univ_eq_interior_union_boundary : (SmoothManifoldWithCorners.interior I M) ∪ @@ -94,7 +94,7 @@ lemma interior_boundary_disjoint : rcases hx with ⟨h1, h2⟩ show (extChartAt I x) x ∈ (∅ : Set E) rw [← interior_frontier_disjoint] - exact ⟨(chartAt H x).interior_extend_target_subset_interior_range I h1, h2⟩ + exact ⟨h1, h2⟩ /-- The boundary is the complement of the interior. -/ lemma boundary_eq_complement_interior : @@ -107,7 +107,7 @@ lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPo let r := ((chartAt H x).isOpen_extend_target I).interior_eq have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r - rw [ModelWithCorners.IsInteriorPoint, r] + rw [ModelWithCorners.isInteriorPoint_iff, r] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) /-- If `I` is boundaryless, `M` has full interior interior. -/ From 094ca884a761282e2465a857a6e2f94f7885c708 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 12:31:24 -0800 Subject: [PATCH 064/203] compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4ea4ffb50239a..a23bb5d44c0e7 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -280,7 +280,10 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin apply HasFDerivWithinAt.hasFDerivAt (s := range I) nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] apply hasFDerivWithinAt_tangentCoordChange - · sorry + · rw [LocalEquiv.right_inv _ hf3', LocalEquiv.trans_source, LocalEquiv.symm_source] + use hf3' + rw [mem_preimage] + exact mem_extChartAt_source .. · rw [mem_nhds_iff] exact ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ @@ -288,7 +291,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin use mem_extChartAt_source .. rw [←mem_preimage] exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - · sorry + · exact mem_extChartAt_source .. /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the From 7754760d85e273f3c59f12327c65557c6767da4c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 13:15:47 -0800 Subject: [PATCH 065/203] range_mem_nhds_isInteriorPoint --- .../Geometry/Manifold/InteriorBoundary.lean | 20 ++++++++++++++----- 1 file changed, 15 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 2d058b5506e29..521e6da33c390 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -102,22 +102,32 @@ lemma boundary_eq_complement_interior : (compl_unique interior_boundary_disjoint univ_eq_interior_union_boundary).symm end SmoothManifoldWithCorners +lemma range_mem_nhds_isInteriorPoint {x : M} (h : I.IsInteriorPoint x) : + range I ∈ nhds (extChartAt I x x) := by + rw [mem_nhds_iff] + exact ⟨interior (range I), interior_subset, isOpen_interior, h⟩ + +namespace ModelWithCorners + +variable [I.Boundaryless] + /-- If `M` has no boundary, every point of `M` is an interior point. -/ -lemma ModelWithCorners.isInteriorPoint [I.Boundaryless] {x : M} : I.IsInteriorPoint x := by +lemma isInteriorPoint {x : M} : I.IsInteriorPoint x := by let r := ((chartAt H x).isOpen_extend_target I).interior_eq have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r - rw [ModelWithCorners.isInteriorPoint_iff, r] + rw [isInteriorPoint_iff, r] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) /-- If `I` is boundaryless, `M` has full interior interior. -/ -lemma ModelWithCorners.interior_eq_univ [I.Boundaryless] : - SmoothManifoldWithCorners.interior I M = univ := by +lemma interior_eq_univ : SmoothManifoldWithCorners.interior I M = univ := by ext refine ⟨fun _ ↦ trivial, fun _ ↦ I.isInteriorPoint⟩ /-- If `I` is boundaryless, `M` has empty boundary. -/ -lemma ModelWithCorners.Boundaryless.boundary_eq_empty [I.Boundaryless] : +lemma Boundaryless.boundary_eq_empty : SmoothManifoldWithCorners.boundary I M = ∅ := by rw [SmoothManifoldWithCorners.boundary_eq_complement_interior, I.interior_eq_univ, compl_empty_iff] + +end ModelWithCorners From 6033099627006efbf9541020230a87426f658a18 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 13:18:25 -0800 Subject: [PATCH 066/203] namespace --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 521e6da33c390..cb6804845b027 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -107,27 +107,27 @@ lemma range_mem_nhds_isInteriorPoint {x : M} (h : I.IsInteriorPoint x) : rw [mem_nhds_iff] exact ⟨interior (range I), interior_subset, isOpen_interior, h⟩ -namespace ModelWithCorners +section boundaryless variable [I.Boundaryless] /-- If `M` has no boundary, every point of `M` is an interior point. -/ -lemma isInteriorPoint {x : M} : I.IsInteriorPoint x := by +lemma ModelWithCorners.isInteriorPoint {x : M} : I.IsInteriorPoint x := by let r := ((chartAt H x).isOpen_extend_target I).interior_eq have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r - rw [isInteriorPoint_iff, r] + rw [ModelWithCorners.isInteriorPoint_iff, r] exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) /-- If `I` is boundaryless, `M` has full interior interior. -/ -lemma interior_eq_univ : SmoothManifoldWithCorners.interior I M = univ := by +lemma ModelWithCorners.interior_eq_univ : SmoothManifoldWithCorners.interior I M = univ := by ext refine ⟨fun _ ↦ trivial, fun _ ↦ I.isInteriorPoint⟩ /-- If `I` is boundaryless, `M` has empty boundary. -/ -lemma Boundaryless.boundary_eq_empty : +lemma ModelWithCorners.Boundaryless.boundary_eq_empty : SmoothManifoldWithCorners.boundary I M = ∅ := by rw [SmoothManifoldWithCorners.boundary_eq_complement_interior, I.interior_eq_univ, compl_empty_iff] -end ModelWithCorners +end boundaryless From d1a8296cbc632051aca1cace04d0b03b715ef7fb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 13:22:49 -0800 Subject: [PATCH 067/203] hI as separate lemma --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 14 +++----------- 1 file changed, 3 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 3d1d052d28acb..a882a7206503e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -42,10 +42,9 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] -variable (I) +variable (I) in def tangentCoordChange (x y : M) := (tangentBundleCore I M).coordChange (achart H x) (achart H y) -variable {I} lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = fderivWithin 𝕜 (extChartAt I y ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl @@ -215,15 +214,10 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin /- express the derivative of the section `v` in the local charts -/ rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv - /- `hI` should be a separate lemma -/ - have hI : range I ∈ nhds (extChartAt I x₀ x₀) := by - rw [mem_nhds_iff] - exact ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), - isOpen_interior, hx⟩ /- use Picard-Lindelöf theorem to extract a solution to the ODE in the chart defined by `v` -/ obtain ⟨f, hf1, ε1, hε1, hf2⟩ := - exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (ContDiffAt.snd (hv.contDiffAt hI)) + exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ + (ContDiffAt.snd (hv.contDiffAt (SmoothManifoldWithCorners.range_mem_nhds_isInteriorPoint hx))) rw [←Real.ball_eq_Ioo] at hf2 /- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, `f t ∈ interior (extChartAt I x₀).target` -/ @@ -249,8 +243,6 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin (v ((extChartAt I x₀).symm (f t)))) t := hf2 t ht1 rw [←tangentCoordChange_def] at h - - have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset /- express the derivative of the integral curve in the local chart -/ rw [HasMFDerivAt] From 7f7a49df6df078dcecfe6d0cbe74c2d739abd45d Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 13:41:04 -0800 Subject: [PATCH 068/203] actually compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a882a7206503e..82af7bbfd29fb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -211,20 +211,20 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ `t₀`.-/ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by - /- express the derivative of the section `v` in the local charts -/ + -- express the differentiability of the section `v` in the local charts rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv - /- use Picard-Lindelöf theorem to extract a solution to the ODE in the chart defined by `v` -/ + -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (ContDiffAt.snd (hv.contDiffAt (SmoothManifoldWithCorners.range_mem_nhds_isInteriorPoint hx))) + (ContDiffAt.snd (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx))) rw [←Real.ball_eq_Ioo] at hf2 - /- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, - `f t ∈ interior (extChartAt I x₀).target` -/ + -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, + -- `f t ∈ interior (extChartAt I x₀).target` -/ have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := - hcont _ (IsOpen.mem_nhds isOpen_interior hx) + hcont _ (IsOpen.mem_nhds isOpen_interior (ModelWithCorners.isInteriorPoint_iff.mp hx)) rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 From ab30700ab6435619846a53b8b7c6d52915f04058 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 26 Nov 2023 13:55:42 -0800 Subject: [PATCH 069/203] minor --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 32 +++++++++----------- 1 file changed, 15 insertions(+), 17 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 82af7bbfd29fb..061be189b20e9 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -220,7 +220,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin (ContDiffAt.snd (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx))) rw [←Real.ball_eq_Ioo] at hf2 -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, - -- `f t ∈ interior (extChartAt I x₀).target` -/ + -- `f t ∈ interior (extChartAt I x₀).target` have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := @@ -228,12 +228,12 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 - /- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve -/ + -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] intros t ht - /- collect useful terms in convenient forms -/ + -- collect useful terms in convenient forms rw [←Real.ball_eq_Ioo] at ht have ht1 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) have ht2 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) @@ -244,22 +244,21 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin t := hf2 t ht1 rw [←tangentCoordChange_def] at h have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset - /- express the derivative of the integral curve in the local chart -/ - rw [HasMFDerivAt] - use ContinuousAt.comp - (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt) - apply HasDerivWithinAt.hasFDerivWithinAt + -- express the derivative of the integral curve in the local chart + refine' ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), + HasDerivWithinAt.hasFDerivWithinAt _⟩ rw [modelWithCornersSelf_coe, range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] - /- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of - coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use - `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires - first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of - coordinates as `D`. -/ - rw [←tangentCoordChange_self (I := I) (M := M) (x := (extChartAt I x₀).symm (f t)) - (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t)))] - rw [←tangentCoordChange_comp (x := x₀)] + -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of + -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use + -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires + -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of + -- coordinates as `D`. + rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) + (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) + (mem_extChartAt_source ..), + ←tangentCoordChange_comp (x := x₀)] apply HasFDerivAt.comp_hasDerivAt _ _ h apply HasFDerivWithinAt.hasFDerivAt (s := range I) nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] @@ -275,7 +274,6 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin use mem_extChartAt_source .. rw [←mem_preimage] exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - · exact mem_extChartAt_source .. /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the From 98ff512fbaf3955b32ab967a5ccc9a79261f12b7 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Mon, 27 Nov 2023 13:46:47 +0100 Subject: [PATCH 070/203] Prefer refine over refine' in new code, per zulip. And combine some apply lines. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 38 ++++++++++---------- 1 file changed, 18 insertions(+), 20 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 061be189b20e9..e9e16bb20479e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -98,7 +98,7 @@ variable {t₀} lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ - refine' ⟨by simp [h1], ε, hε, _⟩ + refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht rw [sub_right_comm, sub_add_eq_add_sub, ←add_mem_Ioo_iff_left] at ht have h2' := h2 (t + dt) ht @@ -106,18 +106,17 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v ←ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] apply HasMFDerivAt.comp t h2' /- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations -/ - refine' ⟨(continuous_add_right _).continuousAt, _⟩ + refine ⟨(continuous_add_right _).continuousAt, ?_⟩ simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] - apply HasFDerivAt.add_const - exact hasFDerivAt_id _ + apply HasFDerivAt.add_const (hasFDerivAt_id _) lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by - refine' ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, _⟩ + refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, ?_⟩ intro hγ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this @@ -128,8 +127,8 @@ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ - refine' ⟨by rw [Function.comp_apply, div_mul_cancel _ (ne_of_gt ha)]; exact h1, ε / a, - div_pos hε ha, _⟩ + refine ⟨by rw [Function.comp_apply, div_mul_cancel _ (ne_of_gt ha)]; exact h1, ε / a, + div_pos hε ha, ?_⟩ intros t ht have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by rw [mem_Ioo, ←div_lt_iff ha, ←lt_div_iff ha, sub_div, add_div] @@ -137,18 +136,17 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt have h2' := h2 (t * a) ht rw [Function.comp_apply, Pi.smul_apply, ←ContinuousLinearMap.smulRight_comp] apply HasMFDerivAt.comp t h2' - refine' ⟨(continuous_mul_right _).continuousAt, _⟩ + refine ⟨(continuous_mul_right _).continuousAt, ?_⟩ simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] - apply HasFDerivAt.mul_const' - exact hasFDerivAt_id _ + apply HasFDerivAt.mul_const' (hasFDerivAt_id _) lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by - refine' ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, _⟩ + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, ?_⟩ intro hγ have := hγ.comp_mul_pos (inv_pos_of_pos ha) rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ←div_mul_eq_div_div_swap, @@ -160,25 +158,24 @@ lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) : IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ - refine' ⟨by simp [h1], ε, hε, _⟩ + refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht rw [←neg_add', neg_add_eq_sub, ←neg_sub, ←neg_mem_Ioo_iff] at ht have h2' := h2 (-t) ht rw [Function.comp_apply, Pi.neg_apply, ←neg_one_smul ℝ (v (γ (-t))), ←ContinuousLinearMap.smulRight_comp] apply HasMFDerivAt.comp t h2' - refine' ⟨continuousAt_neg, _⟩ + refine ⟨continuousAt_neg, ?_⟩ simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] - apply HasDerivAt.hasFDerivAt - exact hasDerivAt_neg _ + apply HasDerivAt.hasFDerivAt (hasDerivAt_neg _) lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by - refine' ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, _⟩ + refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, ?_⟩ intro hγ have := hγ.comp_neg rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this @@ -196,7 +193,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by - refine' ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, _⟩ + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, ?_⟩ intro hγ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ←div_mul_eq_div_div_swap, @@ -229,9 +226,10 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve - refine' ⟨(extChartAt I x₀).symm ∘ f, _, min ε1 ε2, lt_min hε1 hε2, _⟩ + refine ⟨(extChartAt I x₀).symm ∘ f, ?_, min ε1 ε2, lt_min hε1 hε2, ?_⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] + -- xxx: indent! intros t ht -- collect useful terms in convenient forms rw [←Real.ball_eq_Ioo] at ht @@ -245,8 +243,8 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [←tangentCoordChange_def] at h have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset -- express the derivative of the integral curve in the local chart - refine' ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), - HasDerivWithinAt.hasFDerivWithinAt _⟩ + refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), + HasDerivWithinAt.hasFDerivWithinAt ?_⟩ rw [modelWithCornersSelf_coe, range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] From 00f4e77589282540799520455e95a57c6bcfbb85 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Mon, 27 Nov 2023 13:54:28 +0100 Subject: [PATCH 071/203] Use mfld_simps to make proofs more readable. Golf proofs about composition of integral curves. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 52 ++++++-------------- 1 file changed, 15 insertions(+), 37 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index e9e16bb20479e..eb1b26baac741 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -107,22 +107,17 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v apply HasMFDerivAt.comp t h2' /- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations -/ refine ⟨(continuous_add_right _).continuousAt, ?_⟩ - simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, - LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, - LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, - LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, - Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] + simp only [mfld_simps, hasFDerivWithinAt_univ] apply HasFDerivAt.add_const (hasFDerivAt_id _) lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, ?_⟩ - intro hγ + refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this convert this ext - simp + simp only [Function.comp_apply, neg_add_cancel_right] lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by @@ -133,21 +128,14 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by rw [mem_Ioo, ←div_lt_iff ha, ←lt_div_iff ha, sub_div, add_div] exact ht - have h2' := h2 (t * a) ht rw [Function.comp_apply, Pi.smul_apply, ←ContinuousLinearMap.smulRight_comp] - apply HasMFDerivAt.comp t h2' - refine ⟨(continuous_mul_right _).continuousAt, ?_⟩ - simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, - LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, - LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, - LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, - Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] + refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ + simp only [mfld_simps, hasFDerivWithinAt_univ] apply HasFDerivAt.mul_const' (hasFDerivAt_id _) lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, ?_⟩ - intro hγ + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_pos (inv_pos_of_pos ha) rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ←div_mul_eq_div_div_swap, inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this @@ -161,22 +149,15 @@ lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht rw [←neg_add', neg_add_eq_sub, ←neg_sub, ←neg_mem_Ioo_iff] at ht - have h2' := h2 (-t) ht rw [Function.comp_apply, Pi.neg_apply, ←neg_one_smul ℝ (v (γ (-t))), ←ContinuousLinearMap.smulRight_comp] - apply HasMFDerivAt.comp t h2' - refine ⟨continuousAt_neg, ?_⟩ - simp only [writtenInExtChartAt, extChartAt, LocalHomeomorph.extend, - LocalHomeomorph.refl_localEquiv, LocalEquiv.refl_source, - LocalHomeomorph.singletonChartedSpace_chartAt_eq, modelWithCornersSelf_localEquiv, - LocalEquiv.trans_refl, LocalEquiv.refl_coe, LocalEquiv.refl_symm, Function.comp.right_id, - Function.comp.left_id, modelWithCornersSelf_coe, range_id, id_eq, hasFDerivWithinAt_univ] - apply HasDerivAt.hasFDerivAt (hasDerivAt_neg _) + apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ + simp only [mfld_simps, hasFDerivWithinAt_univ] + exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, ?_⟩ - intro hγ + refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ have := hγ.comp_neg rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this exact this @@ -193,8 +174,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, ?_⟩ - intro hγ + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ←div_mul_eq_div_div_swap, inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this @@ -214,14 +194,14 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (ContDiffAt.snd (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx))) + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd rw [←Real.ball_eq_Ioo] at hf2 -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, -- `f t ∈ interior (extChartAt I x₀).target` have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := - hcont _ (IsOpen.mem_nhds isOpen_interior (ModelWithCorners.isInteriorPoint_iff.mp hx)) + hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 @@ -229,7 +209,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin refine ⟨(extChartAt I x₀).symm ∘ f, ?_, min ε1 ε2, lt_min hε1 hε2, ?_⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] - -- xxx: indent! + -- xxx: indent this, as it's a new goal! intros t ht -- collect useful terms in convenient forms rw [←Real.ball_eq_Ioo] at ht @@ -262,9 +242,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] apply hasFDerivWithinAt_tangentCoordChange · rw [LocalEquiv.right_inv _ hf3', LocalEquiv.trans_source, LocalEquiv.symm_source] - use hf3' - rw [mem_preimage] - exact mem_extChartAt_source .. + exact ⟨hf3', mem_extChartAt_source ..⟩ · rw [mem_nhds_iff] exact ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ From e86759de1ae55144ab746099922fc6d398781fad Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Mon, 27 Nov 2023 17:17:57 +0100 Subject: [PATCH 072/203] Small golf, use mfld_simps more. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index eb1b26baac741..b902d9cadcec6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -225,9 +225,9 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin -- express the derivative of the integral curve in the local chart refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), HasDerivWithinAt.hasFDerivWithinAt ?_⟩ - rw [modelWithCornersSelf_coe, range_id, hasDerivWithinAt_univ, ext_chart_model_space_apply, - writtenInExtChartAt, Function.comp_apply, Function.comp.assoc, extChartAt_model_space_eq_id, - LocalEquiv.refl_symm, LocalEquiv.refl_coe, Function.comp.right_id, ←Function.comp.assoc] + simp only [mfld_simps, hasDerivWithinAt_univ] + show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) + (v ((extChartAt I x₀).symm (f t))) t -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires From 85fbe5dda67bfcbc611260b422fcdd2108ca15b4 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 14:08:58 -0800 Subject: [PATCH 073/203] hasFDerivWithinAt_tangentCoordChange --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 88 ++++++++++---------- 1 file changed, 45 insertions(+), 43 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b902d9cadcec6..b2bcd8d766a9c 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -64,10 +64,13 @@ lemma tangentCoordChange_comp {w x y z : M} {v : E} exact h lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} - (h : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source) : + (h : z ∈ (extChartAt I x).source ∩ (extChartAt I y).source) : HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) (range I) (extChartAt I x z) := - ((contDiffWithinAt_ext_coord_change I y x h).differentiableWithinAt (by simp)).hasFDerivWithinAt + have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by + rw [LocalEquiv.trans_source'', LocalEquiv.symm_symm, LocalEquiv.symm_target] + exact mem_image_of_mem _ h + ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt end @@ -209,47 +212,46 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin refine ⟨(extChartAt I x₀).symm ∘ f, ?_, min ε1 ε2, lt_min hε1 hε2, ?_⟩ · apply Eq.symm rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] - -- xxx: indent this, as it's a new goal! - intros t ht - -- collect useful terms in convenient forms - rw [←Real.ball_eq_Ioo] at ht - have ht1 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) - have ht2 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) - have h : HasDerivAt f - ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) - (v ((extChartAt I x₀).symm (f t)))) - t := hf2 t ht1 - rw [←tangentCoordChange_def] at h - have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset - -- express the derivative of the integral curve in the local chart - refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), - HasDerivWithinAt.hasFDerivWithinAt ?_⟩ - simp only [mfld_simps, hasDerivWithinAt_univ] - show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) - (v ((extChartAt I x₀).symm (f t))) t - -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of - -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use - -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires - -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of - -- coordinates as `D`. - rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) - (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) - (mem_extChartAt_source ..), - ←tangentCoordChange_comp (x := x₀)] - apply HasFDerivAt.comp_hasDerivAt _ _ h - apply HasFDerivWithinAt.hasFDerivAt (s := range I) - nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] - apply hasFDerivWithinAt_tangentCoordChange - · rw [LocalEquiv.right_inv _ hf3', LocalEquiv.trans_source, LocalEquiv.symm_source] - exact ⟨hf3', mem_extChartAt_source ..⟩ - · rw [mem_nhds_iff] - exact ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ - · rw [inter_right_comm, inter_self, mem_inter_iff] - use mem_extChartAt_source .. - rw [←mem_preimage] - exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + · intros t ht + -- collect useful terms in convenient forms + rw [←Real.ball_eq_Ioo] at ht + have ht1 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + have ht2 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) + have h : HasDerivAt f + ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (v ((extChartAt I x₀).symm (f t)))) + t := hf2 t ht1 + rw [←tangentCoordChange_def] at h + have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset + -- express the derivative of the integral curve in the local chart + refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), + HasDerivWithinAt.hasFDerivWithinAt ?_⟩ + simp only [mfld_simps, hasDerivWithinAt_univ] + show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) + (v ((extChartAt I x₀).symm (f t))) t + -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of + -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use + -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires + -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of + -- coordinates as `D`. + rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) + (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) + (mem_extChartAt_source ..), + ←tangentCoordChange_comp (x := x₀)] + apply HasFDerivAt.comp_hasDerivAt _ _ h + apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| + mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ + nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] + apply hasFDerivWithinAt_tangentCoordChange + · rw [mem_inter_iff, ←mem_preimage] + exact ⟨mem_of_mem_of_subset hf3' (LocalEquiv.target_subset_preimage_source _), + mem_extChartAt_source ..⟩ + · rw [inter_right_comm, inter_self, mem_inter_iff] + use mem_extChartAt_source .. + rw [←mem_preimage] + exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the From d92489082d890acadc6b9eea84d1dd7813b3255f Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 22:16:30 -0800 Subject: [PATCH 074/203] some golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 81 +++++++++----------- 1 file changed, 38 insertions(+), 43 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b2bcd8d766a9c..6b936ba66f0fb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -209,49 +209,44 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve - refine ⟨(extChartAt I x₀).symm ∘ f, ?_, min ε1 ε2, lt_min hε1 hε2, ?_⟩ - · apply Eq.symm - rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)] - · intros t ht - -- collect useful terms in convenient forms - rw [←Real.ball_eq_Ioo] at ht - have ht1 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) - have ht2 := mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) - have h : HasDerivAt f - ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) - (v ((extChartAt I x₀).symm (f t)))) - t := hf2 t ht1 - rw [←tangentCoordChange_def] at h - have hf3' := mem_of_mem_of_subset (hf3 t ht2) interior_subset - -- express the derivative of the integral curve in the local chart - refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') ((hf2 t ht1).continuousAt), - HasDerivWithinAt.hasFDerivWithinAt ?_⟩ - simp only [mfld_simps, hasDerivWithinAt_univ] - show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) - (v ((extChartAt I x₀).symm (f t))) t - -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of - -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use - -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires - -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of - -- coordinates as `D`. - rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) - (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) - (mem_extChartAt_source ..), - ←tangentCoordChange_comp (x := x₀)] - apply HasFDerivAt.comp_hasDerivAt _ _ h - apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| - mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3 _ ht2⟩ - nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] - apply hasFDerivWithinAt_tangentCoordChange - · rw [mem_inter_iff, ←mem_preimage] - exact ⟨mem_of_mem_of_subset hf3' (LocalEquiv.target_subset_preimage_source _), - mem_extChartAt_source ..⟩ - · rw [inter_right_comm, inter_self, mem_inter_iff] - use mem_extChartAt_source .. - rw [←mem_preimage] - exact mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + refine ⟨(extChartAt I x₀).symm ∘ f, + Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), + min ε1 ε2, lt_min hε1 hε2, ?_⟩ + intros t ht + -- collect useful terms in convenient forms + rw [←Real.ball_eq_Ioo] at ht + have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) + have h : HasDerivAt f + ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (v ((extChartAt I x₀).symm (f t)))) + t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + rw [←tangentCoordChange_def] at h + have hf3' := mem_of_mem_of_subset hf3 interior_subset + -- express the derivative of the integral curve in the local chart + refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), + HasDerivWithinAt.hasFDerivWithinAt ?_⟩ + simp only [mfld_simps, hasDerivWithinAt_univ] + show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) + (v ((extChartAt I x₀).symm (f t))) t + -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of + -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use + -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires + -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of + -- coordinates as `D`. + have hft1 := mem_preimage.mp <| + mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) + rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) + (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, + ←tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] + apply HasFDerivAt.comp_hasDerivAt _ _ h + apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| + mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), + isOpen_interior, hf3⟩ + nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] + exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the From 4cf8097f226ca71224eab3c6aa40160c3483c5ef Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 22:22:04 -0800 Subject: [PATCH 075/203] shorten comment --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 15 ++++++--------- 1 file changed, 6 insertions(+), 9 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 6b936ba66f0fb..736445adc2e80 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -223,20 +223,17 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) rw [←tangentCoordChange_def] at h have hf3' := mem_of_mem_of_subset hf3 interior_subset + have hft1 := mem_preimage.mp <| + mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) -- express the derivative of the integral curve in the local chart refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), HasDerivWithinAt.hasFDerivWithinAt ?_⟩ simp only [mfld_simps, hasDerivWithinAt_univ] show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) - (v ((extChartAt I x₀).symm (f t))) t - -- `h` gives the derivative of `f` at `t` as `↑D (v (γ t))`, where `D` is the change of - -- coordinates from the chart at `γ t` to the chart at `x₀`. we wish to use - -- `HasFDerivAt.comp_hasDerivAt` to get the derivative of `γ` at `t` as `v (γ t)`, which requires - -- first expressing `v (γ t)` as `↑D_inv ↑D (v (γ t))`, where `D_inv` is the opposite change of - -- coordinates as `D`. - have hft1 := mem_preimage.mp <| - mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) + (v ((extChartAt I x₀).symm (f t))) t + -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use + -- `HasFDerivAt.comp_hasDerivAt` on `h` rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, ←tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] From e8bfdb0a75e5e0320d4f9f18ad9cd790a1c338a7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 22:31:03 -0800 Subject: [PATCH 076/203] shorten --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 736445adc2e80..3c5f460cc9aa3 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -201,7 +201,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin rw [←Real.ball_eq_Ioo] at hf2 -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, -- `f t ∈ interior (extChartAt I x₀).target` - have hcont := (hf2 t₀ (Real.ball_eq_Ioo .. ▸ Metric.mem_ball_self hε1)).continuousAt + have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) From de44a9c332d81c84f1527251667ccc2b35cc14c3 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 22:55:16 -0800 Subject: [PATCH 077/203] move tangentCoordChange --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 43 ------------------- .../Manifold/VectorBundle/Tangent.lean | 41 ++++++++++++++++++ 2 files changed, 41 insertions(+), 43 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 3c5f460cc9aa3..f911eccdcd25b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -31,49 +31,6 @@ integral curve, vector field, local existence open scoped Manifold open Set -section - -variable - {𝕜 : Type*} [NontriviallyNormedField 𝕜] - {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] - {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} - {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] - {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] - {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} - {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] - -variable (I) in -def tangentCoordChange (x y : M) := (tangentBundleCore I M).coordChange (achart H x) (achart H y) - -lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = - fderivWithin 𝕜 (extChartAt I y ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl - -lemma tangentCoordChange_self {x z : M} {v : E} (h : z ∈ (extChartAt I x).source) : - tangentCoordChange I x x z v = v := by - apply (tangentBundleCore I M).coordChange_self - rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] - exact h - --- continuousOn? - -lemma tangentCoordChange_comp {w x y z : M} {v : E} - (h : z ∈ (extChartAt I w).source ∩ (extChartAt I x).source ∩ (extChartAt I y).source) : - tangentCoordChange I x y z (tangentCoordChange I w x z v) = tangentCoordChange I w y z v := by - apply (tangentBundleCore I M).coordChange_comp - simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] - exact h - -lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} - (h : z ∈ (extChartAt I x).source ∩ (extChartAt I y).source) : - HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) - (range I) (extChartAt I x z) := - have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by - rw [LocalEquiv.trans_source'', LocalEquiv.symm_symm, LocalEquiv.symm_target] - exact mem_image_of_mem _ h - ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt - -end - variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index aeb31731a39da..5d15ee7be9fb3 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -124,6 +124,47 @@ theorem tangentBundleCore_coordChange_achart (x x' z : M) : rfl #align tangent_bundle_core_coord_change_achart tangentBundleCore_coordChange_achart +section tangentCoordChange + +/-- Since `tangentBundleCore` has the same base sets as the preferred charts of the base manifold, + we define `tangentCoordChange I x y` as a convenient abbreviation for a change of coordinates + between the preferred charts at `x y : M`. -/ +abbrev tangentCoordChange (x y : M) := (tangentBundleCore I M).coordChange (achart H x) (achart H y) + +variable {I} + +lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = + fderivWithin 𝕜 (extChartAt I y ∘ (extChartAt I x).symm) (range I) (extChartAt I x z) := rfl + +lemma tangentCoordChange_self {x z : M} {v : E} (h : z ∈ (extChartAt I x).source) : + tangentCoordChange I x x z v = v := by + apply (tangentBundleCore I M).coordChange_self + rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + exact h + +lemma tangentCoordChange_comp {w x y z : M} {v : E} + (h : z ∈ (extChartAt I w).source ∩ (extChartAt I x).source ∩ (extChartAt I y).source) : + tangentCoordChange I x y z (tangentCoordChange I w x z v) = tangentCoordChange I w y z v := by + apply (tangentBundleCore I M).coordChange_comp + simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + exact h + +lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} + (h : z ∈ (extChartAt I x).source ∩ (extChartAt I y).source) : + HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) + (range I) (extChartAt I x z) := + have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by + rw [LocalEquiv.trans_source'', LocalEquiv.symm_symm, LocalEquiv.symm_target] + exact mem_image_of_mem _ h + ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt + +lemma continuousOn_tangentCoordChange (x y : M) : ContinuousOn (tangentCoordChange I x y) + ((extChartAt I x).source ∩ (extChartAt I y).source) := by + convert (tangentBundleCore I M).continuousOn_coordChange (achart H x) (achart H y) <;> + simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + +end tangentCoordChange + /-- The tangent space at a point of the manifold `M`. It is just `E`. We could use instead `(tangentBundleCore I M).to_topological_vector_bundle_core.fiber x`, but we use `E` to help the kernel. From a792394f2e14c2834118ca5b0ae652868997df19 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 27 Nov 2023 23:13:20 -0800 Subject: [PATCH 078/203] minor spacing --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index d949315e7f0c2..634f4a948bcce 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -57,7 +57,7 @@ namespace SmoothManifoldWithCorners variable (I M) in /-- The **interior** of a manifold `M` is the set of its interior points. -/ -protected def interior : Set M := { x : M | I.IsInteriorPoint x} +protected def interior : Set M := { x : M | I.IsInteriorPoint x } lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := @@ -66,7 +66,7 @@ lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ -protected def boundary : Set M := { x : M | I.IsBoundaryPoint x} +protected def boundary : Set M := { x : M | I.IsBoundaryPoint x } lemma _root_.ModelWithCorners.isBoundaryPoint_iff {x : M} : I.IsBoundaryPoint x ↔ extChartAt I x x ∈ frontier (range I) := Iff.rfl From 4cc6c69284caf9078007cdbb1ad2acfcbf464a0f Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 28 Nov 2023 01:04:56 -0800 Subject: [PATCH 079/203] isIntegralCurveAt_const --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index f911eccdcd25b..56700f91f14a7 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -142,6 +142,13 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ ext simp [inv_mul_eq_div, div_self ha] +variable (t₀) in +lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ x₀ := by + refine ⟨rfl, 1, zero_lt_one, fun t _ => ?_⟩ + rw [h, ←ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), + ContinuousLinearMap.smulRight_one_one] + exact hasMFDerivAt_const .. + /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around From b4ba5a08cae84183bc87ecd0f58d22da3f73d105 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 28 Nov 2023 01:21:00 -0800 Subject: [PATCH 080/203] remove todo --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 56700f91f14a7..9524c3c5b7cfa 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -47,12 +47,6 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) -/- -TODO: -* split the theorem below into smaller lemmas, e.g. involving IsIntegralCurveAt? -* shift and stretch theorems -* constant curve at stationary point of v --/ variable {t₀} lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : From f14affaee86bace85751c5727106fadb65c612a4 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 28 Nov 2023 14:54:20 -0800 Subject: [PATCH 081/203] use implicit lambdas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 9524c3c5b7cfa..ddfc108c3a5c8 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -50,7 +50,7 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ variable {t₀} lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : - IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht @@ -65,7 +65,7 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v apply HasFDerivAt.add_const (hasFDerivAt_id _) lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ x₀ ↔ - IsIntegralCurveAt (γ ∘ (fun t => t + dt)) v (t₀ - dt) x₀ := by + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) x₀ := by refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this @@ -74,7 +74,7 @@ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt simp only [Function.comp_apply, neg_add_cancel_right] lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} - (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by obtain ⟨h1, ε, hε, h2⟩ := hγ refine ⟨by rw [Function.comp_apply, div_mul_cancel _ (ne_of_gt ha)]; exact h1, ε / a, div_pos hε ha, ?_⟩ @@ -88,7 +88,7 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt apply HasFDerivAt.mul_const' (hasFDerivAt_id _) lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_pos (inv_pos_of_pos ha) rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ←div_mul_eq_div_div_swap, @@ -117,17 +117,17 @@ lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : exact this lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} - (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by rw [ne_iff_lt_or_gt] at ha cases' ha with ha ha · apply isIntegralCurveAt_comp_neg.mpr - have : (fun t ↦ t * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp + have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp rw [Function.comp.assoc, this, ←neg_smul, ←div_neg] exact hγ.comp_mul_pos (neg_pos_of_neg ha) · exact hγ.comp_mul_pos ha lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (fun t => t * a)) (a • v) (t₀ / a) x₀ := by + IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ←div_mul_eq_div_div_swap, From 0a3d3597c3079a469455ad95ce6dea428357ccf8 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 29 Nov 2023 15:38:04 +0100 Subject: [PATCH 082/203] Fix style. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 34 ++++++++++---------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index ddfc108c3a5c8..c99b9ac8c043f 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -54,10 +54,10 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v obtain ⟨h1, ε, hε, h2⟩ := hγ refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht - rw [sub_right_comm, sub_add_eq_add_sub, ←add_mem_Ioo_iff_left] at ht + rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] at ht have h2' := h2 (t + dt) ht rw [Function.comp_apply, - ←ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] + ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] apply HasMFDerivAt.comp t h2' /- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations -/ refine ⟨(continuous_add_right _).continuousAt, ?_⟩ @@ -80,9 +80,9 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt div_pos hε ha, ?_⟩ intros t ht have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by - rw [mem_Ioo, ←div_lt_iff ha, ←lt_div_iff ha, sub_div, add_div] + rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] exact ht - rw [Function.comp_apply, Pi.smul_apply, ←ContinuousLinearMap.smulRight_comp] + rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] apply HasFDerivAt.mul_const' (hasFDerivAt_id _) @@ -91,7 +91,7 @@ lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_pos (inv_pos_of_pos ha) - rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ←div_mul_eq_div_div_swap, + rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this convert this ext @@ -102,9 +102,9 @@ lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v obtain ⟨h1, ε, hε, h2⟩ := hγ refine ⟨by simp [h1], ε, hε, ?_⟩ intros t ht - rw [←neg_add', neg_add_eq_sub, ←neg_sub, ←neg_mem_Ioo_iff] at ht - rw [Function.comp_apply, Pi.neg_apply, ←neg_one_smul ℝ (v (γ (-t))), - ←ContinuousLinearMap.smulRight_comp] + rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht + rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), + ← ContinuousLinearMap.smulRight_comp] apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) @@ -122,7 +122,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve cases' ha with ha ha · apply isIntegralCurveAt_comp_neg.mpr have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp - rw [Function.comp.assoc, this, ←neg_smul, ←div_neg] + rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] exact hγ.comp_mul_pos (neg_pos_of_neg ha) · exact hγ.comp_mul_pos ha @@ -130,7 +130,7 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) - rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ←div_mul_eq_div_div_swap, + rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this convert this ext @@ -139,7 +139,7 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ variable (t₀) in lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ x₀ := by refine ⟨rfl, 1, zero_lt_one, fun t _ => ?_⟩ - rw [h, ←ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), + rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. @@ -156,7 +156,7 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd - rw [←Real.ball_eq_Ioo] at hf2 + rw [← Real.ball_eq_Ioo] at hf2 -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, -- `f t ∈ interior (extChartAt I x₀).target` have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt @@ -172,14 +172,14 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin min ε1 ε2, lt_min hε1 hε2, ?_⟩ intros t ht -- collect useful terms in convenient forms - rw [←Real.ball_eq_Ioo] at ht + rw [← Real.ball_eq_Ioo] at ht have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) have h : HasDerivAt f ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) (v ((extChartAt I x₀).symm (f t)))) t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) - rw [←tangentCoordChange_def] at h + rw [← tangentCoordChange_def] at h have hf3' := mem_of_mem_of_subset hf3 interior_subset have hft1 := mem_preimage.mp <| mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source @@ -192,15 +192,15 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin (v ((extChartAt I x₀).symm (f t))) t -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use -- `HasFDerivAt.comp_hasDerivAt` on `h` - rw [←tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) + rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, - ←tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] + ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] apply HasFDerivAt.comp_hasDerivAt _ _ h apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3⟩ - nth_rw 4 [←(extChartAt I x₀).right_inv hf3'] + nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ /-- For any continuously differentiable vector field defined on a manifold without boundary and any From fa1b5f350985ceb1e371215dc4b2590412378ab5 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Wed, 29 Nov 2023 15:38:31 +0100 Subject: [PATCH 083/203] Remove superfluous helper lemma. --- Mathlib/Geometry/Manifold/ChartedSpace.lean | 7 ------- 1 file changed, 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 98bb16723d077..5db698d70957e 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -584,13 +584,6 @@ lemma chart_mem_atlas (H : Type*) {M : Type*} [TopologicalSpace H] [TopologicalS section ChartedSpace -/-- The charts of a charted space cover its domain. -/ -lemma ChartedSpace.covering {H M : Type*} [TopologicalSpace H] [TopologicalSpace M] - [ChartedSpace H M] : ⋃ x : M, (chartAt H x).source = univ := by - refine subset_antisymm (fun y _ ↦ trivial) (fun y _ ↦ ?_) - rw [mem_iUnion] - exact ⟨y, mem_chart_source H y⟩ - /-- 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} From 9e9ee92724fd92f86866b7c59b867956d654b3c0 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 29 Nov 2023 16:14:00 -0800 Subject: [PATCH 084/203] docstring, rename theorem --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 24 ++++++++++++++------ 1 file changed, 17 insertions(+), 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c99b9ac8c043f..c02701b6bfdb9 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -19,9 +19,16 @@ without boundary. ## Main definition -- **IsIntegralCurve γ v t₀ x₀**: If `v : M → TM` is a vector field on `M` and `x : M`, -`IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a differentiable integral curve of `v` with -`γ x₀ = t₀`. +- **`IsIntegralCurveAt γ v t₀ x₀`**: If `v : M → TM` is a vector field on `M` and `x : M`, +`IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a local integral curve of `v` with `γ t₀ = x₀`. +That is, there exists `ε > 0` such that `γ t` is tangent to `v (γ t)` for all +`t ∈ Ioo (t₀ - ε) (t₀ + ε)`. Even though `γ` is defined for all time, its value outside of this +small interval is irrelevant and considered junk. + +## To-do + +- **`IsIntegralCurveOn γ v s`**: `γ t` is tangent to `v (γ t)` for all `t ∈ s`. +- **`IsIntegralCurve γ v`**: `γ` is a global integral curve of `v`, defined for all time. ## Tags @@ -42,7 +49,8 @@ variable (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) /-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means - `γ : ℝ → M` is a differentiable integral curve of `v` with `γ x₀ = t₀`. -/ + `γ : ℝ → M` is a local integral curve of `v` with `γ t₀ = x₀`. That is, there exists `ε > 0` such + that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. -/ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) (x₀ : M) := γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) @@ -59,7 +67,7 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v rw [Function.comp_apply, ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] apply HasMFDerivAt.comp t h2' - /- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations -/ + -- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations refine ⟨(continuous_add_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] apply HasFDerivAt.add_const (hasFDerivAt_id _) @@ -137,6 +145,8 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ simp [inv_mul_eq_div, div_self ha] variable (t₀) in +/-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` + is an integral curve of `v`. -/ lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ x₀ := by refine ⟨rfl, 1, zero_lt_one, fun t _ => ?_⟩ rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), @@ -147,7 +157,7 @@ lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoint x₀) : +theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by -- express the differentiability of the section `v` in the local charts rw [contMDiffAt_iff] at hv @@ -209,4 +219,4 @@ theorem exists_integralCurve_of_contMDiff_tangent_section (hx : I.IsInteriorPoin interval around `t₀`. -/ lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := - exists_integralCurve_of_contMDiff_tangent_section hv I.isInteriorPoint + exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint From 75c2b52bdded7843a2053c4a564531ea7af95d93 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 29 Nov 2023 16:19:32 -0800 Subject: [PATCH 085/203] rename --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c02701b6bfdb9..408d33c8363cf 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -217,6 +217,6 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ -lemma exists_integralCurve_of_contMDiff_tangent_section_boundaryless [I.Boundaryless] : +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint From f019e9bc44510e23b655a2d08840b56601473aa1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 29 Nov 2023 16:28:41 -0800 Subject: [PATCH 086/203] =?UTF-8?q?remove=20`x=E2=82=80`=20from=20`IsInteg?= =?UTF-8?q?ralCurveAt`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 66 +++++++++----------- 1 file changed, 31 insertions(+), 35 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 408d33c8363cf..b864210aa33af 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -19,9 +19,9 @@ without boundary. ## Main definition -- **`IsIntegralCurveAt γ v t₀ x₀`**: If `v : M → TM` is a vector field on `M` and `x : M`, -`IsIntegralCurveAt γ v t₀ x₀` means `γ : ℝ → M` is a local integral curve of `v` with `γ t₀ = x₀`. -That is, there exists `ε > 0` such that `γ t` is tangent to `v (γ t)` for all +- **`IsIntegralCurveAt γ v t₀`**: If `v : M → TM` is a vector field on `M` and `x : M`, +`IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in an open interval of +`t₀`. That is, there exists `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. Even though `γ` is defined for all time, its value outside of this small interval is irrelevant and considered junk. @@ -48,20 +48,19 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) -/-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀ x₀` means - `γ : ℝ → M` is a local integral curve of `v` with `γ t₀ = x₀`. That is, there exists `ε > 0` such - that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. -/ -def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) (x₀ : M) := - γ t₀ = x₀ ∧ ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → +/-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀` means + `γ : ℝ → M` is a local integral curve of `v` in an open interval of `t₀`. That is, there exists + `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. -/ +def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) := + ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) variable {t₀} -lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) (dt : ℝ) : - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) x₀ := by - obtain ⟨h1, ε, hε, h2⟩ := hγ - refine ⟨by simp [h1], ε, hε, ?_⟩ - intros t ht +lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by + obtain ⟨ε, hε, h2⟩ := hγ + refine ⟨ε, hε, fun t ht => ?_⟩ rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] at ht have h2' := h2 (t + dt) ht rw [Function.comp_apply, @@ -72,8 +71,8 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v simp only [mfld_simps, hasFDerivWithinAt_univ] apply HasFDerivAt.add_const (hasFDerivAt_id _) -lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ x₀ ↔ - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) x₀ := by +lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this @@ -81,12 +80,10 @@ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt ext simp only [Function.comp_apply, neg_add_cancel_right] -lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} - (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by - obtain ⟨h1, ε, hε, h2⟩ := hγ - refine ⟨by rw [Function.comp_apply, div_mul_cancel _ (ne_of_gt ha)]; exact h1, ε / a, - div_pos hε ha, ?_⟩ - intros t ht +lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} + (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by + obtain ⟨ε, hε, h2⟩ := hγ + refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] exact ht @@ -96,7 +93,7 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt apply HasFDerivAt.mul_const' (hasFDerivAt_id _) lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_pos (inv_pos_of_pos ha) rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, @@ -105,11 +102,10 @@ lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : ext simp [inv_mul_eq_div, div_self (ne_of_gt ha)] -lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) : - IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by - obtain ⟨h1, ε, hε, h2⟩ := hγ - refine ⟨by simp [h1], ε, hε, ?_⟩ - intros t ht +lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : + IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by + obtain ⟨ε, hε, h2⟩ := hγ + refine ⟨ε, hε, fun t ht => ?_⟩ rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), ← ContinuousLinearMap.smulRight_comp] @@ -118,14 +114,14 @@ lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) x₀ := by + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ have := hγ.comp_neg rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this exact this -lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀ x₀) {a : ℝ} - (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by +lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} + (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by rw [ne_iff_lt_or_gt] at ha cases' ha with ha ha · apply isIntegralCurveAt_comp_neg.mpr @@ -135,7 +131,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve · exact hγ.comp_mul_pos ha lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : - IsIntegralCurveAt γ v t₀ x₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) x₀ := by + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, @@ -147,8 +143,8 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ variable (t₀) in /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is an integral curve of `v`. -/ -lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ x₀ := by - refine ⟨rfl, 1, zero_lt_one, fun t _ => ?_⟩ +lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ := by + refine ⟨1, zero_lt_one, fun t _ => ?_⟩ rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. @@ -158,7 +154,7 @@ lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : - ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := by + ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by -- express the differentiability of the section `v` in the local charts rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv @@ -218,5 +214,5 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : - ∃ (γ : ℝ → M), IsIntegralCurveAt γ v t₀ x₀ := + ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint From 30c333f7de9944ca6d7ef17bd8442ac73891e5ec Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 29 Nov 2023 20:43:43 -0800 Subject: [PATCH 087/203] end with exact --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b864210aa33af..2cbe91aa922e6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -69,7 +69,7 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v -- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations refine ⟨(continuous_add_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] - apply HasFDerivAt.add_const (hasFDerivAt_id _) + exact HasFDerivAt.add_const (hasFDerivAt_id _) _ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -90,7 +90,7 @@ lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] - apply HasFDerivAt.mul_const' (hasFDerivAt_id _) + exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by From e140c4362472315719e56decf94d7f67274c97b4 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 04:02:57 -0800 Subject: [PATCH 088/203] `IsIntegralCurveOn`, `IsIntegralCurve` --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 63 ++++++++++++++++---- 1 file changed, 50 insertions(+), 13 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 2cbe91aa922e6..165b22cdf4d5d 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -19,16 +19,21 @@ without boundary. ## Main definition -- **`IsIntegralCurveAt γ v t₀`**: If `v : M → TM` is a vector field on `M` and `x : M`, -`IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in an open interval of -`t₀`. That is, there exists `ε > 0` such that `γ t` is tangent to `v (γ t)` for all -`t ∈ Ioo (t₀ - ε) (t₀ + ε)`. Even though `γ` is defined for all time, its value outside of this -small interval is irrelevant and considered junk. +Let `v : M → TM` be a vector field on `M`, and let `γ : ℝ → M`. +- **`IsIntegralCurve γ v`**: `γ t` is tangent to `v (γ t)` for all `t : ℝ`. That is, `γ` is a global +integral curve of `v`. +- **`IsIntegralCurveOn γ v s`**: `γ t` is tangent to `v (γ t)` for all `t ∈ s`, where `s : Set ℝ`. +- **`IsIntegralCurveAt γ v t₀`**: `γ t` is tangent to `v (γ t)` for all `t` in some open interval +around `t₀`. That is, `γ` is a local integral curve of `v`. + +For `IsIntegralCurveOn γ v s` and `IsIntegralCurveAt γ v t₀`, even though `γ` is defined for all +time, its value outside of the set `s` or a small interval around `t₀` is irrelevant and considered +junk. ## To-do -- **`IsIntegralCurveOn γ v s`**: `γ t` is tangent to `v (γ t)` for all `t ∈ s`. -- **`IsIntegralCurve γ v`**: `γ` is a global integral curve of `v`, defined for all time. +- Prove `comp_add`, `comp_smul` , etc. lemmas for `IsIntegralCurveOn`, and then derive versions for +`IsIntegralCurveAt` and `IsIntegralCurve` as corollaries. ## Tags @@ -48,12 +53,44 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) +def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : Set ℝ) := + ∀ (t : ℝ), t ∈ s → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) + /-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in an open interval of `t₀`. That is, there exists `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. -/ -def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) := - ∃ ε > (0 : ℝ), ∀ (t : ℝ), t ∈ Ioo (t₀ - ε) (t₀ + ε) → - HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) +def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) := + ∃ ε > (0 : ℝ), IsIntegralCurveOn γ v (Ioo (t - ε) (t + ε)) + +def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := + ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) + +lemma IsIntegralCurve.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} + (h : IsIntegralCurve γ v) (s : Set ℝ) : IsIntegralCurveOn γ v s := fun t _ => h t + +lemma isIntegralCurve_iff_isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : + IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := + ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ + +lemma IsIntegralCurve.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} + (h : IsIntegralCurve γ v) (t : ℝ) : IsIntegralCurveAt γ v t := + ⟨1, zero_lt_one, fun t _ => h t⟩ + +lemma isIntegralCurve_iff_isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : + IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := + ⟨fun h => h.isIntegralCurveAt, fun h t => by + obtain ⟨ε, hε, h⟩ := h t + exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε)⟩ + +lemma IsIntegralCurveOn.mono {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} + (h : IsIntegralCurveOn γ v s) {s' : Set ℝ} (hs : s' ⊆ s) : IsIntegralCurveOn γ v s' := + fun t ht => h t (mem_of_mem_of_subset ht hs) + +lemma IsIntegralCurveOn.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} + (h : IsIntegralCurveOn γ v s) {t : ℝ} (hs : s ∈ nhds t) : IsIntegralCurveAt γ v t := by + rw [Metric.mem_nhds_iff] at hs + obtain ⟨ε, hε, hmem⟩ := hs + exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ variable {t₀} @@ -142,9 +179,9 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ variable (t₀) in /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` - is an integral curve of `v`. -/ -lemma isIntegralCurveAt_const (h : v x₀ = 0) : IsIntegralCurveAt (fun _ => x₀) v t₀ := by - refine ⟨1, zero_lt_one, fun t _ => ?_⟩ + is a global integral curve of `v`. -/ +lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by + intro t rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. From b810d2a4c4cdd12fb532f9aef03d7ca8e2e04a91 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 12:53:46 -0800 Subject: [PATCH 089/203] docstring --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 165b22cdf4d5d..4e31df0520c9f 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -53,15 +53,21 @@ variable {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) +/-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, + `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` + outside of `s` is irrelevant and considered junk. -/ def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : Set ℝ) := ∀ (t : ℝ), t ∈ s → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) -/-- If `v : M → TM` is a vector field on `M` and `x : M`, `IsIntegralCurveAt γ v t₀` means +/-- If `v : M → TM` is a vector field on `M`, and `t₀ : ℝ`, `IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in an open interval of `t₀`. That is, there exists - `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. -/ + `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. The value of + `γ` outside of this interval is irrelevant and considered junk. -/ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) := ∃ ε > (0 : ℝ), IsIntegralCurveOn γ v (Ioo (t - ε) (t + ε)) +/-- If `v : M → TM` is a vector field on `M`, `IsIntegralCurve γ v` means `γ : ℝ → M` is a global + integral curve of `v`. That is, `γ t` is tangent to `v (γ t)` for all `t : ℝ`. -/ def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) From 9f5b559d0b8da3c6169a4e0663caee99b3995335 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 13:00:25 -0800 Subject: [PATCH 090/203] IsIntegralCurveAt.isIntegralCurveOn --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 326 ++++++++++--------- 1 file changed, 166 insertions(+), 160 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4e31df0520c9f..6e76f5ae7aea3 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -98,164 +98,170 @@ lemma IsIntegralCurveOn.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → Tang obtain ⟨ε, hε, hmem⟩ := hs exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ -variable {t₀} - -lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by - obtain ⟨ε, hε, h2⟩ := hγ - refine ⟨ε, hε, fun t ht => ?_⟩ - rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] at ht - have h2' := h2 (t + dt) ht - rw [Function.comp_apply, - ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] - apply HasMFDerivAt.comp t h2' - -- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations - refine ⟨(continuous_add_right _).continuousAt, ?_⟩ - simp only [mfld_simps, hasFDerivWithinAt_univ] - exact HasFDerivAt.add_const (hasFDerivAt_id _) _ - -lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ - have := hγ.comp_add (-dt) - rw [sub_neg_eq_add, sub_add_cancel] at this - convert this - ext - simp only [Function.comp_apply, neg_add_cancel_right] - -lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} - (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - obtain ⟨ε, hε, h2⟩ := hγ - refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ - have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by - rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] - exact ht - rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] - refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ - simp only [mfld_simps, hasFDerivWithinAt_univ] - exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ - -lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : - IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ - have := hγ.comp_mul_pos (inv_pos_of_pos ha) - rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, - inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this - convert this - ext - simp [inv_mul_eq_div, div_self (ne_of_gt ha)] - -lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : - IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by - obtain ⟨ε, hε, h2⟩ := hγ - refine ⟨ε, hε, fun t ht => ?_⟩ - rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht - rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), - ← ContinuousLinearMap.smulRight_comp] - apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ - simp only [mfld_simps, hasFDerivWithinAt_univ] - exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) - -lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : - IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ - have := hγ.comp_neg - rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this - exact this - -lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} - (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - rw [ne_iff_lt_or_gt] at ha - cases' ha with ha ha - · apply isIntegralCurveAt_comp_neg.mpr - have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp - rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] - exact hγ.comp_mul_pos (neg_pos_of_neg ha) - · exact hγ.comp_mul_pos ha - -lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : - IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ - have := hγ.comp_mul_ne_zero (inv_ne_zero ha) - rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, - inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this - convert this - ext - simp [inv_mul_eq_div, div_self ha] - -variable (t₀) in -/-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` - is a global integral curve of `v`. -/ -lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by - intro t - rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), - ContinuousLinearMap.smulRight_one_one] - exact hasMFDerivAt_const .. - -/-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the - manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector - of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around - `t₀`.-/ -theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : - ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by - -- express the differentiability of the section `v` in the local charts - rw [contMDiffAt_iff] at hv - obtain ⟨_, hv⟩ := hv - -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart - obtain ⟨f, hf1, ε1, hε1, hf2⟩ := - exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd - rw [← Real.ball_eq_Ioo] at hf2 - -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, - -- `f t ∈ interior (extChartAt I x₀).target` - have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt - rw [continuousAt_def, hf1] at hcont - have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := - hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) - rw [Metric.mem_nhds_iff] at hnhds - obtain ⟨ε2, hε2, hf3⟩ := hnhds - simp_rw [subset_def, mem_preimage] at hf3 - -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve - refine ⟨(extChartAt I x₀).symm ∘ f, - Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), - min ε1 ε2, lt_min hε1 hε2, ?_⟩ +lemma IsIntegralCurveAt.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} + (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by intros t ht - -- collect useful terms in convenient forms - rw [← Real.ball_eq_Ioo] at ht - have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) - have h : HasDerivAt f - ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) - (v ((extChartAt I x₀).symm (f t)))) - t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) - rw [← tangentCoordChange_def] at h - have hf3' := mem_of_mem_of_subset hf3 interior_subset - have hft1 := mem_preimage.mp <| - mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) - -- express the derivative of the integral curve in the local chart - refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), - HasDerivWithinAt.hasFDerivWithinAt ?_⟩ - simp only [mfld_simps, hasDerivWithinAt_univ] - show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) - (v ((extChartAt I x₀).symm (f t))) t - -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use - -- `HasFDerivAt.comp_hasDerivAt` on `h` - rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) - (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, - ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] - apply HasFDerivAt.comp_hasDerivAt _ _ h - apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| - mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, - subset_trans interior_subset (extChartAt_target_subset_range ..), - isOpen_interior, hf3⟩ - nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] - exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ - -/-- For any continuously differentiable vector field defined on a manifold without boundary and any - chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the - tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open - interval around `t₀`. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : - ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint + obtain ⟨ε, hε, h⟩ := h t ht + exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) + +-- variable {t₀} + +-- lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : +-- IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by +-- obtain ⟨ε, hε, h2⟩ := hγ +-- refine ⟨ε, hε, fun t ht => ?_⟩ +-- rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] at ht +-- have h2' := h2 (t + dt) ht +-- rw [Function.comp_apply, +-- ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] +-- apply HasMFDerivAt.comp t h2' +-- -- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations +-- refine ⟨(continuous_add_right _).continuousAt, ?_⟩ +-- simp only [mfld_simps, hasFDerivWithinAt_univ] +-- exact HasFDerivAt.add_const (hasFDerivAt_id _) _ + +-- lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ +-- IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by +-- refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ +-- have := hγ.comp_add (-dt) +-- rw [sub_neg_eq_add, sub_add_cancel] at this +-- convert this +-- ext +-- simp only [Function.comp_apply, neg_add_cancel_right] + +-- lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} +-- (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +-- obtain ⟨ε, hε, h2⟩ := hγ +-- refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ +-- have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by +-- rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] +-- exact ht +-- rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] +-- refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ +-- simp only [mfld_simps, hasFDerivWithinAt_univ] +-- exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ + +-- lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : +-- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +-- refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ +-- have := hγ.comp_mul_pos (inv_pos_of_pos ha) +-- rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, +-- inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this +-- convert this +-- ext +-- simp [inv_mul_eq_div, div_self (ne_of_gt ha)] + +-- lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : +-- IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by +-- obtain ⟨ε, hε, h2⟩ := hγ +-- refine ⟨ε, hε, fun t ht => ?_⟩ +-- rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht +-- rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), +-- ← ContinuousLinearMap.smulRight_comp] +-- apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ +-- simp only [mfld_simps, hasFDerivWithinAt_univ] +-- exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) + +-- lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : +-- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by +-- refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ +-- have := hγ.comp_neg +-- rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this +-- exact this + +-- lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} +-- (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +-- rw [ne_iff_lt_or_gt] at ha +-- cases' ha with ha ha +-- · apply isIntegralCurveAt_comp_neg.mpr +-- have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp +-- rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] +-- exact hγ.comp_mul_pos (neg_pos_of_neg ha) +-- · exact hγ.comp_mul_pos ha + +-- lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +-- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +-- refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ +-- have := hγ.comp_mul_ne_zero (inv_ne_zero ha) +-- rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, +-- inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this +-- convert this +-- ext +-- simp [inv_mul_eq_div, div_self ha] + +-- variable (t₀) in +-- /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` +-- is a global integral curve of `v`. -/ +-- lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by +-- intro t +-- rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), +-- ContinuousLinearMap.smulRight_one_one] +-- exact hasMFDerivAt_const .. + +-- /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the +-- manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector +-- of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around +-- `t₀`.-/ +-- theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : +-- ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by +-- -- express the differentiability of the section `v` in the local charts +-- rw [contMDiffAt_iff] at hv +-- obtain ⟨_, hv⟩ := hv +-- -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart +-- obtain ⟨f, hf1, ε1, hε1, hf2⟩ := +-- exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ +-- (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd +-- rw [← Real.ball_eq_Ioo] at hf2 +-- -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, +-- -- `f t ∈ interior (extChartAt I x₀).target` +-- have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt +-- rw [continuousAt_def, hf1] at hcont +-- have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := +-- hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) +-- rw [Metric.mem_nhds_iff] at hnhds +-- obtain ⟨ε2, hε2, hf3⟩ := hnhds +-- simp_rw [subset_def, mem_preimage] at hf3 +-- -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve +-- refine ⟨(extChartAt I x₀).symm ∘ f, +-- Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), +-- min ε1 ε2, lt_min hε1 hε2, ?_⟩ +-- intros t ht +-- -- collect useful terms in convenient forms +-- rw [← Real.ball_eq_Ioo] at ht +-- have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) +-- have h : HasDerivAt f +-- ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) +-- (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) +-- (v ((extChartAt I x₀).symm (f t)))) +-- t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) +-- rw [← tangentCoordChange_def] at h +-- have hf3' := mem_of_mem_of_subset hf3 interior_subset +-- have hft1 := mem_preimage.mp <| +-- mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source +-- have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) +-- -- express the derivative of the integral curve in the local chart +-- refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), +-- HasDerivWithinAt.hasFDerivWithinAt ?_⟩ +-- simp only [mfld_simps, hasDerivWithinAt_univ] +-- show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) +-- (v ((extChartAt I x₀).symm (f t))) t +-- -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use +-- -- `HasFDerivAt.comp_hasDerivAt` on `h` +-- rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) +-- (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, +-- ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] +-- apply HasFDerivAt.comp_hasDerivAt _ _ h +-- apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| +-- mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, +-- subset_trans interior_subset (extChartAt_target_subset_range ..), +-- isOpen_interior, hf3⟩ +-- nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] +-- exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ + +-- /-- For any continuously differentiable vector field defined on a manifold without boundary and any +-- chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the +-- tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open +-- interval around `t₀`. -/ +-- lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : +-- ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := +-- exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint From 044593824351cf24d66f0cd1f62398a9c33d3236 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 13:50:03 -0800 Subject: [PATCH 091/203] `comp_add` lemmas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 356 ++++++++++--------- 1 file changed, 194 insertions(+), 162 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 6e76f5ae7aea3..f10db74be7594 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -51,7 +51,7 @@ variable {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners ℝ E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] {v : (x : M) → TangentSpace I x} {x₀ : M} - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (t₀ : ℝ) + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) /-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` @@ -104,164 +104,196 @@ lemma IsIntegralCurveAt.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → Tang obtain ⟨ε, hε, h⟩ := h t ht exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) --- variable {t₀} - --- lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : --- IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by --- obtain ⟨ε, hε, h2⟩ := hγ --- refine ⟨ε, hε, fun t ht => ?_⟩ --- rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] at ht --- have h2' := h2 (t + dt) ht --- rw [Function.comp_apply, --- ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] --- apply HasMFDerivAt.comp t h2' --- -- this makes me think we need lemmas for `HasMFDerivAt 𝓘(E, E) 𝓘(E, E)` of simple operations --- refine ⟨(continuous_add_right _).continuousAt, ?_⟩ --- simp only [mfld_simps, hasFDerivWithinAt_univ] --- exact HasFDerivAt.add_const (hasFDerivAt_id _) _ - --- lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ --- IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by --- refine ⟨fun hγ => IsIntegralCurveAt.comp_add hγ _, fun hγ ↦ ?_⟩ --- have := hγ.comp_add (-dt) --- rw [sub_neg_eq_add, sub_add_cancel] at this --- convert this --- ext --- simp only [Function.comp_apply, neg_add_cancel_right] - --- lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} --- (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by --- obtain ⟨ε, hε, h2⟩ := hγ --- refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ --- have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by --- rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] --- exact ht --- rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] --- refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ --- simp only [mfld_simps, hasFDerivWithinAt_univ] --- exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ - --- lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : --- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by --- refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ --- have := hγ.comp_mul_pos (inv_pos_of_pos ha) --- rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, --- inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this --- convert this --- ext --- simp [inv_mul_eq_div, div_self (ne_of_gt ha)] - --- lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : --- IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by --- obtain ⟨ε, hε, h2⟩ := hγ --- refine ⟨ε, hε, fun t ht => ?_⟩ --- rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht --- rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), --- ← ContinuousLinearMap.smulRight_comp] --- apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ --- simp only [mfld_simps, hasFDerivWithinAt_univ] --- exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) - --- lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : --- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by --- refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ --- have := hγ.comp_neg --- rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this --- exact this - --- lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} --- (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by --- rw [ne_iff_lt_or_gt] at ha --- cases' ha with ha ha --- · apply isIntegralCurveAt_comp_neg.mpr --- have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp --- rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] --- exact hγ.comp_mul_pos (neg_pos_of_neg ha) --- · exact hγ.comp_mul_pos ha - --- lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : --- IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by --- refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ --- have := hγ.comp_mul_ne_zero (inv_ne_zero ha) --- rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, --- inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this --- convert this --- ext --- simp [inv_mul_eq_div, div_self ha] - --- variable (t₀) in --- /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` --- is a global integral curve of `v`. -/ --- lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by --- intro t --- rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), --- ContinuousLinearMap.smulRight_one_one] --- exact hasMFDerivAt_const .. - --- /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the --- manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector --- of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around --- `t₀`.-/ --- theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : --- ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by --- -- express the differentiability of the section `v` in the local charts --- rw [contMDiffAt_iff] at hv --- obtain ⟨_, hv⟩ := hv --- -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart --- obtain ⟨f, hf1, ε1, hε1, hf2⟩ := --- exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ --- (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd --- rw [← Real.ball_eq_Ioo] at hf2 --- -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, --- -- `f t ∈ interior (extChartAt I x₀).target` --- have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt --- rw [continuousAt_def, hf1] at hcont --- have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := --- hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) --- rw [Metric.mem_nhds_iff] at hnhds --- obtain ⟨ε2, hε2, hf3⟩ := hnhds --- simp_rw [subset_def, mem_preimage] at hf3 --- -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve --- refine ⟨(extChartAt I x₀).symm ∘ f, --- Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), --- min ε1 ε2, lt_min hε1 hε2, ?_⟩ --- intros t ht --- -- collect useful terms in convenient forms --- rw [← Real.ball_eq_Ioo] at ht --- have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) --- have h : HasDerivAt f --- ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) --- (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) --- (v ((extChartAt I x₀).symm (f t)))) --- t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) --- rw [← tangentCoordChange_def] at h --- have hf3' := mem_of_mem_of_subset hf3 interior_subset --- have hft1 := mem_preimage.mp <| --- mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source --- have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) --- -- express the derivative of the integral curve in the local chart --- refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), --- HasDerivWithinAt.hasFDerivWithinAt ?_⟩ --- simp only [mfld_simps, hasDerivWithinAt_univ] --- show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) --- (v ((extChartAt I x₀).symm (f t))) t --- -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use --- -- `HasFDerivAt.comp_hasDerivAt` on `h` --- rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) --- (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, --- ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] --- apply HasFDerivAt.comp_hasDerivAt _ _ h --- apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| --- mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, --- subset_trans interior_subset (extChartAt_target_subset_range ..), --- isOpen_interior, hf3⟩ --- nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] --- exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ - --- /-- For any continuously differentiable vector field defined on a manifold without boundary and any --- chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the --- tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open --- interval around `t₀`. -/ --- lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : --- ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := --- exists_isIntegralCurveAt_of_contMDiffAt hv I.isInteriorPoint +/-! ### Translation lemmas -/ + +variable {s : Set ℝ} {t₀ : ℝ} + +lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : + IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by + intros t ht + rw [Function.comp_apply, + ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))] + apply HasMFDerivAt.comp t (hγ (t + dt) ht) + refine ⟨(continuous_add_right _).continuousAt, ?_⟩ + simp only [mfld_simps, hasFDerivWithinAt_univ] + exact HasFDerivAt.add_const (hasFDerivAt_id _) _ + +lemma isIntegralCurveOn_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveOn γ v s ↔ + IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by + refine ⟨fun hγ => hγ.comp_add _, fun hγ => ?_⟩ + have := hγ.comp_add (-dt) + simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] at this + convert this + ext + simp only [Function.comp_apply, neg_add_cancel_right] + +lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by + obtain ⟨ε, hε, h⟩ := hγ + refine ⟨ε, hε, ?_⟩ + convert h.comp_add dt + ext t' + rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] + rfl + +lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ + IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by + refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ + have := hγ.comp_add (-dt) + rw [sub_neg_eq_add, sub_add_cancel] at this + convert this + ext + simp only [Function.comp_apply, neg_add_cancel_right] + +lemma IsIntegralCurve.comp_add {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (dt : ℝ) : + IsIntegralCurve (γ ∘ (· + dt)) v := by + rw [isIntegralCurve_iff_isIntegralCurveOn] at * + exact hγ.comp_add _ + +lemma isIntegralCurve_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurve γ v ↔ + IsIntegralCurve (γ ∘ (· + dt)) v := by + refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ + convert hγ.comp_add (-dt) + ext + simp only [Function.comp_apply, neg_add_cancel_right] + +/-! ### Scale lemmas -/ + +lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} + (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by + obtain ⟨ε, hε, h2⟩ := hγ + refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ + have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by + rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] + exact ht + rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] + refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ + simp only [mfld_simps, hasFDerivWithinAt_univ] + exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ + +lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ + have := hγ.comp_mul_pos (inv_pos_of_pos ha) + rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, + inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this + convert this + ext + simp [inv_mul_eq_div, div_self (ne_of_gt ha)] + +lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : + IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by + obtain ⟨ε, hε, h2⟩ := hγ + refine ⟨ε, hε, fun t ht => ?_⟩ + rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht + rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), + ← ContinuousLinearMap.smulRight_comp] + apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ + simp only [mfld_simps, hasFDerivWithinAt_univ] + exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) + +lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by + refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ + have := hγ.comp_neg + rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this + exact this + +lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} + (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by + rw [ne_iff_lt_or_gt] at ha + cases' ha with ha ha + · apply isIntegralCurveAt_comp_neg.mpr + have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp + rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] + exact hγ.comp_mul_pos (neg_pos_of_neg ha) + · exact hγ.comp_mul_pos ha + +lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by + refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ + have := hγ.comp_mul_ne_zero (inv_ne_zero ha) + rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, + inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this + convert this + ext + simp [inv_mul_eq_div, div_self ha] + +variable (t₀) in +/-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` + is a global integral curve of `v`. -/ +lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by + intro t + rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), + ContinuousLinearMap.smulRight_one_one] + exact hasMFDerivAt_const .. + +variable (t₀) + +/-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the + manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector + of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around + `t₀`.-/ +theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : + ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by + -- express the differentiability of the section `v` in the local charts + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart + obtain ⟨f, hf1, ε1, hε1, hf2⟩ := + exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd + rw [← Real.ball_eq_Ioo] at hf2 + -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, + -- `f t ∈ interior (extChartAt I x₀).target` + have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt + rw [continuousAt_def, hf1] at hcont + have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := + hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) + rw [Metric.mem_nhds_iff] at hnhds + obtain ⟨ε2, hε2, hf3⟩ := hnhds + simp_rw [subset_def, mem_preimage] at hf3 + -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve + refine ⟨(extChartAt I x₀).symm ∘ f, + Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), + min ε1 ε2, lt_min hε1 hε2, ?_⟩ + intros t ht + -- collect useful terms in convenient forms + rw [← Real.ball_eq_Ioo] at ht + have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) + have h : HasDerivAt f + ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) + (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) + (v ((extChartAt I x₀).symm (f t)))) + t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + rw [← tangentCoordChange_def] at h + have hf3' := mem_of_mem_of_subset hf3 interior_subset + have hft1 := mem_preimage.mp <| + mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source + have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) + -- express the derivative of the integral curve in the local chart + refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), + HasDerivWithinAt.hasFDerivWithinAt ?_⟩ + simp only [mfld_simps, hasDerivWithinAt_univ] + show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) + (v ((extChartAt I x₀).symm (f t))) t + -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use + -- `HasFDerivAt.comp_hasDerivAt` on `h` + rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) + (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, + ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] + apply HasFDerivAt.comp_hasDerivAt _ _ h + apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| + mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, + subset_trans interior_subset (extChartAt_target_subset_range ..), + isOpen_interior, hf3⟩ + nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] + exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ + +/-- For any continuously differentiable vector field defined on a manifold without boundary and any + chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the + tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open + interval around `t₀`. -/ +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : + ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := + exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint From 3b7cca035bc508a34dc47d9c02f494566c7c9b7f Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 15:38:08 -0800 Subject: [PATCH 092/203] refactored `IsIntegralCurveXX` --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 85 +++++++++----------- 1 file changed, 40 insertions(+), 45 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index f10db74be7594..5a6f5a9f3671e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -52,6 +52,8 @@ variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) + {s : Set ℝ} {t₀ : ℝ} + /-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` @@ -106,8 +108,6 @@ lemma IsIntegralCurveAt.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → Tang /-! ### Translation lemmas -/ -variable {s : Set ℝ} {t₀ : ℝ} - lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by intros t ht @@ -159,66 +159,61 @@ lemma isIntegralCurve_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurve γ /-! ### Scale lemmas -/ -lemma IsIntegralCurveAt.comp_mul_pos {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} - (ha : 0 < a) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - obtain ⟨ε, hε, h2⟩ := hγ - refine ⟨ε / a, div_pos hε ha, fun t ht => ?_⟩ - have ht : t * a ∈ Ioo (t₀ - ε) (t₀ + ε) := by - rw [mem_Ioo, ← div_lt_iff ha, ← lt_div_iff ha, sub_div, add_div] - exact ht +lemma IsIntegralCurveOn.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : + IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by + intros t ht rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] - refine HasMFDerivAt.comp t (h2 (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ + refine HasMFDerivAt.comp t (hγ (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩ simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ -lemma isIntegralCurvAt_comp_mul_pos {γ : ℝ → M} {a : ℝ} (ha : 0 < a) : - IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_pos hγ ha, fun hγ ↦ ?_⟩ - have := hγ.comp_mul_pos (inv_pos_of_pos ha) - rw [smul_smul, inv_mul_eq_div, div_self (ne_of_gt ha), one_smul, ← div_mul_eq_div_div_swap, - inv_mul_eq_div, div_self (ne_of_gt ha), div_one, Function.comp.assoc] at this +lemma isIntegralCurvOn_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by + refine ⟨fun hγ => hγ.comp_mul a, fun hγ ↦ ?_⟩ + have := hγ.comp_mul a⁻¹ + simp_rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, mem_setOf_eq, mul_assoc, + inv_mul_eq_div, div_self ha, mul_one, setOf_mem_eq] at this convert this - ext - simp [inv_mul_eq_div, div_self (ne_of_gt ha)] - -lemma IsIntegralCurveAt.comp_neg {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : - IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by - obtain ⟨ε, hε, h2⟩ := hγ - refine ⟨ε, hε, fun t ht => ?_⟩ - rw [← neg_add', neg_add_eq_sub, ← neg_sub, ← neg_mem_Ioo_iff] at ht - rw [Function.comp_apply, Pi.neg_apply, ← neg_one_smul ℝ (v (γ (-t))), - ← ContinuousLinearMap.smulRight_comp] - apply (h2 (-t) ht).comp t ⟨continuousAt_neg, ?_⟩ - simp only [mfld_simps, hasFDerivWithinAt_univ] - exact HasDerivAt.hasFDerivAt (hasDerivAt_neg _) - -lemma isIntegralCurveAt_comp_neg {γ : ℝ → M} : - IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ Neg.neg) (-v) (-t₀) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_neg hγ, fun hγ ↦ ?_⟩ - have := hγ.comp_neg - rw [Function.comp.assoc, neg_comp_neg, neg_neg, neg_neg] at this - exact this + ext t + rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - rw [ne_iff_lt_or_gt] at ha - cases' ha with ha ha - · apply isIntegralCurveAt_comp_neg.mpr - have : (· * a) ∘ Neg.neg = fun t ↦ t * -a := by ext; simp - rw [Function.comp.assoc, this, ← neg_smul, ← div_neg] - exact hγ.comp_mul_pos (neg_pos_of_neg ha) - · exact hγ.comp_mul_pos ha + obtain ⟨ε, hε, h⟩ := hγ + refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ + convert h.comp_mul a + ext t + rw [Ioo, Ioo, mem_setOf_eq, mem_setOf_eq, mem_setOf_eq] + by_cases ha' : 0 < a + · rw [abs_eq_self.mpr (le_of_lt ha'), ← sub_div, ← add_div, div_lt_iff ha', lt_div_iff ha'] + · rw [abs_eq_neg_self.mpr (not_lt.mp ha'), div_neg, sub_neg_eq_add, ← sub_eq_add_neg, ← sub_div, + ← add_div, div_lt_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), + lt_div_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), and_comm] lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - refine ⟨fun hγ => IsIntegralCurveAt.comp_mul_ne_zero hγ ha, fun hγ ↦ ?_⟩ + refine ⟨fun hγ => hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this convert this - ext + ext t simp [inv_mul_eq_div, div_self ha] +lemma IsIntegralCurve.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (a : ℝ) : + IsIntegralCurve (γ ∘ (· * a)) (a • v) := by + rw [isIntegralCurve_iff_isIntegralCurveOn] at * + exact hγ.comp_mul _ + +lemma isIntegralCurve_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by + refine ⟨fun hγ => hγ.comp_mul _, fun hγ => ?_⟩ + have := hγ.comp_mul a⁻¹ + rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul] at this + convert this + ext t + rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + variable (t₀) in /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ From 082574b8fa5b63e98fc02503d89ca4621131fd9b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 30 Nov 2023 15:39:43 -0800 Subject: [PATCH 093/203] unused arg --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5a6f5a9f3671e..5b50e128262e1 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -214,7 +214,6 @@ lemma isIntegralCurve_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) ext t rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] -variable (t₀) in /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by From ac903450ad97e1f84caf8af09824910210edf164 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 1 Dec 2023 00:48:49 -0800 Subject: [PATCH 094/203] statement --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5b50e128262e1..913b1ca0bb36c 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -291,3 +291,15 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint + +example {γ γ' : ℝ → M} (h : γ t₀ = γ' t₀) (ht : I.IsInteriorPoint (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) : + ∃ t₁ > t₀, ∀ t ∈ Icc t₀ t₁, γ t = γ' t := by + set v' : E → E := fun x => + tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) + (v ((extChartAt I (γ t₀)).symm x)) with hv' + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht)).snd From 19a2e818ec5662496cdc1446012149dbf14bf639 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 1 Dec 2023 01:33:09 -0800 Subject: [PATCH 095/203] gronwall LipschitzOnWith --- Mathlib/Analysis/ODE/Gronwall.lean | 26 ++++++++++++++------------ 1 file changed, 14 insertions(+), 12 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 26d120f4918b3..be341b9402ffd 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -150,8 +150,8 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ} - (hv : ∀ t, ∀ᵉ (x ∈ s t) (y ∈ s t), dist (v t x) (v t y) ≤ K * dist x y) +theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g f' g' : ℝ → E} {a b : ℝ} {εf εg δ : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -166,8 +166,11 @@ theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : intro t ht have := dist_triangle4_right (f' t) (g' t) (v t (f t)) (v t (g t)) rw [dist_eq_norm] at this - refine' this.trans ((add_le_add (add_le_add (f_bound t ht) (g_bound t ht)) - (hv t (f t) (hfs t ht) (g t) (hgs t ht))).trans _) + have hv := hv t (x := f t) (hfs t ht) (y := g t) (hgs t ht) + rw [edist_dist, edist_dist, ← ENNReal.ofReal_coe_nnreal, + ← ENNReal.ofReal_mul (NNReal.coe_nonneg _), + ENNReal.ofReal_le_ofReal_iff (mul_nonneg (NNReal.coe_nonneg _) dist_nonneg)] at hv + refine this.trans ((add_le_add (add_le_add (f_bound t ht) (g_bound t ht)) hv).trans ?_) rw [dist_eq_norm, add_comm] set_option linter.uppercaseLean3 false in #align dist_le_of_approx_trajectories_ODE_of_mem_set dist_le_of_approx_trajectories_ODE_of_mem_set @@ -185,7 +188,7 @@ theorem dist_le_of_approx_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ gronwallBound δ K (εf + εg) (t - a) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_approx_trajectories_ODE_of_mem_set (fun t x _ y _ => (hv t).dist_le_mul x y) hf hf' + dist_le_of_approx_trajectories_ODE_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' f_bound hfs hg hg' g_bound (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align dist_le_of_approx_trajectories_ODE dist_le_of_approx_trajectories_ODE @@ -196,8 +199,8 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ} - (hv : ∀ t, ∀ᵉ (x ∈ s t) (y ∈ s t), dist (v t x) (v t y) ≤ K * dist x y) +theorem dist_le_of_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} {δ : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) @@ -223,7 +226,7 @@ theorem dist_le_of_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} (hv : (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ δ * exp (K * (t - a)) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_trajectories_ODE_of_mem_set (fun t x _ y _ => (hv t).dist_le_mul x y) hf hf' hfs hg + dist_le_of_trajectories_ODE_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align dist_le_of_trajectories_ODE dist_le_of_trajectories_ODE @@ -231,8 +234,8 @@ set_option linter.uppercaseLean3 false in /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, and we consider only solutions included in `s`. -/ -theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ} - (hv : ∀ t, ∀ᵉ (x ∈ s t) (y ∈ s t), dist (v t x) (v t y) ≤ K * dist x y) +theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) @@ -250,8 +253,7 @@ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, Lip (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : ∀ t ∈ Icc a b, f t = g t := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - ODE_solution_unique_of_mem_set (fun t x _ y _ => (hv t).dist_le_mul x y) hf hf' hfs hg hg' + ODE_solution_unique_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align ODE_solution_unique ODE_solution_unique - From b08e2432a1e4d587549745f2c448093ca6665c69 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 1 Dec 2023 01:58:11 -0800 Subject: [PATCH 096/203] partial --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 913b1ca0bb36c..82467c61f8574 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Winston Yin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Winston Yin -/ +import Mathlib.Analysis.ODE.Gronwall import Mathlib.Analysis.ODE.PicardLindelof import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.MFDeriv @@ -296,6 +297,18 @@ example {γ γ' : ℝ → M} (h : γ t₀ = γ' t₀) (ht : I.IsInteriorPoint ( (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) : ∃ t₁ > t₀, ∀ t ∈ Icc t₀ t₁, γ t = γ' t := by + obtain ⟨ε, hε, hγ⟩ := hγ + obtain ⟨ε', hε', hγ'⟩ := hγ' + set t₁ := t₀ + (min ε ε') / 2 with ht₁ + have hf : ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Icc t₀ t₁) := sorry + have hf' : ContinuousOn ((extChartAt I (γ' t₀)) ∘ γ) (Icc t₀ t₁) := sorry + + + + -- need to shrink ε further to fit inside s + + + set v' : E → E := fun x => tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) (v ((extChartAt I (γ t₀)).symm x)) with hv' @@ -303,3 +316,7 @@ example {γ γ' : ℝ → M} (h : γ t₀ = γ' t₀) (ht : I.IsInteriorPoint ( obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht)).snd + have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun t => hlip + + + -- have := ODE_solution_unique_of_mem_set hlip From 08d1e12172d3e1083f7a5a2f6b3fc6da1243ac38 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 4 Dec 2023 02:41:31 -0800 Subject: [PATCH 097/203] some corollaries in Gronwall --- Mathlib/Analysis/ODE/Gronwall.lean | 38 +++++++++++++++++++- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 +-- 2 files changed, 39 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index be341b9402ffd..2135d0dd9dd9b 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -239,12 +239,48 @@ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : ∀ t ∈ Icc a b, f t = g t := fun t ht ↦ by + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in #align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set +-- Ico rather than Icc +theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ + ODE_solution_unique_of_mem_set hv + (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) + (fun _ ht' ↦ hfs _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) + (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) + (fun _ ht' ↦ hgs _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) ha ⟨ht.1, le_refl _⟩ + +-- starting point t₀ can be inside [a, b] +example {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) + (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) + (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + have hu : Ioo a b ⊆ Ioc a t₀ ∪ Ico t₀ b := sorry + apply EqOn.mono hu + apply EqOn.union + · sorry -- need to reverse time + · exact ODE_solution_unique_of_mem_set_Ioo hv + (hf.mono (Ico_subset_Ioo_left ht.1)) + (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ <| + mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) + (fun _ ht' => hfs _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) + (hg.mono (Ico_subset_Ioo_left ht.1)) + (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hg' _ <| + mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) + (fun _ ht' => hgs _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) ha + +-- another Ioo version t₀ ∈ (a, b) + /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that RHS is Lipschitz continuous in `x`. -/ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 82467c61f8574..30ef573f717a7 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -293,9 +293,9 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint -example {γ γ' : ℝ → M} (h : γ t₀ = γ' t₀) (ht : I.IsInteriorPoint (γ t₀)) +example {γ γ' : ℝ → M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) - (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) : + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ t₁ > t₀, ∀ t ∈ Icc t₀ t₁, γ t = γ' t := by obtain ⟨ε, hε, hγ⟩ := hγ obtain ⟨ε', hε', hγ'⟩ := hγ' From f80b3376be262e0a7de9f577cf19181f40446f8e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 5 Dec 2023 00:42:20 -0800 Subject: [PATCH 098/203] Ioo version --- Mathlib/Analysis/ODE/Gronwall.lean | 38 ++++++++++++++++++++++++------ 1 file changed, 31 insertions(+), 7 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 2135d0dd9dd9b..1aa2a6b9a5796 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -246,7 +246,7 @@ set_option linter.uppercaseLean3 false in #align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set -- Ico rather than Icc -theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -258,18 +258,42 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) (fun _ ht' ↦ hgs _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) ha ⟨ht.1, le_refl _⟩ --- starting point t₀ can be inside [a, b] -example {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +-- starting point t₀ can be inside (a, b) +theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by - have hu : Ioo a b ⊆ Ioc a t₀ ∪ Ico t₀ b := sorry - apply EqOn.mono hu + apply EqOn.mono (Ioo_subset_Ioc_union_Ico (b := t₀)) apply EqOn.union - · sorry -- need to reverse time - · exact ODE_solution_unique_of_mem_set_Ioo hv + · have hv' : ∀ t, LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := fun t ↦ by + rw [← one_mul K] + apply LipschitzWith.comp_lipschitzOnWith (LipschitzWith.neg LipschitzWith.id) + exact hv _ + have : EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) := by + apply ODE_solution_unique_of_mem_set_Ico hv' + (hf.comp continuousOn_neg + fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) _ + (fun _ ht' ↦ hfs _ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) + (hg.comp continuousOn_neg + fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) _ + (fun _ ht' ↦ hgs _ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) + (by simp [ha]) + · intros t' ht' + apply HasDerivAt.hasDerivWithinAt + have := hf' (-t') ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ + convert HasFDerivAt.comp_hasDerivAt t' this (hasDerivAt_neg t') + simp + · intros t' ht' + apply HasDerivAt.hasDerivWithinAt + have := hg' (-t') ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ + convert HasFDerivAt.comp_hasDerivAt t' this (hasDerivAt_neg t') + simp + rw [eqOn_comp_right_iff] at this + convert this + simp + · exact ODE_solution_unique_of_mem_set_Ico hv (hf.mono (Ico_subset_Ioo_left ht.1)) (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) From 49c5acb939f17432443fabaa75df548cbb2ab967 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 6 Dec 2023 00:48:03 -0800 Subject: [PATCH 099/203] stuck --- Mathlib/Analysis/ODE/Gronwall.lean | 2 - Mathlib/Geometry/Manifold/IntegralCurve.lean | 150 ++++++++++++++++-- .../Geometry/Manifold/InteriorBoundary.lean | 4 + 3 files changed, 137 insertions(+), 19 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 1aa2a6b9a5796..b0117c66434be 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -303,8 +303,6 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) (fun _ ht' => hgs _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) ha --- another Ioo version t₀ ∈ (a, b) - /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that RHS is Lipschitz continuous in `x`. -/ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 30ef573f717a7..7477cbaa229b6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -42,6 +42,9 @@ integral curve, vector field, local existence -/ open scoped Manifold +scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x +scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) + open Set variable @@ -223,6 +226,23 @@ lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. +lemma IsIntegralCurveOn.continuousAt {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ∈ s) : + ContinuousAt γ t₀ := (hγ t₀ ht).1 + +lemma IsIntegralCurveOn.continuousOn {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) : + ContinuousOn γ s := fun t ht => (hγ t ht).1.continuousWithinAt + +lemma IsIntegralCurveAt.continuousAt {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : + ContinuousAt γ t₀ := by + obtain ⟨ε, hε, hγ⟩ := hγ + apply hγ.continuousAt + rw [← Real.ball_eq_Ioo] + exact Metric.mem_ball_self hε + +lemma IsIntegralCurve.continuous {γ : ℝ → M} (hγ : IsIntegralCurve γ v) : + Continuous γ := continuous_iff_continuousAt.mpr + fun _ => (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) + variable (t₀) /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the @@ -293,22 +313,11 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint -example {γ γ' : ℝ → M} +example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : - ∃ t₁ > t₀, ∀ t ∈ Icc t₀ t₁, γ t = γ' t := by - obtain ⟨ε, hε, hγ⟩ := hγ - obtain ⟨ε', hε', hγ'⟩ := hγ' - set t₁ := t₀ + (min ε ε') / 2 with ht₁ - have hf : ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Icc t₀ t₁) := sorry - have hf' : ContinuousOn ((extChartAt I (γ' t₀)) ∘ γ) (Icc t₀ t₁) := sorry - - - - -- need to shrink ε further to fit inside s - - - + ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by + -- extract set `s` on which `v` is Lipschitz set v' : E → E := fun x => tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) (v ((extChartAt I (γ t₀)).symm x)) with hv' @@ -317,6 +326,113 @@ example {γ γ' : ℝ → M} obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun t => hlip - - - -- have := ODE_solution_unique_of_mem_set hlip + -- extract `εs` so that `(extChartAt I (γ t₀)) (γ t) ∈ s` for all `t ∈ Ioo (t₀ - εs) (t₀ + εs)` + have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := + ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt + rw [continuousAt_def] at hcont + have hnhds := hcont _ hs + rw [Metric.mem_nhds_iff] at hnhds + obtain ⟨εs, hεs, hmem⟩ := hnhds + simp_rw [subset_def, mem_preimage] at hmem + -- `εs'` for `γ'` + have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := + ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt + rw [continuousAt_def] at hcont' + have hnhds' := hcont' _ (h ▸ hs) + rw [Metric.mem_nhds_iff] at hnhds' + obtain ⟨εs', hεs', hmem'⟩ := hnhds' + simp_rw [subset_def, mem_preimage] at hmem' + -- extract `εe` so `γ t` stays within the interior of the chart around `γ t₀` + have := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) + rw [Metric.mem_nhds_iff] at this + obtain ⟨εe, hεe, hsrc⟩ := this + simp_rw [subset_def, mem_preimage] at hsrc + have := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) + rw [Metric.mem_nhds_iff] at this + obtain ⟨εe', hεe', hsrc'⟩ := this + simp_rw [subset_def, mem_preimage] at hsrc' + -- extract `εγ` from local existence of integral curve + obtain ⟨εγ, hεγ, hγ⟩ := hγ + obtain ⟨εγ', hεγ', hγ'⟩ := hγ' + let ε := min (min εe εe') <| min (min εs εs') (min εγ εγ') + have hf : ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) := by + apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) + · apply hγ.continuousOn.mono + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + · apply MapsTo.mono_left hsrc + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + have hf' : ContinuousOn ((extChartAt I (γ' t₀)) ∘ γ') (Ioo (t₀ - ε) (t₀ + ε)) := by + apply ContinuousOn.comp (continuousOn_extChartAt I (γ' t₀)) + · apply hγ'.continuousOn.mono + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + · apply MapsTo.mono_left hsrc' + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + have hε : 0 < ε := lt_min (lt_min hεe hεe') (lt_min (lt_min hεs hεs') (lt_min hεγ hεγ')) + refine ⟨ε, hε, ?_⟩ + have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') + (Ioo (t₀ - ε) (t₀ + ε)) := by + apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) _ hf _ _ hf' + · intros t ht + have ht' : t ∈ Ioo (t₀ - εγ') (t₀ + εγ') := by + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + have := (hγ' t ht').2 + rw [writtenInExtChartAt, modelWithCornersSelf_coe, range_id, extChartAt_self_apply, + modelWithCornersSelf_coe, id.def, extChartAt_coe_symm, modelWithCornersSelf_coe_symm, + Function.comp.right_id, chartAt_self_eq, LocalHomeomorph.refl_symm, + LocalHomeomorph.refl_apply, Function.comp.right_id, hasFDerivWithinAt_univ, + ← hasDerivAt_iff_hasFDerivAt] at this + have hrw : HasDerivAt ((extChartAt I (γ' t₀)) ∘ γ') + (tangentCoordChange I (γ' t) (γ' t₀) (γ' t) (v (γ' t))) t := by + -- maybe have to use `HasFDerivAt.of_local_left_inverse` here? + sorry + have hsub : (fun x ↦ v') t ((↑(extChartAt I (γ' t₀)) ∘ γ') t) = + (tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t)) := by + dsimp only + rw [h, Function.comp_apply, LocalEquiv.left_inv] + apply hsrc' + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + rw [hsub] + exact hrw + · intros t ht + apply hmem' + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + · simp [h] + · rw [← Real.ball_eq_Ioo] + exact Metric.mem_ball_self hε + · sorry + · intros t ht + apply hmem + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) + · intros t ht + rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv] + apply hsrc + rw [← Real.ball_eq_Ioo] at ht + apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) + simp + · intros t ht + rw [Function.comp_apply, Function.comp_apply, h, LocalEquiv.left_inv] + apply hsrc' + rw [← Real.ball_eq_Ioo] at ht + apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) + simp diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 634f4a948bcce..45a9ed5e2a11a 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -38,6 +38,10 @@ manifold, interior, boundary open Set +open scoped Manifold +scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x +scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) + -- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] From e52f2ff75138e8ec162d1e54eeb3f62fdf18c859 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 6 Dec 2023 02:32:27 -0800 Subject: [PATCH 100/203] can't believe it compiles lol --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 98 ++++++++++++++++++-- 1 file changed, 89 insertions(+), 9 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7477cbaa229b6..f42ea2384598e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -386,16 +386,43 @@ example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] apply Metric.ball_subset_ball simp - have := (hγ' t ht').2 - rw [writtenInExtChartAt, modelWithCornersSelf_coe, range_id, extChartAt_self_apply, - modelWithCornersSelf_coe, id.def, extChartAt_coe_symm, modelWithCornersSelf_coe_symm, - Function.comp.right_id, chartAt_self_eq, LocalHomeomorph.refl_symm, - LocalHomeomorph.refl_apply, Function.comp.right_id, hasFDerivWithinAt_univ, - ← hasDerivAt_iff_hasFDerivAt] at this have hrw : HasDerivAt ((extChartAt I (γ' t₀)) ∘ γ') (tangentCoordChange I (γ' t) (γ' t₀) (γ' t) (v (γ' t))) t := by - -- maybe have to use `HasFDerivAt.of_local_left_inverse` here? - sorry + -- turn `HasDerivAt` into comp of `HasMFDerivAt` + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` + have := (hasMFDerivAt_extChartAt I (x := γ' t₀) (y := γ' t) (by + rw [← extChartAt_source I] + apply hsrc' + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + )).comp t (hγ' t ht') + have h2 : ContinuousLinearMap.comp + (mfderiv I I (↑(chartAt H (γ' t₀))) (γ' t)) + (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ' t))) = + ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) + ((tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t))) := by + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, + ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] + congr + rw [tangentCoordChange_def, mfderiv] + have hdiff : MDifferentiableAt I I (↑(chartAt H (γ' t₀))) (γ' t) := by + apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) + rw [← extChartAt_source I] + apply hsrc' + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + simp only [hdiff, if_true] + rfl + rw [← h2] + exact this have hsub : (fun x ↦ v') t ((↑(extChartAt I (γ' t₀)) ∘ γ') t) = (tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t)) := by dsimp only @@ -416,7 +443,60 @@ example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) · simp [h] · rw [← Real.ball_eq_Ioo] exact Metric.mem_ball_self hε - · sorry + · intros t ht + have ht' : t ∈ Ioo (t₀ - εγ) (t₀ + εγ) := by + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + have hrw : HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by + -- turn `HasDerivAt` into comp of `HasMFDerivAt` + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` + have := (hasMFDerivAt_extChartAt I (x := γ t₀) (y := γ t) (by + rw [← extChartAt_source I] + apply hsrc + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + )).comp t (hγ t ht') + have h2 : ContinuousLinearMap.comp + (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) + (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = + ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, + ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] + congr + rw [tangentCoordChange_def, mfderiv] + have hdiff : MDifferentiableAt I I (↑(chartAt H (γ t₀))) (γ t) := by + apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) + rw [← extChartAt_source I] + apply hsrc + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + simp only [hdiff, if_true] + rfl + rw [← h2] + exact this + have hsub : (fun x ↦ v') t ((↑(extChartAt I (γ t₀)) ∘ γ) t) = + (tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t)) := by + dsimp only + rw [Function.comp_apply, LocalEquiv.left_inv] + apply hsrc + apply mem_of_mem_of_subset ht + rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + rw [hsub] + exact hrw · intros t ht apply hmem apply mem_of_mem_of_subset ht From 5e4df2a7f4909334859c46dab5290de6bafbfec8 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 6 Dec 2023 17:10:51 -0800 Subject: [PATCH 101/203] draft --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index f42ea2384598e..442267efe272f 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -313,7 +313,7 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint -example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) +theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by @@ -486,7 +486,7 @@ example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) rfl rw [← h2] exact this - have hsub : (fun x ↦ v') t ((↑(extChartAt I (γ t₀)) ∘ γ) t) = + have hsub : (fun _ ↦ v') t ((↑(extChartAt I (γ t₀)) ∘ γ) t) = (tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t)) := by dsimp only rw [Function.comp_apply, LocalEquiv.left_inv] @@ -516,3 +516,12 @@ example {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) rw [← Real.ball_eq_Ioo] at ht apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) simp + +example {γ γ' : ℝ → M} {a b : ℝ} (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) + (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by sorry + /- + strategy: + * Lee P.213, just need to translate "S is closed in J" to type theory language + -/ From 33d9a79c1b9c7478af8f0fe6c7ce660f41818947 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 7 Dec 2023 14:51:40 -0800 Subject: [PATCH 102/203] uniqueness sorry-free! --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 94 +++++++++++++++++++- 1 file changed, 92 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 442267efe272f..a9338f38f38df 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -517,11 +517,101 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInt apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) simp -example {γ γ' : ℝ → M} {a b : ℝ} (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) +theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} + {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) - (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by sorry + (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by /- strategy: * Lee P.213, just need to translate "S is closed in J" to type theory language -/ + set s := {t | γ t = γ' t} ∩ Ioo a b with hs + have hsub : Ioo a b ⊆ s := by + apply isPreconnected_Ioo.subset_of_closure_inter_subset + · rw [isOpen_iff_mem_nhds] + intro t₁ ht₁ + rw [mem_nhds_iff] + have : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := by + apply isIntegralCurveAt_eqOn_of_contMDiffAt (I := I) + · apply hip + exact ht₁.2 + · exact hv.contMDiffAt + · apply hγ.isIntegralCurveAt + exact Ioo_mem_nhds ht₁.2.1 ht₁.2.2 + · apply hγ'.isIntegralCurveAt + exact Ioo_mem_nhds ht₁.2.1 ht₁.2.2 + · exact ht₁.1 + obtain ⟨ε, hε, heqon⟩ := this + use Ioo (max a (t₁ - ε)) (min b (t₁ + ε)) + refine ⟨?_, isOpen_Ioo, ?_⟩ + · apply subset_inter + · intros t ht + apply @heqon t + apply mem_of_mem_of_subset ht + apply Ioo_subset_Ioo (by simp) (by simp) + · apply Ioo_subset_Ioo (by simp) (by simp) + · rw [mem_Ioo] + constructor + · apply max_lt ht₁.2.1 + simp [hε] + · apply lt_min ht₁.2.2 + simp [hε] + · use t₀ + rw [mem_inter_iff] + use ht₀ + use h + · -- is this really the most convenient way to pass to subtype topology? + rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, + image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] + intros t ht + rw [mem_preimage, ← closure_subtype] at ht + revert ht t + apply IsClosed.closure_subset + apply isClosed_eq + · rw [continuous_iff_continuousAt] + rintro ⟨t, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ.continuousAt ht + · rw [continuous_iff_continuousAt] + rintro ⟨t, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ'.continuousAt ht + intros t ht + exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 + +theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} + (hip : ∀ t, I.IsInteriorPoint (γ t)) + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by + ext t + have : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by + use 2 * |t - t₀| + 1 + constructor + · apply add_pos_of_nonneg_of_pos _ zero_lt_one + simp + · rw [mem_Ioo] + by_cases ht : t - t₀ < 0 + · rw [abs_of_neg ht] + constructor <;> linarith + · rw [not_lt] at ht + rw [abs_of_nonneg ht] + constructor <;> linarith + obtain ⟨T, hT, ht⟩ := this + apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff + (I := I) (a := t₀ - T) (b := t₀ + T) (t₀ := t₀) + · rw [← Real.ball_eq_Ioo] + exact Metric.mem_ball_self hT + · intros t _ + exact hip t + · exact hv + · rw [isIntegralCurve_iff_isIntegralCurveOn] at hγ + exact IsIntegralCurveOn.mono hγ (subset_univ _) + · rw [isIntegralCurve_iff_isIntegralCurveOn] at hγ' + exact IsIntegralCurveOn.mono hγ' (subset_univ _) + · exact h + · exact ht From f22da927679a6b43e3d86d000c88eca20a8add58 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 7 Dec 2023 18:37:50 -0800 Subject: [PATCH 103/203] docstring --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 39 ++++++++++++++++---- 1 file changed, 31 insertions(+), 8 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a9338f38f38df..5d1c4ab749844 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -11,12 +11,14 @@ import Mathlib.Geometry.Manifold.MFDeriv /-! # Integral curves of vector fields on a manifold -For any continuously differentiable vector field on a manifold `M` and any chosen interior point -`x₀ : M`, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of -`γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. +Let `M` be a manifold and `v : (x : M) → TangentSpace I x` be a vector field on `M`. An integral +curve of `v` is a function `γ : ℝ → M` such that the derivative of `γ` at `t` equals `v (γ t)`. The +integral curve may only be defined for all `t` within some subset of `ℝ`. -As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold -without boundary. +Assume `v` is continuously differentiable. The existence theorem for solutions to ODEs implies that +a unique local integral curve exists for any continuously differentiable vector field `v`. The +uniqueness theorem for solutions to ODEs implies that integral curves of `v` are unique. These are +the main results of this file. ## Main definition @@ -31,14 +33,25 @@ For `IsIntegralCurveOn γ v s` and `IsIntegralCurveAt γ v t₀`, even though ` time, its value outside of the set `s` or a small interval around `t₀` is irrelevant and considered junk. +## Implementation notes + +For the existence and uniqueness theorems, we assume that the image of the integral curve lies in +the interior of the manifold. The case where the integral curve may lie on the boundary of the +manifold requires special treatment, and we leave it as a to-do. + +The uniqueness theorem requires the manifold to be Hausdorff (T2), so that the set on which two +continuous functions agree is closed. + +We state simpler versions of the theorem for manifolds without boundary as corollaries. + ## To-do -- Prove `comp_add`, `comp_smul` , etc. lemmas for `IsIntegralCurveOn`, and then derive versions for -`IsIntegralCurveAt` and `IsIntegralCurve` as corollaries. +- The case where the integral curve may venture to the boundary of the manifold. See Theorem 9.34, + J. M. Lee. May require submanifolds. ## Tags -integral curve, vector field, local existence +integral curve, vector field, local existence, uniqueness -/ open scoped Manifold @@ -313,6 +326,10 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint +/-- Local integral curves are unique. + + If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` + at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : @@ -517,6 +534,11 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInt apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) simp +/-- Integral curves are unique on open intervals. + + If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` + on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` + agree on `Ioo a b`. -/ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) @@ -583,6 +605,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] intros t ht exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 +/-- Global integral curves are unique. -/ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} (hip : ∀ t, I.IsInteriorPoint (γ t)) From a9d4cae6bd97d34f9fc2c615ace53a463e261752 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 01:06:54 -0800 Subject: [PATCH 104/203] space --- Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 5d15ee7be9fb3..568b8eba9bcfc 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -139,14 +139,14 @@ lemma tangentCoordChange_def {x y z : M} : tangentCoordChange I x y z = lemma tangentCoordChange_self {x z : M} {v : E} (h : z ∈ (extChartAt I x).source) : tangentCoordChange I x x z v = v := by apply (tangentBundleCore I M).coordChange_self - rw [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + rw [tangentBundleCore_baseSet, coe_achart, ← extChartAt_source I] exact h lemma tangentCoordChange_comp {w x y z : M} {v : E} (h : z ∈ (extChartAt I w).source ∩ (extChartAt I x).source ∩ (extChartAt I y).source) : tangentCoordChange I x y z (tangentCoordChange I w x z v) = tangentCoordChange I w y z v := by apply (tangentBundleCore I M).coordChange_comp - simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + simp only [tangentBundleCore_baseSet, coe_achart, ← extChartAt_source I] exact h lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} @@ -161,7 +161,7 @@ lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} lemma continuousOn_tangentCoordChange (x y : M) : ContinuousOn (tangentCoordChange I x y) ((extChartAt I x).source ∩ (extChartAt I y).source) := by convert (tangentBundleCore I M).continuousOn_coordChange (achart H x) (achart H y) <;> - simp only [tangentBundleCore_baseSet, coe_achart, ←extChartAt_source I] + simp only [tangentBundleCore_baseSet, coe_achart, ← extChartAt_source I] end tangentCoordChange From 860f52f119f9b1ee80715ac882bcae67694e2b1e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 02:57:13 -0800 Subject: [PATCH 105/203] shorten --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 39 +++++++------------- 1 file changed, 14 insertions(+), 25 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5d1c4ab749844..963274bddb348 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -326,6 +326,8 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint +variable (I) + /-- Local integral curves are unique. If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` @@ -342,7 +344,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInt obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht)).snd - have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun t => hlip + have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip -- extract `εs` so that `(extChartAt I (γ t₀)) (γ t) ∈ s` for all `t ∈ Ioo (t₀ - εs) (t₀ + εs)` have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt @@ -613,28 +615,15 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t have : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by - use 2 * |t - t₀| + 1 - constructor - · apply add_pos_of_nonneg_of_pos _ zero_lt_one - simp - · rw [mem_Ioo] - by_cases ht : t - t₀ < 0 - · rw [abs_of_neg ht] - constructor <;> linarith - · rw [not_lt] at ht - rw [abs_of_nonneg ht] - constructor <;> linarith + refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ + rw [mem_Ioo] + by_cases ht : t - t₀ < 0 + · rw [abs_of_neg ht] + constructor <;> linarith + · rw [abs_of_nonneg (not_lt.mp ht)] + constructor <;> linarith obtain ⟨T, hT, ht⟩ := this - apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff - (I := I) (a := t₀ - T) (b := t₀ + T) (t₀ := t₀) - · rw [← Real.ball_eq_Ioo] - exact Metric.mem_ball_self hT - · intros t _ - exact hip t - · exact hv - · rw [isIntegralCurve_iff_isIntegralCurveOn] at hγ - exact IsIntegralCurveOn.mono hγ (subset_univ _) - · rw [isIntegralCurve_iff_isIntegralCurveOn] at hγ' - exact IsIntegralCurveOn.mono hγ' (subset_univ _) - · exact h - · exact ht + exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ + (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv + (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) + (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht From d6d70f5c22d7fd178b52aa02cd754772cb32e62d Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 03:33:10 -0800 Subject: [PATCH 106/203] shorten --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 57 ++++++++------------ 1 file changed, 22 insertions(+), 35 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 963274bddb348..7136726df2a4b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -553,39 +553,8 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] -/ set s := {t | γ t = γ' t} ∩ Ioo a b with hs have hsub : Ioo a b ⊆ s := by - apply isPreconnected_Ioo.subset_of_closure_inter_subset - · rw [isOpen_iff_mem_nhds] - intro t₁ ht₁ - rw [mem_nhds_iff] - have : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := by - apply isIntegralCurveAt_eqOn_of_contMDiffAt (I := I) - · apply hip - exact ht₁.2 - · exact hv.contMDiffAt - · apply hγ.isIntegralCurveAt - exact Ioo_mem_nhds ht₁.2.1 ht₁.2.2 - · apply hγ'.isIntegralCurveAt - exact Ioo_mem_nhds ht₁.2.1 ht₁.2.2 - · exact ht₁.1 - obtain ⟨ε, hε, heqon⟩ := this - use Ioo (max a (t₁ - ε)) (min b (t₁ + ε)) - refine ⟨?_, isOpen_Ioo, ?_⟩ - · apply subset_inter - · intros t ht - apply @heqon t - apply mem_of_mem_of_subset ht - apply Ioo_subset_Ioo (by simp) (by simp) - · apply Ioo_subset_Ioo (by simp) (by simp) - · rw [mem_Ioo] - constructor - · apply max_lt ht₁.2.1 - simp [hε] - · apply lt_min ht₁.2.2 - simp [hε] - · use t₀ - rw [mem_inter_iff] - use ht₀ - use h + apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ + ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ · -- is this really the most convenient way to pass to subtype topology? rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] @@ -604,6 +573,25 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] apply ContinuousAt.comp _ continuousAt_subtype_val rw [Subtype.coe_mk] exact hγ'.continuousAt ht + · rw [isOpen_iff_mem_nhds] + intro t₁ ht₁ + rw [mem_nhds_iff] + obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := + isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt + (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + ht₁.1 + refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), + subset_inter + (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) + (Ioo_subset_Ioo (by simp) (by simp)), + isOpen_Ioo, ?_⟩ + rw [mem_Ioo] + constructor + · apply max_lt ht₁.2.1 + simp [hε] + · apply lt_min ht₁.2.2 + simp [hε] intros t ht exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 @@ -614,7 +602,7 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t - have : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by + obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ rw [mem_Ioo] by_cases ht : t - t₀ < 0 @@ -622,7 +610,6 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte constructor <;> linarith · rw [abs_of_nonneg (not_lt.mp ht)] constructor <;> linarith - obtain ⟨T, hT, ht⟩ := this exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) From 6ea93382bb282d96a29e2387b2e0afdfd2ca524c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 03:33:37 -0800 Subject: [PATCH 107/203] unused --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7136726df2a4b..99d9eac5cd14b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -564,12 +564,12 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] apply IsClosed.closure_subset apply isClosed_eq · rw [continuous_iff_continuousAt] - rintro ⟨t, ht⟩ + rintro ⟨_, ht⟩ apply ContinuousAt.comp _ continuousAt_subtype_val rw [Subtype.coe_mk] exact hγ.continuousAt ht · rw [continuous_iff_continuousAt] - rintro ⟨t, ht⟩ + rintro ⟨_, ht⟩ apply ContinuousAt.comp _ continuousAt_subtype_val rw [Subtype.coe_mk] exact hγ'.continuousAt ht From 6d4af73d702891b0393352d4e353e02875df3ab2 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Fri, 8 Dec 2023 19:09:05 +0100 Subject: [PATCH 108/203] Small things: sectioning; re-use variables. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 72 +++++++++++--------- 1 file changed, 38 insertions(+), 34 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5b50e128262e1..4edc3c06760f1 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -30,7 +30,7 @@ For `IsIntegralCurveOn γ v s` and `IsIntegralCurveAt γ v t₀`, even though ` time, its value outside of the set `s` or a small interval around `t₀` is irrelevant and considered junk. -## To-do +## TODO - Prove `comp_add`, `comp_smul` , etc. lemmas for `IsIntegralCurveOn`, and then derive versions for `IsIntegralCurveAt` and `IsIntegralCurve` as corollaries. @@ -50,10 +50,6 @@ variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners ℝ E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] - {v : (x : M) → TangentSpace I x} {x₀ : M} - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) - {s : Set ℝ} {t₀ : ℝ} - /-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` @@ -73,42 +69,46 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) -lemma IsIntegralCurve.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} - (h : IsIntegralCurve γ v) (s : Set ℝ) : IsIntegralCurveOn γ v s := fun t _ => h t +variable {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {x₀ : M} + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) + {s : Set ℝ} {t₀ : ℝ} + +lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s) : + IsIntegralCurveOn γ v s := fun t _ => h t -lemma isIntegralCurve_iff_isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : - IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := +lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ -lemma IsIntegralCurve.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} - (h : IsIntegralCurve γ v) (t : ℝ) : IsIntegralCurveAt γ v t := +lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : + IsIntegralCurveAt γ v t := ⟨1, zero_lt_one, fun t _ => h t⟩ -lemma isIntegralCurve_iff_isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : +lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := ⟨fun h => h.isIntegralCurveAt, fun h t => by obtain ⟨ε, hε, h⟩ := h t exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε)⟩ -lemma IsIntegralCurveOn.mono {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : IsIntegralCurveOn γ v s) {s' : Set ℝ} (hs : s' ⊆ s) : IsIntegralCurveOn γ v s' := +lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) {s' : Set ℝ} (hs : s' ⊆ s) : + IsIntegralCurveOn γ v s' := fun t ht => h t (mem_of_mem_of_subset ht hs) -lemma IsIntegralCurveOn.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : IsIntegralCurveOn γ v s) {t : ℝ} (hs : s ∈ nhds t) : IsIntegralCurveAt γ v t := by +lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) {t : ℝ} (hs : s ∈ nhds t) : + IsIntegralCurveAt γ v t := by rw [Metric.mem_nhds_iff] at hs obtain ⟨ε, hε, hmem⟩ := hs exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ -lemma IsIntegralCurveAt.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by +lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : + IsIntegralCurveOn γ v s := by intros t ht obtain ⟨ε, hε, h⟩ := h t ht exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) /-! ### Translation lemmas -/ +section Translation -lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : +lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by intros t ht rw [Function.comp_apply, @@ -118,8 +118,8 @@ lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.add_const (hasFDerivAt_id _) _ -lemma isIntegralCurveOn_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveOn γ v s ↔ - IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by +lemma isIntegralCurveOn_comp_add {dt : ℝ} : + IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by refine ⟨fun hγ => hγ.comp_add _, fun hγ => ?_⟩ have := hγ.comp_add (-dt) simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] at this @@ -127,7 +127,7 @@ lemma isIntegralCurveOn_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveOn ext simp only [Function.comp_apply, neg_add_cancel_right] -lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : +lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε, hε, ?_⟩ @@ -136,8 +136,8 @@ lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] rfl -lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by +lemma isIntegralCurveAt_comp_add {dt : ℝ} : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this @@ -145,21 +145,23 @@ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt ext simp only [Function.comp_apply, neg_add_cancel_right] -lemma IsIntegralCurve.comp_add {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (dt : ℝ) : +lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : IsIntegralCurve (γ ∘ (· + dt)) v := by rw [isIntegralCurve_iff_isIntegralCurveOn] at * exact hγ.comp_add _ -lemma isIntegralCurve_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurve γ v ↔ - IsIntegralCurve (γ ∘ (· + dt)) v := by +lemma isIntegralCurve_comp_add {dt : ℝ} : + IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· + dt)) v := by refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) ext simp only [Function.comp_apply, neg_add_cancel_right] +end Translation -/-! ### Scale lemmas -/ +/-! ### Scaling lemmas -/ +section Scaling -lemma IsIntegralCurveOn.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : +lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by intros t ht rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] @@ -167,7 +169,7 @@ lemma IsIntegralCurveOn.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ -lemma isIntegralCurvOn_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by refine ⟨fun hγ => hγ.comp_mul a, fun hγ ↦ ?_⟩ have := hγ.comp_mul a⁻¹ @@ -177,7 +179,7 @@ lemma isIntegralCurvOn_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0 ext t rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] -lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} +lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ @@ -190,7 +192,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve ← add_div, div_lt_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), lt_div_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), and_comm] -lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ => hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) @@ -200,12 +202,12 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ ext t simp [inv_mul_eq_div, div_self ha] -lemma IsIntegralCurve.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (a : ℝ) : +lemma IsIntegralCurve.comp_mul (hγ : IsIntegralCurve γ v) (a : ℝ) : IsIntegralCurve (γ ∘ (· * a)) (a • v) := by rw [isIntegralCurve_iff_isIntegralCurveOn] at * exact hγ.comp_mul _ -lemma isIntegralCurve_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by refine ⟨fun hγ => hγ.comp_mul _, fun hγ => ?_⟩ have := hγ.comp_mul a⁻¹ @@ -222,6 +224,8 @@ lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. +end Scaling + variable (t₀) /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the From ffa3c4bb56c9f777a872d5c1c37a2f89cfd72219 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Fri, 8 Dec 2023 19:24:26 +0100 Subject: [PATCH 109/203] Tinygolf. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4edc3c06760f1..476f64f10710e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -156,11 +156,11 @@ lemma isIntegralCurve_comp_add {dt : ℝ} : convert hγ.comp_add (-dt) ext simp only [Function.comp_apply, neg_add_cancel_right] + end Translation /-! ### Scaling lemmas -/ section Scaling - lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by intros t ht @@ -238,8 +238,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart - obtain ⟨f, hf1, ε1, hε1, hf2⟩ := - exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ + obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd rw [← Real.ball_eq_Ioo] at hf2 -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, @@ -247,7 +246,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := - hcont _ (isOpen_interior.mem_nhds (ModelWithCorners.isInteriorPoint_iff.mp hx)) + hcont _ (isOpen_interior.mem_nhds ((I.isInteriorPoint_iff).mp hx)) rw [Metric.mem_nhds_iff] at hnhds obtain ⟨ε2, hε2, hf3⟩ := hnhds simp_rw [subset_def, mem_preimage] at hf3 @@ -270,7 +269,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) -- express the derivative of the integral curve in the local chart - refine ⟨ContinuousAt.comp (continuousAt_extChartAt_symm'' _ _ hf3') (h.continuousAt), + refine ⟨(continuousAt_extChartAt_symm'' _ _ hf3').comp h.continuousAt, HasDerivWithinAt.hasFDerivWithinAt ?_⟩ simp only [mfld_simps, hasDerivWithinAt_univ] show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) From 7dde305b304d2bd433f4bf1b7a5bcaec925ee41b Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Fri, 8 Dec 2023 19:26:52 +0100 Subject: [PATCH 110/203] Line breaks. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 476f64f10710e..1b2325e017803 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -73,8 +73,8 @@ variable {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {x₀ : M} (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) {s : Set ℝ} {t₀ : ℝ} -lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s) : - IsIntegralCurveOn γ v s := fun t _ => h t +lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s) : IsIntegralCurveOn γ v s := + fun t _ => h t lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ @@ -105,9 +105,9 @@ lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ obtain ⟨ε, hε, h⟩ := h t ht exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) + /-! ### Translation lemmas -/ section Translation - lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by intros t ht From b6a1adae9e1e86d6d6d040fc6343e7dee3fce662 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 12:37:57 -0800 Subject: [PATCH 111/203] golf a bit --- Mathlib/Analysis/ODE/Gronwall.lean | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index b0117c66434be..e2a18efdd43f8 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -165,13 +165,10 @@ theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : apply norm_le_gronwallBound_of_norm_deriv_right_le (hf.sub hg) h_deriv ha intro t ht have := dist_triangle4_right (f' t) (g' t) (v t (f t)) (v t (g t)) - rw [dist_eq_norm] at this - have hv := hv t (x := f t) (hfs t ht) (y := g t) (hgs t ht) - rw [edist_dist, edist_dist, ← ENNReal.ofReal_coe_nnreal, - ← ENNReal.ofReal_mul (NNReal.coe_nonneg _), - ENNReal.ofReal_le_ofReal_iff (mul_nonneg (NNReal.coe_nonneg _) dist_nonneg)] at hv + have hv := (hv t).dist_le_mul _ (hfs t ht) _ (hgs t ht) + rw [← dist_eq_norm, ← dist_eq_norm] refine this.trans ((add_le_add (add_le_add (f_bound t ht) (g_bound t ht)) hv).trans ?_) - rw [dist_eq_norm, add_comm] + rw [add_comm] set_option linter.uppercaseLean3 false in #align dist_le_of_approx_trajectories_ODE_of_mem_set dist_le_of_approx_trajectories_ODE_of_mem_set From f5f7f70be9d17af2bb0628604a5aa5ff47af92a1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 13:22:53 -0800 Subject: [PATCH 112/203] golf and docstring --- Mathlib/Analysis/ODE/Gronwall.lean | 66 +++++++++++++++++------------- 1 file changed, 37 insertions(+), 29 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index e2a18efdd43f8..668aafd5e3900 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -230,7 +230,9 @@ set_option linter.uppercaseLean3 false in /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, -and we consider only solutions included in `s`. -/ +and we consider only solutions included in `s`. + +This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) @@ -242,63 +244,69 @@ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} set_option linter.uppercaseLean3 false in #align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set --- Ico rather than Icc +/-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with +a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, +and we consider only solutions included in `s`. + +This version shows uniqueness in a half-open interval `Ico a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun t ht ↦ + have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) ODE_solution_unique_of_mem_set hv - (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) - (fun _ ht' ↦ hfs _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) - (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) - (fun _ ht' ↦ hgs _ ⟨ht'.1, lt_trans ht'.2 ht.2⟩) ha ⟨ht.1, le_refl _⟩ + (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) + (fun _ ht' ↦ hfs _ (hss ht')) + (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ (hss ht')) + (fun _ ht' ↦ hgs _ (hss ht')) ha ⟨ht.1, le_refl _⟩ + +/-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with +a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, +and we consider only solutions included in `s`. --- starting point t₀ can be inside (a, b) +This version shows uniqueness in an open interval `Ioo a b` that contains the initial time `t₀`. -/ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + -- split `Ioo a b = Ioc a t₀ ∪ Ico t₀ b` and treat the two cases separately apply EqOn.mono (Ioo_subset_Ioc_union_Ico (b := t₀)) apply EqOn.union - · have hv' : ∀ t, LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := fun t ↦ by + · -- case `t ≤ t₀`: show `fun t ↦ f_or_g (-t)` solves `-v (-t)` within `s (-t)` by composition, + -- so we can use the `Ico` version of the uniqueness lemma backwards in time + have hv' : ∀ t, LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := fun t ↦ by rw [← one_mul K] - apply LipschitzWith.comp_lipschitzOnWith (LipschitzWith.neg LipschitzWith.id) - exact hv _ + exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) have : EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) := by + have hmt : MapsTo Neg.neg (Ico (-t₀) (-a)) (Ioo a b) := + fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ apply ODE_solution_unique_of_mem_set_Ico hv' - (hf.comp continuousOn_neg - fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) _ - (fun _ ht' ↦ hfs _ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) - (hg.comp continuousOn_neg - fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) _ - (fun _ ht' ↦ hgs _ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩) - (by simp [ha]) + (hf.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hfs _ (hmt ht')) + (hg.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hgs _ (hmt ht')) (by simp [ha]) · intros t' ht' apply HasDerivAt.hasDerivWithinAt - have := hf' (-t') ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ - convert HasFDerivAt.comp_hasDerivAt t' this (hasDerivAt_neg t') + convert HasFDerivAt.comp_hasDerivAt t' (hf' (-t') (hmt ht')) (hasDerivAt_neg t') simp · intros t' ht' apply HasDerivAt.hasDerivWithinAt - have := hg' (-t') ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ - convert HasFDerivAt.comp_hasDerivAt t' this (hasDerivAt_neg t') + convert HasFDerivAt.comp_hasDerivAt t' (hg' (-t') (hmt ht')) (hasDerivAt_neg t') simp rw [eqOn_comp_right_iff] at this convert this simp - · exact ODE_solution_unique_of_mem_set_Ico hv + · -- case `t ≥ t₀`: follows trivially from the `Ico` version of the uniqueness lemma + have hss : Ico t₀ b ⊆ Ioo a b := fun _ ht' ↦ mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1) + exact ODE_solution_unique_of_mem_set_Ico hv (hf.mono (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ <| - mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => hfs _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) + (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ (hss ht')) + (fun _ ht' => hfs _ <| (hss ht')) (hg.mono (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hg' _ <| - mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => hgs _ <| mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1)) ha + (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hg' _ (hss ht')) + (fun _ ht' => hgs _ <| (hss ht')) ha /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that RHS is Lipschitz continuous in `x`. -/ From 4b2a738f651b0687ce2c417af1fc0bc8d2e67b20 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 13:41:45 -0800 Subject: [PATCH 113/203] usued args --- Mathlib/Analysis/ODE/Gronwall.lean | 2 +- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 668aafd5e3900..0c968423992d3 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -254,7 +254,7 @@ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Se {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun t ht ↦ + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) ODE_solution_unique_of_mem_set hv (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 99d9eac5cd14b..43b17c6c73f35 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -442,7 +442,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInt rfl rw [← h2] exact this - have hsub : (fun x ↦ v') t ((↑(extChartAt I (γ' t₀)) ∘ γ') t) = + have hsub : (fun _ ↦ v') t ((↑(extChartAt I (γ' t₀)) ∘ γ') t) = (tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t)) := by dsimp only rw [h, Function.comp_apply, LocalEquiv.left_inv] From 1898133eab1dff5baf3a34de0742413cfe8f712a Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 14:02:05 -0800 Subject: [PATCH 114/203] remove temp notation --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 12 ++++++++++-- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 2 -- 2 files changed, 10 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 43b17c6c73f35..851d22160edbd 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -55,8 +55,6 @@ integral curve, vector field, local existence, uniqueness -/ open scoped Manifold -scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x -scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) open Set @@ -614,3 +612,13 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht + +example (s t : Set ℝ) : closure (s ∩ t) ∩ t ⊆ s ∩ t := by + rw [← Subtype.image_preimage_val, ← Subtype.image_preimage_val, + image_subset_image_iff Subtype.val_injective] + intros t ht + rw [mem_preimage, ← closure_subtype] at ht + revert ht t + apply IsClosed.closure_subset + -- goal state: `IsClosed (Subtype.val ⁻¹' s)` + sorry diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 45a9ed5e2a11a..19fbaf1db2452 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -39,8 +39,6 @@ manifold, interior, boundary open Set open scoped Manifold -scoped[Manifold] notation "𝓔(" I ", " x ")" => extChartAt I x -scoped[Manifold] notation "𝓔⁻¹(" I ", " x ")" => LocalEquiv.symm (𝓔(I, x)) -- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] From fa1e07f7cc4dc94558afbdf4bef87c51f8bce36e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 8 Dec 2023 14:11:04 -0800 Subject: [PATCH 115/203] oops --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 10 ---------- 1 file changed, 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 851d22160edbd..920cc08e091ff 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -612,13 +612,3 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht - -example (s t : Set ℝ) : closure (s ∩ t) ∩ t ⊆ s ∩ t := by - rw [← Subtype.image_preimage_val, ← Subtype.image_preimage_val, - image_subset_image_iff Subtype.val_injective] - intros t ht - rw [mem_preimage, ← closure_subtype] at ht - revert ht t - apply IsClosed.closure_subset - -- goal state: `IsClosed (Subtype.val ⁻¹' s)` - sorry From 780373e74c0b850b2f6fafd55616722c6557807a Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 9 Dec 2023 02:18:47 -0800 Subject: [PATCH 116/203] small golf --- Mathlib/Analysis/ODE/Gronwall.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 0c968423992d3..3f947ca93df58 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -299,7 +299,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se convert this simp · -- case `t ≥ t₀`: follows trivially from the `Ico` version of the uniqueness lemma - have hss : Ico t₀ b ⊆ Ioo a b := fun _ ht' ↦ mem_of_mem_of_subset ht' (Ico_subset_Ioo_left ht.1) + have hss := Ico_subset_Ioo_left (b := b) ht.1 exact ODE_solution_unique_of_mem_set_Ico hv (hf.mono (Ico_subset_Ioo_left ht.1)) (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ (hss ht')) From 8cfee4c7efa2599c1c77d675bf0eb005c0af2023 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 9 Dec 2023 02:19:43 -0800 Subject: [PATCH 117/203] EqOn --- Mathlib/Analysis/ODE/Gronwall.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 3f947ca93df58..cf8421995e903 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -314,7 +314,7 @@ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, Lip {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : - ∀ t ∈ Icc a b, f t = g t := + EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial ODE_solution_unique_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha From 64942d13911563aed4267fe72bc5a223689bb493 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 9 Dec 2023 02:23:03 -0800 Subject: [PATCH 118/203] style --- Mathlib/Analysis/ODE/Gronwall.lean | 37 ++++++++++++++++-------------- 1 file changed, 20 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index cf8421995e903..7eff7d25207e5 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -234,11 +234,12 @@ and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} + (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) + (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in @@ -250,11 +251,12 @@ and we consider only solutions included in `s`. This version shows uniqueness in a half-open interval `Ico a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} + (hf : ContinuousOn f (Ico a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) + (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) ODE_solution_unique_of_mem_set hv (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) @@ -268,11 +270,12 @@ and we consider only solutions included in `s`. This version shows uniqueness in an open interval `Ioo a b` that contains the initial time `t₀`. -/ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) - (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) + (hf : ContinuousOn f (Ioo a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) + (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) - (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) + (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by -- split `Ioo a b = Ioc a t₀ ∪ Ico t₀ b` and treat the two cases separately apply EqOn.mono (Ioo_subset_Ioc_union_Ico (b := t₀)) apply EqOn.union @@ -310,11 +313,11 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that RHS is Lipschitz continuous in `x`. -/ -theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) - (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : - EqOn f g (Icc a b) := +theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} + (hv : ∀ t, LipschitzWith K (v t)) {f g : ℝ → E} {a b : ℝ} + (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) + (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial ODE_solution_unique_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha From cb2fab29a0b2b4740ce8fc2ddace62625ce7360e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 9 Dec 2023 02:25:46 -0800 Subject: [PATCH 119/203] style --- Mathlib/Analysis/ODE/Gronwall.lean | 38 ++++++++++++++---------------- 1 file changed, 18 insertions(+), 20 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 7eff7d25207e5..cbcfa9650011e 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -234,12 +234,12 @@ and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} - (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) - (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) - (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : + EqOn f g (Icc a b) := fun t ht ↦ by have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in @@ -251,12 +251,11 @@ and we consider only solutions included in `s`. This version shows uniqueness in a half-open interval `Ico a b`, where `a` is the initial time. -/ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} - (hf : ContinuousOn f (Ico a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) - (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) - (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) ODE_solution_unique_of_mem_set hv (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) @@ -270,12 +269,11 @@ and we consider only solutions included in `s`. This version shows uniqueness in an open interval `Ioo a b` that contains the initial time `t₀`. -/ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) - (hf : ContinuousOn f (Ioo a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) - (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) + (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) - (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) - (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by -- split `Ioo a b = Ioc a t₀ ∪ Ico t₀ b` and treat the two cases separately apply EqOn.mono (Ioo_subset_Ioc_union_Ico (b := t₀)) apply EqOn.union @@ -313,11 +311,11 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that RHS is Lipschitz continuous in `x`. -/ -theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} - (hv : ∀ t, LipschitzWith K (v t)) {f g : ℝ → E} {a b : ℝ} - (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) - (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (ha : f a = g a) : EqOn f g (Icc a b) := +theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) + {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : + EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial ODE_solution_unique_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha From 3a6a98d246e6b590b820fc87f5cdaf988c2a45b5 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 9 Dec 2023 02:27:24 -0800 Subject: [PATCH 120/203] have parentheses --- Mathlib/Analysis/ODE/Gronwall.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index cbcfa9650011e..b6a84c2b7c7b6 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -279,7 +279,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se apply EqOn.union · -- case `t ≤ t₀`: show `fun t ↦ f_or_g (-t)` solves `-v (-t)` within `s (-t)` by composition, -- so we can use the `Ico` version of the uniqueness lemma backwards in time - have hv' : ∀ t, LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := fun t ↦ by + have hv' (t) : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by rw [← one_mul K] exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) have : EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) := by From 06674b38e89fd5a35c6d75aeb08d318c3217b9d1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 12 Dec 2023 20:28:34 -0800 Subject: [PATCH 121/203] style --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 86 ++++++++++---------- 1 file changed, 45 insertions(+), 41 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 920cc08e091ff..994d8993536fb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -65,10 +65,6 @@ variable {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners ℝ E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] - {v : (x : M) → TangentSpace I x} {x₀ : M} - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) - {s : Set ℝ} {t₀ : ℝ} - /-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` @@ -88,42 +84,50 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) -lemma IsIntegralCurve.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} - (h : IsIntegralCurve γ v) (s : Set ℝ) : IsIntegralCurveOn γ v s := fun t _ => h t +variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} + +lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) : + IsIntegralCurveOn γ v s := fun t _ => h t -lemma isIntegralCurve_iff_isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : +lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ -lemma IsIntegralCurve.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} - (h : IsIntegralCurve γ v) (t : ℝ) : IsIntegralCurveAt γ v t := - ⟨1, zero_lt_one, fun t _ => h t⟩ +lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : + IsIntegralCurveAt γ v t := ⟨1, zero_lt_one, fun t _ => h t⟩ -lemma isIntegralCurve_iff_isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} : +lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := ⟨fun h => h.isIntegralCurveAt, fun h t => by obtain ⟨ε, hε, h⟩ := h t exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε)⟩ -lemma IsIntegralCurveOn.mono {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : IsIntegralCurveOn γ v s) {s' : Set ℝ} (hs : s' ⊆ s) : IsIntegralCurveOn γ v s' := - fun t ht => h t (mem_of_mem_of_subset ht hs) +lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : + IsIntegralCurveOn γ v s' := fun t ht => h t (mem_of_mem_of_subset ht hs) + +lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : + IsIntegralCurveOn γ v (s ∪ s') := by + intros t ht + rw [mem_union] at ht + cases' ht with ht ht + · exact h _ ht + · exact h' _ ht -lemma IsIntegralCurveOn.isIntegralCurveAt {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : IsIntegralCurveOn γ v s) {t : ℝ} (hs : s ∈ nhds t) : IsIntegralCurveAt γ v t := by +lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : + IsIntegralCurveAt γ v t₀ := by rw [Metric.mem_nhds_iff] at hs obtain ⟨ε, hε, hmem⟩ := hs exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ -lemma IsIntegralCurveAt.isIntegralCurveOn {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {s : Set ℝ} - (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by +lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : + IsIntegralCurveOn γ v s := by intros t ht obtain ⟨ε, hε, h⟩ := h t ht exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) /-! ### Translation lemmas -/ -lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : +lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by intros t ht rw [Function.comp_apply, @@ -133,8 +137,8 @@ lemma IsIntegralCurveOn.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.add_const (hasFDerivAt_id _) _ -lemma isIntegralCurveOn_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveOn γ v s ↔ - IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by +lemma isIntegralCurveOn_comp_add {dt : ℝ} : + IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by refine ⟨fun hγ => hγ.comp_add _, fun hγ => ?_⟩ have := hγ.comp_add (-dt) simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] at this @@ -142,17 +146,17 @@ lemma isIntegralCurveOn_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveOn ext simp only [Function.comp_apply, neg_add_cancel_right] -lemma IsIntegralCurveAt.comp_add {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : +lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε, hε, ?_⟩ convert h.comp_add dt - ext t' + ext rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] rfl -lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ - IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by +lemma isIntegralCurveAt_comp_add {dt : ℝ} : + IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this @@ -160,13 +164,13 @@ lemma isIntegralCurveAt_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurveAt ext simp only [Function.comp_apply, neg_add_cancel_right] -lemma IsIntegralCurve.comp_add {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (dt : ℝ) : +lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : IsIntegralCurve (γ ∘ (· + dt)) v := by rw [isIntegralCurve_iff_isIntegralCurveOn] at * exact hγ.comp_add _ -lemma isIntegralCurve_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurve γ v ↔ - IsIntegralCurve (γ ∘ (· + dt)) v := by +lemma isIntegralCurve_comp_add {dt : ℝ} : + IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· + dt)) v := by refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) ext @@ -174,7 +178,7 @@ lemma isIntegralCurve_comp_add {γ : ℝ → M} {dt : ℝ} : IsIntegralCurve γ /-! ### Scale lemmas -/ -lemma IsIntegralCurveOn.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : +lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by intros t ht rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp] @@ -182,7 +186,7 @@ lemma IsIntegralCurveOn.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ -lemma isIntegralCurvOn_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by refine ⟨fun hγ => hγ.comp_mul a, fun hγ ↦ ?_⟩ have := hγ.comp_mul a⁻¹ @@ -192,8 +196,8 @@ lemma isIntegralCurvOn_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0 ext t rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] -lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} - (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ convert h.comp_mul a @@ -205,7 +209,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero {γ : ℝ → M} (hγ : IsIntegralCurve ← add_div, div_lt_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), lt_div_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), and_comm] -lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ => hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) @@ -215,12 +219,12 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ ext t simp [inv_mul_eq_div, div_self ha] -lemma IsIntegralCurve.comp_mul {γ : ℝ → M} (hγ : IsIntegralCurve γ v) (a : ℝ) : +lemma IsIntegralCurve.comp_mul (hγ : IsIntegralCurve γ v) (a : ℝ) : IsIntegralCurve (γ ∘ (· * a)) (a • v) := by rw [isIntegralCurve_iff_isIntegralCurveOn] at * exact hγ.comp_mul _ -lemma isIntegralCurve_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by refine ⟨fun hγ => hγ.comp_mul _, fun hγ => ?_⟩ have := hγ.comp_mul a⁻¹ @@ -231,30 +235,30 @@ lemma isIntegralCurve_comp_mul_ne_zero {γ : ℝ → M} {a : ℝ} (ha : a ≠ 0) /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ -lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by +lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ => x) v := by intro t rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. -lemma IsIntegralCurveOn.continuousAt {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ∈ s) : +lemma IsIntegralCurveOn.continuousAt (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ∈ s) : ContinuousAt γ t₀ := (hγ t₀ ht).1 -lemma IsIntegralCurveOn.continuousOn {γ : ℝ → M} (hγ : IsIntegralCurveOn γ v s) : +lemma IsIntegralCurveOn.continuousOn (hγ : IsIntegralCurveOn γ v s) : ContinuousOn γ s := fun t ht => (hγ t ht).1.continuousWithinAt -lemma IsIntegralCurveAt.continuousAt {γ : ℝ → M} (hγ : IsIntegralCurveAt γ v t₀) : +lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : ContinuousAt γ t₀ := by obtain ⟨ε, hε, hγ⟩ := hγ apply hγ.continuousAt rw [← Real.ball_eq_Ioo] exact Metric.mem_ball_self hε -lemma IsIntegralCurve.continuous {γ : ℝ → M} (hγ : IsIntegralCurve γ v) : +lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : Continuous γ := continuous_iff_continuousAt.mpr fun _ => (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) -variable (t₀) +variable (t₀) {x₀ : M} /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector From 5c5f070e721246a87de96927f4eb7bd41d3efc1c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 12 Dec 2023 20:29:43 -0800 Subject: [PATCH 122/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 994d8993536fb..7883bc11ac0eb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -264,7 +264,9 @@ variable (t₀) {x₀ : M} manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : +theorem exists_isIntegralCurveAt_of_contMDiffAt + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) + (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by -- express the differentiability of the section `v` in the local charts rw [contMDiffAt_iff] at hv @@ -324,9 +326,10 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint + exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint variable (I) From ccd18300d388360967345cfac223320451e817c7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 12 Dec 2023 20:39:31 -0800 Subject: [PATCH 123/203] style --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 82 +++++++++++++------- 1 file changed, 52 insertions(+), 30 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 1b2325e017803..ee0a7cab6dc6c 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -10,12 +10,14 @@ import Mathlib.Geometry.Manifold.MFDeriv /-! # Integral curves of vector fields on a manifold -For any continuously differentiable vector field on a manifold `M` and any chosen interior point -`x₀ : M`, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of -`γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. +Let `M` be a manifold and `v : (x : M) → TangentSpace I x` be a vector field on `M`. An integral +curve of `v` is a function `γ : ℝ → M` such that the derivative of `γ` at `t` equals `v (γ t)`. The +integral curve may only be defined for all `t` within some subset of `ℝ`. -As a corollary, such an integral curve exists for any starting point `x₀` if `M` is a manifold -without boundary. +Assume `v` is continuously differentiable. The existence theorem for solutions to ODEs implies that +a unique local integral curve exists for any continuously differentiable vector field `v`. The +uniqueness theorem for solutions to ODEs implies that integral curves of `v` are unique. These are +the main results of this file. ## Main definition @@ -30,17 +32,29 @@ For `IsIntegralCurveOn γ v s` and `IsIntegralCurveAt γ v t₀`, even though ` time, its value outside of the set `s` or a small interval around `t₀` is irrelevant and considered junk. -## TODO +## Implementation notes -- Prove `comp_add`, `comp_smul` , etc. lemmas for `IsIntegralCurveOn`, and then derive versions for -`IsIntegralCurveAt` and `IsIntegralCurve` as corollaries. +For the existence and uniqueness theorems, we assume that the image of the integral curve lies in +the interior of the manifold. The case where the integral curve may lie on the boundary of the +manifold requires special treatment, and we leave it as a to-do. + +The uniqueness theorem requires the manifold to be Hausdorff (T2), so that the set on which two +continuous functions agree is closed. + +We state simpler versions of the theorem for manifolds without boundary as corollaries. + +## To-do + +- The case where the integral curve may venture to the boundary of the manifold. See Theorem 9.34, + J. M. Lee. May require submanifolds. ## Tags -integral curve, vector field, local existence +integral curve, vector field, local existence, uniqueness -/ open scoped Manifold + open Set variable @@ -69,19 +83,17 @@ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) -variable {γ : ℝ → M} {v : (x : M) → TangentSpace I x} {x₀ : M} - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) - {s : Set ℝ} {t₀ : ℝ} +variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} -lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s) : IsIntegralCurveOn γ v s := - fun t _ => h t +lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) : + IsIntegralCurveOn γ v s := fun t _ => h t -lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := +lemma isIntegralCurve_iff_isIntegralCurveOn : + IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : - IsIntegralCurveAt γ v t := - ⟨1, zero_lt_one, fun t _ => h t⟩ + IsIntegralCurveAt γ v t := ⟨1, zero_lt_one, fun t _ => h t⟩ lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := @@ -89,12 +101,19 @@ lemma isIntegralCurve_iff_isIntegralCurveAt : obtain ⟨ε, hε, h⟩ := h t exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε)⟩ -lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) {s' : Set ℝ} (hs : s' ⊆ s) : - IsIntegralCurveOn γ v s' := - fun t ht => h t (mem_of_mem_of_subset ht hs) +lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : + IsIntegralCurveOn γ v s' := fun t ht => h t (mem_of_mem_of_subset ht hs) + +lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : + IsIntegralCurveOn γ v (s ∪ s') := by + intros t ht + rw [mem_union] at ht + cases' ht with ht ht + · exact h _ ht + · exact h' _ ht -lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) {t : ℝ} (hs : s ∈ nhds t) : - IsIntegralCurveAt γ v t := by +lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : + IsIntegralCurveAt γ v t₀ := by rw [Metric.mem_nhds_iff] at hs obtain ⟨ε, hε, hmem⟩ := hs exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ @@ -132,7 +151,7 @@ lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε, hε, ?_⟩ convert h.comp_add dt - ext t' + ext rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] rfl @@ -179,8 +198,8 @@ lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : ext t rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] -lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} - (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by +lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : + IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ convert h.comp_mul a @@ -218,7 +237,7 @@ lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ -lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v := by +lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ => x) v := by intro t rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] @@ -226,13 +245,15 @@ lemma isIntegralCurve_const (h : v x₀ = 0) : IsIntegralCurve (fun _ => x₀) v end Scaling -variable (t₀) +variable (t₀) {x₀ : M} /-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ -theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : +theorem exists_isIntegralCurveAt_of_contMDiffAt + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) + (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by -- express the differentiability of the section `v` in the local charts rw [contMDiffAt_iff] at hv @@ -291,6 +312,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hx : I.IsInteriorPoint x₀) : chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] : +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt hv t₀ I.isInteriorPoint + exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint From 122546dc8f2548a1c7baf9bd655879a9cf10971b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 12 Dec 2023 20:39:49 -0800 Subject: [PATCH 124/203] spacing --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index ee0a7cab6dc6c..01d75ee8521a9 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -179,7 +179,9 @@ lemma isIntegralCurve_comp_add {dt : ℝ} : end Translation /-! ### Scaling lemmas -/ + section Scaling + lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by intros t ht From c1c57c9444cef35d6e8d56c067f3145101775cce Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 12 Dec 2023 20:41:40 -0800 Subject: [PATCH 125/203] spacing --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index d7cecf7e3702d..372a21a304d61 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -180,7 +180,9 @@ lemma isIntegralCurve_comp_add {dt : ℝ} : end Translation /-! ### Scaling lemmas -/ + section Scaling + lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by intros t ht From 0ed0eec35b81119df1a61ee3736f184c56d4e88b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 03:01:36 -0800 Subject: [PATCH 126/203] split aux lemmas out --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 340 +++++++++---------- 1 file changed, 162 insertions(+), 178 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 372a21a304d61..9308f29f18e93 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -339,213 +339,197 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] variable (I) -/-- Local integral curves are unique. - - If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` - at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt {γ γ' : ℝ → M} (ht : I.IsInteriorPoint (γ t₀)) - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) - (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : - ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by - -- extract set `s` on which `v` is Lipschitz - set v' : E → E := fun x => - tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) - (v ((extChartAt I (γ t₀)).symm x)) with hv' - rw [contMDiffAt_iff] at hv - obtain ⟨_, hv⟩ := hv - obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht)).snd - have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - -- extract `εs` so that `(extChartAt I (γ t₀)) (γ t) ∈ s` for all `t ∈ Ioo (t₀ - εs) (t₀ + εs)` +lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux + (hγ : IsIntegralCurveAt γ v t₀) {se : Set E} (hse : se ∈ nhds (extChartAt I (γ t₀) (γ t₀))) : + ∃ ε > 0, (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ + (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), γ x ∈ (extChartAt I (γ t₀)).source) ∧ + (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), + HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ + ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) := by have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt rw [continuousAt_def] at hcont - have hnhds := hcont _ hs + have hnhds := hcont _ hse rw [Metric.mem_nhds_iff] at hnhds - obtain ⟨εs, hεs, hmem⟩ := hnhds + obtain ⟨εmem, hεmem, hmem⟩ := hnhds simp_rw [subset_def, mem_preimage] at hmem - -- `εs'` for `γ'` - have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := - ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt - rw [continuousAt_def] at hcont' - have hnhds' := hcont' _ (h ▸ hs) - rw [Metric.mem_nhds_iff] at hnhds' - obtain ⟨εs', hεs', hmem'⟩ := hnhds' - simp_rw [subset_def, mem_preimage] at hmem' - -- extract `εe` so `γ t` stays within the interior of the chart around `γ t₀` + -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` have := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) rw [Metric.mem_nhds_iff] at this - obtain ⟨εe, hεe, hsrc⟩ := this + obtain ⟨εsrc, hεsrc, hsrc⟩ := this simp_rw [subset_def, mem_preimage] at hsrc - have := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) - rw [Metric.mem_nhds_iff] at this - obtain ⟨εe', hεe', hsrc'⟩ := this - simp_rw [subset_def, mem_preimage] at hsrc' -- extract `εγ` from local existence of integral curve obtain ⟨εγ, hεγ, hγ⟩ := hγ - obtain ⟨εγ', hεγ', hγ'⟩ := hγ' - let ε := min (min εe εe') <| min (min εs εs') (min εγ εγ') - have hf : ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) := by - apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) - · apply hγ.continuousOn.mono - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - · apply MapsTo.mono_left hsrc + let ε := min εmem <| min εsrc εγ + have hε : 0 < ε := lt_min hεmem (lt_min hεsrc hεγ) + refine ⟨ε, hε, ?_, ?_, ?_, ?_⟩ + · -- some mono lemmas for each of these here would be perfect. an API for "shrinking ε"? + intros t ht + apply hmem t + rw [← Real.ball_eq_Ioo] at ht + revert ht + apply Metric.ball_subset_ball + exact min_le_left _ _ + · intros t ht + apply hsrc t + rw [← Real.ball_eq_Ioo] at ht + revert ht + apply Metric.ball_subset_ball + apply le_trans (min_le_right _ _) + exact min_le_left _ _ + · intros t ht + -- turn `HasDerivAt` into comp of `HasMFDerivAt` + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` + -- see, `htsrc` and `htγ` should be packed away in mono lemmas elsewhere + have htsrc : t ∈ Ioo (t₀ - εsrc) (t₀ + εsrc) := by rw [← Real.ball_eq_Ioo] + rw [← Real.ball_eq_Ioo] at ht + revert ht apply Metric.ball_subset_ball - simp - have hf' : ContinuousOn ((extChartAt I (γ' t₀)) ∘ γ') (Ioo (t₀ - ε) (t₀ + ε)) := by - apply ContinuousOn.comp (continuousOn_extChartAt I (γ' t₀)) - · apply hγ'.continuousOn.mono - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - · apply MapsTo.mono_left hsrc' + apply le_trans (min_le_right _ _) + exact min_le_left _ _ + have htγ : t ∈ Ioo (t₀ - εγ) (t₀ + εγ) := by rw [← Real.ball_eq_Ioo] + rw [← Real.ball_eq_Ioo] at ht + revert ht apply Metric.ball_subset_ball - simp - have hε : 0 < ε := lt_min (lt_min hεe hεe') (lt_min (lt_min hεs hεs') (lt_min hεγ hεγ')) - refine ⟨ε, hε, ?_⟩ - have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') - (Ioo (t₀ - ε) (t₀ + ε)) := by - apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) _ hf _ _ hf' - · intros t ht - have ht' : t ∈ Ioo (t₀ - εγ') (t₀ + εγ') := by - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - have hrw : HasDerivAt ((extChartAt I (γ' t₀)) ∘ γ') - (tangentCoordChange I (γ' t) (γ' t₀) (γ' t) (v (γ' t))) t := by - -- turn `HasDerivAt` into comp of `HasMFDerivAt` - rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` - have := (hasMFDerivAt_extChartAt I (x := γ' t₀) (y := γ' t) (by - rw [← extChartAt_source I] - apply hsrc' - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - )).comp t (hγ' t ht') - have h2 : ContinuousLinearMap.comp - (mfderiv I I (↑(chartAt H (γ' t₀))) (γ' t)) - (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ' t))) = - ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) - ((tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t))) := by - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, - ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] - congr - rw [tangentCoordChange_def, mfderiv] - have hdiff : MDifferentiableAt I I (↑(chartAt H (γ' t₀))) (γ' t) := by - apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) - rw [← extChartAt_source I] - apply hsrc' - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - simp only [hdiff, if_true] - rfl - rw [← h2] - exact this - have hsub : (fun _ ↦ v') t ((↑(extChartAt I (γ' t₀)) ∘ γ') t) = - (tangentCoordChange I (γ' t) (γ' t₀) (γ' t)) (v (γ' t)) := by - dsimp only - rw [h, Function.comp_apply, LocalEquiv.left_inv] - apply hsrc' + apply le_trans (min_le_right _ _) + exact min_le_right _ _ + have := (hasMFDerivAt_extChartAt I (x := γ t₀) (y := γ t) (by + rw [← extChartAt_source I] + apply hsrc + rw [Real.ball_eq_Ioo] + exact htsrc + )).comp t (hγ t htγ) + -- can we generalise the following to some lemmas? + have h2 : ContinuousLinearMap.comp + (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) + (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = + ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, + ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] + congr + rw [tangentCoordChange_def, mfderiv] + have hdiff : MDifferentiableAt I I (↑(chartAt H (γ t₀))) (γ t) := by + apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) + rw [← extChartAt_source I] + apply hsrc apply mem_of_mem_of_subset ht rw [← Real.ball_eq_Ioo] apply Metric.ball_subset_ball simp - rw [hsub] - exact hrw - · intros t ht - apply hmem' - apply mem_of_mem_of_subset ht + simp only [hdiff, if_true] + rfl + rw [← h2] + exact this + · apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) + -- see, these are again mono + · apply hγ.continuousOn.mono + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball + simp + · apply MapsTo.mono_left hsrc rw [← Real.ball_eq_Ioo] apply Metric.ball_subset_ball simp + +-- what would an api look like to synthesise this by composing individual mono lemmas? +lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono {se : Set E} {ε ε' : ℝ} (hle : ε' ≤ ε) + (h : (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ + (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), γ x ∈ (extChartAt I (γ t₀)).source) ∧ + (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), + HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ + ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) ) : + (∀ x ∈ Ioo (t₀ - ε') (t₀ + ε'), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ + (∀ x ∈ Ioo (t₀ - ε') (t₀ + ε'), γ x ∈ (extChartAt I (γ t₀)).source) ∧ + (∀ t ∈ Ioo (t₀ - ε') (t₀ + ε'), + HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ + ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε') (t₀ + ε')) := by + obtain ⟨h1, h2, h3, h4⟩ := h + -- all of these should be standard mono lemmas + refine ⟨?_, ?_, ?_, ?_⟩ + · intros t ht + apply h1 t + rw [← Real.ball_eq_Ioo] + rw [← Real.ball_eq_Ioo] at ht + revert ht + apply Metric.ball_subset_ball hle + · intros t ht + apply h2 t + rw [← Real.ball_eq_Ioo] + rw [← Real.ball_eq_Ioo] at ht + revert ht + apply Metric.ball_subset_ball hle + · intros t ht + apply h3 t + rw [← Real.ball_eq_Ioo] + rw [← Real.ball_eq_Ioo] at ht + revert ht + apply Metric.ball_subset_ball hle + · apply h4.mono + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball hle + +/-- Local integral curves are unique. + + If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` + at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ +theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : + ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by + -- extract set `s` on which `v` is Lipschitz + set v' : E → E := fun x => + tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) + (v ((extChartAt I (γ t₀)).symm x)) with hv' + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd + have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip + obtain ⟨ε, hε, haux⟩ := + isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ hs + obtain ⟨ε', hε', haux'⟩ := + isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ' (h ▸ hs) + have hpos : 0 < min ε ε' := lt_min_iff.mpr ⟨hε, hε'⟩ + obtain ⟨hmem, hsrc, hγ, hcont⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono + I t₀ (min_le_left _ _) haux + obtain ⟨hmem', hsrc', hγ', hcont'⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono + I t₀ (min_le_right _ _) haux' + refine ⟨min ε ε', hpos, ?_⟩ + have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') + (Ioo (t₀ - (min ε ε')) (t₀ + (min ε ε'))) := by + apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) _ hcont _ hmem hcont' _ hmem' · simp [h] · rw [← Real.ball_eq_Ioo] - exact Metric.mem_ball_self hε + exact Metric.mem_ball_self hpos · intros t ht - have ht' : t ∈ Ioo (t₀ - εγ) (t₀ + εγ) := by - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - have hrw : HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by - -- turn `HasDerivAt` into comp of `HasMFDerivAt` - rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` - have := (hasMFDerivAt_extChartAt I (x := γ t₀) (y := γ t) (by - rw [← extChartAt_source I] - apply hsrc - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - )).comp t (hγ t ht') - have h2 : ContinuousLinearMap.comp - (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) - (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = - ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, - ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] - congr - rw [tangentCoordChange_def, mfderiv] - have hdiff : MDifferentiableAt I I (↑(chartAt H (γ t₀))) (γ t) := by - apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) - rw [← extChartAt_source I] - apply hsrc - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - simp only [hdiff, if_true] - rfl - rw [← h2] - exact this - have hsub : (fun _ ↦ v') t ((↑(extChartAt I (γ t₀)) ∘ γ) t) = - (tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t)) := by - dsimp only + rw [hv'] + apply (hγ t ht).hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? + have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, LocalEquiv.left_inv] - apply hsrc - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - rw [hsub] - exact hrw + exact hsrc t ht + rw [this] · intros t ht - apply hmem - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp + rw [hv', h] + apply (hγ' t ht).hasFDerivAt.congr_fderiv + have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by + rw [Function.comp_apply, LocalEquiv.left_inv] + exact hsrc' t ht + rw [this] refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) · intros t ht - rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv] - apply hsrc - rw [← Real.ball_eq_Ioo] at ht - apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) - simp + rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv _ (hsrc _ ht)] · intros t ht - rw [Function.comp_apply, Function.comp_apply, h, LocalEquiv.left_inv] - apply hsrc' - rw [← Real.ball_eq_Ioo] at ht - apply mem_of_mem_of_subset ht (Metric.ball_subset_ball _) - simp + rw [Function.comp_apply, Function.comp_apply, h, LocalEquiv.left_inv _ (hsrc' _ ht)] /-- Integral curves are unique on open intervals. From 729d5476b17aa211b062bfc0cfd793eb58aa2d06 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 18:36:10 -0800 Subject: [PATCH 127/203] refactored with shrinkable --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 277 ++++++++++--------- 1 file changed, 149 insertions(+), 128 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 9308f29f18e93..8fd2751f8b072 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -339,14 +339,93 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] variable (I) +section shrinkable + +def Shrinkable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε') (_ : ε' ≤ ε) (_ : p ε), p ε' + +def Growable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε) (_ : ε ≤ ε') (_ : p ε), p ε' + +lemma shrinkable_min {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) {εp εq : ℝ} + (hεp : 0 < εp) (hεq : 0 < εq) (hp : p εp) (hq : q εq) : + p (min εp εq) ∧ q (min εp εq) := + ⟨hsp (lt_min hεp hεq) (min_le_left _ _) hp, hsq (lt_min hεp hεq) (min_le_right _ _) hq⟩ + +lemma shrinkable_exists_and {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) + (hp : ∃ ε > 0, p ε) (hq : ∃ ε > 0, q ε) : ∃ ε > 0, p ε ∧ q ε := by + obtain ⟨εp, hεp, hp⟩ := hp + obtain ⟨εq, hεq, hq⟩ := hq + exact ⟨min εp εq, lt_min hεp hεq, shrinkable_min hsp hsq hεp hεq hp hq⟩ + +lemma shrinkable_and {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) : + Shrinkable (fun ε => p ε ∧ q ε) := by + intros ε ε' hpos hle h + exact ⟨hsp hpos hle h.1, hsq hpos hle h.2⟩ + +lemma shrinkable_forall_mem_set {α : Type*} {s : ℝ → Set α} (p : α → Prop) + (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => ∀ x ∈ s ε, p x) := by + intros ε ε' hpos hle hp x hx + apply hp + exact hs x hpos hle hx + +lemma shrinkable_eqOn {α : Type*} {s : ℝ → Set ℝ} (f g : ℝ → α) + (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => EqOn f g (s ε)) := by + intros ε ε' hpos hle hp + apply hp.mono + intros x hx + exact hs _ hpos hle hx + +lemma shrinkable_continuousOn {α : Type*} [TopologicalSpace α] {s : ℝ → Set ℝ} (f : ℝ → α) + (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => ContinuousOn f (s ε)) := by + intros ε ε' hpos hle hp + apply hp.mono + intros x hx + exact hs _ hpos hle hx + +lemma growable_mem_Ioo (x : ℝ) : Growable fun ε ↦ x ∈ Ioo (t₀ - ε) (t₀ + ε) := by + intros ε ε' _ hle + simp_rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball hle + +end shrinkable + +lemma IsIntegralCurveOn.hasDerivAt {ε : ℝ} + (hγ : IsIntegralCurveOn γ v (Ioo (t₀ - ε) (t₀ + ε))) {t : ℝ} (ht : t ∈ Ioo (t₀ - ε) (t₀ + ε)) + (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : + HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by + -- turn `HasDerivAt` into comp of `HasMFDerivAt` + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + have hsub : ContinuousLinearMap.comp + (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) + (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = + ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, + ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] + congr + have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) + ((extChartAt_source I (γ t₀)) ▸ hsrc) + rw [tangentCoordChange_def, mfderiv, if_pos this] + rfl + rw [← hsub] + apply HasMFDerivAt.comp t _ (hγ _ ht) + apply hasMFDerivAt_extChartAt + rw [← extChartAt_source I] + exact hsrc + lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux (hγ : IsIntegralCurveAt γ v t₀) {se : Set E} (hse : se ∈ nhds (extChartAt I (γ t₀) (γ t₀))) : - ∃ ε > 0, (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ - (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), γ x ∈ (extChartAt I (γ t₀)).source) ∧ + ∃ ε > 0, (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) t ∈ se) ∧ + (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), γ t ∈ (extChartAt I (γ t₀)).source) ∧ (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) := by + + -- extract `εmem` so `γ t` when expressed in the chart stays within `se` have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt rw [continuousAt_def] at hcont @@ -354,129 +433,63 @@ lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux rw [Metric.mem_nhds_iff] at hnhds obtain ⟨εmem, hεmem, hmem⟩ := hnhds simp_rw [subset_def, mem_preimage] at hmem + -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` have := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) rw [Metric.mem_nhds_iff] at this obtain ⟨εsrc, hεsrc, hsrc⟩ := this simp_rw [subset_def, mem_preimage] at hsrc + -- extract `εγ` from local existence of integral curve obtain ⟨εγ, hεγ, hγ⟩ := hγ - let ε := min εmem <| min εsrc εγ - have hε : 0 < ε := lt_min hεmem (lt_min hεsrc hεγ) - refine ⟨ε, hε, ?_, ?_, ?_, ?_⟩ - · -- some mono lemmas for each of these here would be perfect. an API for "shrinking ε"? - intros t ht - apply hmem t - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball - exact min_le_left _ _ - · intros t ht - apply hsrc t - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball - apply le_trans (min_le_right _ _) - exact min_le_left _ _ - · intros t ht - -- turn `HasDerivAt` into comp of `HasMFDerivAt` - rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - -- finagle to use `HasMFDerivAt.comp` on `hasMFDerivAt_extChartAt` and `this` - -- see, `htsrc` and `htγ` should be packed away in mono lemmas elsewhere - have htsrc : t ∈ Ioo (t₀ - εsrc) (t₀ + εsrc) := by - rw [← Real.ball_eq_Ioo] - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball - apply le_trans (min_le_right _ _) - exact min_le_left _ _ - have htγ : t ∈ Ioo (t₀ - εγ) (t₀ + εγ) := by - rw [← Real.ball_eq_Ioo] - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball - apply le_trans (min_le_right _ _) - exact min_le_right _ _ - have := (hasMFDerivAt_extChartAt I (x := γ t₀) (y := γ t) (by - rw [← extChartAt_source I] - apply hsrc - rw [Real.ball_eq_Ioo] - exact htsrc - )).comp t (hγ t htγ) - -- can we generalise the following to some lemmas? - have h2 : ContinuousLinearMap.comp - (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) - (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = - ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, - ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] - congr - rw [tangentCoordChange_def, mfderiv] - have hdiff : MDifferentiableAt I I (↑(chartAt H (γ t₀))) (γ t) := by - apply mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) - rw [← extChartAt_source I] - apply hsrc - apply mem_of_mem_of_subset ht - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - simp only [hdiff, if_true] - rfl - rw [← h2] - exact this - · apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) - -- see, these are again mono - · apply hγ.continuousOn.mono - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - · apply MapsTo.mono_left hsrc - rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball - simp - --- what would an api look like to synthesise this by composing individual mono lemmas? -lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono {se : Set E} {ε ε' : ℝ} (hle : ε' ≤ ε) - (h : (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ - (∀ x ∈ Ioo (t₀ - ε) (t₀ + ε), γ x ∈ (extChartAt I (γ t₀)).source) ∧ - (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), - HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ - ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) ) : - (∀ x ∈ Ioo (t₀ - ε') (t₀ + ε'), ((extChartAt I (γ t₀)) ∘ γ) x ∈ se) ∧ - (∀ x ∈ Ioo (t₀ - ε') (t₀ + ε'), γ x ∈ (extChartAt I (γ t₀)).source) ∧ - (∀ t ∈ Ioo (t₀ - ε') (t₀ + ε'), - HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ - ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε') (t₀ + ε')) := by - obtain ⟨h1, h2, h3, h4⟩ := h - -- all of these should be standard mono lemmas - refine ⟨?_, ?_, ?_, ?_⟩ - · intros t ht - apply h1 t - rw [← Real.ball_eq_Ioo] - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball hle - · intros t ht - apply h2 t - rw [← Real.ball_eq_Ioo] - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball hle - · intros t ht - apply h3 t - rw [← Real.ball_eq_Ioo] - rw [← Real.ball_eq_Ioo] at ht - revert ht - apply Metric.ball_subset_ball hle - · apply h4.mono + + -- all this shrinkable stuff could be automatically synthesised, I feel + have hs1 : + Shrinkable (fun ε => (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) t ∈ se)) := + shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) + have hs2 : + Shrinkable (fun ε => ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), γ t ∈ (extChartAt I (γ t₀)).source) := + shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) + have hs3 : + Shrinkable (fun ε => ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) := + shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) + have hs4 : + Shrinkable (fun ε => ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε))) := by + intros ε ε' _ hle h + apply h.mono rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball hle + exact Metric.ball_subset_ball hle + have hs5 : -- this is a repeat + Shrinkable (fun ε => ContinuousOn γ (Ioo (t₀ - ε) (t₀ + ε))) := by + intros ε ε' _ hle h + apply h.mono + rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] + exact Metric.ball_subset_ball hle + + apply shrinkable_exists_and hs1 (shrinkable_and hs2 (shrinkable_and hs3 hs4)) + · refine ⟨εmem, hεmem, ?_⟩ + simp_rw [← Real.ball_eq_Ioo] + exact hmem + · apply shrinkable_exists_and hs2 (shrinkable_and hs3 hs4) + · refine ⟨εsrc, hεsrc, ?_⟩ + simp_rw [← Real.ball_eq_Ioo] + exact hsrc + · refine ⟨min εsrc εγ, lt_min hεsrc hεγ, ?_⟩ -- these two could use shrinkable lemmas + constructor + · intros _ ht + apply hγ.hasDerivAt + · revert ht + simp_rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball (min_le_right _ _) + · apply hsrc + revert ht + simp_rw [← Real.ball_eq_Ioo] + apply Metric.ball_subset_ball (min_le_left _ _) + · apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) + · exact hs5 (lt_min hεsrc hεγ) (min_le_right _ _) hγ.continuousOn + · rw [Real.ball_eq_Ioo] at hsrc + apply hs2 (lt_min hεsrc hεγ) (min_le_left _ _) hsrc /-- Local integral curves are unique. @@ -495,22 +508,29 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - obtain ⟨ε, hε, haux⟩ := + obtain ⟨ε, hε, hmem, hsrc, hγ, hcont⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ hs - obtain ⟨ε', hε', haux'⟩ := + obtain ⟨ε', hε', hmem', hsrc', hγ', hcont'⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ' (h ▸ hs) - have hpos : 0 < min ε ε' := lt_min_iff.mpr ⟨hε, hε'⟩ - obtain ⟨hmem, hsrc, hγ, hcont⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono - I t₀ (min_le_left _ _) haux - obtain ⟨hmem', hsrc', hγ', hcont'⟩ := isIntegralCurveAt_eqOn_of_contMDiffAt_aux_mono - I t₀ (min_le_right _ _) haux' + + have hpos := lt_min hε hε' refine ⟨min ε ε', hpos, ?_⟩ + -- some tactic here where I can just say "shrink ε in hcont to min ε ε'", + -- and it auto generates the goals + have hmem := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hmem + have hmem' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hmem' + have hsrc := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hsrc + have hsrc' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hsrc' + have hγ := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hγ + have hγ' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hγ' + have hcont := shrinkable_continuousOn _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hcont + have hcont' := shrinkable_continuousOn _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hcont' + have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') (Ioo (t₀ - (min ε ε')) (t₀ + (min ε ε'))) := by - apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) _ hcont _ hmem hcont' _ hmem' - · simp [h] - · rw [← Real.ball_eq_Ioo] - exact Metric.mem_ball_self hpos + + apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) + (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hpos) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht rw [hv'] apply (hγ t ht).hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? @@ -525,6 +545,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t rw [Function.comp_apply, LocalEquiv.left_inv] exact hsrc' t ht rw [this] + refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) · intros t ht rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv _ (hsrc _ ht)] From 5e45ca213d5e1767bbd131650209b16e5697b702 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 21:35:32 -0800 Subject: [PATCH 128/203] use filters, eventually --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 94 ++++++++------------ 1 file changed, 38 insertions(+), 56 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 8fd2751f8b072..f7398c2c2ee6d 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -341,8 +341,10 @@ variable (I) section shrinkable +/-- -/ def Shrinkable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε') (_ : ε' ≤ ε) (_ : p ε), p ε' +/-- -/ def Growable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε) (_ : ε ≤ ε') (_ : p ε), p ε' lemma shrinkable_min {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) {εp εq : ℝ} @@ -430,66 +432,46 @@ lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt rw [continuousAt_def] at hcont have hnhds := hcont _ hse - rw [Metric.mem_nhds_iff] at hnhds - obtain ⟨εmem, hεmem, hmem⟩ := hnhds - simp_rw [subset_def, mem_preimage] at hmem + rw [← eventually_mem_nhds] at hnhds -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` - have := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) - rw [Metric.mem_nhds_iff] at this - obtain ⟨εsrc, hεsrc, hsrc⟩ := this - simp_rw [subset_def, mem_preimage] at hsrc + have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) + rw [← eventually_mem_nhds] at hsrc -- extract `εγ` from local existence of integral curve - obtain ⟨εγ, hεγ, hγ⟩ := hγ - - -- all this shrinkable stuff could be automatically synthesised, I feel - have hs1 : - Shrinkable (fun ε => (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) t ∈ se)) := - shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) - have hs2 : - Shrinkable (fun ε => ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), γ t ∈ (extChartAt I (γ t₀)).source) := - shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) - have hs3 : - Shrinkable (fun ε => ∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) := - shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) - have hs4 : - Shrinkable (fun ε => ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε))) := by - intros ε ε' _ hle h - apply h.mono - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - exact Metric.ball_subset_ball hle - have hs5 : -- this is a repeat - Shrinkable (fun ε => ContinuousOn γ (Ioo (t₀ - ε) (t₀ + ε))) := by - intros ε ε' _ hle h - apply h.mono - rw [← Real.ball_eq_Ioo, ← Real.ball_eq_Ioo] - exact Metric.ball_subset_ball hle - - apply shrinkable_exists_and hs1 (shrinkable_and hs2 (shrinkable_and hs3 hs4)) - · refine ⟨εmem, hεmem, ?_⟩ - simp_rw [← Real.ball_eq_Ioo] - exact hmem - · apply shrinkable_exists_and hs2 (shrinkable_and hs3 hs4) - · refine ⟨εsrc, hεsrc, ?_⟩ - simp_rw [← Real.ball_eq_Ioo] - exact hsrc - · refine ⟨min εsrc εγ, lt_min hεsrc hεγ, ?_⟩ -- these two could use shrinkable lemmas - constructor - · intros _ ht - apply hγ.hasDerivAt - · revert ht - simp_rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball (min_le_right _ _) - · apply hsrc - revert ht - simp_rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball (min_le_left _ _) - · apply ContinuousOn.comp (continuousOn_extChartAt I (γ t₀)) - · exact hs5 (lt_min hεsrc hεγ) (min_le_right _ _) hγ.continuousOn - · rw [Real.ball_eq_Ioo] at hsrc - apply hs2 (lt_min hεsrc hεγ) (min_le_left _ _) hsrc + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hγ + + obtain ⟨ε, hε, h⟩ := Metric.eventually_nhds_iff_ball.mp ((hnhds.and hsrc).and hγ) + + have h1 := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (h t ht).1.1 + have h2 := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (h t ht).1.2 + have h3 := fun t ht => (h t ht).2 + + simp_rw [← Real.ball_eq_Ioo] + refine ⟨ε, hε, h1, h2, ?_, ?_⟩ + · intros t ht + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + have hsub : ContinuousLinearMap.comp + (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) + (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = + ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, + ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] + congr + have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas (γ t₀)) + (extChartAt_source I (γ t₀) ▸ h2 t ht) + rw [tangentCoordChange_def, mfderiv, if_pos this] + rfl + rw [← hsub] + apply HasMFDerivAt.comp t _ (h3 t ht) + apply hasMFDerivAt_extChartAt + rw [← extChartAt_source I] + exact h2 t ht + · exact (continuousOn_extChartAt I (γ t₀)).comp (IsIntegralCurveOn.continuousOn h3) h2 /-- Local integral curves are unique. From 842284d9f760ce9fa6e27826138f8c9220976ccf Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 22:53:14 -0800 Subject: [PATCH 129/203] all using filters now --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 183 ++++++------------- 1 file changed, 56 insertions(+), 127 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index f7398c2c2ee6d..6472aa68efbe3 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -339,57 +339,6 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] variable (I) -section shrinkable - -/-- -/ -def Shrinkable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε') (_ : ε' ≤ ε) (_ : p ε), p ε' - -/-- -/ -def Growable (p : ℝ → Prop) : Prop := ∀ {ε ε' : ℝ} (_ : 0 < ε) (_ : ε ≤ ε') (_ : p ε), p ε' - -lemma shrinkable_min {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) {εp εq : ℝ} - (hεp : 0 < εp) (hεq : 0 < εq) (hp : p εp) (hq : q εq) : - p (min εp εq) ∧ q (min εp εq) := - ⟨hsp (lt_min hεp hεq) (min_le_left _ _) hp, hsq (lt_min hεp hεq) (min_le_right _ _) hq⟩ - -lemma shrinkable_exists_and {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) - (hp : ∃ ε > 0, p ε) (hq : ∃ ε > 0, q ε) : ∃ ε > 0, p ε ∧ q ε := by - obtain ⟨εp, hεp, hp⟩ := hp - obtain ⟨εq, hεq, hq⟩ := hq - exact ⟨min εp εq, lt_min hεp hεq, shrinkable_min hsp hsq hεp hεq hp hq⟩ - -lemma shrinkable_and {p q : ℝ → Prop} (hsp : Shrinkable p) (hsq : Shrinkable q) : - Shrinkable (fun ε => p ε ∧ q ε) := by - intros ε ε' hpos hle h - exact ⟨hsp hpos hle h.1, hsq hpos hle h.2⟩ - -lemma shrinkable_forall_mem_set {α : Type*} {s : ℝ → Set α} (p : α → Prop) - (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => ∀ x ∈ s ε, p x) := by - intros ε ε' hpos hle hp x hx - apply hp - exact hs x hpos hle hx - -lemma shrinkable_eqOn {α : Type*} {s : ℝ → Set ℝ} (f g : ℝ → α) - (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => EqOn f g (s ε)) := by - intros ε ε' hpos hle hp - apply hp.mono - intros x hx - exact hs _ hpos hle hx - -lemma shrinkable_continuousOn {α : Type*} [TopologicalSpace α] {s : ℝ → Set ℝ} (f : ℝ → α) - (hs : ∀ x, Growable fun ε => x ∈ s ε) : Shrinkable (fun ε => ContinuousOn f (s ε)) := by - intros ε ε' hpos hle hp - apply hp.mono - intros x hx - exact hs _ hpos hle hx - -lemma growable_mem_Ioo (x : ℝ) : Growable fun ε ↦ x ∈ Ioo (t₀ - ε) (t₀ + ε) := by - intros ε ε' _ hle - simp_rw [← Real.ball_eq_Ioo] - apply Metric.ball_subset_ball hle - -end shrinkable - lemma IsIntegralCurveOn.hasDerivAt {ε : ℝ} (hγ : IsIntegralCurveOn γ v (Ioo (t₀ - ε) (t₀ + ε))) {t : ℝ} (ht : t ∈ Ioo (t₀ - ε) (t₀ + ε)) (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : @@ -418,61 +367,6 @@ lemma IsIntegralCurveOn.hasDerivAt {ε : ℝ} rw [← extChartAt_source I] exact hsrc -lemma isIntegralCurveAt_eqOn_of_contMDiffAt_aux - (hγ : IsIntegralCurveAt γ v t₀) {se : Set E} (hse : se ∈ nhds (extChartAt I (γ t₀) (γ t₀))) : - ∃ ε > 0, (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), ((extChartAt I (γ t₀)) ∘ γ) t ∈ se) ∧ - (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), γ t ∈ (extChartAt I (γ t₀)).source) ∧ - (∀ t ∈ Ioo (t₀ - ε) (t₀ + ε), - HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t) ∧ - ContinuousOn ((extChartAt I (γ t₀)) ∘ γ) (Ioo (t₀ - ε) (t₀ + ε)) := by - - -- extract `εmem` so `γ t` when expressed in the chart stays within `se` - have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := - ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt - rw [continuousAt_def] at hcont - have hnhds := hcont _ hse - rw [← eventually_mem_nhds] at hnhds - - -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` - have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) - rw [← eventually_mem_nhds] at hsrc - - -- extract `εγ` from local existence of integral curve - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hγ - - obtain ⟨ε, hε, h⟩ := Metric.eventually_nhds_iff_ball.mp ((hnhds.and hsrc).and hγ) - - have h1 := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (h t ht).1.1 - have h2 := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (h t ht).1.2 - have h3 := fun t ht => (h t ht).2 - - simp_rw [← Real.ball_eq_Ioo] - refine ⟨ε, hε, h1, h2, ?_, ?_⟩ - · intros t ht - rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - have hsub : ContinuousLinearMap.comp - (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) - (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = - ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, - ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] - congr - have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas (γ t₀)) - (extChartAt_source I (γ t₀) ▸ h2 t ht) - rw [tangentCoordChange_def, mfderiv, if_pos this] - rfl - rw [← hsub] - apply HasMFDerivAt.comp t _ (h3 t ht) - apply hasMFDerivAt_extChartAt - rw [← extChartAt_source I] - exact h2 t ht - · exact (continuousOn_extChartAt I (γ t₀)).comp (IsIntegralCurveOn.continuousOn h3) h2 - /-- Local integral curves are unique. If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` @@ -490,39 +384,74 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - obtain ⟨ε, hε, hmem, hsrc, hγ, hcont⟩ := - isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ hs - obtain ⟨ε', hε', hmem', hsrc', hγ', hcont'⟩ := - isIntegralCurveAt_eqOn_of_contMDiffAt_aux I t₀ hγ' (h ▸ hs) - - have hpos := lt_min hε hε' - refine ⟨min ε ε', hpos, ?_⟩ - -- some tactic here where I can just say "shrink ε in hcont to min ε ε'", - -- and it auto generates the goals - have hmem := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hmem - have hmem' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hmem' - have hsrc := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hsrc - have hsrc' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hsrc' - have hγ := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hγ - have hγ' := shrinkable_forall_mem_set _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hγ' - have hcont := shrinkable_continuousOn _ (growable_mem_Ioo t₀) hpos (min_le_left _ _) hcont - have hcont' := shrinkable_continuousOn _ (growable_mem_Ioo t₀) hpos (min_le_right _ _) hcont' + + -- extract `εmem` so `γ t` when expressed in the chart stays within `se` + have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := + ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt + rw [continuousAt_def] at hcont + have hnhds := hcont _ hs + rw [← eventually_mem_nhds] at hnhds + + -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` + have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) + rw [← eventually_mem_nhds] at hsrc + + -- extract `εγ` from local existence of integral curve + have hmfd := hγ + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hmfd + + -- extract `εmem` so `γ t` when expressed in the chart stays within `se` + have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := + ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt + rw [continuousAt_def] at hcont' + have hnhds' := hcont' _ (h ▸ hs) + rw [← eventually_mem_nhds] at hnhds' + + -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` + have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) + rw [← eventually_mem_nhds] at hsrc' + + -- extract `εγ` from local existence of integral curve + have hmfd' := hγ' + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hmfd' + + have haux := hnhds.and <| hsrc.and <| hmfd.and <| hnhds'.and <| hsrc'.and hmfd' + rw [Metric.eventually_nhds_iff_ball] at haux + + obtain ⟨ε, hε, haux⟩ := haux + refine ⟨ε, hε, ?_⟩ + + have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 + have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 + have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 + have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 + have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 + have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 + + have hcont := (continuousOn_extChartAt I (γ t₀)).comp + (IsIntegralCurveOn.continuousOn hmfd) hsrc + have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp + (IsIntegralCurveOn.continuousOn hmfd') hsrc' + + simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') - (Ioo (t₀ - (min ε ε')) (t₀ + (min ε ε'))) := by + (Ioo (t₀ - ε) (t₀ + ε)) := by apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) - (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hpos) hcont _ hmem hcont' _ hmem' (by simp [h]) + (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht rw [hv'] - apply (hγ t ht).hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? + have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) + apply this.hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, LocalEquiv.left_inv] exact hsrc t ht rw [this] · intros t ht rw [hv', h] - apply (hγ' t ht).hasFDerivAt.congr_fderiv + have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) + apply this.hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, LocalEquiv.left_inv] exact hsrc' t ht From abbcbd8cb43e723038e88414ea5671bf1e9f2098 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 22:57:16 -0800 Subject: [PATCH 130/203] style --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 6472aa68efbe3..f1ad061acc23d 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -398,7 +398,8 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t -- extract `εγ` from local existence of integral curve have hmfd := hγ - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hmfd + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, + ← Metric.eventually_nhds_iff_ball] at hmfd -- extract `εmem` so `γ t` when expressed in the chart stays within `se` have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := @@ -413,7 +414,8 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t -- extract `εγ` from local existence of integral curve have hmfd' := hγ' - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hmfd' + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, + ← Metric.eventually_nhds_iff_ball] at hmfd' have haux := hnhds.and <| hsrc.and <| hmfd.and <| hnhds'.and <| hsrc'.and hmfd' rw [Metric.eventually_nhds_iff_ball] at haux @@ -451,7 +453,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [hv', h] have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) - apply this.hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? + apply this.hasFDerivAt.congr_fderiv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, LocalEquiv.left_inv] exact hsrc' t ht From 8d914203d02fb7cf51bc65ecd74e2a8497fef1e7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 13 Dec 2023 23:27:19 -0800 Subject: [PATCH 131/203] doc --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 36 +++++++++++--------- 1 file changed, 20 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index f1ad061acc23d..2425ce8087a25 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -339,8 +339,7 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] variable (I) -lemma IsIntegralCurveOn.hasDerivAt {ε : ℝ} - (hγ : IsIntegralCurveOn γ v (Ioo (t₀ - ε) (t₀ + ε))) {t : ℝ} (ht : t ∈ Ioo (t₀ - ε) (t₀ + ε)) +lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by @@ -375,54 +374,55 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by - -- extract set `s` on which `v` is Lipschitz + -- first define `v'` as the vector field expressed in the local chart around `γ t₀` + -- this is basically what the function looks like when `hv` is unfolded set v' : E → E := fun x => tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) (v ((extChartAt I (γ t₀)).symm x)) with hv' + + -- extract set `s` on which `v'` is Lipschitz rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - -- extract `εmem` so `γ t` when expressed in the chart stays within `se` + -- `γ t` when expressed in the local chart should remain inside `s` have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := - ContinuousAt.comp (continuousAt_extChartAt ..) hγ.continuousAt + (continuousAt_extChartAt ..).comp hγ.continuousAt rw [continuousAt_def] at hcont have hnhds := hcont _ hs rw [← eventually_mem_nhds] at hnhds - -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` + -- `γ t` should remain inside the domain of the local chart around `γ t₀` have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) rw [← eventually_mem_nhds] at hsrc - -- extract `εγ` from local existence of integral curve - have hmfd := hγ + -- `γ` is tangent to `v` in some neighbourhood of `t₀` simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, - ← Metric.eventually_nhds_iff_ball] at hmfd + ← Metric.eventually_nhds_iff_ball] at hγ - -- extract `εmem` so `γ t` when expressed in the chart stays within `se` + -- same as above but for `γ'` have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt rw [continuousAt_def] at hcont' have hnhds' := hcont' _ (h ▸ hs) rw [← eventually_mem_nhds] at hnhds' - -- extract `εsrc` so `γ t` stays within the interior of the chart around `γ t₀` have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) rw [← eventually_mem_nhds] at hsrc' - -- extract `εγ` from local existence of integral curve - have hmfd' := hγ' simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Real.ball_eq_Ioo, - ← Metric.eventually_nhds_iff_ball] at hmfd' + ← Metric.eventually_nhds_iff_ball] at hγ' - have haux := hnhds.and <| hsrc.and <| hmfd.and <| hnhds'.and <| hsrc'.and hmfd' + -- there exists a neighbourhood around `t₀` in which all of the above hold + have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' rw [Metric.eventually_nhds_iff_ball] at haux obtain ⟨ε, hε, haux⟩ := haux refine ⟨ε, hε, ?_⟩ + -- break out all the conditions again have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 @@ -430,16 +430,19 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 + -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood have hcont := (continuousOn_extChartAt I (γ t₀)).comp (IsIntegralCurveOn.continuousOn hmfd) hsrc have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp (IsIntegralCurveOn.continuousOn hmfd') hsrc' + -- todo: make up your mind whether to use `ball` or `Ioo` simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' + -- `γ` and `γ'` are have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') (Ioo (t₀ - ε) (t₀ + ε)) := by - + -- uniqueness of ODE solutions in an open interval apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht @@ -459,6 +462,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t exact hsrc' t ht rw [this] + -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) · intros t ht rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv _ (hsrc _ ht)] From 2c4d77d7fd709133772709da9ffb1d698ecbcd0a Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 02:14:07 -0800 Subject: [PATCH 132/203] restate existence theorem using filters --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 26 ++++++++++---------- 1 file changed, 13 insertions(+), 13 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 01d75ee8521a9..7af140175f7ec 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -257,36 +257,36 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by - -- express the differentiability of the section `v` in the local charts + -- express the differentiability of the vector field `v` in the local chart rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart - obtain ⟨f, hf1, ε1, hε1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd - rw [← Real.ball_eq_Ioo] at hf2 - -- use continuity of `f` to extract `ε2` so that for `t ∈ Real.ball t₀ ε2`, - -- `f t ∈ interior (extChartAt I x₀).target` - have hcont := (hf2 t₀ (Metric.mem_ball_self hε1)).continuousAt + obtain ⟨f, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd + simp_rw [← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hf2 + -- use continuity of `f` so that `f t` remains inside `interior (extChartAt I x₀).target` + have ⟨a, ha, hf2'⟩ := Metric.eventually_nhds_iff_ball.mp hf2 + have hcont := (hf2' t₀ (Metric.mem_ball_self ha)).continuousAt rw [continuousAt_def, hf1] at hcont have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := hcont _ (isOpen_interior.mem_nhds ((I.isInteriorPoint_iff).mp hx)) - rw [Metric.mem_nhds_iff] at hnhds - obtain ⟨ε2, hε2, hf3⟩ := hnhds - simp_rw [subset_def, mem_preimage] at hf3 + rw [← eventually_mem_nhds] at hnhds + -- obtain an `ε` so that the above conditions both hold in an `ε`-neighbourhood of `t₀` + obtain ⟨ε, hε, haux⟩ := Metric.eventually_nhds_iff_ball.mp (hf2.and hnhds) -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve refine ⟨(extChartAt I x₀).symm ∘ f, Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), - min ε1 ε2, lt_min hε1 hε2, ?_⟩ + ε, hε, ?_⟩ intros t ht -- collect useful terms in convenient forms rw [← Real.ball_eq_Ioo] at ht - have hf3 := hf3 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_right ..)) have h : HasDerivAt f ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) (v ((extChartAt I x₀).symm (f t)))) - t := hf2 t <| mem_of_mem_of_subset ht (Metric.ball_subset_ball (min_le_left ..)) + t := (haux t ht).1 rw [← tangentCoordChange_def] at h + have hf3 := mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2 have hf3' := mem_of_mem_of_subset hf3 interior_subset have hft1 := mem_preimage.mp <| mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source From 5f7bb14386b1bac046afaaa61fa913fe0969bac3 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:05:37 -0800 Subject: [PATCH 133/203] define `IsIntegralCurveAt` using `Eventually` --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 57 +++++++++++--------- 1 file changed, 31 insertions(+), 26 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7af140175f7ec..b041477ce0983 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -76,7 +76,7 @@ def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : S `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. The value of `γ` outside of this interval is irrelevant and considered junk. -/ def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) := - ∃ ε > (0 : ℝ), IsIntegralCurveOn γ v (Ioo (t - ε) (t + ε)) + ∃ s ∈ nhds t, IsIntegralCurveOn γ v s /-- If `v : M → TM` is a vector field on `M`, `IsIntegralCurve γ v` means `γ : ℝ → M` is a global integral curve of `v`. That is, `γ t` is tangent to `v (γ t)` for all `t : ℝ`. -/ @@ -93,13 +93,13 @@ lemma isIntegralCurve_iff_isIntegralCurveOn : ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : - IsIntegralCurveAt γ v t := ⟨1, zero_lt_one, fun t _ => h t⟩ + IsIntegralCurveAt γ v t := ⟨univ, Filter.univ_mem, fun t _ => h t⟩ lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := ⟨fun h => h.isIntegralCurveAt, fun h t => by - obtain ⟨ε, hε, h⟩ := h t - exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε)⟩ + obtain ⟨s, hs, h⟩ := h t + exact h t (mem_of_mem_nhds hs)⟩ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : IsIntegralCurveOn γ v s' := fun t ht => h t (mem_of_mem_of_subset ht hs) @@ -112,21 +112,30 @@ lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegral · exact h _ ht · exact h' _ ht +lemma isIntegralCurveAt_iff : + IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by + constructor + · intro h + obtain ⟨s, hs, h⟩ := h + obtain ⟨ε, hε, hsub⟩ := Metric.mem_nhds_iff.mp hs + exact ⟨ε, hε, h.mono hsub⟩ + · intro h + obtain ⟨ε, hε, h⟩ := h + refine ⟨Metric.ball t₀ ε, Metric.ball_mem_nhds _ hε, h⟩ + lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : - IsIntegralCurveAt γ v t₀ := by - rw [Metric.mem_nhds_iff] at hs - obtain ⟨ε, hε, hmem⟩ := hs - exact ⟨ε, hε, Real.ball_eq_Ioo _ _ ▸ h.mono hmem⟩ + IsIntegralCurveAt γ v t₀ := ⟨s, hs, h⟩ lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by intros t ht - obtain ⟨ε, hε, h⟩ := h t ht - exact h t (Real.ball_eq_Ioo _ _ ▸ Metric.mem_ball_self hε) - + obtain ⟨s, hs, h⟩ := h t ht + exact h t (mem_of_mem_nhds hs) /-! ### Translation lemmas -/ + section Translation + lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by intros t ht @@ -148,12 +157,12 @@ lemma isIntegralCurveOn_comp_add {dt : ℝ} : lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by + rw [isIntegralCurveAt_iff] at hγ obtain ⟨ε, hε, h⟩ := hγ - refine ⟨ε, hε, ?_⟩ + refine ⟨Metric.ball (t₀ - dt) ε, Metric.isOpen_ball.mem_nhds (Metric.mem_ball_self hε), ?_⟩ convert h.comp_add dt - ext - rw [sub_right_comm, sub_add_eq_add_sub, ← add_mem_Ioo_iff_left] - rfl + ext t + rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, dist_sub_eq_dist_add_right] lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -202,16 +211,13 @@ lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - obtain ⟨ε, hε, h⟩ := hγ + obtain ⟨ε, hε, h⟩ := isIntegralCurveAt_iff.mp hγ + rw [isIntegralCurveAt_iff] refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ convert h.comp_mul a ext t - rw [Ioo, Ioo, mem_setOf_eq, mem_setOf_eq, mem_setOf_eq] - by_cases ha' : 0 < a - · rw [abs_eq_self.mpr (le_of_lt ha'), ← sub_div, ← add_div, div_lt_iff ha', lt_div_iff ha'] - · rw [abs_eq_neg_self.mpr (not_lt.mp ha'), div_neg, sub_neg_eq_add, ← sub_eq_add_neg, ← sub_div, - ← add_div, div_lt_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), - lt_div_iff_of_neg (ha.lt_of_le (not_lt.mp ha')), and_comm] + rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, Real.dist_eq, Real.dist_eq, + lt_div_iff (abs_pos.mpr ha), ← abs_mul, sub_mul, div_mul_cancel _ ha] lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by @@ -271,15 +277,14 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := hcont _ (isOpen_interior.mem_nhds ((I.isInteriorPoint_iff).mp hx)) rw [← eventually_mem_nhds] at hnhds - -- obtain an `ε` so that the above conditions both hold in an `ε`-neighbourhood of `t₀` - obtain ⟨ε, hε, haux⟩ := Metric.eventually_nhds_iff_ball.mp (hf2.and hnhds) + -- obtain a neighbourhood `s` so that the above conditions both hold in `s` + obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve refine ⟨(extChartAt I x₀).symm ∘ f, Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), - ε, hε, ?_⟩ + s, hs, ?_⟩ intros t ht -- collect useful terms in convenient forms - rw [← Real.ball_eq_Ioo] at ht have h : HasDerivAt f ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) From 63410724b9a1813f9b791b13a4490a666d2df7ec Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:40:18 -0800 Subject: [PATCH 134/203] PartialEquiv --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b041477ce0983..b1934c039d9ae 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -281,7 +281,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve refine ⟨(extChartAt I x₀).symm ∘ f, - Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), + Eq.symm (by rw [Function.comp_apply, hf1, PartialEquiv.left_inv _ (mem_extChartAt_source ..)]), s, hs, ?_⟩ intros t ht -- collect useful terms in convenient forms From 79111219c232fd6296066980b4bcd70b88b2bfc4 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:41:01 -0800 Subject: [PATCH 135/203] PartialEquiv --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index feaea1d7db03e..b68106d383501 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -298,7 +298,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve refine ⟨(extChartAt I x₀).symm ∘ f, - Eq.symm (by rw [Function.comp_apply, hf1, LocalEquiv.left_inv _ (mem_extChartAt_source ..)]), + Eq.symm (by rw [Function.comp_apply, hf1, PartialEquiv.left_inv _ (mem_extChartAt_source ..)]), s, hs, ?_⟩ intros t ht -- collect useful terms in convenient forms @@ -452,7 +452,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) apply this.hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by - rw [Function.comp_apply, LocalEquiv.left_inv] + rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc t ht rw [this] · intros t ht @@ -460,16 +460,16 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) apply this.hasFDerivAt.congr_fderiv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by - rw [Function.comp_apply, LocalEquiv.left_inv] + rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc' t ht rw [this] -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) · intros t ht - rw [Function.comp_apply, Function.comp_apply, LocalEquiv.left_inv _ (hsrc _ ht)] + rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] · intros t ht - rw [Function.comp_apply, Function.comp_apply, h, LocalEquiv.left_inv _ (hsrc' _ ht)] + rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] /-- Integral curves are unique on open intervals. From 8e5421220216bbe805201e807e9edbb912fe9c96 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:42:19 -0800 Subject: [PATCH 136/203] PartialEquiv --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 634f4a948bcce..3a97c37742087 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -116,7 +116,7 @@ lemma ModelWithCorners.isInteriorPoint {x : M} : I.IsInteriorPoint x := by have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r rw [ModelWithCorners.isInteriorPoint_iff, r] - exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) + exact PartialEquiv.map_source _ (mem_extChartAt_source _ _) /-- If `I` is boundaryless, `M` has full interior interior. -/ lemma ModelWithCorners.interior_eq_univ : SmoothManifoldWithCorners.interior I M = univ := by From 7b1bc43aac102ce3db5d97d4e95becdc551f5cbb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:42:39 -0800 Subject: [PATCH 137/203] PartialEquiv --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 19fbaf1db2452..2b868d7d0a6d1 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -118,7 +118,7 @@ lemma ModelWithCorners.isInteriorPoint {x : M} : I.IsInteriorPoint x := by have : extChartAt I x = (chartAt H x).extend I := rfl rw [← this] at r rw [ModelWithCorners.isInteriorPoint_iff, r] - exact LocalEquiv.map_source _ (mem_extChartAt_source _ _) + exact PartialEquiv.map_source _ (mem_extChartAt_source _ _) /-- If `I` is boundaryless, `M` has full interior interior. -/ lemma ModelWithCorners.interior_eq_univ : SmoothManifoldWithCorners.interior I M = univ := by From 4efc1ad6f26d8c537618aa50decc6b3b8c441b4b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:46:36 -0800 Subject: [PATCH 138/203] PartialEquiv --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 2 +- Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 3a97c37742087..cf2ce7f54933b 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -62,7 +62,7 @@ protected def interior : Set M := { x : M | I.IsInteriorPoint x } lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := ⟨fun h ↦ (chartAt H x).mem_interior_extend_target _ (mem_chart_target H x) h, - fun h ↦ LocalHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ + fun h ↦ PartialHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 7fed865cae04e..06777179b0a05 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -154,7 +154,7 @@ lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) (range I) (extChartAt I x z) := have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by - rw [LocalEquiv.trans_source'', LocalEquiv.symm_symm, LocalEquiv.symm_target] + rw [PartialEquiv.trans_source'', PartialEquiv.symm_symm, PartialEquiv.symm_target] exact mem_image_of_mem _ h ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt From 2508953416c3ade8d9316dabb2b677e1e3ba8ad2 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 03:47:10 -0800 Subject: [PATCH 139/203] Partial --- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 2 +- Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 2b868d7d0a6d1..a0b67b6d3657b 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -64,7 +64,7 @@ protected def interior : Set M := { x : M | I.IsInteriorPoint x } lemma _root_.ModelWithCorners.isInteriorPoint_iff {x : M} : I.IsInteriorPoint x ↔ extChartAt I x x ∈ interior (extChartAt I x).target := ⟨fun h ↦ (chartAt H x).mem_interior_extend_target _ (mem_chart_target H x) h, - fun h ↦ LocalHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ + fun h ↦ PartialHomeomorph.interior_extend_target_subset_interior_range _ _ h⟩ variable (I M) in /-- The **boundary** of a manifold `M` is the set of its boundary points. -/ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 7fed865cae04e..06777179b0a05 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -154,7 +154,7 @@ lemma hasFDerivWithinAt_tangentCoordChange {x y z : M} HasFDerivWithinAt ((extChartAt I y) ∘ (extChartAt I x).symm) (tangentCoordChange I x y z) (range I) (extChartAt I x z) := have h' : extChartAt I x z ∈ ((extChartAt I x).symm ≫ (extChartAt I y)).source := by - rw [LocalEquiv.trans_source'', LocalEquiv.symm_symm, LocalEquiv.symm_target] + rw [PartialEquiv.trans_source'', PartialEquiv.symm_symm, PartialEquiv.symm_target] exact mem_image_of_mem _ h ((contDiffWithinAt_ext_coord_change I y x h').differentiableWithinAt (by simp)).hasFDerivWithinAt From 7b804f9346604c8e82863aed0603eb7cf0fca127 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Thu, 14 Dec 2023 14:18:06 +0100 Subject: [PATCH 140/203] Replace => by \mapsto, per the style guide. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 28 ++++++++++---------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b1934c039d9ae..17c28a260003e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -86,23 +86,23 @@ def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) : - IsIntegralCurveOn γ v s := fun t _ => h t + IsIntegralCurveOn γ v s := fun t _ ↦ h t lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := - ⟨fun h => h.isIntegralCurveOn _, fun h t => h t (mem_univ _)⟩ + ⟨fun h ↦ h.isIntegralCurveOn _, fun h t ↦ h t (mem_univ _)⟩ lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : - IsIntegralCurveAt γ v t := ⟨univ, Filter.univ_mem, fun t _ => h t⟩ + IsIntegralCurveAt γ v t := ⟨univ, Filter.univ_mem, fun t _ ↦ h t⟩ lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := - ⟨fun h => h.isIntegralCurveAt, fun h t => by + ⟨fun h ↦ h.isIntegralCurveAt, fun h t ↦ by obtain ⟨s, hs, h⟩ := h t exact h t (mem_of_mem_nhds hs)⟩ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : - IsIntegralCurveOn γ v s' := fun t ht => h t (mem_of_mem_of_subset ht hs) + IsIntegralCurveOn γ v s' := fun t ht ↦ h t (mem_of_mem_of_subset ht hs) lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : IsIntegralCurveOn γ v (s ∪ s') := by @@ -148,7 +148,7 @@ lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : lemma isIntegralCurveOn_comp_add {dt : ℝ} : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by - refine ⟨fun hγ => hγ.comp_add _, fun hγ => ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] at this convert this @@ -166,7 +166,7 @@ lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by - refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ have := hγ.comp_add (-dt) rw [sub_neg_eq_add, sub_add_cancel] at this convert this @@ -180,7 +180,7 @@ lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : lemma isIntegralCurve_comp_add {dt : ℝ} : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· + dt)) v := by - refine ⟨fun hγ => hγ.comp_add _, fun hγ ↦ ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) ext simp only [Function.comp_apply, neg_add_cancel_right] @@ -201,7 +201,7 @@ lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by - refine ⟨fun hγ => hγ.comp_mul a, fun hγ ↦ ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_mul a, fun hγ ↦ ?_⟩ have := hγ.comp_mul a⁻¹ simp_rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, mem_setOf_eq, mul_assoc, inv_mul_eq_div, div_self ha, mul_one, setOf_mem_eq] at this @@ -221,7 +221,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - refine ⟨fun hγ => hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ have := hγ.comp_mul_ne_zero (inv_ne_zero ha) rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this @@ -236,7 +236,7 @@ lemma IsIntegralCurve.comp_mul (hγ : IsIntegralCurve γ v) (a : ℝ) : lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by - refine ⟨fun hγ => hγ.comp_mul _, fun hγ => ?_⟩ + refine ⟨fun hγ ↦ hγ.comp_mul _, fun hγ ↦ ?_⟩ have := hγ.comp_mul a⁻¹ rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul] at this convert this @@ -245,7 +245,7 @@ lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ -lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ => x) v := by +lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ ↦ x) v := by intro t rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), ContinuousLinearMap.smulRight_one_one] @@ -260,7 +260,7 @@ variable (t₀) {x₀ : M} of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`.-/ theorem exists_isIntegralCurveAt_of_contMDiffAt - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) + (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) (hx : I.IsInteriorPoint x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by -- express the differentiability of the vector field `v` in the local chart @@ -320,6 +320,6 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) x₀) : + (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint From 22d5cebc4ca2dbbd909a595e9ff7704ac65fc26c Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Thu, 14 Dec 2023 14:20:59 +0100 Subject: [PATCH 141/203] Small golfs. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 16 +++++----------- 1 file changed, 5 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 17c28a260003e..86b974a590d4b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -88,8 +88,7 @@ variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) : IsIntegralCurveOn γ v s := fun t _ ↦ h t -lemma isIntegralCurve_iff_isIntegralCurveOn : - IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := +lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h ↦ h.isIntegralCurveOn _, fun h t ↦ h t (mem_univ _)⟩ lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : @@ -107,21 +106,16 @@ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : IsIntegralCurveOn γ v (s ∪ s') := by intros t ht - rw [mem_union] at ht cases' ht with ht ht · exact h _ ht · exact h' _ ht lemma isIntegralCurveAt_iff : IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by - constructor - · intro h - obtain ⟨s, hs, h⟩ := h - obtain ⟨ε, hε, hsub⟩ := Metric.mem_nhds_iff.mp hs - exact ⟨ε, hε, h.mono hsub⟩ - · intro h - obtain ⟨ε, hε, h⟩ := h - refine ⟨Metric.ball t₀ ε, Metric.ball_mem_nhds _ hε, h⟩ + refine ⟨?_, fun ⟨ε, hε, h⟩ ↦ ⟨Metric.ball t₀ ε, Metric.ball_mem_nhds _ hε, h⟩⟩ + rintro ⟨s, hs, h⟩ + obtain ⟨ε, hε, hsub⟩ := Metric.mem_nhds_iff.mp hs + exact ⟨ε, hε, h.mono hsub⟩ lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : IsIntegralCurveAt γ v t₀ := ⟨s, hs, h⟩ From c4dd6bc416fadce9f47834a84a83915891f0cea3 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 17:03:41 -0800 Subject: [PATCH 142/203] small golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +------ 1 file changed, 1 insertion(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b68106d383501..21825422c7ccb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -438,7 +438,6 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp (IsIntegralCurveOn.continuousOn hmfd') hsrc' - -- todo: make up your mind whether to use `ball` or `Ioo` simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' -- `γ` and `γ'` are @@ -522,11 +521,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] (Ioo_subset_Ioo (by simp) (by simp)), isOpen_Ioo, ?_⟩ rw [mem_Ioo] - constructor - · apply max_lt ht₁.2.1 - simp [hε] - · apply lt_min ht₁.2.2 - simp [hε] + exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ intros t ht exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 From d32e895b45d5954d96c5e36555db2d2c9803aff0 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 17:12:47 -0800 Subject: [PATCH 143/203] congr_deriv --- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 3b25240ada6f8..2c9c2b6137886 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -614,6 +614,17 @@ theorem Filter.EventuallyEq.hasDerivWithinAt_iff_of_mem (h₁ : f₁ =ᶠ[𝓝[s ⟨fun h' ↦ h'.congr_of_eventuallyEq_of_mem h₁.symm hx, fun h' ↦ h'.congr_of_eventuallyEq_of_mem h₁ hx⟩ +theorem HasStrictDerivAt.congr_deriv (h : HasStrictDerivAt f f' x) (h' : f' = g') : + HasStrictDerivAt f g' x := + h' ▸ h + +theorem HasDerivAt.congr_deriv (h : HasDerivAt f f' x) (h' : f' = g') : HasDerivAt f g' x := + h' ▸ h + +theorem HasDerivWithinAt.congr_deriv (h : HasDerivWithinAt f f' s x) (h' : f' = g') : + HasDerivWithinAt f g' s x := + h' ▸ h + theorem HasDerivAt.congr_of_eventuallyEq (h : HasDerivAt f f' x) (h₁ : f₁ =ᶠ[𝓝 x] f) : HasDerivAt f₁ f' x := HasDerivAtFilter.congr_of_eventuallyEq h h₁ (mem_of_mem_nhds h₁ : _) From fff3d19675eaad3e1f97b4998f4befe38132ac5c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 17:17:20 -0800 Subject: [PATCH 144/203] use new lemmas --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 21825422c7ccb..5269f2d63f0cc 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -449,7 +449,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [hv'] have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) - apply this.hasFDerivAt.congr_fderiv -- missing `hasDerivAt.congr_deriv` ? + apply this.congr_fderiv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc t ht From 4559d572abbf12fe9b6c894ce212cc4b927448f1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 17:17:32 -0800 Subject: [PATCH 145/203] missed one --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5269f2d63f0cc..060ee954c511a 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -457,7 +457,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [hv', h] have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) - apply this.hasFDerivAt.congr_fderiv + apply this.congr_fderiv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc' t ht From 17cb91b9cb0b0bfb3b5cbe1128248f2c5b6ff944 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 17:44:55 -0800 Subject: [PATCH 146/203] fderiv -> deriv --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 060ee954c511a..a56aae678fecc 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -343,6 +343,8 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] variable (I) +/-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when + expressed in the local chart around the initial point `γ t₀`. -/ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) @@ -449,7 +451,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [hv'] have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) - apply this.congr_fderiv + apply this.congr_deriv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc t ht @@ -457,7 +459,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [hv', h] have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) - apply this.congr_fderiv + apply this.congr_deriv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] exact hsrc' t ht From 8712cb899210dbe081b8d53ac29e78bd9a01ddd1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 18:09:42 -0800 Subject: [PATCH 147/203] congr_mfderiv --- Mathlib/Geometry/Manifold/MFDeriv.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Geometry/Manifold/MFDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv.lean index 9b5ba6a0fb688..1c35a31614efb 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv.lean @@ -840,6 +840,14 @@ theorem HasMFDerivWithinAt.congr_mono (h : HasMFDerivWithinAt I I' f s x f') (h.mono h₁).congr_of_eventuallyEq (Filter.mem_inf_of_right ht) hx #align has_mfderiv_within_at.congr_mono HasMFDerivWithinAt.congr_mono +theorem HasMFDerivAt.congr_mfderiv (h : HasMFDerivAt I I' f x f') (h' : f' = f₁') : + HasMFDerivAt I I' f x f₁' := + h' ▸ h + +theorem HasMFDerivWithinAt.congr_mfderiv (h : HasMFDerivWithinAt I I' f s x f') (h' : f' = f₁') : + HasMFDerivWithinAt I I' f s x f₁' := + h' ▸ h + theorem HasMFDerivAt.congr_of_eventuallyEq (h : HasMFDerivAt I I' f x f') (h₁ : f₁ =ᶠ[𝓝 x] f) : HasMFDerivAt I I' f₁ x f' := by rw [← hasMFDerivWithinAt_univ] at h ⊢ From d4c5b052a0d170ead799f34d782587f5657d6cfb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 18:39:47 -0800 Subject: [PATCH 148/203] mfderiv_chartAt_eq_tangentCoordChange --- Mathlib/Geometry/Manifold/MFDeriv.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/Geometry/Manifold/MFDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv.lean index 1c35a31614efb..7f6fc4c88a48a 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv.lean @@ -1895,6 +1895,13 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} exact ((chartAt H (TotalSpace.proj p)).right_inv h).symm #align tangent_map_chart_symm tangentMap_chart_symm +lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : + mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by + have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) + ((extChartAt_source I y) ▸ hsrc) + rw [tangentCoordChange_def, mfderiv, if_pos this] + rfl + end Charts end SpecificFunctions From a5c75c150281d0a5beaeb8ad83fd537cf72af573 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 18:51:10 -0800 Subject: [PATCH 149/203] golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 380 +++++++++---------- 1 file changed, 185 insertions(+), 195 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a56aae678fecc..7239dd089bd6b 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -351,198 +351,188 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by -- turn `HasDerivAt` into comp of `HasMFDerivAt` rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - have hsub : ContinuousLinearMap.comp - (mfderiv I I (↑(chartAt H (γ t₀))) (γ t)) - (ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) (v (γ t))) = - ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) := by - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, - ContinuousLinearMap.one_apply, ContinuousLinearMap.map_smul_of_tower, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply] - congr - have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) - ((extChartAt_source I (γ t₀)) ▸ hsrc) - rw [tangentCoordChange_def, mfderiv, if_pos this] - rfl - rw [← hsub] - apply HasMFDerivAt.comp t _ (hγ _ ht) - apply hasMFDerivAt_extChartAt - rw [← extChartAt_source I] - exact hsrc - -/-- Local integral curves are unique. - - If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` - at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) - (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : - ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by - -- first define `v'` as the vector field expressed in the local chart around `γ t₀` - -- this is basically what the function looks like when `hv` is unfolded - set v' : E → E := fun x => - tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) - (v ((extChartAt I (γ t₀)).symm x)) with hv' - - -- extract set `s` on which `v'` is Lipschitz - rw [contMDiffAt_iff] at hv - obtain ⟨_, hv⟩ := hv - obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd - have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - - -- `γ t` when expressed in the local chart should remain inside `s` - have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := - (continuousAt_extChartAt ..).comp hγ.continuousAt - rw [continuousAt_def] at hcont - have hnhds := hcont _ hs - rw [← eventually_mem_nhds] at hnhds - - -- `γ t` should remain inside the domain of the local chart around `γ t₀` - have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) - rw [← eventually_mem_nhds] at hsrc - - -- `γ` is tangent to `v` in some neighbourhood of `t₀` - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ - - -- same as above but for `γ'` - have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := - ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt - rw [continuousAt_def] at hcont' - have hnhds' := hcont' _ (h ▸ hs) - rw [← eventually_mem_nhds] at hnhds' - - have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) - rw [← eventually_mem_nhds] at hsrc' - - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ' - - -- there exists a neighbourhood around `t₀` in which all of the above hold - have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' - rw [Metric.eventually_nhds_iff_ball] at haux - - obtain ⟨ε, hε, haux⟩ := haux - refine ⟨ε, hε, ?_⟩ - - -- break out all the conditions again - have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 - have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 - have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 - have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 - have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 - have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 - - -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood - have hcont := (continuousOn_extChartAt I (γ t₀)).comp - (IsIntegralCurveOn.continuousOn hmfd) hsrc - have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp - (IsIntegralCurveOn.continuousOn hmfd') hsrc' - - simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' - - -- `γ` and `γ'` are - have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') - (Ioo (t₀ - ε) (t₀ + ε)) := by - -- uniqueness of ODE solutions in an open interval - apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) - (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) - · intros t ht - rw [hv'] - have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) - apply this.congr_deriv - have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hsrc t ht - rw [this] - · intros t ht - rw [hv', h] - have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) - apply this.congr_deriv - have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hsrc' t ht - rw [this] - - -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` - refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) - · intros t ht - rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] - · intros t ht - rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] - -/-- Integral curves are unique on open intervals. - - If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` - on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` - agree on `Ioo a b`. -/ -theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} - {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) - (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) - (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by - /- - strategy: - * Lee P.213, just need to translate "S is closed in J" to type theory language - -/ - set s := {t | γ t = γ' t} ∩ Ioo a b with hs - have hsub : Ioo a b ⊆ s := by - apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ - ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ - · -- is this really the most convenient way to pass to subtype topology? - rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, - image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] - intros t ht - rw [mem_preimage, ← closure_subtype] at ht - revert ht t - apply IsClosed.closure_subset - apply isClosed_eq - · rw [continuous_iff_continuousAt] - rintro ⟨_, ht⟩ - apply ContinuousAt.comp _ continuousAt_subtype_val - rw [Subtype.coe_mk] - exact hγ.continuousAt ht - · rw [continuous_iff_continuousAt] - rintro ⟨_, ht⟩ - apply ContinuousAt.comp _ continuousAt_subtype_val - rw [Subtype.coe_mk] - exact hγ'.continuousAt ht - · rw [isOpen_iff_mem_nhds] - intro t₁ ht₁ - rw [mem_nhds_iff] - obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt - (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - ht₁.1 - refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), - subset_inter - (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) - (Ioo_subset_Ioo (by simp) (by simp)), - isOpen_Ioo, ?_⟩ - rw [mem_Ioo] - exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ - intros t ht - exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 - -/-- Global integral curves are unique. -/ -theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} - (hip : ∀ t, I.IsInteriorPoint (γ t)) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) - (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by - ext t - obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by - refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ - rw [mem_Ioo] - by_cases ht : t - t₀ < 0 - · rw [abs_of_neg ht] - constructor <;> linarith - · rw [abs_of_nonneg (not_lt.mp ht)] - constructor <;> linarith - exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ - (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv - (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) - (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht + have := HasMFDerivAt.comp t + (hasMFDerivAt_extChartAt I (extChartAt_source I (γ t₀) ▸ hsrc)) (hγ _ ht) + apply this.congr_mfderiv + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, + mfderiv_chartAt_eq_tangentCoordChange I (extChartAt_source I (γ t₀) ▸ hsrc)] + rfl + + +-- /-- Local integral curves are unique. + +-- If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` +-- at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ +-- theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) +-- (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) +-- (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : +-- ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by +-- -- first define `v'` as the vector field expressed in the local chart around `γ t₀` +-- -- this is basically what the function looks like when `hv` is unfolded +-- set v' : E → E := fun x => +-- tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) +-- (v ((extChartAt I (γ t₀)).symm x)) with hv' + +-- -- extract set `s` on which `v'` is Lipschitz +-- rw [contMDiffAt_iff] at hv +-- obtain ⟨_, hv⟩ := hv +-- obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := +-- ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd +-- have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip + +-- -- `γ t` when expressed in the local chart should remain inside `s` +-- have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := +-- (continuousAt_extChartAt ..).comp hγ.continuousAt +-- rw [continuousAt_def] at hcont +-- have hnhds := hcont _ hs +-- rw [← eventually_mem_nhds] at hnhds + +-- -- `γ t` should remain inside the domain of the local chart around `γ t₀` +-- have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) +-- rw [← eventually_mem_nhds] at hsrc + +-- -- `γ` is tangent to `v` in some neighbourhood of `t₀` +-- simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ + +-- -- same as above but for `γ'` +-- have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := +-- ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt +-- rw [continuousAt_def] at hcont' +-- have hnhds' := hcont' _ (h ▸ hs) +-- rw [← eventually_mem_nhds] at hnhds' + +-- have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) +-- rw [← eventually_mem_nhds] at hsrc' + +-- simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ' + +-- -- there exists a neighbourhood around `t₀` in which all of the above hold +-- have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' +-- rw [Metric.eventually_nhds_iff_ball] at haux + +-- obtain ⟨ε, hε, haux⟩ := haux +-- refine ⟨ε, hε, ?_⟩ + +-- -- break out all the conditions again +-- have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 +-- have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 +-- have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 +-- have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 +-- have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 +-- have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 + +-- -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood +-- have hcont := (continuousOn_extChartAt I (γ t₀)).comp +-- (IsIntegralCurveOn.continuousOn hmfd) hsrc +-- have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp +-- (IsIntegralCurveOn.continuousOn hmfd') hsrc' + +-- simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' + +-- -- `γ` and `γ'` are +-- have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') +-- (Ioo (t₀ - ε) (t₀ + ε)) := by +-- -- uniqueness of ODE solutions in an open interval +-- apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) +-- (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) +-- · intros t ht +-- rw [hv'] +-- have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) +-- apply this.congr_deriv +-- have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by +-- rw [Function.comp_apply, PartialEquiv.left_inv] +-- exact hsrc t ht +-- rw [this] +-- · intros t ht +-- rw [hv', h] +-- have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) +-- apply this.congr_deriv +-- have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by +-- rw [Function.comp_apply, PartialEquiv.left_inv] +-- exact hsrc' t ht +-- rw [this] + +-- -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` +-- refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) +-- · intros t ht +-- rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] +-- · intros t ht +-- rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] + +-- /-- Integral curves are unique on open intervals. + +-- If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` +-- on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` +-- agree on `Ioo a b`. -/ +-- theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] +-- [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} +-- {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) +-- (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) +-- (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) +-- (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by +-- /- +-- strategy: +-- * Lee P.213, just need to translate "S is closed in J" to type theory language +-- -/ +-- set s := {t | γ t = γ' t} ∩ Ioo a b with hs +-- have hsub : Ioo a b ⊆ s := by +-- apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ +-- ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ +-- · -- is this really the most convenient way to pass to subtype topology? +-- rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, +-- image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] +-- intros t ht +-- rw [mem_preimage, ← closure_subtype] at ht +-- revert ht t +-- apply IsClosed.closure_subset +-- apply isClosed_eq +-- · rw [continuous_iff_continuousAt] +-- rintro ⟨_, ht⟩ +-- apply ContinuousAt.comp _ continuousAt_subtype_val +-- rw [Subtype.coe_mk] +-- exact hγ.continuousAt ht +-- · rw [continuous_iff_continuousAt] +-- rintro ⟨_, ht⟩ +-- apply ContinuousAt.comp _ continuousAt_subtype_val +-- rw [Subtype.coe_mk] +-- exact hγ'.continuousAt ht +-- · rw [isOpen_iff_mem_nhds] +-- intro t₁ ht₁ +-- rw [mem_nhds_iff] +-- obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := +-- isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt +-- (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) +-- (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) +-- ht₁.1 +-- refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), +-- subset_inter +-- (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) +-- (Ioo_subset_Ioo (by simp) (by simp)), +-- isOpen_Ioo, ?_⟩ +-- rw [mem_Ioo] +-- exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ +-- intros t ht +-- exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 + +-- /-- Global integral curves are unique. -/ +-- theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] +-- [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} +-- (hip : ∀ t, I.IsInteriorPoint (γ t)) +-- (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) +-- (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by +-- ext t +-- obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by +-- refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ +-- rw [mem_Ioo] +-- by_cases ht : t - t₀ < 0 +-- · rw [abs_of_neg ht] +-- constructor <;> linarith +-- · rw [abs_of_nonneg (not_lt.mp ht)] +-- constructor <;> linarith +-- exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ +-- (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv +-- (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) +-- (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht From f17b87d40fc11de37c2fcafd087f29ab99ff35e1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 18:53:48 -0800 Subject: [PATCH 150/203] uncomment --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 355 +++++++++---------- Mathlib/Geometry/Manifold/MFDeriv.lean | 2 +- 2 files changed, 178 insertions(+), 179 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7239dd089bd6b..dd759d8c637ed 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -350,189 +350,188 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by -- turn `HasDerivAt` into comp of `HasMFDerivAt` + have hsrc := extChartAt_source I (γ t₀) ▸ hsrc rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - have := HasMFDerivAt.comp t - (hasMFDerivAt_extChartAt I (extChartAt_source I (γ t₀) ▸ hsrc)) (hγ _ ht) - apply this.congr_mfderiv + apply (HasMFDerivAt.comp t + (hasMFDerivAt_extChartAt I hsrc) (hγ _ ht)).congr_mfderiv rw [ContinuousLinearMap.ext_iff] intro a rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, - mfderiv_chartAt_eq_tangentCoordChange I (extChartAt_source I (γ t₀) ▸ hsrc)] + mfderiv_chartAt_eq_tangentCoordChange I hsrc] rfl +/-- Local integral curves are unique. + + If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` + at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ +theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : + ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by + -- first define `v'` as the vector field expressed in the local chart around `γ t₀` + -- this is basically what the function looks like when `hv` is unfolded + set v' : E → E := fun x => + tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) + (v ((extChartAt I (γ t₀)).symm x)) with hv' + + -- extract set `s` on which `v'` is Lipschitz + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd + have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip + + -- `γ t` when expressed in the local chart should remain inside `s` + have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := + (continuousAt_extChartAt ..).comp hγ.continuousAt + rw [continuousAt_def] at hcont + have hnhds := hcont _ hs + rw [← eventually_mem_nhds] at hnhds --- /-- Local integral curves are unique. - --- If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` --- at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ --- theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) --- (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) --- (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : --- ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by --- -- first define `v'` as the vector field expressed in the local chart around `γ t₀` --- -- this is basically what the function looks like when `hv` is unfolded --- set v' : E → E := fun x => --- tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) --- (v ((extChartAt I (γ t₀)).symm x)) with hv' - --- -- extract set `s` on which `v'` is Lipschitz --- rw [contMDiffAt_iff] at hv --- obtain ⟨_, hv⟩ := hv --- obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := --- ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd --- have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip - --- -- `γ t` when expressed in the local chart should remain inside `s` --- have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := --- (continuousAt_extChartAt ..).comp hγ.continuousAt --- rw [continuousAt_def] at hcont --- have hnhds := hcont _ hs --- rw [← eventually_mem_nhds] at hnhds - --- -- `γ t` should remain inside the domain of the local chart around `γ t₀` --- have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) --- rw [← eventually_mem_nhds] at hsrc - --- -- `γ` is tangent to `v` in some neighbourhood of `t₀` --- simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ - --- -- same as above but for `γ'` --- have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := --- ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt --- rw [continuousAt_def] at hcont' --- have hnhds' := hcont' _ (h ▸ hs) --- rw [← eventually_mem_nhds] at hnhds' - --- have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) --- rw [← eventually_mem_nhds] at hsrc' - --- simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ' - --- -- there exists a neighbourhood around `t₀` in which all of the above hold --- have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' --- rw [Metric.eventually_nhds_iff_ball] at haux - --- obtain ⟨ε, hε, haux⟩ := haux --- refine ⟨ε, hε, ?_⟩ - --- -- break out all the conditions again --- have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 --- have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 --- have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 --- have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 --- have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 --- have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 - --- -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood --- have hcont := (continuousOn_extChartAt I (γ t₀)).comp --- (IsIntegralCurveOn.continuousOn hmfd) hsrc --- have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp --- (IsIntegralCurveOn.continuousOn hmfd') hsrc' - --- simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' - --- -- `γ` and `γ'` are --- have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') --- (Ioo (t₀ - ε) (t₀ + ε)) := by --- -- uniqueness of ODE solutions in an open interval --- apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) --- (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) --- · intros t ht --- rw [hv'] --- have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) --- apply this.congr_deriv --- have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by --- rw [Function.comp_apply, PartialEquiv.left_inv] --- exact hsrc t ht --- rw [this] --- · intros t ht --- rw [hv', h] --- have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) --- apply this.congr_deriv --- have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by --- rw [Function.comp_apply, PartialEquiv.left_inv] --- exact hsrc' t ht --- rw [this] - --- -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` --- refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) --- · intros t ht --- rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] --- · intros t ht --- rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] - --- /-- Integral curves are unique on open intervals. - --- If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` --- on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` --- agree on `Ioo a b`. -/ --- theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] --- [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} --- {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) --- (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) --- (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) --- (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by --- /- --- strategy: --- * Lee P.213, just need to translate "S is closed in J" to type theory language --- -/ --- set s := {t | γ t = γ' t} ∩ Ioo a b with hs --- have hsub : Ioo a b ⊆ s := by --- apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ --- ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ --- · -- is this really the most convenient way to pass to subtype topology? --- rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, --- image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] --- intros t ht --- rw [mem_preimage, ← closure_subtype] at ht --- revert ht t --- apply IsClosed.closure_subset --- apply isClosed_eq --- · rw [continuous_iff_continuousAt] --- rintro ⟨_, ht⟩ --- apply ContinuousAt.comp _ continuousAt_subtype_val --- rw [Subtype.coe_mk] --- exact hγ.continuousAt ht --- · rw [continuous_iff_continuousAt] --- rintro ⟨_, ht⟩ --- apply ContinuousAt.comp _ continuousAt_subtype_val --- rw [Subtype.coe_mk] --- exact hγ'.continuousAt ht --- · rw [isOpen_iff_mem_nhds] --- intro t₁ ht₁ --- rw [mem_nhds_iff] --- obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := --- isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt --- (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) --- (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) --- ht₁.1 --- refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), --- subset_inter --- (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) --- (Ioo_subset_Ioo (by simp) (by simp)), --- isOpen_Ioo, ?_⟩ --- rw [mem_Ioo] --- exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ --- intros t ht --- exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 - --- /-- Global integral curves are unique. -/ --- theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] --- [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} --- (hip : ∀ t, I.IsInteriorPoint (γ t)) --- (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) --- (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by --- ext t --- obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by --- refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ --- rw [mem_Ioo] --- by_cases ht : t - t₀ < 0 --- · rw [abs_of_neg ht] --- constructor <;> linarith --- · rw [abs_of_nonneg (not_lt.mp ht)] --- constructor <;> linarith --- exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ --- (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv --- (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) --- (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht + -- `γ t` should remain inside the domain of the local chart around `γ t₀` + have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) + rw [← eventually_mem_nhds] at hsrc + + -- `γ` is tangent to `v` in some neighbourhood of `t₀` + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ + + -- same as above but for `γ'` + have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := + ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt + rw [continuousAt_def] at hcont' + have hnhds' := hcont' _ (h ▸ hs) + rw [← eventually_mem_nhds] at hnhds' + + have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) + rw [← eventually_mem_nhds] at hsrc' + + simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ' + + -- there exists a neighbourhood around `t₀` in which all of the above hold + have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' + rw [Metric.eventually_nhds_iff_ball] at haux + + obtain ⟨ε, hε, haux⟩ := haux + refine ⟨ε, hε, ?_⟩ + + -- break out all the conditions again + have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 + have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 + have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 + have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 + have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 + have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 + + -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood + have hcont := (continuousOn_extChartAt I (γ t₀)).comp + (IsIntegralCurveOn.continuousOn hmfd) hsrc + have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp + (IsIntegralCurveOn.continuousOn hmfd') hsrc' + + simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' + + -- `γ` and `γ'` are + have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') + (Ioo (t₀ - ε) (t₀ + ε)) := by + -- uniqueness of ODE solutions in an open interval + apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) + (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) + · intros t ht + rw [hv'] + have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) + apply this.congr_deriv + have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by + rw [Function.comp_apply, PartialEquiv.left_inv] + exact hsrc t ht + rw [this] + · intros t ht + rw [hv', h] + have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) + apply this.congr_deriv + have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by + rw [Function.comp_apply, PartialEquiv.left_inv] + exact hsrc' t ht + rw [this] + + -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` + refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) + · intros t ht + rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] + · intros t ht + rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] + +/-- Integral curves are unique on open intervals. + + If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` + on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` + agree on `Ioo a b`. -/ +theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} + {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) + (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by + /- + strategy: + * Lee P.213, just need to translate "S is closed in J" to type theory language + -/ + set s := {t | γ t = γ' t} ∩ Ioo a b with hs + have hsub : Ioo a b ⊆ s := by + apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ + ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ + · -- is this really the most convenient way to pass to subtype topology? + rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, + image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] + intros t ht + rw [mem_preimage, ← closure_subtype] at ht + revert ht t + apply IsClosed.closure_subset + apply isClosed_eq + · rw [continuous_iff_continuousAt] + rintro ⟨_, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ.continuousAt ht + · rw [continuous_iff_continuousAt] + rintro ⟨_, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ'.continuousAt ht + · rw [isOpen_iff_mem_nhds] + intro t₁ ht₁ + rw [mem_nhds_iff] + obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := + isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt + (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + ht₁.1 + refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), + subset_inter + (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) + (Ioo_subset_Ioo (by simp) (by simp)), + isOpen_Ioo, ?_⟩ + rw [mem_Ioo] + exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ + intros t ht + exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 + +/-- Global integral curves are unique. -/ +theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} + (hip : ∀ t, I.IsInteriorPoint (γ t)) + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by + ext t + obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by + refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ + rw [mem_Ioo] + by_cases ht : t - t₀ < 0 + · rw [abs_of_neg ht] + constructor <;> linarith + · rw [abs_of_nonneg (not_lt.mp ht)] + constructor <;> linarith + exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ + (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv + (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) + (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht diff --git a/Mathlib/Geometry/Manifold/MFDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv.lean index 7f6fc4c88a48a..f2edd72bc9547 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv.lean @@ -1898,7 +1898,7 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) - ((extChartAt_source I y) ▸ hsrc) + (extChartAt_source I y ▸ hsrc) rw [tangentCoordChange_def, mfderiv, if_pos this] rfl From 9acde4a13bd3c017b63a140f90ffd9a1e372dba9 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 14 Dec 2023 18:58:24 -0800 Subject: [PATCH 151/203] reorder --- Mathlib/Geometry/Manifold/MFDeriv.lean | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) diff --git a/Mathlib/Geometry/Manifold/MFDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv.lean index f2edd72bc9547..b1397e8dae250 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv.lean @@ -821,6 +821,13 @@ theorem MDifferentiable.prod_mk_space {f : M → E'} {g : M → E''} (hf : MDiff /-! ### Congruence lemmas for derivatives on manifolds -/ +theorem HasMFDerivAt.congr_mfderiv (h : HasMFDerivAt I I' f x f') (h' : f' = f₁') : + HasMFDerivAt I I' f x f₁' := + h' ▸ h + +theorem HasMFDerivWithinAt.congr_mfderiv (h : HasMFDerivWithinAt I I' f s x f') (h' : f' = f₁') : + HasMFDerivWithinAt I I' f s x f₁' := + h' ▸ h theorem HasMFDerivWithinAt.congr_of_eventuallyEq (h : HasMFDerivWithinAt I I' f s x f') (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : HasMFDerivWithinAt I I' f₁ s x f' := by @@ -840,14 +847,6 @@ theorem HasMFDerivWithinAt.congr_mono (h : HasMFDerivWithinAt I I' f s x f') (h.mono h₁).congr_of_eventuallyEq (Filter.mem_inf_of_right ht) hx #align has_mfderiv_within_at.congr_mono HasMFDerivWithinAt.congr_mono -theorem HasMFDerivAt.congr_mfderiv (h : HasMFDerivAt I I' f x f') (h' : f' = f₁') : - HasMFDerivAt I I' f x f₁' := - h' ▸ h - -theorem HasMFDerivWithinAt.congr_mfderiv (h : HasMFDerivWithinAt I I' f s x f') (h' : f' = f₁') : - HasMFDerivWithinAt I I' f s x f₁' := - h' ▸ h - theorem HasMFDerivAt.congr_of_eventuallyEq (h : HasMFDerivAt I I' f x f') (h₁ : f₁ =ᶠ[𝓝 x] f) : HasMFDerivAt I I' f₁ x f' := by rw [← hasMFDerivWithinAt_univ] at h ⊢ From fa459a52b70eb1e82d98d32addd467b39f69d576 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 15 Dec 2023 13:18:14 -0800 Subject: [PATCH 152/203] implicit variables --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index dd759d8c637ed..e23d2385f2f74 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -341,8 +341,6 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint -variable (I) - /-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when expressed in the local chart around the initial point `γ t₀`. -/ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) @@ -361,6 +359,8 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h mfderiv_chartAt_eq_tangentCoordChange I hsrc] rfl +variable {t₀} + /-- Local integral curves are unique. If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` @@ -439,7 +439,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht rw [hv'] - have := hmfd.hasDerivAt I t₀ ht (hsrc t ht) + have := hmfd.hasDerivAt t₀ ht (hsrc t ht) apply this.congr_deriv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] @@ -447,7 +447,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t rw [this] · intros t ht rw [hv', h] - have := hmfd'.hasDerivAt I t₀ ht (hsrc' t ht) + have := hmfd'.hasDerivAt t₀ ht (hsrc' t ht) apply this.congr_deriv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] @@ -502,7 +502,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] intro t₁ ht₁ rw [mem_nhds_iff] obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt I _ (hip _ ht₁.2) hv.contMDiffAt + isIntegralCurveAt_eqOn_of_contMDiffAt (hip _ ht₁.2) hv.contMDiffAt (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) ht₁.1 @@ -531,7 +531,7 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte constructor <;> linarith · rw [abs_of_nonneg (not_lt.mp ht)] constructor <;> linarith - exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff I t₀ + exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht From e34a30a85ddf52fcf2ef34599a773515ec2596cb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 15 Dec 2023 13:27:10 -0800 Subject: [PATCH 153/203] implicit variable --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b750bcee75c73..0ebb6d2c8c565 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -326,6 +326,8 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ +variable {t₀} + /-- For any continuously differentiable vector field defined on a manifold without boundary and any chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open @@ -353,8 +355,6 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h mfderiv_chartAt_eq_tangentCoordChange I hsrc] rfl -variable {t₀} - /-- Local integral curves are unique. If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` @@ -433,7 +433,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht rw [hv'] - have := hmfd.hasDerivAt t₀ ht (hsrc t ht) + have := hmfd.hasDerivAt ht (hsrc t ht) apply this.congr_deriv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] @@ -441,7 +441,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t rw [this] · intros t ht rw [hv', h] - have := hmfd'.hasDerivAt t₀ ht (hsrc' t ht) + have := hmfd'.hasDerivAt ht (hsrc' t ht) apply this.congr_deriv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] From 3ba365fc8cfbab02e43a853964af6b7934e88d11 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 15 Dec 2023 13:54:19 -0800 Subject: [PATCH 154/203] organise variables, docstring --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 42 +++++++++++++++----- 1 file changed, 31 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 0ebb6d2c8c565..a73b13122d6c2 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -359,7 +359,7 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t₀)) +theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by @@ -373,7 +373,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint ht₀)).snd + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip -- `γ t` when expressed in the local chart should remain inside `s` @@ -455,14 +455,21 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (ht₀ : I.IsInteriorPoint (γ t · intros t ht rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] +theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [I.Boundaryless] + (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : + ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := + isIntegralCurveAt_eqOn_of_contMDiffAt I.isInteriorPoint hv hγ hγ' h + +variable [T2Space M] {a b : ℝ} + /-- Integral curves are unique on open intervals. If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` agree on `Ioo a b`. -/ -theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} - {a b : ℝ} (ht₀ : t₀ ∈ Ioo a b) (hip : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) +theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) + (hγt : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by @@ -496,7 +503,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] intro t₁ ht₁ rw [mem_nhds_iff] obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt (hip _ ht₁.2) hv.contMDiffAt + isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) ht₁.1 @@ -510,10 +517,18 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff {M : Type*} [TopologicalSpace M] intros t ht exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 -/-- Global integral curves are unique. -/ -theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] [T2Space M] {v : (x : M) → TangentSpace I x} {γ γ' : ℝ → M} - (hip : ∀ t, I.IsInteriorPoint (γ t)) +theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [I.Boundaryless] + (ht₀ : t₀ ∈ Ioo a b) + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) + (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := + isIntegralCurveOn_Ioo_eqOn_of_contMDiff ht₀ (fun _ _ => I.isInteriorPoint) hv hγ hγ' h + +/-- Global integral curves are unique. + + If a continuously differentiable vector field `v` admits two global integral curves + `γ γ' : ℝ → M`, and `γ t₀ = γ' t₀` for some `t₀`, then `γ` and `γ'` are equal. -/ +theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t @@ -526,6 +541,11 @@ theorem isIntegralCurve_eq_of_contMDiff {M : Type*} [TopologicalSpace M] [Charte · rw [abs_of_nonneg (not_lt.mp ht)] constructor <;> linarith exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff - (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hip t) hv + (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hγt t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht + +theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [I.Boundaryless] + (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := + isIntegralCurve_eq_of_contMDiff (fun _ => I.isInteriorPoint) hv hγ hγ' h From e784d9397071b6fe0cf27aa9c1940574b6ff38ae Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 17 Dec 2023 14:09:50 -0800 Subject: [PATCH 155/203] Update Mathlib/Geometry/Manifold/IntegralCurve.lean Co-authored-by: thorimur <68410468+thorimur@users.noreply.github.com> --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 86b974a590d4b..524fff3aa4cfa 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -104,11 +104,7 @@ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : IsIntegralCurveOn γ v s' := fun t ht ↦ h t (mem_of_mem_of_subset ht hs) lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : - IsIntegralCurveOn γ v (s ∪ s') := by - intros t ht - cases' ht with ht ht - · exact h _ ht - · exact h' _ ht + IsIntegralCurveOn γ v (s ∪ s') := fun _ ↦ fun | .inl ht => h _ ht | .inr ht => h' _ ht lemma isIntegralCurveAt_iff : IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by From 47e137db0df1a2e040dbf80cdcd6e34ac656ff15 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 17 Dec 2023 14:13:45 -0800 Subject: [PATCH 156/203] Update Mathlib/Geometry/Manifold/IntegralCurve.lean Co-authored-by: thorimur <68410468+thorimur@users.noreply.github.com> --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 21 ++++++++++---------- 1 file changed, 11 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 524fff3aa4cfa..31d39efd6a1f0 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -275,33 +275,34 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt s, hs, ?_⟩ intros t ht -- collect useful terms in convenient forms - have h : HasDerivAt f - ((fderivWithin ℝ ((extChartAt I x₀) ∘ (extChartAt I ((extChartAt I x₀).symm (f t))).symm) - (range I) (extChartAt I ((extChartAt I x₀).symm (f t)) ((extChartAt I x₀).symm (f t)))) - (v ((extChartAt I x₀).symm (f t)))) - t := (haux t ht).1 + let xₜ : M := (extChartAt I x₀).symm (f t) -- `xₜ := γ t` + have h : HasDerivAt f (x := t) <| fderivWithin ℝ + (extChartAt I x₀ ∘ (extChartAt I xₜ).symm) + (range I) + (extChartAt I xₜ xₜ) + (v xₜ) := + (haux t ht).1 rw [← tangentCoordChange_def] at h have hf3 := mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2 have hf3' := mem_of_mem_of_subset hf3 interior_subset have hft1 := mem_preimage.mp <| mem_of_mem_of_subset hf3' (extChartAt I x₀).target_subset_preimage_source - have hft2 := mem_extChartAt_source I ((extChartAt I x₀).symm (f t)) + have hft2 := mem_extChartAt_source I xₜ -- express the derivative of the integral curve in the local chart refine ⟨(continuousAt_extChartAt_symm'' _ _ hf3').comp h.continuousAt, HasDerivWithinAt.hasFDerivWithinAt ?_⟩ simp only [mfld_simps, hasDerivWithinAt_univ] - show HasDerivAt (((extChartAt I ((extChartAt I x₀).symm (f t))) ∘ (extChartAt I x₀).symm) ∘ f) - (v ((extChartAt I x₀).symm (f t))) t + show HasDerivAt ((extChartAt I xₜ ∘ (extChartAt I x₀).symm) ∘ f) (v xₜ) t -- express `v (γ t)` as `D⁻¹ D (v (γ t))`, where `D` is a change of coordinates, so we can use -- `HasFDerivAt.comp_hasDerivAt` on `h` - rw [← tangentCoordChange_self (I := I) (x := (extChartAt I x₀).symm (f t)) - (z := (extChartAt I x₀).symm (f t)) (v := v ((extChartAt I x₀).symm (f t))) hft2, + rw [← tangentCoordChange_self (I := I) (x := xₜ) (z := xₜ) (v := v xₜ) hft2, ← tangentCoordChange_comp (x := x₀) ⟨⟨hft2, hft1⟩, hft2⟩] apply HasFDerivAt.comp_hasDerivAt _ _ h apply HasFDerivWithinAt.hasFDerivAt (s := range I) _ <| mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3⟩ + dsimp only [xₜ] nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ From 2e0b9bf8d3f918593d7b7a23a45ae0b67624d02b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 17 Dec 2023 16:55:41 -0800 Subject: [PATCH 157/203] Update Mathlib/Geometry/Manifold/IntegralCurve.lean Co-authored-by: thorimur <68410468+thorimur@users.noreply.github.com> --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 31d39efd6a1f0..d1e080e7d65b2 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -69,7 +69,7 @@ variable `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` outside of `s` is irrelevant and considered junk. -/ def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : Set ℝ) := - ∀ (t : ℝ), t ∈ s → HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) + ∀ t ∈ s, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight <| v (γ t)) /-- If `v : M → TM` is a vector field on `M`, and `t₀ : ℝ`, `IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in an open interval of `t₀`. That is, there exists From 26be30a056846654f47621c13e8ffcc15d263f58 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sun, 17 Dec 2023 17:01:50 -0800 Subject: [PATCH 158/203] remove newlines --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index d1e080e7d65b2..62db6ffff5042 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -276,12 +276,8 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt intros t ht -- collect useful terms in convenient forms let xₜ : M := (extChartAt I x₀).symm (f t) -- `xₜ := γ t` - have h : HasDerivAt f (x := t) <| fderivWithin ℝ - (extChartAt I x₀ ∘ (extChartAt I xₜ).symm) - (range I) - (extChartAt I xₜ xₜ) - (v xₜ) := - (haux t ht).1 + have h : HasDerivAt f (x := t) <| fderivWithin ℝ (extChartAt I x₀ ∘ (extChartAt I xₜ).symm) + (range I) (extChartAt I xₜ xₜ) (v xₜ) := (haux t ht).1 rw [← tangentCoordChange_def] at h have hf3 := mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2 have hf3' := mem_of_mem_of_subset hf3 interior_subset From 5aab4262765b6ef62f6a425d8a8dc6f20e03f101 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 19 Dec 2023 15:50:39 -0800 Subject: [PATCH 159/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 62db6ffff5042..7aa35163abed5 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -258,7 +258,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd + (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hx)).snd simp_rw [← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hf2 -- use continuity of `f` so that `f t` remains inside `interior (extChartAt I x₀).target` have ⟨a, ha, hf2'⟩ := Metric.eventually_nhds_iff_ball.mp hf2 From 1ea3813e2b2b24f6e3ff1125958ce2f727cdd876 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 19 Dec 2023 16:08:54 -0800 Subject: [PATCH 160/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 6ce1a3941b22d..fe32fc72e3c78 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -366,7 +366,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hγt₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip -- `γ t` when expressed in the local chart should remain inside `s` From 749ea8caba0cd2e454e80d53d2a98a97bae2978e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 29 Dec 2023 12:36:56 -0800 Subject: [PATCH 161/203] IsIntegralCurveAt.hasMFDerivAt --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7aa35163abed5..e0c87bd839761 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -113,6 +113,11 @@ lemma isIntegralCurveAt_iff : obtain ⟨ε, hε, hsub⟩ := Metric.mem_nhds_iff.mp hs exact ⟨ε, hε, h.mono hsub⟩ +lemma IsIntegralCurveAt.hasMFDerivAt (h : IsIntegralCurveAt γ v t₀) : + HasMFDerivAt 𝓘(ℝ, ℝ) I γ t₀ ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t₀))) := + have ⟨_, hs, h⟩ := h + h t₀ (mem_of_mem_nhds hs) + lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : IsIntegralCurveAt γ v t₀ := ⟨s, hs, h⟩ From c74e8306550645e4172e93e3300d5581a9394957 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 5 Jan 2024 11:39:29 -0800 Subject: [PATCH 162/203] BoundarylessManifold --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index e0c87bd839761..7adf7b4b59cfb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -311,7 +311,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint + exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) From ec1df62648f29fcb08bce15f356c43fb236d3001 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 5 Jan 2024 11:44:43 -0800 Subject: [PATCH 163/203] BoundarylessManifold --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 27631913229a3..668ae60788685 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -333,7 +333,7 @@ variable {t₀} lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := - exists_isIntegralCurveAt_of_contMDiffAt t₀ hv I.isInteriorPoint + exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) /-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when expressed in the local chart around the initial point `γ t₀`. -/ @@ -457,7 +457,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [I.Boundaryless] (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt I.isInteriorPoint hv hγ hγ' h + isIntegralCurveAt_eqOn_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h variable [T2Space M] {a b : ℝ} @@ -520,7 +520,8 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [I.Boundaryless] (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := - isIntegralCurveOn_Ioo_eqOn_of_contMDiff ht₀ (fun _ _ => I.isInteriorPoint) hv hγ hγ' h + isIntegralCurveOn_Ioo_eqOn_of_contMDiff + ht₀ (fun _ _ => BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h /-- Global integral curves are unique. @@ -546,4 +547,5 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [I.Boundaryless] (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := - isIntegralCurve_eq_of_contMDiff (fun _ => I.isInteriorPoint) hv hγ hγ' h + isIntegralCurve_eq_of_contMDiff + (fun _ => BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h From 9db3abfeb71c8848ed7a42db17fc5c72d50942e8 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 5 Jan 2024 11:47:50 -0800 Subject: [PATCH 164/203] BoundarylessManifold --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 668ae60788685..85ae1498e8452 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -330,7 +330,7 @@ variable {t₀} chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around `t₀`. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [I.Boundaryless] +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) @@ -453,7 +453,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t · intros t ht rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] -theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [I.Boundaryless] +theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := @@ -515,7 +515,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) intros t ht exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 -theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [I.Boundaryless] +theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifold I M] (ht₀ : t₀ ∈ Ioo a b) (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) @@ -544,7 +544,7 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht -theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [I.Boundaryless] +theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I M] (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := isIntegralCurve_eq_of_contMDiff From 353f3c9a4967b2e3707a396b0f8c23514926a1b7 Mon Sep 17 00:00:00 2001 From: Michael Rothgang Date: Mon, 8 Jan 2024 14:13:25 +0100 Subject: [PATCH 165/203] Make I implicit again. --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- Mathlib/Geometry/Manifold/InteriorBoundary.lean | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7adf7b4b59cfb..641ad56dd26b0 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -263,7 +263,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hx)).snd + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd simp_rw [← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hf2 -- use continuity of `f` so that `f t` remains inside `interior (extChartAt I x₀).target` have ⟨a, ha, hf2'⟩ := Metric.eventually_nhds_iff_ball.mp hf2 diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index 970413b2edd42..9ceff811e0798 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -98,6 +98,7 @@ lemma boundary_eq_complement_interior : I.boundary M = (I.interior M)ᶜ := by apply (compl_unique ?_ I.interior_union_boundary_eq_univ).symm exact disjoint_iff_inter_eq_empty.mp (I.disjoint_interior_boundary) +variable {I} in lemma _root_.range_mem_nhds_isInteriorPoint {x : M} (h : I.IsInteriorPoint x) : range I ∈ nhds (extChartAt I x x) := by rw [mem_nhds_iff] From 7d7a7169c0c35b35c150a198dd8eaa2a47578e1d Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 14:39:03 -0800 Subject: [PATCH 166/203] docs, golf, add `isIntegralCurveOn_iff_isIntegralCurveAt` --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 110 +++++++++---------- 1 file changed, 52 insertions(+), 58 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 641ad56dd26b0..8aaca58b976ab 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -14,46 +14,44 @@ Let `M` be a manifold and `v : (x : M) → TangentSpace I x` be a vector field o curve of `v` is a function `γ : ℝ → M` such that the derivative of `γ` at `t` equals `v (γ t)`. The integral curve may only be defined for all `t` within some subset of `ℝ`. -Assume `v` is continuously differentiable. The existence theorem for solutions to ODEs implies that -a unique local integral curve exists for any continuously differentiable vector field `v`. The -uniqueness theorem for solutions to ODEs implies that integral curves of `v` are unique. These are -the main results of this file. - -## Main definition +## Main definitions Let `v : M → TM` be a vector field on `M`, and let `γ : ℝ → M`. -- **`IsIntegralCurve γ v`**: `γ t` is tangent to `v (γ t)` for all `t : ℝ`. That is, `γ` is a global +* `IsIntegralCurve γ v`: `γ t` is tangent to `v (γ t)` for all `t : ℝ`. That is, `γ` is a global integral curve of `v`. -- **`IsIntegralCurveOn γ v s`**: `γ t` is tangent to `v (γ t)` for all `t ∈ s`, where `s : Set ℝ`. -- **`IsIntegralCurveAt γ v t₀`**: `γ t` is tangent to `v (γ t)` for all `t` in some open interval +* `IsIntegralCurveOn γ v s`: `γ t` is tangent to `v (γ t)` for all `t ∈ s`, where `s : Set ℝ`. +* `IsIntegralCurveAt γ v t₀`: `γ t` is tangent to `v (γ t)` for all `t` in some open interval around `t₀`. That is, `γ` is a local integral curve of `v`. For `IsIntegralCurveOn γ v s` and `IsIntegralCurveAt γ v t₀`, even though `γ` is defined for all time, its value outside of the set `s` or a small interval around `t₀` is irrelevant and considered junk. +## Main results + +* `exists_isIntegralCurveAt_of_contMDiffAt_boundaryless`: Existence of local integral curves for a +$C^1$ vector field. This follows from the existence theorem for solutions to ODEs +(`exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt`). + ## Implementation notes For the existence and uniqueness theorems, we assume that the image of the integral curve lies in the interior of the manifold. The case where the integral curve may lie on the boundary of the -manifold requires special treatment, and we leave it as a to-do. +manifold requires special treatment, and we leave it as a TODO. -The uniqueness theorem requires the manifold to be Hausdorff (T2), so that the set on which two -continuous functions agree is closed. +We state simpler versions of the theorem for boundaryless manifolds as corollaries. -We state simpler versions of the theorem for manifolds without boundary as corollaries. +## TODO -## To-do - -- The case where the integral curve may venture to the boundary of the manifold. See Theorem 9.34, - J. M. Lee. May require submanifolds. +* The case where the integral curve may venture to the boundary of the manifold. See Theorem 9.34, +J. M. Lee. May require submanifolds. ## Tags -integral curve, vector field, local existence, uniqueness +integral curve, vector field, local existence -/ -open scoped Manifold +open scoped Manifold Topology open Set @@ -61,26 +59,22 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] - {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] - {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners ℝ E' H'} - {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] - -/-- If `γ : ℝ → M`, `v : M → TM` is a vector field on `M`, and `s ∈ Set ℝ`, - `IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` - outside of `s` is irrelevant and considered junk. -/ -def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : Set ℝ) := + +/-- If `γ : ℝ → M` is $C^1$ on `s : Set ℝ` and `v` is a vector field on `M`, +`IsIntegralCurveOn γ v s` means `γ t` is tangent to `v (γ t)` for all `t ∈ s`. The value of `γ` +outside of `s` is irrelevant and considered junk. -/ +def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : Set ℝ) : Prop := ∀ t ∈ s, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight <| v (γ t)) -/-- If `v : M → TM` is a vector field on `M`, and `t₀ : ℝ`, `IsIntegralCurveAt γ v t₀` means - `γ : ℝ → M` is a local integral curve of `v` in an open interval of `t₀`. That is, there exists - `ε > 0` such that `γ t` is tangent to `v (γ t)` for all `t ∈ Ioo (t₀ - ε) (t₀ + ε)`. The value of - `γ` outside of this interval is irrelevant and considered junk. -/ -def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) := - ∃ s ∈ nhds t, IsIntegralCurveOn γ v s +/-- If `v` is a vector field on `M` and `t₀ : ℝ`, `IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a +local integral curve of `v` in a neighbourhood containing `t₀`. The value of `γ` outside of this +interval is irrelevant and considered junk. -/ +def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) : Prop := + ∃ s ∈ 𝓝 t, IsIntegralCurveOn γ v s /-- If `v : M → TM` is a vector field on `M`, `IsIntegralCurve γ v` means `γ : ℝ → M` is a global - integral curve of `v`. That is, `γ t` is tangent to `v (γ t)` for all `t : ℝ`. -/ -def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) := +integral curve of `v`. That is, `γ t` is tangent to `v (γ t)` for all `t : ℝ`. -/ +def IsIntegralCurve (γ : ℝ → M) (v : (x : M) → TangentSpace I x) : Prop := ∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t))) variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} @@ -106,6 +100,8 @@ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : IsIntegralCurveOn γ v (s ∪ s') := fun _ ↦ fun | .inl ht => h _ ht | .inr ht => h' _ ht +/-- `γ` is an integral curve for `v` at `t₀` iff `γ` is an integral curve on some interval +containing `t₀`. -/ lemma isIntegralCurveAt_iff : IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by refine ⟨?_, fun ⟨ε, hε, h⟩ ↦ ⟨Metric.ball t₀ ε, Metric.ball_mem_nhds _ hε, h⟩⟩ @@ -118,15 +114,20 @@ lemma IsIntegralCurveAt.hasMFDerivAt (h : IsIntegralCurveAt γ v t₀) : have ⟨_, hs, h⟩ := h h t₀ (mem_of_mem_nhds hs) -lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ nhds t₀) : +lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ 𝓝 t₀) : IsIntegralCurveAt γ v t₀ := ⟨s, hs, h⟩ +/-- If `γ` is an integral curve at each `t ∈ s`, it is an integral curve on `s`. -/ lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by intros t ht obtain ⟨s, hs, h⟩ := h t ht exact h t (mem_of_mem_nhds hs) +lemma isIntegralCurveOn_iff_isIntegralCurveAt (hs : IsOpen s) : + IsIntegralCurveOn γ v s ↔ ∀ t ∈ s, IsIntegralCurveAt γ v t := + ⟨fun h _ ht ↦ h.isIntegralCurveAt (hs.mem_nhds ht), IsIntegralCurveAt.isIntegralCurveOn⟩ + /-! ### Translation lemmas -/ section Translation @@ -144,11 +145,10 @@ lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) : lemma isIntegralCurveOn_comp_add {dt : ℝ} : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ - have := hγ.comp_add (-dt) - simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] at this - convert this + convert hγ.comp_add (-dt) ext simp only [Function.comp_apply, neg_add_cancel_right] + aesop lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -162,11 +162,10 @@ lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ - have := hγ.comp_add (-dt) - rw [sub_neg_eq_add, sub_add_cancel] at this - convert this + convert hγ.comp_add (-dt) ext simp only [Function.comp_apply, neg_add_cancel_right] + aesop lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : IsIntegralCurve (γ ∘ (· + dt)) v := by @@ -208,7 +207,7 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by obtain ⟨ε, hε, h⟩ := isIntegralCurveAt_iff.mp hγ rw [isIntegralCurveAt_iff] - refine ⟨ε / |a|, div_pos hε (abs_pos.mpr ha), ?_⟩ + refine ⟨ε / |a|, by positivity, ?_⟩ convert h.comp_mul a ext t rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, Real.dist_eq, Real.dist_eq, @@ -239,7 +238,7 @@ lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` - is a global integral curve of `v`. -/ +is a global integral curve of `v`. -/ lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ ↦ x) v := by intro t rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ), @@ -250,26 +249,24 @@ end Scaling variable (t₀) {x₀ : M} -/-- For any continuously differentiable vector field and any chosen non-boundary point `x₀` on the - manifold, there exists an integral curve `γ : ℝ → M` such that `γ t₀ = x₀` and the tangent vector - of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open interval around - `t₀`.-/ +/-- Existence of local integral curves for a $C^1$ vector field at interior points of a smooth +manifold. -/ theorem exists_isIntegralCurveAt_of_contMDiffAt (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) (hx : I.IsInteriorPoint x₀) : - ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by + ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by -- express the differentiability of the vector field `v` in the local chart rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd + (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hx)).snd simp_rw [← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hf2 -- use continuity of `f` so that `f t` remains inside `interior (extChartAt I x₀).target` have ⟨a, ha, hf2'⟩ := Metric.eventually_nhds_iff_ball.mp hf2 have hcont := (hf2' t₀ (Metric.mem_ball_self ha)).continuousAt rw [continuousAt_def, hf1] at hcont - have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ nhds t₀ := + have hnhds : f ⁻¹' (interior (extChartAt I x₀).target) ∈ 𝓝 t₀ := hcont _ (isOpen_interior.mem_nhds ((I.isInteriorPoint_iff).mp hx)) rw [← eventually_mem_nhds] at hnhds -- obtain a neighbourhood `s` so that the above conditions both hold in `s` @@ -303,15 +300,12 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt mem_nhds_iff.mpr ⟨interior (extChartAt I x₀).target, subset_trans interior_subset (extChartAt_target_subset_range ..), isOpen_interior, hf3⟩ - dsimp only [xₜ] - nth_rw 4 [← (extChartAt I x₀).right_inv hf3'] + rw [← (extChartAt I x₀).right_inv hf3'] exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩ -/-- For any continuously differentiable vector field defined on a manifold without boundary and any - chosen starting point `x₀ : M`, an integral curve `γ : ℝ → M` exists such that `γ t₀ = x₀` and the - tangent vector of `γ` at `t` coincides with the vector field at `γ t` for all `t` within an open - interval around `t₀`. -/ +/-- Existence of local integral curves for a $C^1$ vector field on a smooth manifold without +boundary. -/ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : - ∃ (γ : ℝ → M), γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := + ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) From db893bd366188878c07257b52b6ae47e5ebe5856 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 15:22:33 -0800 Subject: [PATCH 167/203] use `Eventually` in `IsIntegralCurveAt` --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 48 +++++++++++--------- 1 file changed, 27 insertions(+), 21 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 8aaca58b976ab..4e48a18fe6456 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -69,8 +69,8 @@ def IsIntegralCurveOn (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (s : S /-- If `v` is a vector field on `M` and `t₀ : ℝ`, `IsIntegralCurveAt γ v t₀` means `γ : ℝ → M` is a local integral curve of `v` in a neighbourhood containing `t₀`. The value of `γ` outside of this interval is irrelevant and considered junk. -/ -def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t : ℝ) : Prop := - ∃ s ∈ 𝓝 t, IsIntegralCurveOn γ v s +def IsIntegralCurveAt (γ : ℝ → M) (v : (x : M) → TangentSpace I x) (t₀ : ℝ) : Prop := + ∀ᶠ t in 𝓝 t₀, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight <| v (γ t)) /-- If `v : M → TM` is a vector field on `M`, `IsIntegralCurve γ v` means `γ : ℝ → M` is a global integral curve of `v`. That is, `γ t` is tangent to `v (γ t)` for all `t : ℝ`. -/ @@ -85,13 +85,27 @@ lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) lemma isIntegralCurve_iff_isIntegralCurveOn : IsIntegralCurve γ v ↔ IsIntegralCurveOn γ v univ := ⟨fun h ↦ h.isIntegralCurveOn _, fun h t ↦ h t (mem_univ _)⟩ +lemma isIntegralCurveAt_iff : + IsIntegralCurveAt γ v t₀ ↔ ∃ s ∈ 𝓝 t₀, IsIntegralCurveOn γ v s := by + simp_rw [IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] + rfl + +/-- `γ` is an integral curve for `v` at `t₀` iff `γ` is an integral curve on some interval +containing `t₀`. -/ +lemma isIntegralCurveAt_iff' : + IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by + simp_rw [IsIntegralCurveOn, ← Metric.eventually_nhds_iff_ball] + rfl + lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : - IsIntegralCurveAt γ v t := ⟨univ, Filter.univ_mem, fun t _ ↦ h t⟩ + IsIntegralCurveAt γ v t := by + rw [isIntegralCurveAt_iff] + exact ⟨univ, Filter.univ_mem, fun t _ ↦ h t⟩ lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := ⟨fun h ↦ h.isIntegralCurveAt, fun h t ↦ by - obtain ⟨s, hs, h⟩ := h t + obtain ⟨s, hs, h⟩ := isIntegralCurveAt_iff.mp (h t) exact h t (mem_of_mem_nhds hs)⟩ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : @@ -100,28 +114,19 @@ lemma IsIntegralCurveOn.mono (h : IsIntegralCurveOn γ v s) (hs : s' ⊆ s) : lemma IsIntegralCurveOn.of_union (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') : IsIntegralCurveOn γ v (s ∪ s') := fun _ ↦ fun | .inl ht => h _ ht | .inr ht => h' _ ht -/-- `γ` is an integral curve for `v` at `t₀` iff `γ` is an integral curve on some interval -containing `t₀`. -/ -lemma isIntegralCurveAt_iff : - IsIntegralCurveAt γ v t₀ ↔ ∃ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε) := by - refine ⟨?_, fun ⟨ε, hε, h⟩ ↦ ⟨Metric.ball t₀ ε, Metric.ball_mem_nhds _ hε, h⟩⟩ - rintro ⟨s, hs, h⟩ - obtain ⟨ε, hε, hsub⟩ := Metric.mem_nhds_iff.mp hs - exact ⟨ε, hε, h.mono hsub⟩ - lemma IsIntegralCurveAt.hasMFDerivAt (h : IsIntegralCurveAt γ v t₀) : HasMFDerivAt 𝓘(ℝ, ℝ) I γ t₀ ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t₀))) := - have ⟨_, hs, h⟩ := h + have ⟨_, hs, h⟩ := isIntegralCurveAt_iff.mp h h t₀ (mem_of_mem_nhds hs) lemma IsIntegralCurveOn.isIntegralCurveAt (h : IsIntegralCurveOn γ v s) (hs : s ∈ 𝓝 t₀) : - IsIntegralCurveAt γ v t₀ := ⟨s, hs, h⟩ + IsIntegralCurveAt γ v t₀ := isIntegralCurveAt_iff.mpr ⟨s, hs, h⟩ /-- If `γ` is an integral curve at each `t ∈ s`, it is an integral curve on `s`. -/ lemma IsIntegralCurveAt.isIntegralCurveOn (h : ∀ t ∈ s, IsIntegralCurveAt γ v t) : IsIntegralCurveOn γ v s := by intros t ht - obtain ⟨s, hs, h⟩ := h t ht + obtain ⟨s, hs, h⟩ := isIntegralCurveAt_iff.mp (h t ht) exact h t (mem_of_mem_nhds hs) lemma isIntegralCurveOn_iff_isIntegralCurveAt (hs : IsOpen s) : @@ -152,9 +157,9 @@ lemma isIntegralCurveOn_comp_add {dt : ℝ} : lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by - rw [isIntegralCurveAt_iff] at hγ + rw [isIntegralCurveAt_iff'] at * obtain ⟨ε, hε, h⟩ := hγ - refine ⟨Metric.ball (t₀ - dt) ε, Metric.isOpen_ball.mem_nhds (Metric.mem_ball_self hε), ?_⟩ + refine ⟨ε, hε, ?_⟩ convert h.comp_add dt ext t rw [mem_setOf_eq, Metric.mem_ball, Metric.mem_ball, dist_sub_eq_dist_add_right] @@ -205,8 +210,8 @@ lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by - obtain ⟨ε, hε, h⟩ := isIntegralCurveAt_iff.mp hγ - rw [isIntegralCurveAt_iff] + rw [isIntegralCurveAt_iff'] at * + obtain ⟨ε, hε, h⟩ := hγ refine ⟨ε / |a|, by positivity, ?_⟩ convert h.comp_mul a ext t @@ -260,7 +265,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt obtain ⟨_, hv⟩ := hv -- use Picard-Lindelöf theorem to extract a solution to the ODE in the local chart obtain ⟨f, hf1, hf2⟩ := exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt t₀ - (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hx)).snd + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hx)).snd simp_rw [← Real.ball_eq_Ioo, ← Metric.eventually_nhds_iff_ball] at hf2 -- use continuity of `f` so that `f t` remains inside `interior (extChartAt I x₀).target` have ⟨a, ha, hf2'⟩ := Metric.eventually_nhds_iff_ball.mp hf2 @@ -272,6 +277,7 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt -- obtain a neighbourhood `s` so that the above conditions both hold in `s` obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve + simp_rw [isIntegralCurveAt_iff] refine ⟨(extChartAt I x₀).symm ∘ f, Eq.symm (by rw [Function.comp_apply, hf1, PartialEquiv.left_inv _ (mem_extChartAt_source ..)]), s, hs, ?_⟩ From ad6d0a69e33ba0c9b0ab673cb1bf414c738b0c13 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 15:26:03 -0800 Subject: [PATCH 168/203] minor --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 4e48a18fe6456..ffdd2635dc046 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -98,9 +98,7 @@ lemma isIntegralCurveAt_iff' : rfl lemma IsIntegralCurve.isIntegralCurveAt (h : IsIntegralCurve γ v) (t : ℝ) : - IsIntegralCurveAt γ v t := by - rw [isIntegralCurveAt_iff] - exact ⟨univ, Filter.univ_mem, fun t _ ↦ h t⟩ + IsIntegralCurveAt γ v t := isIntegralCurveAt_iff.mpr ⟨univ, Filter.univ_mem, fun t _ ↦ h t⟩ lemma isIntegralCurve_iff_isIntegralCurveAt : IsIntegralCurve γ v ↔ ∀ t : ℝ, IsIntegralCurveAt γ v t := @@ -277,10 +275,9 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt -- obtain a neighbourhood `s` so that the above conditions both hold in `s` obtain ⟨s, hs, haux⟩ := (hf2.and hnhds).exists_mem -- prove that `γ := (extChartAt I x₀).symm ∘ f` is a desired integral curve - simp_rw [isIntegralCurveAt_iff] refine ⟨(extChartAt I x₀).symm ∘ f, Eq.symm (by rw [Function.comp_apply, hf1, PartialEquiv.left_inv _ (mem_extChartAt_source ..)]), - s, hs, ?_⟩ + isIntegralCurveAt_iff.mpr ⟨s, hs, ?_⟩⟩ intros t ht -- collect useful terms in convenient forms let xₜ : M := (extChartAt I x₀).symm (f t) -- `xₜ := γ t` From 5c7eca6337a03aa3ffdc95f0b2ac92832a2ad6a6 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 17:36:16 -0800 Subject: [PATCH 169/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 16 ++++++---------- 1 file changed, 6 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 1106ad5e19222..d790938376514 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -256,10 +256,9 @@ lemma IsIntegralCurveOn.continuousOn (hγ : IsIntegralCurveOn γ v s) : ContinuousOn γ s := fun t ht => (hγ t ht).1.continuousWithinAt lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : - ContinuousAt γ t₀ := by - obtain ⟨ε, hε, hγ⟩ := hγ - apply hγ.continuousAt - exact mem_of_mem_nhds hε + ContinuousAt γ t₀ := + have ⟨_, hs, hγ⟩ := isIntegralCurveAt_iff.mp hγ + hγ.continuousAt <| mem_of_mem_nhds hs lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : Continuous γ := continuous_iff_continuousAt.mpr @@ -330,6 +329,8 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) +variable {t₀} + /-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when expressed in the local chart around the initial point `γ t₀`. -/ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) @@ -366,7 +367,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint I hγt₀)).snd + ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip -- `γ t` when expressed in the local chart should remain inside `s` @@ -380,9 +381,6 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) rw [← eventually_mem_nhds] at hsrc - -- `γ` is tangent to `v` in some neighbourhood of `t₀` - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ - -- same as above but for `γ'` have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt @@ -393,8 +391,6 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) rw [← eventually_mem_nhds] at hsrc' - simp_rw [IsIntegralCurveAt, IsIntegralCurveOn, ← Filter.eventually_iff_exists_mem] at hγ' - -- there exists a neighbourhood around `t₀` in which all of the above hold have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' rw [Metric.eventually_nhds_iff_ball] at haux From fed6c6874aced6ceb1b2723dabe379e9b21d323f Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 18:06:21 -0800 Subject: [PATCH 170/203] golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 191 +++++++++--------- .../Geometry/Manifold/InteriorBoundary.lean | 2 - 2 files changed, 98 insertions(+), 95 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index d790938376514..15d42c876338a 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -33,6 +33,9 @@ junk. * `exists_isIntegralCurveAt_of_contMDiffAt_boundaryless`: Existence of local integral curves for a $C^1$ vector field. This follows from the existence theorem for solutions to ODEs (`exists_forall_hasDerivAt_Ioo_eq_of_contDiffAt`). +* `isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless`: Uniqueness of local integral curves for a +$C^1$ vector field. This follows from the uniqueness theorem for solutions to ODEs +(`ODE_solution_unique_of_mem_set_Ioo`). This requires the manifold to be Hausdorff (`T2Space`). ## Implementation notes @@ -45,7 +48,10 @@ We state simpler versions of the theorem for boundaryless manifolds as corollari ## TODO * The case where the integral curve may venture to the boundary of the manifold. See Theorem 9.34, -J. M. Lee. May require submanifolds. +Lee. May require submanifolds. + +## Reference +* Lee, J. M. (2012). _Introduction to Smooth Manifolds_. Springer New York. ## Tags @@ -132,6 +138,24 @@ lemma isIntegralCurveOn_iff_isIntegralCurveAt (hs : IsOpen s) : IsIntegralCurveOn γ v s ↔ ∀ t ∈ s, IsIntegralCurveAt γ v t := ⟨fun h _ ht ↦ h.isIntegralCurveAt (hs.mem_nhds ht), IsIntegralCurveAt.isIntegralCurveOn⟩ +/-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when + expressed in the local chart around the initial point `γ t₀`. -/ +lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) + (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : + HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by + -- turn `HasDerivAt` into comp of `HasMFDerivAt` + have hsrc := extChartAt_source I (γ t₀) ▸ hsrc + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + apply (HasMFDerivAt.comp t + (hasMFDerivAt_extChartAt I hsrc) (hγ _ ht)).congr_mfderiv + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, + mfderiv_chartAt_eq_tangentCoordChange I hsrc] + rfl + /-! ### Translation lemmas -/ section Translation @@ -253,7 +277,7 @@ lemma IsIntegralCurveOn.continuousAt (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ContinuousAt γ t₀ := (hγ t₀ ht).1 lemma IsIntegralCurveOn.continuousOn (hγ : IsIntegralCurveOn γ v s) : - ContinuousOn γ s := fun t ht => (hγ t ht).1.continuousWithinAt + ContinuousOn γ s := fun t ht ↦ (hγ t ht).1.continuousWithinAt lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : ContinuousAt γ t₀ := @@ -262,10 +286,12 @@ lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : Continuous γ := continuous_iff_continuousAt.mpr - fun _ => (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) + fun _ ↦ (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) end Scaling +section ExistUnique + variable (t₀) {x₀ : M} /-- Existence of local integral curves for a $C^1$ vector field at interior points of a smooth @@ -331,44 +357,26 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold variable {t₀} -/-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when - expressed in the local chart around the initial point `γ t₀`. -/ -lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) - (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : - HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by - -- turn `HasDerivAt` into comp of `HasMFDerivAt` - have hsrc := extChartAt_source I (γ t₀) ▸ hsrc - rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - apply (HasMFDerivAt.comp t - (hasMFDerivAt_extChartAt I hsrc) (hγ _ ht)).congr_mfderiv - rw [ContinuousLinearMap.ext_iff] - intro a - rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, - ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, - mfderiv_chartAt_eq_tangentCoordChange I hsrc] - rfl - /-- Local integral curves are unique. - If a continuously differentiable vector field `v` admits two local integral curves `γ γ' : ℝ → M` - at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval around `t₀` -/ +If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with +`γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by -- first define `v'` as the vector field expressed in the local chart around `γ t₀` -- this is basically what the function looks like when `hv` is unfolded - set v' : E → E := fun x => + set v' : E → E := fun x ↦ tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) (v ((extChartAt I (γ t₀)).symm x)) with hv' - -- extract set `s` on which `v'` is Lipschitz + -- extract a set `s` on which `v'` is Lipschitz rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - ContDiffAt.exists_lipschitzOnWith (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd - have hlip : ∀ t : ℝ, LipschitzOnWith K ((fun _ => v') t) ((fun _ => s) t) := fun _ => hlip + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd.exists_lipschitzOnWith + have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip -- `γ t` when expressed in the local chart should remain inside `s` have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := @@ -399,20 +407,20 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t refine ⟨ε, hε, ?_⟩ -- break out all the conditions again - have hmem := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 - have hsrc := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 - have hmfd : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.1 - have hmem' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 - have hsrc' := fun t ht => mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 - have hmfd' : IsIntegralCurveOn _ _ _ := fun t ht => (haux t ht).2.2.2.2.2 + have hmem := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 + have hsrc := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 + have hcrv : IsIntegralCurveOn _ _ _ := fun t ht ↦ (haux t ht).2.2.1 + have hmem' := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 + have hsrc' := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 + have hcrv' : IsIntegralCurveOn _ _ _ := fun t ht ↦ (haux t ht).2.2.2.2.2 -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood have hcont := (continuousOn_extChartAt I (γ t₀)).comp - (IsIntegralCurveOn.continuousOn hmfd) hsrc + (IsIntegralCurveOn.continuousOn hcrv) hsrc have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp - (IsIntegralCurveOn.continuousOn hmfd') hsrc' + (IsIntegralCurveOn.continuousOn hcrv') hsrc' - simp_rw [Real.ball_eq_Ioo] at hmem hsrc hmfd hcont hmem' hsrc' hmfd' hcont' + simp_rw [Real.ball_eq_Ioo] at hmem hsrc hcrv hcont hmem' hsrc' hcrv' hcont' -- `γ` and `γ'` are have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') @@ -422,7 +430,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) · intros t ht rw [hv'] - have := hmfd.hasDerivAt ht (hsrc t ht) + have := hcrv.hasDerivAt ht (hsrc t ht) apply this.congr_deriv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] @@ -430,7 +438,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t rw [this] · intros t ht rw [hv', h] - have := hmfd'.hasDerivAt ht (hsrc' t ht) + have := hcrv'.hasDerivAt ht (hsrc' t ht) apply this.congr_deriv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] @@ -445,7 +453,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold I M] - (hv : ContMDiffAt I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := isIntegralCurveAt_eqOn_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h @@ -454,72 +462,68 @@ variable [T2Space M] {a b : ℝ} /-- Integral curves are unique on open intervals. - If a continuously differentiable vector field `v` admits two integral curves `γ γ' : ℝ → M` - on some open interval `Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` - agree on `Ioo a b`. -/ +If a $C^1$ vector field `v` admits two integral curves `γ γ' : ℝ → M` on some open interval +`Ioo a b`, and `γ t₀ = γ' t₀` for some `t ∈ Ioo a b`, then `γ` and `γ'` agree on `Ioo a b`. -/ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) (hγt : ∀ t ∈ Ioo a b, I.IsInteriorPoint (γ t)) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := by - /- - strategy: - * Lee P.213, just need to translate "S is closed in J" to type theory language - -/ set s := {t | γ t = γ' t} ∩ Ioo a b with hs - have hsub : Ioo a b ⊆ s := by - apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ - ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ - · -- is this really the most convenient way to pass to subtype topology? - rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, - image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] - intros t ht - rw [mem_preimage, ← closure_subtype] at ht - revert ht t - apply IsClosed.closure_subset - apply isClosed_eq - · rw [continuous_iff_continuousAt] - rintro ⟨_, ht⟩ - apply ContinuousAt.comp _ continuousAt_subtype_val - rw [Subtype.coe_mk] - exact hγ.continuousAt ht - · rw [continuous_iff_continuousAt] - rintro ⟨_, ht⟩ - apply ContinuousAt.comp _ continuousAt_subtype_val - rw [Subtype.coe_mk] - exact hγ'.continuousAt ht - · rw [isOpen_iff_mem_nhds] - intro t₁ ht₁ - rw [mem_nhds_iff] - obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt - (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - ht₁.1 - refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), - subset_inter - (fun t ht => @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) - (Ioo_subset_Ioo (by simp) (by simp)), - isOpen_Ioo, ?_⟩ - rw [mem_Ioo] - exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ - intros t ht - exact mem_setOf.mp ((subset_def ▸ hsub) t ht).1 + -- since `Ioo a b` is connected, we get `s = Ioo a b` by showing that `s` is clopen in `Ioo a b` + -- in the subtype toplogy (`s` is also non-empty by assumption) + -- here we use a slightly weaker alternative theorem + suffices hsub : Ioo a b ⊆ s from fun t ht ↦ mem_setOf.mp ((subset_def ▸ hsub) t ht).1 + apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ + ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ + · -- is this really the most convenient way to pass to subtype topology? + rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, + image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] + intros t ht + rw [mem_preimage, ← closure_subtype] at ht + revert ht t + apply IsClosed.closure_subset + apply isClosed_eq + · rw [continuous_iff_continuousAt] + rintro ⟨_, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ.continuousAt ht + · rw [continuous_iff_continuousAt] + rintro ⟨_, ht⟩ + apply ContinuousAt.comp _ continuousAt_subtype_val + rw [Subtype.coe_mk] + exact hγ'.continuousAt ht + · rw [isOpen_iff_mem_nhds] + intro t₁ ht₁ + rw [mem_nhds_iff] + obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := + isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt + (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) + ht₁.1 + refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), + subset_inter + (fun t ht ↦ @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) + (Ioo_subset_Ioo (by simp) (by simp)), + isOpen_Ioo, ?_⟩ + rw [mem_Ioo] + exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifold I M] (ht₀ : t₀ ∈ Ioo a b) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurveOn γ v (Ioo a b)) (hγ' : IsIntegralCurveOn γ' v (Ioo a b)) (h : γ t₀ = γ' t₀) : EqOn γ γ' (Ioo a b) := isIntegralCurveOn_Ioo_eqOn_of_contMDiff - ht₀ (fun _ _ => BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + ht₀ (fun _ _ ↦ BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h /-- Global integral curves are unique. If a continuously differentiable vector field `v` admits two global integral curves `γ γ' : ℝ → M`, and `γ t₀ = γ' t₀` for some `t₀`, then `γ` and `γ'` are equal. -/ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by @@ -531,12 +535,13 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) · rw [abs_of_nonneg (not_lt.mp ht)] constructor <;> linarith exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff - (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ => hγt t) hv + (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ ↦ hγt t) hv (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I M] - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := - isIntegralCurve_eq_of_contMDiff - (fun _ => BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + isIntegralCurve_eq_of_contMDiff (fun _ ↦ BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + +end ExistUnique diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index bb339d6848a5f..9ceff811e0798 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -40,8 +40,6 @@ this requires a definition of submanifolds open Set -open scoped Manifold - -- Let `M` be a manifold with corners over the pair `(E, H)`. variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] From 89e02979d10fee84c84ccbb339c7f3562a15961c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 18:29:10 -0800 Subject: [PATCH 171/203] golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 19 ++++++++----------- 1 file changed, 8 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 15d42c876338a..85126e546667e 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -415,14 +415,12 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t have hcrv' : IsIntegralCurveOn _ _ _ := fun t ht ↦ (haux t ht).2.2.2.2.2 -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood - have hcont := (continuousOn_extChartAt I (γ t₀)).comp - (IsIntegralCurveOn.continuousOn hcrv) hsrc - have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp - (IsIntegralCurveOn.continuousOn hcrv') hsrc' + have hcont := (continuousOn_extChartAt I (γ t₀)).comp hcrv.continuousOn hsrc + have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp hcrv'.continuousOn hsrc' simp_rw [Real.ball_eq_Ioo] at hmem hsrc hcrv hcont hmem' hsrc' hcrv' hcont' - -- `γ` and `γ'` are + -- `γ` and `γ'` are equal on `Ioo (t₀ - ε) (t₀ + ε)` when expressed in the local chart have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') (Ioo (t₀ - ε) (t₀ + ε)) := by -- uniqueness of ODE solutions in an open interval @@ -507,7 +505,6 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) (fun t ht ↦ @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) (Ioo_subset_Ioo (by simp) (by simp)), isOpen_Ioo, ?_⟩ - rw [mem_Ioo] exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifold I M] @@ -520,14 +517,14 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifo /-- Global integral curves are unique. - If a continuously differentiable vector field `v` admits two global integral curves - `γ γ' : ℝ → M`, and `γ t₀ = γ' t₀` for some `t₀`, then `γ` and `γ'` are equal. -/ +If a continuously differentiable vector field `v` admits two global integral curves +`γ γ' : ℝ → M`, and `γ t₀ = γ' t₀` for some `t₀`, then `γ` and `γ'` are equal. -/ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by - refine ⟨2 * |t - t₀| + 1, add_pos_of_nonneg_of_pos (by simp) zero_lt_one, ?_⟩ + refine ⟨2 * |t - t₀| + 1, by positivity, ?_⟩ rw [mem_Ioo] by_cases ht : t - t₀ < 0 · rw [abs_of_neg ht] @@ -536,8 +533,8 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) constructor <;> linarith exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ ↦ hγt t) hv - (IsIntegralCurveOn.mono (hγ.isIntegralCurveOn _) (subset_univ _)) - (IsIntegralCurveOn.mono (hγ'.isIntegralCurveOn _) (subset_univ _)) h ht + ((hγ.isIntegralCurveOn _).mono (subset_univ _)) + ((hγ'.isIntegralCurveOn _).mono (subset_univ _)) h ht theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I M] (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) From a00ed88f6a51283611680db5f2a7e9dd50716b5b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 19:28:12 -0800 Subject: [PATCH 172/203] golf some earlier proofs --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 37 ++++++++++---------- 1 file changed, 18 insertions(+), 19 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7c0963bbb0833..5e51a07a55958 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -174,8 +174,9 @@ lemma isIntegralCurveOn_comp_add {dt : ℝ} : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) - ext - simp [Function.comp_apply, neg_add_cancel_right] + ext t + · simp only [Function.comp_apply, neg_add_cancel_right] + · simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by @@ -190,8 +191,9 @@ lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) - ext - simp only [Function.comp_apply, neg_add_cancel_right, sub_neg_eq_add, sub_add_cancel] + ext t + · simp only [Function.comp_apply, neg_add_cancel_right] + · simp only [sub_neg_eq_add, sub_add_cancel] lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : IsIntegralCurve (γ ∘ (· + dt)) v := by @@ -202,7 +204,7 @@ lemma isIntegralCurve_comp_add {dt : ℝ} : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· + dt)) v := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) - ext + ext t simp only [Function.comp_apply, neg_add_cancel_right] end Translation @@ -222,12 +224,11 @@ lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_mul a, fun hγ ↦ ?_⟩ - have := hγ.comp_mul a⁻¹ - simp_rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, mem_setOf_eq, mul_assoc, - inv_mul_eq_div, div_self ha, mul_one, setOf_mem_eq] at this - convert this + convert hγ.comp_mul a⁻¹ ext t - rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] + · simp only [mem_setOf_eq, mul_assoc, inv_mul_eq_div, div_self ha, mul_one, setOf_mem_eq] lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by @@ -242,12 +243,11 @@ lemma IsIntegralCurveAt.comp_mul_ne_zero (hγ : IsIntegralCurveAt γ v t₀) {a lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ ↦ hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ - have := hγ.comp_mul_ne_zero (inv_ne_zero ha) - rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul, ← div_mul_eq_div_div_swap, - inv_mul_eq_div, div_self ha, div_one, Function.comp.assoc] at this - convert this + convert hγ.comp_mul_ne_zero (inv_ne_zero ha) ext t - simp [inv_mul_eq_div, div_self ha] + · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] + · simp only [div_inv_eq_mul, div_mul_cancel _ ha] lemma IsIntegralCurve.comp_mul (hγ : IsIntegralCurve γ v) (a : ℝ) : IsIntegralCurve (γ ∘ (· * a)) (a • v) := by @@ -257,11 +257,10 @@ lemma IsIntegralCurve.comp_mul (hγ : IsIntegralCurve γ v) (a : ℝ) : lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by refine ⟨fun hγ ↦ hγ.comp_mul _, fun hγ ↦ ?_⟩ - have := hγ.comp_mul a⁻¹ - rw [smul_smul, inv_mul_eq_div, div_self ha, one_smul] at this - convert this + convert hγ.comp_mul a⁻¹ ext t - rw [Function.comp_apply, Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` is a global integral curve of `v`. -/ From 8493b010ad96a3b8b37af051aefe85aaf91e40c6 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 19:48:56 -0800 Subject: [PATCH 173/203] minor --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 5e51a07a55958..b61edacc42199 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -472,6 +472,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) apply isPreconnected_Ioo.subset_of_closure_inter_subset (s := Ioo a b) (u := s) _ ⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩ · -- is this really the most convenient way to pass to subtype topology? + -- TODO: shorten this when better API around subtype topology exists rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val, image_subset_image_iff Subtype.val_injective, preimage_setOf_eq] intros t ht @@ -521,8 +522,7 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by - refine ⟨2 * |t - t₀| + 1, by positivity, ?_⟩ - rw [mem_Ioo] + refine ⟨|t - t₀| + 1, by positivity, ?_⟩ by_cases ht : t - t₀ < 0 · rw [abs_of_neg ht] constructor <;> linarith From 15f7e4f1d3d6e1472f30cedd4a3d5af73e901eaa Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 20:22:16 -0800 Subject: [PATCH 174/203] doc --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index b61edacc42199..41593e7fb1901 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -55,7 +55,7 @@ Lee. May require submanifolds. ## Tags -integral curve, vector field, local existence +integral curve, vector field, local existence, uniqueness -/ open scoped Manifold Topology From 1c91b79d970fad579a36adc0177036582a20b4bb Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 8 Jan 2024 21:14:08 -0800 Subject: [PATCH 175/203] section header --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 41593e7fb1901..7ca869ab11942 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -287,6 +287,8 @@ lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : end Scaling +/-! ### Existence and uniqueness -/ + section ExistUnique variable (t₀) {x₀ : M} From fbd2199d0b480e02d8a9b35af059fc1d60947f10 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 11 Jan 2024 02:02:34 -0800 Subject: [PATCH 176/203] typo --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 7ca869ab11942..eb3950e3a8cfb 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -221,7 +221,7 @@ lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) : simp only [mfld_simps, hasFDerivWithinAt_univ] exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _ -lemma isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : +lemma isIntegralCurveOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_mul a, fun hγ ↦ ?_⟩ convert hγ.comp_mul a⁻¹ From a889e9fa40ca44bddf824b4557ca101bd011bd94 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 16 Jan 2024 22:42:27 -0800 Subject: [PATCH 177/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 5 +++++ 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a2c86dd9bb122..d163ae1315fb8 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -6,7 +6,7 @@ Authors: Winston Yin import Mathlib.Analysis.ODE.Gronwall import Mathlib.Analysis.ODE.PicardLindelof import Mathlib.Geometry.Manifold.InteriorBoundary -import Mathlib.Geometry.Manifold.MFDeriv.SpecificFunctions +import Mathlib.Geometry.Manifold.MFDeriv.Atlas /-! # Integral curves of vector fields on a manifold diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 7f991331b1d10..4230640795e6a 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -782,4 +782,9 @@ theorem tangentMap_comp (hg : MDifferentiable I' I'' g) (hf : MDifferentiable I ext p : 1; exact tangentMap_comp_at _ (hg _) (hf _) #align tangent_map_comp tangentMap_comp +lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : + mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by + have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) (extChartAt_source I y ▸ hsrc) + simp [mfderiv, if_pos this, Function.comp.assoc] + end DerivativesProperties From a44856a0f79562f71411673837e18a273de951ee Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 16 Jan 2024 22:44:38 -0800 Subject: [PATCH 178/203] gronwall doc --- Mathlib/Analysis/ODE/Gronwall.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 6fb481ffacb55..53e0404ff6146 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -245,7 +245,7 @@ set_option linter.uppercaseLean3 false in #align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with -a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, +a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, and we consider only solutions included in `s`. This version shows uniqueness in a half-open interval `Ico a b`, where `a` is the initial time. -/ @@ -263,7 +263,7 @@ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Se (fun _ ht' ↦ hgs _ (hss ht')) ha ⟨ht.1, le_refl _⟩ /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with -a given initial value provided that RHS is Lipschitz continuous in `x` within `s`, +a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, and we consider only solutions included in `s`. This version shows uniqueness in an open interval `Ioo a b` that contains the initial time `t₀`. -/ @@ -309,7 +309,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se (fun _ ht' => hgs _ <| (hss ht')) ha /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with -a given initial value provided that RHS is Lipschitz continuous in `x`. -/ +a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) From 92df5b96f98c38a2a7ad7cb2c9ef62444df41533 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 16 Jan 2024 22:47:26 -0800 Subject: [PATCH 179/203] minor --- Mathlib/Analysis/ODE/Gronwall.lean | 2 +- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 53e0404ff6146..f71bb6abe8403 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -278,7 +278,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se apply EqOn.union · -- case `t ≤ t₀`: show `fun t ↦ f_or_g (-t)` solves `-v (-t)` within `s (-t)` by composition, -- so we can use the `Ico` version of the uniqueness lemma backwards in time - have hv' (t) : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by + have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by rw [← one_mul K] exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) have : EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) := by diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index d163ae1315fb8..280708ac513c3 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -143,7 +143,7 @@ lemma isIntegralCurveOn_iff_isIntegralCurveAt (hs : IsOpen s) : lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) (hsrc : γ t ∈ (extChartAt I (γ t₀)).source) : HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) - ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t := by + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by -- turn `HasDerivAt` into comp of `HasMFDerivAt` have hsrc := extChartAt_source I (γ t₀) ▸ hsrc rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] From 06d6fa4a1870217a23ddd9fb9bd654862c86ca2e Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 16 Jan 2024 22:49:54 -0800 Subject: [PATCH 180/203] move lemma --- Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 5 +++++ Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 5 ----- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index 50834a5e556f9..40dfbe815cad2 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -167,6 +167,11 @@ theorem tangentMap_chart_symm {p : TangentBundle I M} {q : TangentBundle I H} exact ((chartAt H (TotalSpace.proj p)).right_inv h).symm #align tangent_map_chart_symm tangentMap_chart_symm +lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : + mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by + have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) (extChartAt_source I y ▸ hsrc) + simp [mfderiv, if_pos this, Function.comp.assoc] + end Charts diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 4230640795e6a..7f991331b1d10 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -782,9 +782,4 @@ theorem tangentMap_comp (hg : MDifferentiable I' I'' g) (hf : MDifferentiable I ext p : 1; exact tangentMap_comp_at _ (hg _) (hf _) #align tangent_map_comp tangentMap_comp -lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y).source) : - mfderiv I I (chartAt H y) x = tangentCoordChange I x y x := by - have := mdifferentiableAt_atlas I (ChartedSpace.chart_mem_atlas _) (extChartAt_source I y ▸ hsrc) - simp [mfderiv, if_pos this, Function.comp.assoc] - end DerivativesProperties From 092bf9f0eac0e2bbd52effc2bcbfc9b7ac2821f3 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 19 Jan 2024 17:34:40 -0800 Subject: [PATCH 181/203] ext t --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 280708ac513c3..12c71dbc7e23f 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -174,8 +174,8 @@ lemma isIntegralCurveOn_comp_add {dt : ℝ} : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) - ext t - · simp only [Function.comp_apply, neg_add_cancel_right] + · ext t + simp only [Function.comp_apply, neg_add_cancel_right] · simp only [mem_setOf_eq, neg_add_cancel_right, setOf_mem_eq] lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) : @@ -191,8 +191,8 @@ lemma isIntegralCurveAt_comp_add {dt : ℝ} : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩ convert hγ.comp_add (-dt) - ext t - · simp only [Function.comp_apply, neg_add_cancel_right] + · ext t + simp only [Function.comp_apply, neg_add_cancel_right] · simp only [sub_neg_eq_add, sub_add_cancel] lemma IsIntegralCurve.comp_add (hγ : IsIntegralCurve γ v) (dt : ℝ) : @@ -225,8 +225,8 @@ lemma isIntegralCurveOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by refine ⟨fun hγ ↦ hγ.comp_mul a, fun hγ ↦ ?_⟩ convert hγ.comp_mul a⁻¹ - ext t - · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · ext t + simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] · simp only [mem_setOf_eq, mul_assoc, inv_mul_eq_div, div_self ha, mul_one, setOf_mem_eq] @@ -244,8 +244,8 @@ lemma isIntegralCurveAt_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurveAt γ v t₀ ↔ IsIntegralCurveAt (γ ∘ (· * a)) (a • v) (t₀ / a) := by refine ⟨fun hγ ↦ hγ.comp_mul_ne_zero ha, fun hγ ↦ ?_⟩ convert hγ.comp_mul_ne_zero (inv_ne_zero ha) - ext t - · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · ext t + simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] · simp only [div_inv_eq_mul, div_mul_cancel _ ha] @@ -258,8 +258,8 @@ lemma isIntegralCurve_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) : IsIntegralCurve γ v ↔ IsIntegralCurve (γ ∘ (· * a)) (a • v) := by refine ⟨fun hγ ↦ hγ.comp_mul _, fun hγ ↦ ?_⟩ convert hγ.comp_mul a⁻¹ - ext t - · simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] + · ext t + simp only [Function.comp_apply, mul_assoc, inv_mul_eq_div, div_self ha, mul_one] · simp only [smul_smul, inv_mul_eq_div, div_self ha, one_smul] /-- If the vector field `v` vanishes at `x₀`, then the constant curve at `x₀` From e7f4498661b81a7f168c853981b101de8c846535 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 19 Jan 2024 17:40:56 -0800 Subject: [PATCH 182/203] _Icc --- Mathlib/Analysis/ODE/Gronwall.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index f71bb6abe8403..b2643253a5036 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -233,7 +233,7 @@ a given initial value provided that the RHS is Lipschitz continuous in `x` withi and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ -theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -242,7 +242,7 @@ theorem ODE_solution_unique_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in -#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set +#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set_Icc /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, @@ -256,7 +256,7 @@ theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Se (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) - ODE_solution_unique_of_mem_set hv + ODE_solution_unique_of_mem_set_Icc hv (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) (fun _ ht' ↦ hfs _ (hss ht')) (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ (hss ht')) @@ -316,7 +316,7 @@ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, Lip (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - ODE_solution_unique_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' + ODE_solution_unique_of_mem_set_Icc (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align ODE_solution_unique ODE_solution_unique From 06050fd5debdeeebdd77b5da42474e269a3bf3ad Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 19 Jan 2024 17:59:57 -0800 Subject: [PATCH 183/203] cleanup with suffices --- Mathlib/Analysis/ODE/Gronwall.lean | 34 +++++++++++++++--------------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index b2643253a5036..1586b49d8ab58 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -281,23 +281,23 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by rw [← one_mul K] exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) - have : EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) := by - have hmt : MapsTo Neg.neg (Ico (-t₀) (-a)) (Ioo a b) := - fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ - apply ODE_solution_unique_of_mem_set_Ico hv' - (hf.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hfs _ (hmt ht')) - (hg.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hgs _ (hmt ht')) (by simp [ha]) - · intros t' ht' - apply HasDerivAt.hasDerivWithinAt - convert HasFDerivAt.comp_hasDerivAt t' (hf' (-t') (hmt ht')) (hasDerivAt_neg t') - simp - · intros t' ht' - apply HasDerivAt.hasDerivWithinAt - convert HasFDerivAt.comp_hasDerivAt t' (hg' (-t') (hmt ht')) (hasDerivAt_neg t') - simp - rw [eqOn_comp_right_iff] at this - convert this - simp + have hmt : MapsTo Neg.neg (Ico (-t₀) (-a)) (Ioo a b) := + fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ + suffices EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) by + rw [eqOn_comp_right_iff] at this + convert this + simp + apply ODE_solution_unique_of_mem_set_Ico hv' + (hf.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hfs _ (hmt ht')) + (hg.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hgs _ (hmt ht')) (by simp [ha]) + · intros t' ht' + apply HasDerivAt.hasDerivWithinAt + convert HasFDerivAt.comp_hasDerivAt t' (hf' (-t') (hmt ht')) (hasDerivAt_neg t') + simp + · intros t' ht' + apply HasDerivAt.hasDerivWithinAt + convert HasFDerivAt.comp_hasDerivAt t' (hg' (-t') (hmt ht')) (hasDerivAt_neg t') + simp · -- case `t ≥ t₀`: follows trivially from the `Ico` version of the uniqueness lemma have hss := Ico_subset_Ioo_left (b := b) ht.1 exact ODE_solution_unique_of_mem_set_Ico hv From 6dc062f7e43cdbd14cad594c737ab64d79f7a4b8 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 19 Jan 2024 20:44:35 -0800 Subject: [PATCH 184/203] refactor --- Mathlib/Analysis/ODE/Gronwall.lean | 138 ++++++++++++++++------------- 1 file changed, 78 insertions(+), 60 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 1586b49d8ab58..cfa6e8f0ae6ba 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -233,7 +233,7 @@ a given initial value provided that the RHS is Lipschitz continuous in `x` withi and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ -theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_set_Icc_right {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -242,71 +242,89 @@ theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Se have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in -#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set_Icc - -/-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with -a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, -and we consider only solutions included in `s`. +#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set_Icc_right -This version shows uniqueness in a half-open interval `Ico a b`, where `a` is the initial time. -/ -theorem ODE_solution_unique_of_mem_set_Ico {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +/-- A time-reversed version of `ODE_solution_unique_of_mem_Icc_right`. Uniqueness is shown in a +closed interval `Icc a b`, where `b` is the "initial" time. -/ +theorem ODE_solution_unique_of_mem_set_Icc_left {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Ico a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) - (hg : ContinuousOn g (Ico a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Ico a b) := fun _ ht ↦ - have hss := Ico_subset_Ico_right (a := a) (le_of_lt ht.2) - ODE_solution_unique_of_mem_set_Icc hv - (hf.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hf' _ (hss ht')) - (fun _ ht' ↦ hfs _ (hss ht')) - (hg.mono (Icc_subset_Ico_right ht.2)) (fun _ ht' ↦ hg' _ (hss ht')) - (fun _ ht' ↦ hgs _ (hss ht')) ha ⟨ht.1, le_refl _⟩ - -/-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with -a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, -and we consider only solutions included in `s`. - -This version shows uniqueness in an open interval `Ioo a b` that contains the initial time `t₀`. -/ + {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ioc a b, HasDerivWithinAt f (v t (f t)) (Iic t) t) (hfs : ∀ t ∈ Ioc a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ioc a b, HasDerivWithinAt g (v t (g t)) (Iic t) t) + (hgs : ∀ t ∈ Ioc a b, g t ∈ s t) (hb : f b = g b) : EqOn f g (Icc a b) := by + have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by + rw [← one_mul K] + exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) + have hmt1 : MapsTo Neg.neg (Icc (-b) (-a)) (Icc a b) := + fun _ ht ↦ ⟨le_neg.mp ht.2, neg_le.mp ht.1⟩ + have hmt2 : MapsTo Neg.neg (Ico (-b) (-a)) (Ioc a b) := + fun _ ht ↦ ⟨lt_neg.mp ht.2, neg_le.mp ht.1⟩ + have hmt3 (t : ℝ) : MapsTo Neg.neg (Ici t) (Iic (-t)) := + fun _ ht' ↦ mem_Iic.mpr <| neg_le_neg ht' + suffices EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Icc (-b) (-a)) by + rw [eqOn_comp_right_iff] at this + convert this + simp + apply ODE_solution_unique_of_mem_set_Icc_right hv' + (hf.comp continuousOn_neg hmt1) _ (fun _ ht ↦ hfs _ (hmt2 ht)) + (hg.comp continuousOn_neg hmt1) _ (fun _ ht ↦ hgs _ (hmt2 ht)) (by simp [hb]) + · intros t ht + convert HasFDerivWithinAt.comp_hasDerivWithinAt t (hf' (-t) (hmt2 ht)) + (hasDerivAt_neg t).hasDerivWithinAt (hmt3 t) + simp + · intros t ht + convert HasFDerivWithinAt.comp_hasDerivWithinAt t (hg' (-t) (hmt2 ht)) + (hasDerivAt_neg t).hasDerivWithinAt (hmt3 t) + simp + +/-- A version of `ODE_solution_unique_of_mem_Icc_right` for uniqueness in a closed interval whose +interior contains the initial time. -/ +theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (heq : f t₀ = g t₀) : EqOn f g (Icc a b) := by + rw [← Icc_union_Icc_eq_Icc (le_of_lt ht.1) (le_of_lt ht.2)] + apply EqOn.union + · have hss : Ioc a t₀ ⊆ Ioo a b := Ioc_subset_Ioo_right ht.2 + exact ODE_solution_unique_of_mem_set_Icc_left hv + (hf.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) + (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) + (hg.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) + (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq + · have hss : Ico t₀ b ⊆ Ioo a b := Ico_subset_Ioo_left ht.1 + exact ODE_solution_unique_of_mem_set_Icc_right hv + (hf.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) + (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) + (hg.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) + (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq + +/-- A version of `ODE_solution_unique_of_mem_set_Icc` for uniqueness in an open interval. -/ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) - (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (ha : f t₀ = g t₀) : EqOn f g (Ioo a b) := by - -- split `Ioo a b = Ioc a t₀ ∪ Ico t₀ b` and treat the two cases separately - apply EqOn.mono (Ioo_subset_Ioc_union_Ico (b := t₀)) - apply EqOn.union - · -- case `t ≤ t₀`: show `fun t ↦ f_or_g (-t)` solves `-v (-t)` within `s (-t)` by composition, - -- so we can use the `Ico` version of the uniqueness lemma backwards in time - have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by - rw [← one_mul K] - exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) - have hmt : MapsTo Neg.neg (Ico (-t₀) (-a)) (Ioo a b) := - fun _ ht' ↦ ⟨lt_neg.mp ht'.2, lt_of_le_of_lt (neg_le.mp ht'.1) ht.2⟩ - suffices EqOn (f ∘ Neg.neg) (g ∘ Neg.neg) (Ico (-t₀) (-a)) by - rw [eqOn_comp_right_iff] at this - convert this - simp - apply ODE_solution_unique_of_mem_set_Ico hv' - (hf.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hfs _ (hmt ht')) - (hg.comp continuousOn_neg hmt) _ (fun _ ht' ↦ hgs _ (hmt ht')) (by simp [ha]) - · intros t' ht' - apply HasDerivAt.hasDerivWithinAt - convert HasFDerivAt.comp_hasDerivAt t' (hf' (-t') (hmt ht')) (hasDerivAt_neg t') - simp - · intros t' ht' - apply HasDerivAt.hasDerivWithinAt - convert HasFDerivAt.comp_hasDerivAt t' (hg' (-t') (hmt ht')) (hasDerivAt_neg t') - simp - · -- case `t ≥ t₀`: follows trivially from the `Ico` version of the uniqueness lemma - have hss := Ico_subset_Ioo_left (b := b) ht.1 - exact ODE_solution_unique_of_mem_set_Ico hv - (hf.mono (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hf' _ (hss ht')) - (fun _ ht' => hfs _ <| (hss ht')) - (hg.mono (Ico_subset_Ioo_left ht.1)) - (fun _ ht' => HasDerivAt.hasDerivWithinAt <| hg' _ (hss ht')) - (fun _ ht' => hgs _ <| (hss ht')) ha + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (heq : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + intros t' ht' + rcases lt_or_le t' t₀ with (h | h) + · have hss : Ioc t' t₀ ⊆ Ioo a b := + fun _ ht'' ↦ ⟨lt_trans ht'.1 ht''.1, lt_of_le_of_lt ht''.2 ht.2⟩ + exact ODE_solution_unique_of_mem_set_Icc_left hv + (hf.mono <| Icc_subset_Ioo ht'.1 ht.2) + (fun _ ht'' ↦ (hf' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hfs _ (hss ht''))) + (hg.mono <| Icc_subset_Ioo ht'.1 ht.2) + (fun _ ht'' ↦ (hg' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hgs _ (hss ht''))) heq + ⟨le_rfl, le_of_lt h⟩ + · have hss : Ico t₀ t' ⊆ Ioo a b := + fun _ ht'' ↦ ⟨lt_of_lt_of_le ht.1 ht''.1, lt_trans ht''.2 ht'.2⟩ + exact ODE_solution_unique_of_mem_set_Icc_right hv + (hf.mono <| Icc_subset_Ioo ht.1 ht'.2) + (fun _ ht'' ↦ (hf' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hfs _ (hss ht''))) + (hg.mono <| Icc_subset_Ioo ht.1 ht'.2) + (fun _ ht'' ↦ (hg' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hgs _ (hss ht''))) heq + ⟨h, le_rfl⟩ /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ @@ -316,7 +334,7 @@ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, Lip (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - ODE_solution_unique_of_mem_set_Icc (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' + ODE_solution_unique_of_mem_set_Icc_right (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align ODE_solution_unique ODE_solution_unique From 266b14bca9370159d22098b19c51cec29dea4dd8 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 19 Jan 2024 21:06:44 -0800 Subject: [PATCH 185/203] eventually --- Mathlib/Analysis/ODE/Gronwall.lean | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index cfa6e8f0ae6ba..9333ca008f7c0 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -326,6 +326,27 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se (fun _ ht'' ↦ (hg' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hgs _ (hss ht''))) heq ⟨h, le_rfl⟩ +/-- Local unqueness of ODE solutions. -/ +theorem ODE_solution_unique_of_mem_set_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {t₀ : ℝ} + (hf : ∀ᶠ t in 𝓝 t₀, ContinuousAt f t) + (hf' : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t) + (hfs : ∀ᶠ t in 𝓝 t₀, f t ∈ s t) + (hg : ∀ᶠ t in 𝓝 t₀, ContinuousAt g t) + (hg' : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t) + (hgs : ∀ᶠ t in 𝓝 t₀, g t ∈ s t) + (heq : f t₀ = g t₀) : f =ᶠ[𝓝 t₀] g := by + obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp <| + hf.and <| hf'.and <| hfs.and <| hg.and <| hg'.and <| hgs + rw [Filter.eventuallyEq_iff_exists_mem] + refine ⟨ball t₀ ε, ball_mem_nhds _ hε, ?_⟩ + simp_rw [Real.ball_eq_Ioo] at * + exact ODE_solution_unique_of_mem_set_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) + (ContinuousAt.continuousOn fun _ ht ↦ (h _ ht).1) + (fun _ ht ↦ (h _ ht).2.1) (fun _ ht ↦ (h _ ht).2.2.1) + (ContinuousAt.continuousOn fun _ ht ↦ (h _ ht).2.2.2.1) + (fun _ ht ↦ (h _ ht).2.2.2.2.1) (fun _ ht ↦ (h _ ht).2.2.2.2.2) heq + /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) From 831d9b3e4b67c5a07e71a5236e4948c22113025c Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 02:43:53 -0800 Subject: [PATCH 186/203] EventuallyEq version compiles --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 127 ++++++++++++++++--- 1 file changed, 110 insertions(+), 17 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 12c71dbc7e23f..0b351589eb2c0 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -138,6 +138,21 @@ lemma isIntegralCurveOn_iff_isIntegralCurveAt (hs : IsOpen s) : IsIntegralCurveOn γ v s ↔ ∀ t ∈ s, IsIntegralCurveAt γ v t := ⟨fun h _ ht ↦ h.isIntegralCurveAt (hs.mem_nhds ht), IsIntegralCurveAt.isIntegralCurveOn⟩ +lemma IsIntegralCurveOn.continuousAt (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ∈ s) : + ContinuousAt γ t₀ := (hγ t₀ ht).1 + +lemma IsIntegralCurveOn.continuousOn (hγ : IsIntegralCurveOn γ v s) : + ContinuousOn γ s := fun t ht ↦ (hγ t ht).1.continuousWithinAt + +lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : + ContinuousAt γ t₀ := + have ⟨_, hs, hγ⟩ := isIntegralCurveAt_iff.mp hγ + hγ.continuousAt <| mem_of_mem_nhds hs + +lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : + Continuous γ := continuous_iff_continuousAt.mpr + fun _ ↦ (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) + /-- If `γ` is an integral curve of a vector field `v`, then `γ t` is tangent to `v (γ t)` when expressed in the local chart around the initial point `γ t₀`. -/ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (ht : t ∈ s) @@ -156,6 +171,26 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h mfderiv_chartAt_eq_tangentCoordChange I hsrc] rfl +-- TODO: cleanup +lemma IsIntegralCurveAt.eventually_hasDerivAt (hγ : IsIntegralCurveAt γ v t₀) : + ∀ᶠ t in 𝓝 t₀, HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) + (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by + have := eventually_mem_nhds.mpr + (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I _)) |>.and hγ + apply this.mono + rintro t ⟨ht1, ht2⟩ + have hsrc := mem_of_mem_nhds ht1 + rw [mem_preimage, extChartAt_source I (γ t₀)] at hsrc + rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] + apply (HasMFDerivAt.comp t + (hasMFDerivAt_extChartAt I hsrc) ht2).congr_mfderiv + rw [ContinuousLinearMap.ext_iff] + intro a + rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, + ← ContinuousLinearMap.one_apply (R₁ := ℝ) a, ← ContinuousLinearMap.smulRight_apply, + mfderiv_chartAt_eq_tangentCoordChange I hsrc] + rfl + /-! ### Translation lemmas -/ section Translation @@ -270,21 +305,6 @@ lemma isIntegralCurve_const {x : M} (h : v x = 0) : IsIntegralCurve (fun _ ↦ x ContinuousLinearMap.smulRight_one_one] exact hasMFDerivAt_const .. -lemma IsIntegralCurveOn.continuousAt (hγ : IsIntegralCurveOn γ v s) (ht : t₀ ∈ s) : - ContinuousAt γ t₀ := (hγ t₀ ht).1 - -lemma IsIntegralCurveOn.continuousOn (hγ : IsIntegralCurveOn γ v s) : - ContinuousOn γ s := fun t ht ↦ (hγ t ht).1.continuousWithinAt - -lemma IsIntegralCurveAt.continuousAt (hγ : IsIntegralCurveAt γ v t₀) : - ContinuousAt γ t₀ := - have ⟨_, hs, hγ⟩ := isIntegralCurveAt_iff.mp hγ - hγ.continuousAt <| mem_of_mem_nhds hs - -lemma IsIntegralCurve.continuous (hγ : IsIntegralCurve γ v) : - Continuous γ := continuous_iff_continuousAt.mpr - fun _ ↦ (hγ.isIntegralCurveOn univ).continuousAt (mem_univ _) - end Scaling /-! ### Existence and uniqueness -/ @@ -356,11 +376,84 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold variable {t₀} +-- TODO: cleanup /-- Local integral curves are unique. If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) + (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) + (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : + γ =ᶠ[𝓝 t₀] γ' := by + -- first define `v'` as the vector field expressed in the local chart around `γ t₀` + -- this is basically what the function looks like when `hv` is unfolded + set v' : E → E := fun x ↦ + tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) + (v ((extChartAt I (γ t₀)).symm x)) with hv' + + -- extract a set `s` on which `v'` is Lipschitz + rw [contMDiffAt_iff] at hv + obtain ⟨_, hv⟩ := hv + obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := + (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd.exists_lipschitzOnWith + have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip + + have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) + rw [← eventually_mem_nhds] at hsrc + have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) + rw [← eventually_mem_nhds] at hsrc' + + suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' by + have hh := this.fun_comp (extChartAt I (γ t₀)).symm + refine Filter.EventuallyEq.trans ?_ (Filter.EventuallyEq.trans hh ?_) + · apply hsrc.mono -- extract lemma? + intros t ht + rw [Function.comp_apply, Function.comp_apply, + PartialEquiv.left_inv _ (mem_preimage.mp <| mem_of_mem_nhds ht)] + · apply hsrc'.mono + intros t ht + rw [Function.comp_apply, Function.comp_apply, h, + PartialEquiv.left_inv _ (mem_preimage.mp <| mem_of_mem_nhds ht)] + apply ODE_solution_unique_of_mem_set_eventually hlip + · -- extract lemma `IsIntegralCurveAt.eventually_continuousAt`? + apply (hsrc.and hγ).mono + rintro t ⟨ht1, ht2⟩ + exact (continuousAt_extChartAt' _ _ (mem_preimage.mp <| mem_of_mem_nhds ht1)).comp + ht2.continuousAt + · apply (hsrc.and hγ.eventually_hasDerivAt).mono + rintro t ⟨ht1, ht2⟩ + rw [hv'] + apply ht2.congr_deriv + have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by + rw [Function.comp_apply, PartialEquiv.left_inv] + exact mem_preimage.mp <| mem_of_mem_nhds ht1 + rw [this] + · apply ((continuousAt_extChartAt I (γ t₀)).comp hγ.continuousAt).preimage_mem_nhds + rw [Function.comp_apply] + exact hs + -- repeat + · apply (hsrc'.and hγ').mono + rintro t ⟨ht1, ht2⟩ + exact (continuousAt_extChartAt' _ _ (mem_preimage.mp <| mem_of_mem_nhds ht1)).comp + ht2.continuousAt + · apply (hsrc'.and hγ'.eventually_hasDerivAt).mono + rintro t ⟨ht1, ht2⟩ + rw [hv', h] + apply ht2.congr_deriv + have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by + rw [Function.comp_apply, PartialEquiv.left_inv] + exact mem_preimage.mp <| mem_of_mem_nhds ht1 + rw [this] + · apply ((continuousAt_extChartAt I (γ' t₀)).comp hγ'.continuousAt).preimage_mem_nhds + rw [Function.comp_apply, ← h] + exact hs + · rw [Function.comp_apply, Function.comp_apply, h] + +/-- Local integral curves are unique. + +If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with +`γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ +theorem isIntegralCurveAt_eqOn_of_contMDiffAt' (hγt₀ : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by @@ -453,7 +546,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + isIntegralCurveAt_eqOn_of_contMDiffAt' (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h variable [T2Space M] {a b : ℝ} @@ -496,7 +589,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) intro t₁ ht₁ rw [mem_nhds_iff] obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt + isIntegralCurveAt_eqOn_of_contMDiffAt' (hγt _ ht₁.2) hv.contMDiffAt (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) ht₁.1 From b4bd4f7f514402d1ba750620c3e571951f24fd87 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 02:59:56 -0800 Subject: [PATCH 187/203] use new theorem --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 21 +++++++------------- 1 file changed, 7 insertions(+), 14 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 0b351589eb2c0..232aa433cee05 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -545,8 +545,8 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt' (hγt₀ : I.IsInteriorPoint (γ theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : - ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt' (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + γ =ᶠ[𝓝 t₀] γ' := + isIntegralCurveAt_eqOn_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h variable [T2Space M] {a b : ℝ} @@ -587,18 +587,11 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) exact hγ'.continuousAt ht · rw [isOpen_iff_mem_nhds] intro t₁ ht₁ - rw [mem_nhds_iff] - obtain ⟨ε, hε, heqon⟩ : ∃ ε > 0, EqOn γ γ' (Ioo (t₁ - ε) (t₁ + ε)) := - isIntegralCurveAt_eqOn_of_contMDiffAt' (hγt _ ht₁.2) hv.contMDiffAt - (hγ.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - (hγ'.isIntegralCurveAt <| Ioo_mem_nhds ht₁.2.1 ht₁.2.2) - ht₁.1 - refine ⟨Ioo (max a (t₁ - ε)) (min b (t₁ + ε)), - subset_inter - (fun t ht ↦ @heqon t <| mem_of_mem_of_subset ht <| Ioo_subset_Ioo (by simp) (by simp)) - (Ioo_subset_Ioo (by simp) (by simp)), - isOpen_Ioo, ?_⟩ - exact ⟨max_lt ht₁.2.1 (by simp [hε]), lt_min ht₁.2.2 (by simp [hε])⟩ + have hmem := Ioo_mem_nhds ht₁.2.1 ht₁.2.2 + have heq : γ =ᶠ[𝓝 t₁] γ' := isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt + (hγ.isIntegralCurveAt hmem) (hγ'.isIntegralCurveAt hmem) ht₁.1 + apply (heq.and hmem).mono + exact fun _ ht ↦ ht theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless [BoundarylessManifold I M] (ht₀ : t₀ ∈ Ioo a b) From ce92a983c0c04af3f9d57c021bb506ed90b08bd5 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 03:38:57 -0800 Subject: [PATCH 188/203] golf a bit --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 24 ++++++++++---------- 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 232aa433cee05..356f43aa87c22 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -171,19 +171,16 @@ lemma IsIntegralCurveOn.hasDerivAt (hγ : IsIntegralCurveOn γ v s) {t : ℝ} (h mfderiv_chartAt_eq_tangentCoordChange I hsrc] rfl --- TODO: cleanup lemma IsIntegralCurveAt.eventually_hasDerivAt (hγ : IsIntegralCurveAt γ v t₀) : ∀ᶠ t in 𝓝 t₀, HasDerivAt ((extChartAt I (γ t₀)) ∘ γ) (tangentCoordChange I (γ t) (γ t₀) (γ t) (v (γ t))) t := by - have := eventually_mem_nhds.mpr - (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I _)) |>.and hγ - apply this.mono + apply eventually_mem_nhds.mpr + (hγ.continuousAt.preimage_mem_nhds (extChartAt_source_mem_nhds I _)) |>.and hγ |>.mono rintro t ⟨ht1, ht2⟩ have hsrc := mem_of_mem_nhds ht1 rw [mem_preimage, extChartAt_source I (γ t₀)] at hsrc rw [hasDerivAt_iff_hasFDerivAt, ← hasMFDerivAt_iff_hasFDerivAt] - apply (HasMFDerivAt.comp t - (hasMFDerivAt_extChartAt I hsrc) ht2).congr_mfderiv + apply (HasMFDerivAt.comp t (hasMFDerivAt_extChartAt I hsrc) ht2).congr_mfderiv rw [ContinuousLinearMap.ext_iff] intro a rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.smulRight_apply, map_smul, @@ -403,22 +400,25 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) rw [← eventually_mem_nhds] at hsrc' + have hgt {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : + g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht + suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' by have hh := this.fun_comp (extChartAt I (γ t₀)).symm refine Filter.EventuallyEq.trans ?_ (Filter.EventuallyEq.trans hh ?_) · apply hsrc.mono -- extract lemma? intros t ht rw [Function.comp_apply, Function.comp_apply, - PartialEquiv.left_inv _ (mem_preimage.mp <| mem_of_mem_nhds ht)] + PartialEquiv.left_inv _ (hgt ht)] · apply hsrc'.mono intros t ht rw [Function.comp_apply, Function.comp_apply, h, - PartialEquiv.left_inv _ (mem_preimage.mp <| mem_of_mem_nhds ht)] + PartialEquiv.left_inv _ (hgt ht)] apply ODE_solution_unique_of_mem_set_eventually hlip · -- extract lemma `IsIntegralCurveAt.eventually_continuousAt`? apply (hsrc.and hγ).mono rintro t ⟨ht1, ht2⟩ - exact (continuousAt_extChartAt' _ _ (mem_preimage.mp <| mem_of_mem_nhds ht1)).comp + exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt · apply (hsrc.and hγ.eventually_hasDerivAt).mono rintro t ⟨ht1, ht2⟩ @@ -426,7 +426,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t apply ht2.congr_deriv have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by rw [Function.comp_apply, PartialEquiv.left_inv] - exact mem_preimage.mp <| mem_of_mem_nhds ht1 + exact hgt ht1 rw [this] · apply ((continuousAt_extChartAt I (γ t₀)).comp hγ.continuousAt).preimage_mem_nhds rw [Function.comp_apply] @@ -434,7 +434,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t -- repeat · apply (hsrc'.and hγ').mono rintro t ⟨ht1, ht2⟩ - exact (continuousAt_extChartAt' _ _ (mem_preimage.mp <| mem_of_mem_nhds ht1)).comp + exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt · apply (hsrc'.and hγ'.eventually_hasDerivAt).mono rintro t ⟨ht1, ht2⟩ @@ -442,7 +442,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t apply ht2.congr_deriv have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by rw [Function.comp_apply, PartialEquiv.left_inv] - exact mem_preimage.mp <| mem_of_mem_nhds ht1 + exact hgt ht1 rw [this] · apply ((continuousAt_extChartAt I (γ' t₀)).comp hγ'.continuousAt).preimage_mem_nhds rw [Function.comp_apply, ← h] From f47dc111568ba4ffef95b5c1a92c513c28d708aa Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 03:55:34 -0800 Subject: [PATCH 189/203] golf some more --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 31 ++++++++------------ 1 file changed, 13 insertions(+), 18 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 356f43aa87c22..36371e9d382a9 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -395,10 +395,9 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd.exists_lipschitzOnWith have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip - have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) - rw [← eventually_mem_nhds] at hsrc - have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) - rw [← eventually_mem_nhds] at hsrc' + have hsrc {g : ℝ → M} (hg : IsIntegralCurveAt g v t₀) : + ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds.mpr <| + continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds I (g t₀) have hgt {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht @@ -406,44 +405,40 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' by have hh := this.fun_comp (extChartAt I (γ t₀)).symm refine Filter.EventuallyEq.trans ?_ (Filter.EventuallyEq.trans hh ?_) - · apply hsrc.mono -- extract lemma? + · apply (hsrc hγ).mono -- extract lemma? intros t ht rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hgt ht)] - · apply hsrc'.mono + · apply (hsrc hγ').mono intros t ht rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hgt ht)] apply ODE_solution_unique_of_mem_set_eventually hlip · -- extract lemma `IsIntegralCurveAt.eventually_continuousAt`? - apply (hsrc.and hγ).mono + apply (hsrc hγ |>.and hγ).mono rintro t ⟨ht1, ht2⟩ exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt - · apply (hsrc.and hγ.eventually_hasDerivAt).mono + · apply (hsrc hγ |>.and hγ.eventually_hasDerivAt).mono rintro t ⟨ht1, ht2⟩ rw [hv'] apply ht2.congr_deriv - have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hgt ht1 - rw [this] + congr <;> + rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] · apply ((continuousAt_extChartAt I (γ t₀)).comp hγ.continuousAt).preimage_mem_nhds rw [Function.comp_apply] exact hs -- repeat - · apply (hsrc'.and hγ').mono + · apply (hsrc hγ' |>.and hγ').mono rintro t ⟨ht1, ht2⟩ exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt - · apply (hsrc'.and hγ'.eventually_hasDerivAt).mono + · apply (hsrc hγ' |>.and hγ'.eventually_hasDerivAt).mono rintro t ⟨ht1, ht2⟩ rw [hv', h] apply ht2.congr_deriv - have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hgt ht1 - rw [this] + congr <;> + rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] · apply ((continuousAt_extChartAt I (γ' t₀)).comp hγ'.continuousAt).preimage_mem_nhds rw [Function.comp_apply, ← h] exact hs From 6e4fd1a72ba08f4f0fe94a512d35f8c749dbcf31 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 07:43:41 -0800 Subject: [PATCH 190/203] golf --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 65 ++++++++------------ 1 file changed, 26 insertions(+), 39 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 36371e9d382a9..daafd083a86ad 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -394,55 +394,42 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd.exists_lipschitzOnWith have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip - - have hsrc {g : ℝ → M} (hg : IsIntegralCurveAt g v t₀) : + -- internal lemmas to reduce code duplication + have hsrc {g} (hg : IsIntegralCurveAt g v t₀) : ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds.mpr <| continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds I (g t₀) - have hgt {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht - - suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' by - have hh := this.fun_comp (extChartAt I (γ t₀)).symm - refine Filter.EventuallyEq.trans ?_ (Filter.EventuallyEq.trans hh ?_) - · apply (hsrc hγ).mono -- extract lemma? - intros t ht - rw [Function.comp_apply, Function.comp_apply, - PartialEquiv.left_inv _ (hgt ht)] - · apply (hsrc hγ').mono - intros t ht - rw [Function.comp_apply, Function.comp_apply, h, - PartialEquiv.left_inv _ (hgt ht)] - apply ODE_solution_unique_of_mem_set_eventually hlip - · -- extract lemma `IsIntegralCurveAt.eventually_continuousAt`? - apply (hsrc hγ |>.and hγ).mono - rintro t ⟨ht1, ht2⟩ - exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp - ht2.continuousAt - · apply (hsrc hγ |>.and hγ.eventually_hasDerivAt).mono - rintro t ⟨ht1, ht2⟩ - rw [hv'] - apply ht2.congr_deriv - congr <;> - rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] - · apply ((continuousAt_extChartAt I (γ t₀)).comp hγ.continuousAt).preimage_mem_nhds - rw [Function.comp_apply] - exact hs - -- repeat - · apply (hsrc hγ' |>.and hγ').mono + have heq {g} (hg : IsIntegralCurveAt g v t₀) : + g =ᶠ[𝓝 t₀] (extChartAt I (g t₀)).symm ∘ ↑(extChartAt I (g t₀)) ∘ g := by + apply (hsrc hg).mono + intros t ht + rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hgt ht)] + have hcont {g} (hg : IsIntegralCurveAt g v t₀) : + ∀ᶠ t in 𝓝 t₀, ContinuousAt ((extChartAt I (g t₀)) ∘ g) t := by + apply (hsrc hg |>.and hg).mono rintro t ⟨ht1, ht2⟩ - exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp - ht2.continuousAt - · apply (hsrc hγ' |>.and hγ'.eventually_hasDerivAt).mono + exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt + have hdrv {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : + ∀ᶠ (t : ℝ) in 𝓝 t₀, HasDerivAt ((extChartAt I (g t₀)) ∘ g) + ((fun _ ↦ v') t (((extChartAt I (g t₀)) ∘ g) t)) t := by + apply (hsrc hg |>.and hg.eventually_hasDerivAt).mono rintro t ⟨ht1, ht2⟩ - rw [hv', h] + rw [hv', h'] apply ht2.congr_deriv congr <;> rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] - · apply ((continuousAt_extChartAt I (γ' t₀)).comp hγ'.continuousAt).preimage_mem_nhds - rw [Function.comp_apply, ← h] + have hmem {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : + ∀ᶠ t in 𝓝 t₀, ((extChartAt I (g t₀)) ∘ g) t ∈ (fun _ ↦ s) t := by + apply ((continuousAt_extChartAt I (g t₀)).comp hg.continuousAt).preimage_mem_nhds + rw [Function.comp_apply, ← h'] exact hs - · rw [Function.comp_apply, Function.comp_apply, h] + -- main proof + suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from + (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) + apply ODE_solution_unique_of_mem_set_eventually hlip + (hcont hγ) (hdrv hγ rfl) (hmem hγ rfl) (hcont hγ') (hdrv hγ' h) (hmem hγ' h) + rw [Function.comp_apply, Function.comp_apply, h] /-- Local integral curves are unique. From ba7629cb796afa2782fda6890cb7af9cd6762643 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 07:50:08 -0800 Subject: [PATCH 191/203] naming --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index daafd083a86ad..c77265fcbd158 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -378,7 +378,7 @@ variable {t₀} If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) +theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : γ =ᶠ[𝓝 t₀] γ' := by @@ -435,7 +435,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with `γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt' (hγt₀ : I.IsInteriorPoint (γ t₀)) +theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by @@ -528,7 +528,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : γ =ᶠ[𝓝 t₀] γ' := - isIntegralCurveAt_eqOn_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h + isIntegralCurveAt_eventuallyEq_of_contMDiffAt (BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h variable [T2Space M] {a b : ℝ} @@ -570,7 +570,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) · rw [isOpen_iff_mem_nhds] intro t₁ ht₁ have hmem := Ioo_mem_nhds ht₁.2.1 ht₁.2.2 - have heq : γ =ᶠ[𝓝 t₁] γ' := isIntegralCurveAt_eqOn_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt + have heq : γ =ᶠ[𝓝 t₁] γ' := isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt (hγ.isIntegralCurveAt hmem) (hγ'.isIntegralCurveAt hmem) ht₁.1 apply (heq.and hmem).mono exact fun _ ht ↦ ht From f29fc5d7aabeae03eb058d42533717ae4a6afc55 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 07:51:01 -0800 Subject: [PATCH 192/203] naming --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c77265fcbd158..c73b234746441 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -524,7 +524,7 @@ theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t · intros t ht rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] -theorem isIntegralCurveAt_eqOn_of_contMDiffAt_boundaryless [BoundarylessManifold I M] +theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : γ =ᶠ[𝓝 t₀] γ' := From 98c3c82a2c9173ac6a9b1e83f96fadc994d78ccd Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 07:54:26 -0800 Subject: [PATCH 193/203] shorten hypothesis --- Mathlib/Analysis/ODE/Gronwall.lean | 12 ++++-------- 1 file changed, 4 insertions(+), 8 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 9333ca008f7c0..87d1164bae2a4 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -329,15 +329,11 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se /-- Local unqueness of ODE solutions. -/ theorem ODE_solution_unique_of_mem_set_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {t₀ : ℝ} - (hf : ∀ᶠ t in 𝓝 t₀, ContinuousAt f t) - (hf' : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t) - (hfs : ∀ᶠ t in 𝓝 t₀, f t ∈ s t) - (hg : ∀ᶠ t in 𝓝 t₀, ContinuousAt g t) - (hg' : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t) - (hgs : ∀ᶠ t in 𝓝 t₀, g t ∈ s t) + (h : ∀ᶠ t in 𝓝 t₀, + ContinuousAt f t ∧ HasDerivAt f (v t (f t)) t ∧ f t ∈ s t ∧ + ContinuousAt g t ∧ HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) (heq : f t₀ = g t₀) : f =ᶠ[𝓝 t₀] g := by - obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp <| - hf.and <| hf'.and <| hfs.and <| hg.and <| hg'.and <| hgs + obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp h rw [Filter.eventuallyEq_iff_exists_mem] refine ⟨ball t₀ ε, ball_mem_nhds _ hε, ?_⟩ simp_rw [Real.ball_eq_Ioo] at * From fda10a0a3555dbe8dee361a5d9bad5f591a159f3 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 07:56:31 -0800 Subject: [PATCH 194/203] 100 --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c73b234746441..55a9ccc1e0bab 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -570,8 +570,8 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) · rw [isOpen_iff_mem_nhds] intro t₁ ht₁ have hmem := Ioo_mem_nhds ht₁.2.1 ht₁.2.2 - have heq : γ =ᶠ[𝓝 t₁] γ' := isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt _ ht₁.2) hv.contMDiffAt - (hγ.isIntegralCurveAt hmem) (hγ'.isIntegralCurveAt hmem) ht₁.1 + have heq : γ =ᶠ[𝓝 t₁] γ' := isIntegralCurveAt_eventuallyEq_of_contMDiffAt + (hγt _ ht₁.2) hv.contMDiffAt (hγ.isIntegralCurveAt hmem) (hγ'.isIntegralCurveAt hmem) ht₁.1 apply (heq.and hmem).mono exact fun _ ht ↦ ht From 3c1cc52a5191c5cfa91779aa0c58d1ab0845b1ad Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 08:47:50 -0800 Subject: [PATCH 195/203] simplify assumptions --- Mathlib/Analysis/ODE/Gronwall.lean | 50 +++++++++++++++--------------- 1 file changed, 25 insertions(+), 25 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 87d1164bae2a4..c64c0fdd55327 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -303,45 +303,45 @@ theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Se /-- A version of `ODE_solution_unique_of_mem_set_Icc` for uniqueness in an open interval. -/ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Ioo a b)) - (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) - (hg : ContinuousOn g (Ioo a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) - (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (heq : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) + (hf : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) + (hg : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) + (heq : f t₀ = g t₀) : EqOn f g (Ioo a b) := by intros t' ht' rcases lt_or_le t' t₀ with (h | h) - · have hss : Ioc t' t₀ ⊆ Ioo a b := - fun _ ht'' ↦ ⟨lt_trans ht'.1 ht''.1, lt_of_le_of_lt ht''.2 ht.2⟩ + · have hss : Icc t' t₀ ⊆ Ioo a b := + fun _ ht'' ↦ ⟨lt_of_lt_of_le ht'.1 ht''.1, lt_of_le_of_lt ht''.2 ht.2⟩ exact ODE_solution_unique_of_mem_set_Icc_left hv - (hf.mono <| Icc_subset_Ioo ht'.1 ht.2) - (fun _ ht'' ↦ (hf' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hfs _ (hss ht''))) - (hg.mono <| Icc_subset_Ioo ht'.1 ht.2) - (fun _ ht'' ↦ (hg' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hgs _ (hss ht''))) heq + (ContinuousAt.continuousOn fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) + (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').1.hasDerivWithinAt) + (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').2) + (ContinuousAt.continuousOn fun _ ht'' ↦ (hg _ <| hss ht'').1.continuousAt) + (fun _ ht'' ↦ (hg _ <| hss <| Ioc_subset_Icc_self ht'').1.hasDerivWithinAt) + (fun _ ht'' ↦ (hg _ <| hss <| Ioc_subset_Icc_self ht'').2) heq ⟨le_rfl, le_of_lt h⟩ - · have hss : Ico t₀ t' ⊆ Ioo a b := - fun _ ht'' ↦ ⟨lt_of_lt_of_le ht.1 ht''.1, lt_trans ht''.2 ht'.2⟩ + · have hss : Icc t₀ t' ⊆ Ioo a b := + fun _ ht'' ↦ ⟨lt_of_lt_of_le ht.1 ht''.1, lt_of_le_of_lt ht''.2 ht'.2⟩ exact ODE_solution_unique_of_mem_set_Icc_right hv - (hf.mono <| Icc_subset_Ioo ht.1 ht'.2) - (fun _ ht'' ↦ (hf' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hfs _ (hss ht''))) - (hg.mono <| Icc_subset_Ioo ht.1 ht'.2) - (fun _ ht'' ↦ (hg' _ (hss ht'')).hasDerivWithinAt) (fun _ ht'' ↦ (hgs _ (hss ht''))) heq + (ContinuousAt.continuousOn fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) + (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').1.hasDerivWithinAt) + (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').2) + (ContinuousAt.continuousOn fun _ ht'' ↦ (hg _ <| hss ht'').1.continuousAt) + (fun _ ht'' ↦ (hg _ <| hss <| Ico_subset_Icc_self ht'').1.hasDerivWithinAt) + (fun _ ht'' ↦ (hg _ <| hss <| Ico_subset_Icc_self ht'').2) heq ⟨h, le_rfl⟩ /-- Local unqueness of ODE solutions. -/ theorem ODE_solution_unique_of_mem_set_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {t₀ : ℝ} - (h : ∀ᶠ t in 𝓝 t₀, - ContinuousAt f t ∧ HasDerivAt f (v t (f t)) t ∧ f t ∈ s t ∧ - ContinuousAt g t ∧ HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) + (hf : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) + (hg : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) (heq : f t₀ = g t₀) : f =ᶠ[𝓝 t₀] g := by - obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp h + obtain ⟨ε, hε, h⟩ := eventually_nhds_iff_ball.mp (hf.and hg) rw [Filter.eventuallyEq_iff_exists_mem] refine ⟨ball t₀ ε, ball_mem_nhds _ hε, ?_⟩ simp_rw [Real.ball_eq_Ioo] at * - exact ODE_solution_unique_of_mem_set_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) - (ContinuousAt.continuousOn fun _ ht ↦ (h _ ht).1) - (fun _ ht ↦ (h _ ht).2.1) (fun _ ht ↦ (h _ ht).2.2.1) - (ContinuousAt.continuousOn fun _ ht ↦ (h _ ht).2.2.2.1) - (fun _ ht ↦ (h _ ht).2.2.2.2.1) (fun _ ht ↦ (h _ ht).2.2.2.2.2) heq + apply ODE_solution_unique_of_mem_set_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) + (fun _ ht ↦ (h _ ht).1) (fun _ ht ↦ (h _ ht).2) heq /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ From d1d4fe1b7c531f1e5f9923a3daa7150d9cd7ea44 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 09:13:00 -0800 Subject: [PATCH 196/203] match gronwall --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 130 ++----------------- 1 file changed, 14 insertions(+), 116 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 55a9ccc1e0bab..c93cac79ac1aa 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -373,7 +373,6 @@ lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold variable {t₀} --- TODO: cleanup /-- Local integral curves are unique. If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with @@ -387,7 +386,6 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi set v' : E → E := fun x ↦ tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) (v ((extChartAt I (γ t₀)).symm x)) with hv' - -- extract a set `s` on which `v'` is Lipschitz rw [contMDiffAt_iff] at hv obtain ⟨_, hv⟩ := hv @@ -405,124 +403,24 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi apply (hsrc hg).mono intros t ht rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hgt ht)] - have hcont {g} (hg : IsIntegralCurveAt g v t₀) : - ∀ᶠ t in 𝓝 t₀, ContinuousAt ((extChartAt I (g t₀)) ∘ g) t := by - apply (hsrc hg |>.and hg).mono - rintro t ⟨ht1, ht2⟩ - exact (continuousAt_extChartAt' _ _ (hgt ht1)).comp ht2.continuousAt - have hdrv {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : - ∀ᶠ (t : ℝ) in 𝓝 t₀, HasDerivAt ((extChartAt I (g t₀)) ∘ g) - ((fun _ ↦ v') t (((extChartAt I (g t₀)) ∘ g) t)) t := by - apply (hsrc hg |>.and hg.eventually_hasDerivAt).mono - rintro t ⟨ht1, ht2⟩ - rw [hv', h'] - apply ht2.congr_deriv - congr <;> - rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] - have hmem {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : - ∀ᶠ t in 𝓝 t₀, ((extChartAt I (g t₀)) ∘ g) t ∈ (fun _ ↦ s) t := by - apply ((continuousAt_extChartAt I (g t₀)).comp hg.continuousAt).preimage_mem_nhds - rw [Function.comp_apply, ← h'] - exact hs + have hdrv {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : ∀ᶠ t in 𝓝 t₀, + HasDerivAt ((extChartAt I (g t₀)) ∘ g) ((fun _ ↦ v') t (((extChartAt I (g t₀)) ∘ g) t)) t ∧ + ((extChartAt I (g t₀)) ∘ g) t ∈ (fun _ ↦ s) t := by + apply Filter.Eventually.and + · apply (hsrc hg |>.and hg.eventually_hasDerivAt).mono + rintro t ⟨ht1, ht2⟩ + rw [hv', h'] + apply ht2.congr_deriv + congr <;> + rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] + · apply ((continuousAt_extChartAt I (g t₀)).comp hg.continuousAt).preimage_mem_nhds + rw [Function.comp_apply, ← h'] + exact hs -- main proof suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) apply ODE_solution_unique_of_mem_set_eventually hlip - (hcont hγ) (hdrv hγ rfl) (hmem hγ rfl) (hcont hγ') (hdrv hγ' h) (hmem hγ' h) - rw [Function.comp_apply, Function.comp_apply, h] - -/-- Local integral curves are unique. - -If a $C^1$ vector field `v` admits two local integral curves `γ γ' : ℝ → M` at `t₀` with -`γ t₀ = γ' t₀`, then `γ` and `γ'` agree on some open interval containing `t₀`. -/ -theorem isIntegralCurveAt_eqOn_of_contMDiffAt (hγt₀ : I.IsInteriorPoint (γ t₀)) - (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) - (hγ : IsIntegralCurveAt γ v t₀) (hγ' : IsIntegralCurveAt γ' v t₀) (h : γ t₀ = γ' t₀) : - ∃ ε > 0, EqOn γ γ' (Ioo (t₀ - ε) (t₀ + ε)) := by - -- first define `v'` as the vector field expressed in the local chart around `γ t₀` - -- this is basically what the function looks like when `hv` is unfolded - set v' : E → E := fun x ↦ - tangentCoordChange I ((extChartAt I (γ t₀)).symm x) (γ t₀) ((extChartAt I (γ t₀)).symm x) - (v ((extChartAt I (γ t₀)).symm x)) with hv' - - -- extract a set `s` on which `v'` is Lipschitz - rw [contMDiffAt_iff] at hv - obtain ⟨_, hv⟩ := hv - obtain ⟨K, s, hs, hlip⟩ : ∃ K, ∃ s ∈ nhds _, LipschitzOnWith K v' s := - (hv.contDiffAt (range_mem_nhds_isInteriorPoint hγt₀)).snd.exists_lipschitzOnWith - have hlip (t : ℝ) : LipschitzOnWith K ((fun _ ↦ v') t) ((fun _ ↦ s) t) := hlip - - -- `γ t` when expressed in the local chart should remain inside `s` - have hcont : ContinuousAt ((extChartAt I (γ t₀)) ∘ γ) t₀ := - (continuousAt_extChartAt ..).comp hγ.continuousAt - rw [continuousAt_def] at hcont - have hnhds := hcont _ hs - rw [← eventually_mem_nhds] at hnhds - - -- `γ t` should remain inside the domain of the local chart around `γ t₀` - have hsrc := continuousAt_def.mp hγ.continuousAt _ <| extChartAt_source_mem_nhds I (γ t₀) - rw [← eventually_mem_nhds] at hsrc - - -- same as above but for `γ'` - have hcont' : ContinuousAt ((extChartAt I (γ' t₀)) ∘ γ') t₀ := - ContinuousAt.comp (continuousAt_extChartAt ..) hγ'.continuousAt - rw [continuousAt_def] at hcont' - have hnhds' := hcont' _ (h ▸ hs) - rw [← eventually_mem_nhds] at hnhds' - - have hsrc' := continuousAt_def.mp hγ'.continuousAt _ <| extChartAt_source_mem_nhds I (γ' t₀) - rw [← eventually_mem_nhds] at hsrc' - - -- there exists a neighbourhood around `t₀` in which all of the above hold - have haux := hnhds.and <| hsrc.and <| hγ.and <| hnhds'.and <| hsrc'.and hγ' - rw [Metric.eventually_nhds_iff_ball] at haux - - obtain ⟨ε, hε, haux⟩ := haux - refine ⟨ε, hε, ?_⟩ - - -- break out all the conditions again - have hmem := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).1 - have hsrc := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.1 - have hcrv : IsIntegralCurveOn _ _ _ := fun t ht ↦ (haux t ht).2.2.1 - have hmem' := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.1 - have hsrc' := fun t ht ↦ mem_preimage.mp <| mem_of_mem_nhds (haux t ht).2.2.2.2.1 - have hcrv' : IsIntegralCurveOn _ _ _ := fun t ht ↦ (haux t ht).2.2.2.2.2 - - -- `γ` and `γ'` when expressed in the local chart are continuous on this neighbourhood - have hcont := (continuousOn_extChartAt I (γ t₀)).comp hcrv.continuousOn hsrc - have hcont' := (continuousOn_extChartAt I (γ' t₀)).comp hcrv'.continuousOn hsrc' - - simp_rw [Real.ball_eq_Ioo] at hmem hsrc hcrv hcont hmem' hsrc' hcrv' hcont' - - -- `γ` and `γ'` are equal on `Ioo (t₀ - ε) (t₀ + ε)` when expressed in the local chart - have heqon : EqOn ((extChartAt I (γ t₀)) ∘ γ) ((extChartAt I (γ' t₀)) ∘ γ') - (Ioo (t₀ - ε) (t₀ + ε)) := by - -- uniqueness of ODE solutions in an open interval - apply ODE_solution_unique_of_mem_set_Ioo hlip (t₀ := t₀) - (Real.ball_eq_Ioo _ _ ▸ (Metric.mem_ball_self hε)) hcont _ hmem hcont' _ hmem' (by simp [h]) - · intros t ht - rw [hv'] - have := hcrv.hasDerivAt ht (hsrc t ht) - apply this.congr_deriv - have : γ t = (extChartAt I (γ t₀)).symm (((extChartAt I (γ t₀)) ∘ γ) t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hsrc t ht - rw [this] - · intros t ht - rw [hv', h] - have := hcrv'.hasDerivAt ht (hsrc' t ht) - apply this.congr_deriv - have : γ' t = (extChartAt I (γ' t₀)).symm (((extChartAt I (γ' t₀)) ∘ γ') t) := by - rw [Function.comp_apply, PartialEquiv.left_inv] - exact hsrc' t ht - rw [this] - - -- finally show `EqOn γ γ' _` by composing with the inverse of the local chart around `γ t₀` - refine EqOn.trans ?_ (EqOn.trans (heqon.comp_left (g := (extChartAt I (γ t₀)).symm)) ?_) - · intros t ht - rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hsrc _ ht)] - · intros t ht - rw [Function.comp_apply, Function.comp_apply, h, PartialEquiv.left_inv _ (hsrc' _ ht)] + (hdrv hγ rfl) (hdrv hγ' h) (by rw [Function.comp_apply, Function.comp_apply, h]) theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) (γ t₀)) From aea843d4edf70e0a7a4890738a5fc574e3ef2759 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Sat, 20 Jan 2024 09:16:32 -0800 Subject: [PATCH 197/203] minor --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index c93cac79ac1aa..a1c9812cf5c27 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -396,13 +396,8 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi have hsrc {g} (hg : IsIntegralCurveAt g v t₀) : ∀ᶠ t in 𝓝 t₀, g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t := eventually_mem_nhds.mpr <| continuousAt_def.mp hg.continuousAt _ <| extChartAt_source_mem_nhds I (g t₀) - have hgt {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : + have hmem {g : ℝ → M} {t} (ht : g ⁻¹' (extChartAt I (g t₀)).source ∈ 𝓝 t) : g t ∈ (extChartAt I (g t₀)).source := mem_preimage.mp <| mem_of_mem_nhds ht - have heq {g} (hg : IsIntegralCurveAt g v t₀) : - g =ᶠ[𝓝 t₀] (extChartAt I (g t₀)).symm ∘ ↑(extChartAt I (g t₀)) ∘ g := by - apply (hsrc hg).mono - intros t ht - rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hgt ht)] have hdrv {g} (hg : IsIntegralCurveAt g v t₀) (h' : γ t₀ = g t₀) : ∀ᶠ t in 𝓝 t₀, HasDerivAt ((extChartAt I (g t₀)) ∘ g) ((fun _ ↦ v') t (((extChartAt I (g t₀)) ∘ g) t)) t ∧ ((extChartAt I (g t₀)) ∘ g) t ∈ (fun _ ↦ s) t := by @@ -412,10 +407,15 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi rw [hv', h'] apply ht2.congr_deriv congr <;> - rw [Function.comp_apply, PartialEquiv.left_inv _ (hgt ht1)] + rw [Function.comp_apply, PartialEquiv.left_inv _ (hmem ht1)] · apply ((continuousAt_extChartAt I (g t₀)).comp hg.continuousAt).preimage_mem_nhds rw [Function.comp_apply, ← h'] exact hs + have heq {g} (hg : IsIntegralCurveAt g v t₀) : + g =ᶠ[𝓝 t₀] (extChartAt I (g t₀)).symm ∘ ↑(extChartAt I (g t₀)) ∘ g := by + apply (hsrc hg).mono + intros t ht + rw [Function.comp_apply, Function.comp_apply, PartialEquiv.left_inv _ (hmem ht)] -- main proof suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) From e68b3a6518a6b19ab96f4aabbff025aa745c1a6b Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 22 Jan 2024 16:10:41 -0800 Subject: [PATCH 198/203] extract abs lemma --- Mathlib/Algebra/Order/Ring/Abs.lean | 9 ++++++++ Mathlib/Geometry/Manifold/IntegralCurve.lean | 23 +++++++++----------- 2 files changed, 19 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 56c757b6b3c26..a97ed03a56b8c 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -73,6 +73,15 @@ theorem abs_le_one_iff_mul_self_le_one : |a| ≤ 1 ↔ a * a ≤ 1 := by simpa only [abs_one, one_mul] using @abs_le_iff_mul_self_le α _ a 1 #align abs_le_one_iff_mul_self_le_one abs_le_one_iff_mul_self_le_one +theorem exists_mem_Ioo_neg_self + {α : Type*} [LinearOrderedRing α] (a : α) : ∃ b > 0, |a| < b := by + refine ⟨|a| + 1, lt_of_lt_of_le zero_lt_one <| by simp, ?_⟩ + cases' le_or_lt 0 a with ht ht + · suffices -1 < a + a by simp [abs_eq_self.mpr ht] + exact lt_of_lt_of_le neg_one_lt_zero (by simpa) + · suffices a + a < 1 by simp [abs_eq_neg_self.mpr ht.le] + exact lt_of_lt_of_le (by simpa) zero_le_one + end LinearOrderedRing section LinearOrderedCommRing diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index a1c9812cf5c27..e6eb94c2b5af6 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -419,7 +419,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi -- main proof suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) - apply ODE_solution_unique_of_mem_set_eventually hlip + exact ODE_solution_unique_of_mem_set_eventually hlip (hdrv hγ rfl) (hdrv hγ' h) (by rw [Function.comp_apply, Function.comp_apply, h]) theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [BoundarylessManifold I M] @@ -453,8 +453,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b) intros t ht rw [mem_preimage, ← closure_subtype] at ht revert ht t - apply IsClosed.closure_subset - apply isClosed_eq + apply IsClosed.closure_subset (isClosed_eq _ _) · rw [continuous_iff_continuousAt] rintro ⟨_, ht⟩ apply ContinuousAt.comp _ continuousAt_subtype_val @@ -489,17 +488,15 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t - obtain ⟨T, hT, ht⟩ : ∃ T > 0, t ∈ Ioo (t₀ - T) (t₀ + T) := by - refine ⟨|t - t₀| + 1, by positivity, ?_⟩ - by_cases ht : t - t₀ < 0 - · rw [abs_of_neg ht] - constructor <;> linarith - · rw [abs_of_nonneg (not_lt.mp ht)] - constructor <;> linarith - exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff - (Real.ball_eq_Ioo t₀ T ▸ Metric.mem_ball_self hT) (fun t _ ↦ hγt t) hv + obtain ⟨T, ht₀, ht⟩ : ∃ T, t ∈ Ioo (-T) T ∧ t₀ ∈ Ioo (-T) T := by + obtain ⟨T, hT₁, hT₂⟩ := exists_mem_Ioo_neg_self t + obtain ⟨hT₂, hT₃⟩ := abs_lt.mp hT₂ + obtain ⟨S, hS₁, hS₂⟩ := exists_mem_Ioo_neg_self t₀ + obtain ⟨hS₂, hS₃⟩ := abs_lt.mp hS₂ + exact ⟨T + S, by constructor <;> constructor <;> linarith⟩ + exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff ht (fun t _ ↦ hγt t) hv ((hγ.isIntegralCurveOn _).mono (subset_univ _)) - ((hγ'.isIntegralCurveOn _).mono (subset_univ _)) h ht + ((hγ'.isIntegralCurveOn _).mono (subset_univ _)) h ht₀ theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I M] (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) From 2dd5739c1f296e913c92f055fb5a24878cea5bf0 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 22 Jan 2024 16:14:47 -0800 Subject: [PATCH 199/203] rename --- Mathlib/Analysis/ODE/Gronwall.lean | 44 +++++++++++++++--------------- 1 file changed, 22 insertions(+), 22 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index c64c0fdd55327..91b25423a02d7 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -150,7 +150,7 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem dist_le_of_approx_trajectories_ODE_of_mem {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g f' g' : ℝ → E} {a b : ℝ} {εf εg δ : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) @@ -170,7 +170,7 @@ theorem dist_le_of_approx_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : refine this.trans ((add_le_add (add_le_add (f_bound t ht) (g_bound t ht)) hv).trans ?_) rw [add_comm] set_option linter.uppercaseLean3 false in -#align dist_le_of_approx_trajectories_ODE_of_mem_set dist_le_of_approx_trajectories_ODE_of_mem_set +#align dist_le_of_approx_trajectories_ODE_of_mem_set dist_le_of_approx_trajectories_ODE_of_mem /-- If `f` and `g` are two approximate solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some @@ -185,7 +185,7 @@ theorem dist_le_of_approx_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ gronwallBound δ K (εf + εg) (t - a) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_approx_trajectories_ODE_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' + dist_le_of_approx_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' f_bound hfs hg hg' g_bound (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align dist_le_of_approx_trajectories_ODE dist_le_of_approx_trajectories_ODE @@ -196,7 +196,7 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem dist_le_of_trajectories_ODE_of_mem {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} {δ : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) @@ -207,10 +207,10 @@ theorem dist_le_of_trajectories_ODE_of_mem_set {v : ℝ → E → E} {s : ℝ have g_bound : ∀ t ∈ Ico a b, dist (v t (g t)) (v t (g t)) ≤ 0 := by intros; rw [dist_self] intro t ht have := - dist_le_of_approx_trajectories_ODE_of_mem_set hv hf hf' f_bound hfs hg hg' g_bound hgs ha t ht + dist_le_of_approx_trajectories_ODE_of_mem hv hf hf' f_bound hfs hg hg' g_bound hgs ha t ht rwa [zero_add, gronwallBound_ε0] at this set_option linter.uppercaseLean3 false in -#align dist_le_of_trajectories_ODE_of_mem_set dist_le_of_trajectories_ODE_of_mem_set +#align dist_le_of_trajectories_ODE_of_mem_set dist_le_of_trajectories_ODE_of_mem /-- If `f` and `g` are two exact solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some @@ -223,7 +223,7 @@ theorem dist_le_of_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} (hv : (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ δ * exp (K * (t - a)) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - dist_le_of_trajectories_ODE_of_mem_set (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg + dist_le_of_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align dist_le_of_trajectories_ODE dist_le_of_trajectories_ODE @@ -233,20 +233,20 @@ a given initial value provided that the RHS is Lipschitz continuous in `x` withi and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ -theorem ODE_solution_unique_of_mem_set_Icc_right {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_Icc_right {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by - have := dist_le_of_trajectories_ODE_of_mem_set hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht + have := dist_le_of_trajectories_ODE_of_mem hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in -#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_set_Icc_right +#align ODE_solution_unique_of_mem_set ODE_solution_unique_of_mem_Icc_right /-- A time-reversed version of `ODE_solution_unique_of_mem_Icc_right`. Uniqueness is shown in a closed interval `Icc a b`, where `b` is the "initial" time. -/ -theorem ODE_solution_unique_of_mem_set_Icc_left {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_Icc_left {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ioc a b, HasDerivWithinAt f (v t (f t)) (Iic t) t) (hfs : ∀ t ∈ Ioc a b, f t ∈ s t) @@ -265,7 +265,7 @@ theorem ODE_solution_unique_of_mem_set_Icc_left {v : ℝ → E → E} {s : ℝ rw [eqOn_comp_right_iff] at this convert this simp - apply ODE_solution_unique_of_mem_set_Icc_right hv' + apply ODE_solution_unique_of_mem_Icc_right hv' (hf.comp continuousOn_neg hmt1) _ (fun _ ht ↦ hfs _ (hmt2 ht)) (hg.comp continuousOn_neg hmt1) _ (fun _ ht ↦ hgs _ (hmt2 ht)) (by simp [hb]) · intros t ht @@ -279,7 +279,7 @@ theorem ODE_solution_unique_of_mem_set_Icc_left {v : ℝ → E → E} {s : ℝ /-- A version of `ODE_solution_unique_of_mem_Icc_right` for uniqueness in a closed interval whose interior contains the initial time. -/ -theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_mem_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) @@ -288,20 +288,20 @@ theorem ODE_solution_unique_of_mem_set_Icc {v : ℝ → E → E} {s : ℝ → Se rw [← Icc_union_Icc_eq_Icc (le_of_lt ht.1) (le_of_lt ht.2)] apply EqOn.union · have hss : Ioc a t₀ ⊆ Ioo a b := Ioc_subset_Ioo_right ht.2 - exact ODE_solution_unique_of_mem_set_Icc_left hv + exact ODE_solution_unique_of_mem_Icc_left hv (hf.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) (hg.mono <| Icc_subset_Icc_right <| le_of_lt ht.2) (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq · have hss : Ico t₀ b ⊆ Ioo a b := Ico_subset_Ioo_left ht.1 - exact ODE_solution_unique_of_mem_set_Icc_right hv + exact ODE_solution_unique_of_mem_Icc_right hv (hf.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) (fun _ ht' ↦ (hf' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hfs _ (hss ht'))) (hg.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq -/-- A version of `ODE_solution_unique_of_mem_set_Icc` for uniqueness in an open interval. -/ -theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +/-- A version of `ODE_solution_unique_of_mem_Icc` for uniqueness in an open interval. -/ +theorem ODE_solution_unique_of_mem_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) @@ -311,7 +311,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se rcases lt_or_le t' t₀ with (h | h) · have hss : Icc t' t₀ ⊆ Ioo a b := fun _ ht'' ↦ ⟨lt_of_lt_of_le ht'.1 ht''.1, lt_of_le_of_lt ht''.2 ht.2⟩ - exact ODE_solution_unique_of_mem_set_Icc_left hv + exact ODE_solution_unique_of_mem_Icc_left hv (ContinuousAt.continuousOn fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').1.hasDerivWithinAt) (fun _ ht'' ↦ (hf _ <| hss <| Ioc_subset_Icc_self ht'').2) @@ -321,7 +321,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se ⟨le_rfl, le_of_lt h⟩ · have hss : Icc t₀ t' ⊆ Ioo a b := fun _ ht'' ↦ ⟨lt_of_lt_of_le ht.1 ht''.1, lt_of_le_of_lt ht''.2 ht'.2⟩ - exact ODE_solution_unique_of_mem_set_Icc_right hv + exact ODE_solution_unique_of_mem_Icc_right hv (ContinuousAt.continuousOn fun _ ht'' ↦ (hf _ <| hss ht'').1.continuousAt) (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').1.hasDerivWithinAt) (fun _ ht'' ↦ (hf _ <| hss <| Ico_subset_Icc_self ht'').2) @@ -331,7 +331,7 @@ theorem ODE_solution_unique_of_mem_set_Ioo {v : ℝ → E → E} {s : ℝ → Se ⟨h, le_rfl⟩ /-- Local unqueness of ODE solutions. -/ -theorem ODE_solution_unique_of_mem_set_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} +theorem ODE_solution_unique_of_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {t₀ : ℝ} (hf : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) (hg : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) @@ -340,7 +340,7 @@ theorem ODE_solution_unique_of_mem_set_eventually {v : ℝ → E → E} {s : ℝ rw [Filter.eventuallyEq_iff_exists_mem] refine ⟨ball t₀ ε, ball_mem_nhds _ hε, ?_⟩ simp_rw [Real.ball_eq_Ioo] at * - apply ODE_solution_unique_of_mem_set_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) + apply ODE_solution_unique_of_mem_Ioo hv (Real.ball_eq_Ioo t₀ ε ▸ mem_ball_self hε) (fun _ ht ↦ (h _ ht).1) (fun _ ht ↦ (h _ ht).2) heq /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with @@ -351,7 +351,7 @@ theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, Lip (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial - ODE_solution_unique_of_mem_set_Icc_right (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' + ODE_solution_unique_of_mem_Icc_right (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha set_option linter.uppercaseLean3 false in #align ODE_solution_unique ODE_solution_unique From 6bc6460d023ead02514a8872b797d7cb5356bb72 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 22 Jan 2024 16:23:26 -0800 Subject: [PATCH 200/203] name --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index 9ece03645e2e1..cf67995e1d5d0 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -487,9 +487,9 @@ theorem isIntegralCurve_eq_of_contMDiff (hγt : ∀ t, I.IsInteriorPoint (γ t)) (hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' := by ext t obtain ⟨T, ht₀, ht⟩ : ∃ T, t ∈ Ioo (-T) T ∧ t₀ ∈ Ioo (-T) T := by - obtain ⟨T, hT₁, hT₂⟩ := exists_mem_Ioo_neg_self t + obtain ⟨T, hT₁, hT₂⟩ := exists_abs_lt t obtain ⟨hT₂, hT₃⟩ := abs_lt.mp hT₂ - obtain ⟨S, hS₁, hS₂⟩ := exists_mem_Ioo_neg_self t₀ + obtain ⟨S, hS₁, hS₂⟩ := exists_abs_lt t₀ obtain ⟨hS₂, hS₃⟩ := abs_lt.mp hS₂ exact ⟨T + S, by constructor <;> constructor <;> linarith⟩ exact isIntegralCurveOn_Ioo_eqOn_of_contMDiff ht (fun t _ ↦ hγt t) hv From 0b0b3d65ebee5084370d879b1a425198bba5c42a Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 22 Jan 2024 22:52:45 -0800 Subject: [PATCH 201/203] fix --- Mathlib/Geometry/Manifold/IntegralCurve.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index cf67995e1d5d0..2cf8f42e28d23 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -417,7 +417,7 @@ theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt (hγt₀ : I.IsInteriorPoi -- main proof suffices (extChartAt I (γ t₀)) ∘ γ =ᶠ[𝓝 t₀] (extChartAt I (γ' t₀)) ∘ γ' from (heq hγ).trans <| (this.fun_comp (extChartAt I (γ t₀)).symm).trans (h ▸ (heq hγ').symm) - exact ODE_solution_unique_of_mem_set_eventually hlip + exact ODE_solution_unique_of_eventually hlip (hdrv hγ rfl) (hdrv hγ' h) (by rw [Function.comp_apply, Function.comp_apply, h]) theorem isIntegralCurveAt_eventuallyEq_of_contMDiffAt_boundaryless [BoundarylessManifold I M] From bceedafe20dd00173bfb43a01d6997df9ef97bf1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Mon, 22 Jan 2024 23:18:22 -0800 Subject: [PATCH 202/203] collect vars --- Mathlib/Analysis/ODE/Gronwall.lean | 117 +++++++++++++++++------------ 1 file changed, 70 insertions(+), 47 deletions(-) diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 91b25423a02d7..72ac9a2db8179 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -144,19 +144,24 @@ theorem norm_le_gronwallBound_of_norm_deriv_right_le {f f' : ℝ → E} {δ K ε (fun x hx _r hr => (hf' x hx).liminf_right_slope_norm_le hr) ha bound #align norm_le_gronwall_bound_of_norm_deriv_right_le norm_le_gronwallBound_of_norm_deriv_right_le +variable {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} {f g f' g' : ℝ → E} {a b t₀ : ℝ} {εf εg δ : ℝ} + (hv : ∀ t, LipschitzOnWith K (v t) (s t)) + /-- If `f` and `g` are two approximate solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_approx_trajectories_ODE_of_mem {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g f' g' : ℝ → E} {a b : ℝ} {εf εg δ : ℝ} (hf : ContinuousOn f (Icc a b)) +theorem dist_le_of_approx_trajectories_ODE_of_mem + (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) - (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) - (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (g' t) (Ici t) t) - (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) (hgs : ∀ t ∈ Ico a b, g t ∈ s t) + (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) + (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (g' t) (Ici t) t) + (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ gronwallBound δ K (εf + εg) (t - a) := by simp only [dist_eq_norm] at ha ⊢ @@ -177,12 +182,15 @@ can't grow faster than exponentially. This is a simple corollary of Grönwall's people call this Grönwall's inequality too. This version assumes all inequalities to be true in the whole space. -/ -theorem dist_le_of_approx_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} - (hv : ∀ t, LipschitzWith K (v t)) {f g f' g' : ℝ → E} {a b : ℝ} {εf εg δ : ℝ} - (hf : ContinuousOn f (Icc a b)) (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) - (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) (hg : ContinuousOn g (Icc a b)) +theorem dist_le_of_approx_trajectories_ODE + (hv : ∀ t, LipschitzWith K (v t)) + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (f' t) (Ici t) t) + (f_bound : ∀ t ∈ Ico a b, dist (f' t) (v t (f t)) ≤ εf) + (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (g' t) (Ici t) t) - (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) (ha : dist (f a) (g a) ≤ δ) : + (g_bound : ∀ t ∈ Ico a b, dist (g' t) (v t (g t)) ≤ εg) + (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ gronwallBound δ K (εf + εg) (t - a) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial dist_le_of_approx_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' @@ -196,10 +204,10 @@ people call this Grönwall's inequality too. This version assumes all inequalities to be true in some time-dependent set `s t`, and assumes that the solutions never leave this set. -/ -theorem dist_le_of_trajectories_ODE_of_mem {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} {δ : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) +theorem dist_le_of_trajectories_ODE_of_mem + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hfs : ∀ t ∈ Ico a b, f t ∈ s t) (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ δ * exp (K * (t - a)) := by @@ -217,10 +225,13 @@ can't grow faster than exponentially. This is a simple corollary of Grönwall's people call this Grönwall's inequality too. This version assumes all inequalities to be true in the whole space. -/ -theorem dist_le_of_trajectories_ODE {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) - {f g : ℝ → E} {a b : ℝ} {δ : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) - (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : dist (f a) (g a) ≤ δ) : +theorem dist_le_of_trajectories_ODE + (hv : ∀ t, LipschitzWith K (v t)) + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) + (ha : dist (f a) (g a) ≤ δ) : ∀ t ∈ Icc a b, dist (f t) (g t) ≤ δ * exp (K * (t - a)) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial dist_le_of_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg @@ -233,12 +244,15 @@ a given initial value provided that the RHS is Lipschitz continuous in `x` withi and we consider only solutions included in `s`. This version shows uniqueness in a closed interval `Icc a b`, where `a` is the initial time. -/ -theorem ODE_solution_unique_of_mem_Icc_right {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hfs : ∀ t ∈ Ico a b, f t ∈ s t) - (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) - (hgs : ∀ t ∈ Ico a b, g t ∈ s t) (ha : f a = g a) : EqOn f g (Icc a b) := fun t ht ↦ by +theorem ODE_solution_unique_of_mem_Icc_right + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hfs : ∀ t ∈ Ico a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) + (hgs : ∀ t ∈ Ico a b, g t ∈ s t) + (ha : f a = g a) : + EqOn f g (Icc a b) := fun t ht ↦ by have := dist_le_of_trajectories_ODE_of_mem hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this set_option linter.uppercaseLean3 false in @@ -246,12 +260,15 @@ set_option linter.uppercaseLean3 false in /-- A time-reversed version of `ODE_solution_unique_of_mem_Icc_right`. Uniqueness is shown in a closed interval `Icc a b`, where `b` is the "initial" time. -/ -theorem ODE_solution_unique_of_mem_Icc_left {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ioc a b, HasDerivWithinAt f (v t (f t)) (Iic t) t) (hfs : ∀ t ∈ Ioc a b, f t ∈ s t) - (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ioc a b, HasDerivWithinAt g (v t (g t)) (Iic t) t) - (hgs : ∀ t ∈ Ioc a b, g t ∈ s t) (hb : f b = g b) : EqOn f g (Icc a b) := by +theorem ODE_solution_unique_of_mem_Icc_left + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ioc a b, HasDerivWithinAt f (v t (f t)) (Iic t) t) + (hfs : ∀ t ∈ Ioc a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ioc a b, HasDerivWithinAt g (v t (g t)) (Iic t) t) + (hgs : ∀ t ∈ Ioc a b, g t ∈ s t) + (hb : f b = g b) : + EqOn f g (Icc a b) := by have hv' t : LipschitzOnWith K (Neg.neg ∘ (v (-t))) (s (-t)) := by rw [← one_mul K] exact LipschitzWith.id.neg.comp_lipschitzOnWith (hv _) @@ -279,12 +296,16 @@ theorem ODE_solution_unique_of_mem_Icc_left {v : ℝ → E → E} {s : ℝ → S /-- A version of `ODE_solution_unique_of_mem_Icc_right` for uniqueness in a closed interval whose interior contains the initial time. -/ -theorem ODE_solution_unique_of_mem_Icc {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) - (hg : ContinuousOn g (Icc a b)) (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) - (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) (heq : f t₀ = g t₀) : EqOn f g (Icc a b) := by +theorem ODE_solution_unique_of_mem_Icc + (ht : t₀ ∈ Ioo a b) + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t) + (hfs : ∀ t ∈ Ioo a b, f t ∈ s t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t) + (hgs : ∀ t ∈ Ioo a b, g t ∈ s t) + (heq : f t₀ = g t₀) : + EqOn f g (Icc a b) := by rw [← Icc_union_Icc_eq_Icc (le_of_lt ht.1) (le_of_lt ht.2)] apply EqOn.union · have hss : Ioc a t₀ ⊆ Ioo a b := Ioc_subset_Ioo_right ht.2 @@ -301,12 +322,12 @@ theorem ODE_solution_unique_of_mem_Icc {v : ℝ → E → E} {s : ℝ → Set E} (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq /-- A version of `ODE_solution_unique_of_mem_Icc` for uniqueness in an open interval. -/ -theorem ODE_solution_unique_of_mem_Ioo {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) - {f g : ℝ → E} {a b t₀ : ℝ} (ht : t₀ ∈ Ioo a b) +theorem ODE_solution_unique_of_mem_Ioo + (ht : t₀ ∈ Ioo a b) (hf : ∀ t ∈ Ioo a b, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) (hg : ∀ t ∈ Ioo a b, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) - (heq : f t₀ = g t₀) : EqOn f g (Ioo a b) := by + (heq : f t₀ = g t₀) : + EqOn f g (Ioo a b) := by intros t' ht' rcases lt_or_le t' t₀ with (h | h) · have hss : Icc t' t₀ ⊆ Ioo a b := @@ -331,8 +352,7 @@ theorem ODE_solution_unique_of_mem_Ioo {v : ℝ → E → E} {s : ℝ → Set E} ⟨h, le_rfl⟩ /-- Local unqueness of ODE solutions. -/ -theorem ODE_solution_unique_of_eventually {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} - (hv : ∀ t, LipschitzOnWith K (v t) (s t)) {f g : ℝ → E} {t₀ : ℝ} +theorem ODE_solution_unique_of_eventually (hf : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) (hg : ∀ᶠ t in 𝓝 t₀, HasDerivAt g (v t (g t)) t ∧ g t ∈ s t) (heq : f t₀ = g t₀) : f =ᶠ[𝓝 t₀] g := by @@ -345,10 +365,13 @@ theorem ODE_solution_unique_of_eventually {v : ℝ → E → E} {s : ℝ → Set /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) with a given initial value provided that the RHS is Lipschitz continuous in `x`. -/ -theorem ODE_solution_unique {v : ℝ → E → E} {K : ℝ≥0} (hv : ∀ t, LipschitzWith K (v t)) - {f g : ℝ → E} {a b : ℝ} (hf : ContinuousOn f (Icc a b)) - (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) (hg : ContinuousOn g (Icc a b)) - (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) (ha : f a = g a) : +theorem ODE_solution_unique + (hv : ∀ t, LipschitzWith K (v t)) + (hf : ContinuousOn f (Icc a b)) + (hf' : ∀ t ∈ Ico a b, HasDerivWithinAt f (v t (f t)) (Ici t) t) + (hg : ContinuousOn g (Icc a b)) + (hg' : ∀ t ∈ Ico a b, HasDerivWithinAt g (v t (g t)) (Ici t) t) + (ha : f a = g a) : EqOn f g (Icc a b) := have hfs : ∀ t ∈ Ico a b, f t ∈ @univ E := fun _ _ => trivial ODE_solution_unique_of_mem_Icc_right (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' From 492f603339d8b17bd6709e3dfd7c5ed4b65e12c7 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Tue, 23 Jan 2024 01:23:46 -0800 Subject: [PATCH 203/203] fix --- Mathlib/Algebra/Order/Ring/Abs.lean | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 5fd1cc6a56713..c5d59c19e19a3 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -158,14 +158,11 @@ lemma sq_eq_sq_iff_abs_eq_abs (a b : α) : a ^ 2 = b ^ 2 ↔ |a| = |b| := by simpa only [one_pow, abs_one] using @sq_lt_sq _ _ 1 a #align one_lt_sq_iff_one_lt_abs one_lt_sq_iff_one_lt_abs -lemma exists_abs_lt - {α : Type*} [LinearOrderedRing α] (a : α) : ∃ b > 0, |a| < b := by +lemma exists_abs_lt {α : Type*} [LinearOrderedRing α] (a : α) : ∃ b > 0, |a| < b := by refine ⟨|a| + 1, lt_of_lt_of_le zero_lt_one <| by simp, ?_⟩ cases' le_or_lt 0 a with ht ht - · suffices -1 < a + a by simp [abs_eq_self.mpr ht] - exact lt_of_lt_of_le neg_one_lt_zero (by simpa) - · suffices a + a < 1 by simp [abs_eq_neg_self.mpr ht.le] - exact lt_of_lt_of_le (by simpa) zero_le_one + · simp only [abs_of_nonneg ht, lt_add_iff_pos_right, zero_lt_one] + · simp only [abs_of_neg ht, lt_add_iff_pos_right, zero_lt_one] end LinearOrderedRing