Skip to content

Commit 132e511

Browse files
winstonyinurkudgrunweg
committed
feat: Integral curves are either injective or periodic (#9343)
Integral curves are either injective, constant, or periodic and non-constant. When we have notions of submanifolds, this'll be useful for showing that the image of an integral curve is a submanifold. - [x] depends on: #8886 Co-authored-by: Yury G. Kudryashov <urkud@urkud.name> Co-authored-by: Michael Rothgang <rothgami@math.hu-berlin.de>
1 parent b4d01dc commit 132e511

File tree

2 files changed

+40
-4
lines changed

2 files changed

+40
-4
lines changed

Mathlib/Algebra/Periodic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -349,6 +349,11 @@ theorem Periodic.lift_coe [AddGroup α] (h : Periodic f c) (a : α) :
349349
rfl
350350
#align function.periodic.lift_coe Function.Periodic.lift_coe
351351

352+
/-- A periodic function `f : R → X` on a semiring (or, more generally, `AddZeroClass`)
353+
of non-zero period is not injective. -/
354+
lemma Periodic.not_injective {R X : Type*} [AddZeroClass R] {f : R → X} {c : R}
355+
(hf : Periodic f c) (hc : c ≠ 0) : ¬ Injective f := fun h ↦ hc <| h hf.eq
356+
352357
/-! ### Antiperiodicity -/
353358

354359
/-- A function `f` is said to be `antiperiodic` with antiperiod `c` if for all `x`,

Mathlib/Geometry/Manifold/IntegralCurve.lean

Lines changed: 35 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -60,7 +60,7 @@ integral curve, vector field, local existence, uniqueness
6060

6161
open scoped Manifold Topology
6262

63-
open Set
63+
open Function Set
6464

6565
variable
6666
{E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E]
@@ -193,8 +193,7 @@ section Translation
193193
lemma IsIntegralCurveOn.comp_add (hγ : IsIntegralCurveOn γ v s) (dt : ℝ) :
194194
IsIntegralCurveOn (γ ∘ (· + dt)) v { t | t + dt ∈ s } := by
195195
intros t ht
196-
rw [Function.comp_apply,
197-
← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))]
196+
rw [comp_apply, ← ContinuousLinearMap.comp_id (ContinuousLinearMap.smulRight 1 (v (γ (t + dt))))]
198197
apply HasMFDerivAt.comp t (hγ (t + dt) ht)
199198
refine ⟨(continuous_add_right _).continuousAt, ?_⟩
200199
simp only [mfld_simps, hasFDerivWithinAt_univ]
@@ -246,7 +245,7 @@ section Scaling
246245
lemma IsIntegralCurveOn.comp_mul (hγ : IsIntegralCurveOn γ v s) (a : ℝ) :
247246
IsIntegralCurveOn (γ ∘ (· * a)) (a • v) { t | t * a ∈ s } := by
248247
intros t ht
249-
rw [Function.comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp]
248+
rw [comp_apply, Pi.smul_apply, ← ContinuousLinearMap.smulRight_comp]
250249
refine HasMFDerivAt.comp t (hγ (t * a) ht) ⟨(continuous_mul_right _).continuousAt, ?_⟩
251250
simp only [mfld_simps, hasFDerivWithinAt_univ]
252251
exact HasFDerivAt.mul_const' (hasFDerivAt_id _) _
@@ -501,4 +500,36 @@ theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I
501500
(hγ : IsIntegralCurve γ v) (hγ' : IsIntegralCurve γ' v) (h : γ t₀ = γ' t₀) : γ = γ' :=
502501
isIntegralCurve_eq_of_contMDiff (fun _ ↦ BoundarylessManifold.isInteriorPoint I) hv hγ hγ' h
503502

503+
/-- For a global integral curve `γ`, if it crosses itself at `a b : ℝ`, then it is periodic with
504+
period `a - b`. -/
505+
lemma IsIntegralCurve.periodic_of_eq [BoundarylessManifold I M]
506+
(hγ : IsIntegralCurve γ v)
507+
(hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M)))
508+
(heq : γ a = γ b) : Periodic γ (a - b) := by
509+
intro t
510+
apply congrFun <|
511+
isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless (t₀ := b) hv (hγ.comp_add _) hγ _
512+
rw [comp_apply, add_sub_cancel'_right, heq]
513+
514+
/-- A global integral curve is injective xor periodic with positive period. -/
515+
lemma IsIntegralCurve.periodic_xor_injective [BoundarylessManifold I M]
516+
(hγ : IsIntegralCurve γ v)
517+
(hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) :
518+
Xor' (∃ T > 0, Periodic γ T) (Injective γ) := by
519+
rw [xor_iff_iff_not]
520+
refine ⟨fun ⟨T, hT, hf⟩ ↦ hf.not_injective (ne_of_gt hT), ?_⟩
521+
intro h
522+
rw [Injective] at h
523+
push_neg at h
524+
obtain ⟨a, b, heq, hne⟩ := h
525+
refine ⟨|a - b|, ?_, ?_⟩
526+
· rw [gt_iff_lt, abs_pos, sub_ne_zero]
527+
exact hne
528+
· by_cases hab : a - b < 0
529+
· rw [abs_of_neg hab, neg_sub]
530+
exact hγ.periodic_of_eq hv heq.symm
531+
· rw [not_lt] at hab
532+
rw [abs_of_nonneg hab]
533+
exact hγ.periodic_of_eq hv heq
534+
504535
end ExistUnique

0 commit comments

Comments
 (0)