feat: More notes

This commit is contained in:
jstoobysmith 2025-02-05 05:44:40 +00:00
parent 256a1c3e94
commit 35445a5be6
5 changed files with 33 additions and 17 deletions

View file

@ -423,10 +423,12 @@ lemma timeOrder_ofFieldOpList_singleton (φ : 𝓕.FieldOp) :
𝓣(ofFieldOpList [φ]) = ofFieldOpList [φ] := by
rw [ofFieldOpList, timeOrder_eq_ι_timeOrderF, timeOrderF_ofFieldOpListF_singleton]
/-- The time order of a list `𝓣(φ₀…φₙ)` is equal to
`𝓢(φᵢ,φ₀…φᵢ₋₁) • φᵢ * 𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` where `φᵢ` is the maximal time field in `φ₀…φₙ`-/
lemma timeOrder_eq_maxTimeField_mul_finset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) :
𝓣(ofFieldOpList (φ :: φs)) = 𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ ⟨(eraseMaxTimeField φ φs).get,
(Finset.filter (fun x =>
(maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs) Finset.univ)⟩) •
(Finset.univ.filter (fun x =>
(maxTimeFieldPosFin φ φs).succAbove x < maxTimeFieldPosFin φ φs))⟩) •
ofFieldOp (maxTimeField φ φs) * 𝓣(ofFieldOpList (eraseMaxTimeField φ φs)) := by
rw [ofFieldOpList, timeOrder_eq_ι_timeOrderF, timeOrderF_eq_maxTimeField_mul_finset]
rfl