Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: local existence of integral curves of vector field #8483

Closed
wants to merge 126 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
126 commits
Select commit Hold shift + click to select a range
4e53e96
doc
winstonyin Nov 13, 2023
f14b378
compiles
winstonyin Nov 17, 2023
22b4f1f
indent
winstonyin Nov 17, 2023
0747174
import, indent
winstonyin Nov 17, 2023
23d691f
docstring
winstonyin Nov 17, 2023
e8e7338
Mathlib.lean
winstonyin Nov 17, 2023
ccc8f77
lake manifest
winstonyin Nov 17, 2023
42341d1
Merge remote-tracking branch 'origin/master' into integral_curve
winstonyin Nov 17, 2023
720ec35
removed temporary notation
winstonyin Nov 18, 2023
25f1065
integralCurveAt, change PicardLindelof statement order
winstonyin Nov 18, 2023
d4139b6
compiles
winstonyin Nov 18, 2023
4b57dfe
Sketch a section on interior and boundary of manifolds.
grunweg Nov 18, 2023
e9d3546
More progress. Tweaks, polish and fleshing out some sorries.
grunweg Nov 19, 2023
cbd22f4
One more lemma.
grunweg Nov 19, 2023
0dbf3b0
One lemma was in mathlib. :-)
grunweg Nov 19, 2023
a978f02
Fix the build; mild clean-ups.
grunweg Nov 19, 2023
1f6dbaa
Small simplifications.
grunweg Nov 19, 2023
9c58652
has -> is, docstring, shift/stretch lemmas
winstonyin Nov 20, 2023
6b55cbc
detailed proof steps
winstonyin Nov 20, 2023
0ada08e
open Set
winstonyin Nov 20, 2023
063822d
`IsIntegralCurveAt` in preamble
winstonyin Nov 20, 2023
fb4196f
IsIntegralCurveAt.comp_add
winstonyin Nov 20, 2023
6f665ad
translation, scale, and reverse lemmas
winstonyin Nov 21, 2023
b59c82f
iff lemmas
winstonyin Nov 21, 2023
3c2e424
Small polish. Note that smoothness is not required;
grunweg Nov 21, 2023
c95243e
The boundary is closed: entails showing that interior and boundary ar…
grunweg Nov 21, 2023
933cc59
Almost-proof that boundary has empty interior: need to adjust definit…
grunweg Nov 21, 2023
2041d30
Correct the definition of boundary:
grunweg Nov 21, 2023
f1fda32
Fix and complete proof that the boundary has empty interior.
grunweg Nov 21, 2023
6a4e39c
Small clean-ups:
grunweg Nov 21, 2023
88bea4c
Clean up lemmas about boundary.
grunweg Nov 21, 2023
73e203f
Small golfs.
grunweg Nov 21, 2023
90090b2
Fix rebase, make it build and fix a docstring typo.
grunweg Nov 21, 2023
338ab9f
Fix casing for real: IsInteriorPoint is correct.
grunweg Nov 21, 2023
378088c
Name lemma in snake_case.
grunweg Nov 21, 2023
efe0880
naming
winstonyin Nov 22, 2023
d35e843
move lemma, style
winstonyin Nov 22, 2023
dce81a6
Define topological manifolds and use them for defining interior and b…
grunweg Nov 21, 2023
79fcb13
Alternative, more minimal fix.
grunweg Nov 21, 2023
e1a9363
sketch boundary argument
grunweg Nov 21, 2023
f38f823
MAYBE this tweak makes the proof easier.
grunweg Nov 21, 2023
07d1dbc
Cleanup, one easy sorry.
grunweg Nov 21, 2023
c331ee2
Another easy sorry.
grunweg Nov 21, 2023
191a5c1
Move a helper upfront; eliminating another sorry.
grunweg Nov 21, 2023
66b1943
Starting at scifi lemma: might be science *fiction* actually...
grunweg Nov 21, 2023
7840b91
Small golfs.
grunweg Nov 22, 2023
61fd70e
Extract lemmas used in the proof so far.
grunweg Nov 22, 2023
5869bd5
Fill sorry in description of boundary.
grunweg Nov 22, 2023
058bb70
Extract one more useful lemma; tiny golf.
grunweg Nov 22, 2023
7e1542d
Simplify: use chart source instead of extChart.source.
grunweg Nov 22, 2023
219f85c
Fill in one more helper sorry.
grunweg Nov 22, 2023
b97611e
Move helpers to better namespace; show a MapsTo version;
grunweg Nov 22, 2023
ef4c232
Reduce interior independence statement to a local lemma.
grunweg Nov 22, 2023
184a062
Give up: remove WIP claim that boundary as empty interior.
grunweg Nov 22, 2023
c634f1a
Rename and golf mapsTo results for extended local homeos.
grunweg Nov 22, 2023
dd9e73f
Small golfs and clean-ups.
grunweg Nov 22, 2023
3d820be
Move complete (and now unused) helper results to their files:
grunweg Nov 22, 2023
2c69553
compiles
winstonyin Nov 22, 2023
7c6bccf
new lemmas and slight golfing
winstonyin Nov 23, 2023
e74b860
Essentials are complete.
grunweg Nov 25, 2023
f8e3340
Clean up; document TODOs/next steps; move all helpers to their proper…
grunweg Nov 25, 2023
4febc74
Remove helper lemmas not needed here.
grunweg Nov 25, 2023
c1f3395
Pare down interior/boundary file to the necesssary basics; sorry-free.
grunweg Nov 25, 2023
25ff469
Merge branch 'MR-manifolds-boundary' into integral_curve
grunweg Nov 25, 2023
1459dfb
Simplify.
grunweg Nov 25, 2023
094ca88
compiles
winstonyin Nov 26, 2023
ca11211
Merge branch 'integral_curve' of https://github.com/leanprover-commun…
winstonyin Nov 26, 2023
7754760
range_mem_nhds_isInteriorPoint
winstonyin Nov 26, 2023
6033099
namespace
winstonyin Nov 26, 2023
429fd8b
Merge remote-tracking branch 'origin/MR-manifolds-boundary' into inte…
winstonyin Nov 26, 2023
d1a8296
hI as separate lemma
winstonyin Nov 26, 2023
7f7a49d
actually compiles
winstonyin Nov 26, 2023
ab30700
minor
winstonyin Nov 26, 2023
98ff512
Prefer refine over refine' in new code, per zulip.
grunweg Nov 27, 2023
00f4e77
Use mfld_simps to make proofs more readable.
grunweg Nov 27, 2023
e86759d
Small golf, use mfld_simps more.
grunweg Nov 27, 2023
85fbe5d
hasFDerivWithinAt_tangentCoordChange
winstonyin Nov 27, 2023
d924890
some golf
winstonyin Nov 28, 2023
4cf8097
shorten comment
winstonyin Nov 28, 2023
e8bfdb0
shorten
winstonyin Nov 28, 2023
de44a9c
move tangentCoordChange
winstonyin Nov 28, 2023
a792394
minor spacing
winstonyin Nov 28, 2023
4cc6c69
isIntegralCurveAt_const
winstonyin Nov 28, 2023
b4ba5a0
remove todo
winstonyin Nov 28, 2023
f14affa
use implicit lambdas
winstonyin Nov 28, 2023
0a3d359
Fix style.
grunweg Nov 29, 2023
fa1b5f3
Remove superfluous helper lemma.
grunweg Nov 29, 2023
9e9ee92
docstring, rename theorem
winstonyin Nov 30, 2023
75c2b52
rename
winstonyin Nov 30, 2023
f019e9b
remove `x₀` from `IsIntegralCurveAt`
winstonyin Nov 30, 2023
30c333f
end with exact
winstonyin Nov 30, 2023
e140c43
`IsIntegralCurveOn`, `IsIntegralCurve`
winstonyin Nov 30, 2023
b810d2a
docstring
winstonyin Nov 30, 2023
9f5b559
IsIntegralCurveAt.isIntegralCurveOn
winstonyin Nov 30, 2023
0445938
`comp_add` lemmas
winstonyin Nov 30, 2023
3b7cca0
refactored `IsIntegralCurveXX`
winstonyin Nov 30, 2023
082574b
unused arg
winstonyin Nov 30, 2023
fbb74c9
Merge remote-tracking branch 'origin/master' into integral_curve
winstonyin Dec 8, 2023
a9d4cae
space
winstonyin Dec 8, 2023
6d4af73
Small things: sectioning; re-use variables.
grunweg Dec 8, 2023
ffa3c4b
Tinygolf.
grunweg Dec 8, 2023
7dde305
Line breaks.
grunweg Dec 8, 2023
ccd1830
style
winstonyin Dec 13, 2023
122546d
spacing
winstonyin Dec 13, 2023
2c4d77d
restate existence theorem using filters
winstonyin Dec 14, 2023
5f7bb14
define `IsIntegralCurveAt` using `Eventually`
winstonyin Dec 14, 2023
29170c6
Merge remote-tracking branch 'origin/master' into integral_curve
winstonyin Dec 14, 2023
6341072
PartialEquiv
winstonyin Dec 14, 2023
8e54212
PartialEquiv
winstonyin Dec 14, 2023
4efc1ad
PartialEquiv
winstonyin Dec 14, 2023
7b804f9
Replace => by \mapsto, per the style guide.
grunweg Dec 14, 2023
22d5ceb
Small golfs.
grunweg Dec 14, 2023
e784d93
Update Mathlib/Geometry/Manifold/IntegralCurve.lean
winstonyin Dec 17, 2023
47e137d
Update Mathlib/Geometry/Manifold/IntegralCurve.lean
winstonyin Dec 17, 2023
2e0b9bf
Update Mathlib/Geometry/Manifold/IntegralCurve.lean
winstonyin Dec 18, 2023
26be30a
remove newlines
winstonyin Dec 18, 2023
90755c3
Merge remote-tracking branch 'origin/master' into integral_curve
winstonyin Dec 19, 2023
5aab426
fix
winstonyin Dec 19, 2023
749ea8c
IsIntegralCurveAt.hasMFDerivAt
winstonyin Dec 29, 2023
ca88b9d
Merge remote-tracking branch 'origin/master' into integral_curve
winstonyin Jan 5, 2024
c74e830
BoundarylessManifold
winstonyin Jan 5, 2024
353f3c9
Make I implicit again.
grunweg Jan 8, 2024
7d7a716
docs, golf, add `isIntegralCurveOn_iff_isIntegralCurveAt`
winstonyin Jan 8, 2024
db893bd
use `Eventually` in `IsIntegralCurveAt`
winstonyin Jan 8, 2024
ad6d0a6
minor
winstonyin Jan 8, 2024
04649d0
Tinygolfs.
grunweg Jan 8, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2134,6 +2134,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.InteriorBoundary
import Mathlib.Geometry.Manifold.LocalDiffeomorph
import Mathlib.Geometry.Manifold.LocalInvariantProperties
Expand Down
13 changes: 7 additions & 6 deletions Mathlib/Analysis/ODE/PicardLindelof.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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
314 changes: 314 additions & 0 deletions Mathlib/Geometry/Manifold/IntegralCurve.lean
winstonyin marked this conversation as resolved.
Show resolved Hide resolved
Original file line number Diff line number Diff line change
@@ -0,0 +1,314 @@
/-
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.InteriorBoundary
import Mathlib.Geometry.Manifold.MFDeriv

/-!
# Integral curves of vector fields on a manifold

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 `ℝ`.

## Main definitions

winstonyin marked this conversation as resolved.
Show resolved Hide resolved
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`.

winstonyin marked this conversation as resolved.
Show resolved Hide resolved
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 TODO.

We state simpler versions of the theorem for boundaryless manifolds as corollaries.

## 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.

## Tags

integral curve, vector field, local existence
-/

open scoped Manifold Topology

open Set

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]

/-- 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` 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 :=
∀ᶠ t in 𝓝 t₀, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight <| v (γ t))
Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This is now stated in terms of Filter.Eventually to make the three definitions more symmetric.


/-- 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) : Prop :=
∀ t : ℝ, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t)))

variable {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ}

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Can we add this lemma (which I would have been tempted to make the base definition):

Suggested change
lemma IsIntegralCurveAt_iff_eventually :
IsIntegralCurveAt γ v t₀ ↔
∀ᶠ t in 𝓝 t₀, HasMFDerivAt 𝓘(ℝ, ℝ) I γ t ((1 : ℝ →L[ℝ] ℝ).smulRight <| v (γ t)) := by
rw [Filter.eventually_iff_exists_mem]
exact ⟨id, id⟩

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I changed the base definition of IsIntegralCurveAt to use Filter.Eventually. See this commit. This makes the definitions more "symmetric", but also adds a few calls to isIntegralCurveAt_iff or isIntegralCurveAt_iff' throughout later proofs. If this looks good, I'll merge in a few hours.

lemma IsIntegralCurve.isIntegralCurveOn (h : IsIntegralCurve γ v) (s : Set ℝ) :
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 _)⟩

lemma isIntegralCurveAt_iff :
Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The lemma @ocfnash suggested above, before I changed the definition of IsIntegralCurveAt.

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 := isIntegralCurveAt_iff.mpr ⟨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⟩ := isIntegralCurveAt_iff.mp (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)

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

lemma IsIntegralCurveAt.hasMFDerivAt (h : IsIntegralCurveAt γ v t₀) :
HasMFDerivAt 𝓘(ℝ, ℝ) I γ t₀ ((1 : ℝ →L[ℝ] ℝ).smulRight (v (γ t₀))) :=
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₀ := 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) :
winstonyin marked this conversation as resolved.
Show resolved Hide resolved
IsIntegralCurveOn γ v s := by
intros t ht
obtain ⟨s, hs, h⟩ := isIntegralCurveAt_iff.mp (h t ht)
exact h t (mem_of_mem_nhds hs)

winstonyin marked this conversation as resolved.
Show resolved Hide resolved
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

lemma IsIntegralCurveOn.comp_add (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 {dt : ℝ} :
IsIntegralCurveOn γ v s ↔ IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by
winstonyin marked this conversation as resolved.
Show resolved Hide resolved
refine ⟨fun hγ ↦ hγ.comp_add _, fun hγ ↦ ?_⟩
convert hγ.comp_add (-dt)
ext
simp only [Function.comp_apply, neg_add_cancel_right]
simp

lemma IsIntegralCurveAt.comp_add (hγ : IsIntegralCurveAt γ v t₀) (dt : ℝ) :
IsIntegralCurveAt (γ ∘ (· + dt)) v (t₀ - dt) := by
rw [isIntegralCurveAt_iff'] at *
obtain ⟨ε, hε, h⟩ := 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]

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]
simp only [sub_neg_eq_add, sub_add_cancel]

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 {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

/-! ### 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
rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp]
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 isIntegralCurvOn_comp_mul_ne_zero {a : ℝ} (ha : a ≠ 0) :
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I have the nagging feeling the next four lemmas could be golfed (all the manipulations around multiplication and division). I might be wrong; perhaps none of simp, ring and (g)congr is useful here. It might be worth extracting an MWE and asking the hive mind on zulip.

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 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
rw [isIntegralCurveAt_iff'] at *
obtain ⟨ε, hε, h⟩ := hγ
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,
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
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 t
simp [inv_mul_eq_div, div_self ha]

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 {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]

/-- 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
intro t
rw [h, ← ContinuousLinearMap.zero_apply (R₁ := ℝ) (R₂ := ℝ) (1 : ℝ),
ContinuousLinearMap.smulRight_one_one]
exact hasMFDerivAt_const ..

end Scaling

variable (t₀) {x₀ : M}

/-- 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
-- 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
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) ∈ 𝓝 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`
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, PartialEquiv.left_inv _ (mem_extChartAt_source ..)]),
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`
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 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 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 := 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⟩
rw [← (extChartAt I x₀).right_inv hf3']
exact hasFDerivWithinAt_tangentCoordChange ⟨hft1, hft2⟩

/-- 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₀ :=
exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I)
1 change: 1 addition & 0 deletions Mathlib/Geometry/Manifold/InteriorBoundary.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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]
Expand Down