Documentation

Mathlib.Geometry.Manifold.IntegralCurve.Basic

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:

Main definitions #

Let v : M → TM be a vector field on M, and let γ : ℝ → M.

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.

Reference #

Tags #

integral curve, vector field

def IsIntegralCurveOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] (γ : 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.

Equations
Instances For
    def IsIntegralCurveAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] (γ : M) (v : (x : M) → TangentSpace I x) (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.

    Equations
    Instances For
      def IsIntegralCurve {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] (γ : M) (v : (x : M) → TangentSpace I x) :

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

      Equations
      Instances For
        theorem IsIntegralCurve.isIntegralCurveOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} (h : IsIntegralCurve γ v) (s : Set ) :
        theorem isIntegralCurve_iff_isIntegralCurveOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} :
        theorem isIntegralCurveAt_iff {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {t₀ : } :
        IsIntegralCurveAt γ v t₀ snhds t₀, IsIntegralCurveOn γ v s
        theorem isIntegralCurveAt_iff' {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {t₀ : } :
        IsIntegralCurveAt γ v t₀ ε > 0, IsIntegralCurveOn γ v (Metric.ball t₀ ε)

        γ is an integral curve for v at t₀ iff γ is an integral curve on some interval containing t₀.

        theorem IsIntegralCurve.isIntegralCurveAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} (h : IsIntegralCurve γ v) (t : ) :
        theorem isIntegralCurve_iff_isIntegralCurveAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} :
        IsIntegralCurve γ v ∀ (t : ), IsIntegralCurveAt γ v t
        theorem IsIntegralCurveOn.mono {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s s' : Set } (h : IsIntegralCurveOn γ v s) (hs : s' s) :
        theorem IsIntegralCurveOn.of_union {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s s' : Set } (h : IsIntegralCurveOn γ v s) (h' : IsIntegralCurveOn γ v s') :
        theorem IsIntegralCurveAt.hasMFDerivAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {t₀ : } (h : IsIntegralCurveAt γ v t₀) :
        theorem IsIntegralCurveOn.isIntegralCurveAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } {t₀ : } (h : IsIntegralCurveOn γ v s) (hs : s nhds t₀) :
        theorem IsIntegralCurveAt.isIntegralCurveOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } (h : ts, IsIntegralCurveAt γ v t) :

        If γ is an integral curve at each t ∈ s, it is an integral curve on s.

        theorem isIntegralCurveOn_iff_isIntegralCurveAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } (hs : IsOpen s) :
        IsIntegralCurveOn γ v s ts, IsIntegralCurveAt γ v t
        theorem IsIntegralCurveOn.continuousAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } {t₀ : } (hγ : IsIntegralCurveOn γ v s) (ht : t₀ s) :
        ContinuousAt γ t₀
        theorem IsIntegralCurveOn.continuousOn {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } (hγ : IsIntegralCurveOn γ v s) :
        theorem IsIntegralCurveAt.continuousAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {t₀ : } (hγ : IsIntegralCurveAt γ v t₀) :
        ContinuousAt γ t₀
        theorem IsIntegralCurve.continuous {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} (hγ : IsIntegralCurve γ v) :
        theorem IsIntegralCurveOn.hasDerivAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {s : Set } {t₀ : } [SmoothManifoldWithCorners I M] (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

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

        theorem IsIntegralCurveAt.eventually_hasDerivAt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {H : Type u_2} [TopologicalSpace H] {I : ModelWithCorners E H} {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] {γ : M} {v : (x : M) → TangentSpace I x} {t₀ : } [SmoothManifoldWithCorners I M] (hγ : IsIntegralCurveAt γ v t₀) :
        ∀ᶠ (t : ) in nhds t₀, HasDerivAt ((extChartAt I (γ t₀)) γ) ((tangentCoordChange I (γ t) (γ t₀) (γ t)) (v (γ t))) t