@@ -60,7 +60,7 @@ integral curve, vector field, local existence, uniqueness
6060
6161open scoped Manifold Topology
6262
63- open Set
63+ open Function Set
6464
6565variable
6666 {E : Type *} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E]
@@ -193,8 +193,7 @@ section Translation
193193lemma 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
246245lemma 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+
504535end ExistUnique
0 commit comments