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 ℝ
.
This is the first of a series of files, organised as follows:
Mathlib/Geometry/Manifold/IntegralCurve/Basic.lean
(this file): Basic definitions and lemmas relating them to each other and to continuity and differentiabilityMathlib/Geometry/Manifold/IntegralCurve/Transform.lean
: Lemmas about translating or scaling the domain of an integral curve by a constantMathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean
: Local existence and uniqueness theorems for integral curves
Main definitions #
Let v : M → TM
be a vector field on M
, and let γ : ℝ → M
.
IsMIntegralCurve γ v
:γ t
is tangent tov (γ t)
for allt : ℝ
. That is,γ
is a global integral curve ofv
.IsMIntegralCurveOn γ v s
:γ t
is tangent tov (γ t)
for allt ∈ s
, wheres : Set ℝ
.IsMIntegralCurveAt γ v t₀
:γ t
is tangent tov (γ t)
for allt
in some open interval aroundt₀
. That is,γ
is a local integral curve ofv
.
For IsMIntegralCurveOn γ v s
and IsMIntegralCurveAt γ 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.
TODO #
- Implement
IsMIntegralCurveWithinAt
.
Reference #
Tags #
integral curve, vector field
If γ : ℝ → M
is $C^1$ on s : Set ℝ
and v
is a vector field on M
,
IsMIntegralCurveOn γ v s
means γ t
is tangent to v (γ t)
for all t ∈ s
. The value of γ
outside of s
is irrelevant and considered junk.
Equations
- IsMIntegralCurveOn γ v s = ∀ t ∈ s, HasMFDerivWithinAt (modelWithCornersSelf ℝ ℝ) I γ s t (ContinuousLinearMap.smulRight 1 (v (γ t)))
Instances For
Alias of IsMIntegralCurveOn
.
If γ : ℝ → M
is $C^1$ on s : Set ℝ
and v
is a vector field on M
,
IsMIntegralCurveOn γ v s
means γ t
is tangent to v (γ t)
for all t ∈ s
. The value of γ
outside of s
is irrelevant and considered junk.
Equations
Instances For
If v
is a vector field on M
and t₀ : ℝ
, IsMIntegralCurveAt γ 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.
Equations
- IsMIntegralCurveAt γ v t₀ = ∀ᶠ (t : ℝ) in nhds t₀, HasMFDerivAt (modelWithCornersSelf ℝ ℝ) I γ t (ContinuousLinearMap.smulRight 1 (v (γ t)))
Instances For
Alias of IsMIntegralCurveAt
.
If v
is a vector field on M
and t₀ : ℝ
, IsMIntegralCurveAt γ 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.
Equations
Instances For
If v : M → TM
is a vector field on M
, IsMIntegralCurve γ v
means γ : ℝ → M
is a global
integral curve of v
. That is, γ t
is tangent to v (γ t)
for all t : ℝ
.
Equations
- IsMIntegralCurve γ v = ∀ (t : ℝ), HasMFDerivAt (modelWithCornersSelf ℝ ℝ) I γ t (ContinuousLinearMap.smulRight 1 (v (γ t)))
Instances For
Alias of IsMIntegralCurve
.
If v : M → TM
is a vector field on M
, IsMIntegralCurve γ v
means γ : ℝ → M
is a global
integral curve of v
. That is, γ t
is tangent to v (γ t)
for all t : ℝ
.
Equations
Instances For
Alias of IsMIntegralCurve.isMIntegralCurveOn
.
Alias of isMIntegralCurve_iff_isMIntegralCurveOn
.
Alias of isMIntegralCurveAt_iff
.
γ
is an integral curve for v
at t₀
iff γ
is an integral curve on some interval
containing t₀
.
Alias of isMIntegralCurveAt_iff'
.
γ
is an integral curve for v
at t₀
iff γ
is an integral curve on some interval
containing t₀
.
Alias of IsMIntegralCurve.isMIntegralCurveAt
.
Alias of isMIntegralCurve_iff_isMIntegralCurveAt
.
Alias of IsMIntegralCurveOn.mono
.
Alias of IsMIntegralCurveAt.hasMFDerivAt
.
Alias of IsMIntegralCurveOn.isMIntegralCurveAt
.
If γ
is an integral curve at each t ∈ s
, it is an integral curve on s
.
Alias of IsMIntegralCurveAt.isMIntegralCurveOn
.
If γ
is an integral curve at each t ∈ s
, it is an integral curve on s
.
Alias of IsMIntegralCurveOn.continuousWithinAt
.
Alias of IsMIntegralCurveOn.continuousWithinAt
.
Alias of IsMIntegralCurveOn.continuousOn
.
Alias of IsMIntegralCurveAt.continuousAt
.
Alias of IsMIntegralCurve.continuous
.
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₀
.
Alias of IsMIntegralCurveOn.hasDerivWithinAt
.
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₀
.
Alias of IsMIntegralCurveAt.eventually_hasDerivAt
.