/- Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.FieldSpecification.CrAnSection import HepLean.PerturbationTheory.FieldSpecification.NormalOrder /-! # Time ordering of states -/ namespace FieldSpecification variable {𝓕 : FieldSpecification} /-! ## Time ordering for states -/ /-- The time ordering relation on states. We have that `timeOrderRel Ο†0 Ο†1` is true if and only if `Ο†1` has a time less-then or equal to `Ο†0`, or `Ο†1` is a negative asymptotic state, or `Ο†0` is a positive asymptotic state. -/ def timeOrderRel : 𝓕.FieldOp β†’ 𝓕.FieldOp β†’ Prop | FieldOp.outAsymp _, _ => True | FieldOp.position Ο†0, FieldOp.position Ο†1 => Ο†1.2 0 ≀ Ο†0.2 0 | FieldOp.position _, FieldOp.inAsymp _ => True | FieldOp.position _, FieldOp.outAsymp _ => False | FieldOp.inAsymp _, FieldOp.outAsymp _ => False | FieldOp.inAsymp _, FieldOp.position _ => False | FieldOp.inAsymp _, FieldOp.inAsymp _ => True /-- The relation `timeOrderRel` is decidable, but not computable so due to `Real.decidableLE`. -/ noncomputable instance : (Ο† Ο†' : 𝓕.FieldOp) β†’ Decidable (timeOrderRel Ο† Ο†') | FieldOp.outAsymp _, _ => isTrue True.intro | FieldOp.position Ο†0, FieldOp.position Ο†1 => inferInstanceAs (Decidable (Ο†1.2 0 ≀ Ο†0.2 0)) | FieldOp.position _, FieldOp.inAsymp _ => isTrue True.intro | FieldOp.position _, FieldOp.outAsymp _ => isFalse (fun a => a) | FieldOp.inAsymp _, FieldOp.outAsymp _ => isFalse (fun a => a) | FieldOp.inAsymp _, FieldOp.position _ => isFalse (fun a => a) | FieldOp.inAsymp _, FieldOp.inAsymp _ => isTrue True.intro /-- Time ordering is total. -/ instance : IsTotal 𝓕.FieldOp 𝓕.timeOrderRel where total a b := by cases a <;> cases b <;> simp only [or_self, or_false, or_true, timeOrderRel, Fin.isValue, implies_true, imp_self, IsEmpty.forall_iff] exact LinearOrder.le_total _ _ /-- Time ordering is transitive. -/ instance : IsTrans 𝓕.FieldOp 𝓕.timeOrderRel where trans a b c := by cases a <;> cases b <;> cases c <;> simp only [timeOrderRel, Fin.isValue, implies_true, imp_self, IsEmpty.forall_iff] exact fun h1 h2 => Preorder.le_trans _ _ _ h2 h1 noncomputable section open FieldStatistic open HepLean.List /-- Given a list `Ο† :: Ο†s` of states, the (zero-based) position of the state which is of maximum time. For example - for the list `[Ο†1(t = 4), Ο†2(t = 5), Ο†3(t = 3), Ο†4(t = 5)]` this would return `1`. This is defined for a list `Ο† :: Ο†s` instead of `Ο†s` to ensure that such a position exists. -/ def maxTimeFieldPos (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : β„• := insertionSortMinPos timeOrderRel Ο† Ο†s lemma maxTimeFieldPos_lt_length (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : maxTimeFieldPos Ο† Ο†s < (Ο† :: Ο†s).length := by simp [maxTimeFieldPos] /-- Given a list `Ο† :: Ο†s` of states, the left-most state of maximum time, if there are more. As an example: - for the list `[Ο†1(t = 4), Ο†2(t = 5), Ο†3(t = 3), Ο†4(t = 5)]` this would return `Ο†2(t = 5)`. It is the state at the position `maxTimeFieldPos Ο† Ο†s`. -/ def maxTimeField (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : 𝓕.FieldOp := insertionSortMin timeOrderRel Ο† Ο†s /-- Given a list `Ο† :: Ο†s` of states, the list with the left-most state of maximum time removed. As an example: - for the list `[Ο†1(t = 4), Ο†2(t = 5), Ο†3(t = 3), Ο†4(t = 5)]` this would return `[Ο†1(t = 4), Ο†3(t = 3), Ο†4(t = 5)]`. -/ def eraseMaxTimeField (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : List 𝓕.FieldOp := insertionSortDropMinPos timeOrderRel Ο† Ο†s @[simp] lemma eraseMaxTimeField_length (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : (eraseMaxTimeField Ο† Ο†s).length = Ο†s.length := by simp [eraseMaxTimeField, insertionSortDropMinPos, eraseIdx_length'] lemma maxTimeFieldPos_lt_eraseMaxTimeField_length_succ (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : maxTimeFieldPos Ο† Ο†s < (eraseMaxTimeField Ο† Ο†s).length.succ := by simp only [eraseMaxTimeField_length, Nat.succ_eq_add_one] exact maxTimeFieldPos_lt_length Ο† Ο†s /-- Given a list `Ο† :: Ο†s` of states, the position of the left-most state of maximum time as an element of `Fin (eraseMaxTimeField Ο† Ο†s).length.succ`. As an example: - for the list `[Ο†1(t = 4), Ο†2(t = 5), Ο†3(t = 3), Ο†4(t = 5)]` this would return `⟨1,...⟩`. -/ def maxTimeFieldPosFin (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : Fin (eraseMaxTimeField Ο† Ο†s).length.succ := insertionSortMinPosFin timeOrderRel Ο† Ο†s lemma lt_maxTimeFieldPosFin_not_timeOrder (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) (i : Fin (eraseMaxTimeField Ο† Ο†s).length) (hi : (maxTimeFieldPosFin Ο† Ο†s).succAbove i < maxTimeFieldPosFin Ο† Ο†s) : Β¬ timeOrderRel ((eraseMaxTimeField Ο† Ο†s)[i.val]) (maxTimeField Ο† Ο†s) := by exact insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt timeOrderRel Ο† Ο†s i hi lemma timeOrder_maxTimeField (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) (i : Fin (eraseMaxTimeField Ο† Ο†s).length) : timeOrderRel (maxTimeField Ο† Ο†s) ((eraseMaxTimeField Ο† Ο†s)[i.val]) := by exact insertionSortMin_lt_mem_insertionSortDropMinPos timeOrderRel Ο† Ο†s _ /-- The sign associated with putting a list of states into time order (with the state of greatest time to the left). We pick up a minus sign for every fermion paired crossed. -/ def timeOrderSign (Ο†s : List 𝓕.FieldOp) : β„‚ := Wick.koszulSign 𝓕.fieldOpStatistic 𝓕.timeOrderRel Ο†s @[simp] lemma timeOrderSign_nil : timeOrderSign (𝓕 := 𝓕) [] = 1 := by simp only [timeOrderSign] rfl lemma timeOrderSign_pair_ordered {Ο† ψ : 𝓕.FieldOp} (h : timeOrderRel Ο† ψ) : timeOrderSign [Ο†, ψ] = 1 := by simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, ite_eq_left_iff, ite_eq_right_iff, and_imp] exact fun h' => False.elim (h' h) lemma timeOrderSign_pair_not_ordered {Ο† ψ : 𝓕.FieldOp} (h : Β¬ timeOrderRel Ο† ψ) : timeOrderSign [Ο†, ψ] = 𝓒(𝓕 |>β‚› Ο†, 𝓕 |>β‚› ψ) := by simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, instCommGroup.eq_1] rw [if_neg h] simp [FieldStatistic.exchangeSign_eq_if] lemma timerOrderSign_of_eraseMaxTimeField (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : timeOrderSign (eraseMaxTimeField Ο† Ο†s) = timeOrderSign (Ο† :: Ο†s) * 𝓒(𝓕 |>β‚› maxTimeField Ο† Ο†s, 𝓕 |>β‚› (Ο† :: Ο†s).take (maxTimeFieldPos Ο† Ο†s)) := by rw [eraseMaxTimeField, insertionSortDropMinPos, timeOrderSign, Wick.koszulSign_eraseIdx_insertionSortMinPos] rw [← timeOrderSign, ← maxTimeField] rfl /-- The time ordering of a list of states. A schematic example is: - `normalOrderList [Ο†1(t = 4), Ο†2(t = 5), Ο†3(t = 3), Ο†4(t = 5)]` is equal to `[Ο†2(t = 5), Ο†4(t = 5), Ο†1(t = 4), Ο†3(t = 3)]` -/ def timeOrderList (Ο†s : List 𝓕.FieldOp) : List 𝓕.FieldOp := List.insertionSort 𝓕.timeOrderRel Ο†s lemma timeOrderList_pair_ordered {Ο† ψ : 𝓕.FieldOp} (h : timeOrderRel Ο† ψ) : timeOrderList [Ο†, ψ] = [Ο†, ψ] := by simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_left_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h' h) lemma timeOrderList_pair_not_ordered {Ο† ψ : 𝓕.FieldOp} (h : Β¬ timeOrderRel Ο† ψ) : timeOrderList [Ο†, ψ] = [ψ, Ο†] := by simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_right_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h h') @[simp] lemma timeOrderList_nil : timeOrderList (𝓕 := 𝓕) [] = [] := by simp [timeOrderList] lemma timeOrderList_eq_maxTimeField_timeOrderList (Ο† : 𝓕.FieldOp) (Ο†s : List 𝓕.FieldOp) : timeOrderList (Ο† :: Ο†s) = maxTimeField Ο† Ο†s :: timeOrderList (eraseMaxTimeField Ο† Ο†s) := by exact insertionSort_eq_insertionSortMin_cons timeOrderRel Ο† Ο†s /-! ## Time ordering for CrAnFieldOp -/ /-! ## timeOrderRel -/ /-- For a field specification `𝓕`, `𝓕.crAnTimeOrderRel` is a relation on `𝓕.CrAnFieldOp` representing time ordering. It is defined such that `𝓕.crAnTimeOrderRel Ο†β‚€ φ₁` is true if and only if one of the following holds - `Ο†β‚€` is an *outgoing* asymptotic operator - `φ₁` is an *incoming* asymptotic field operator - `Ο†β‚€` and `φ₁` are both position field operators where the `SpaceTime` point of `Ο†β‚€` has a time *greater* then or equal to that of `φ₁`. Thus, colloquially `𝓕.crAnTimeOrderRel Ο†β‚€ φ₁` if `Ο†β‚€` has time *greater* then or equal to `φ₁`. The use of *greater* then rather then *less* then is because on ordering lists of operators it is needed that the operator with the greatest time is to the left. -/ def crAnTimeOrderRel (a b : 𝓕.CrAnFieldOp) : Prop := 𝓕.timeOrderRel a.1 b.1 /-- The relation `crAnTimeOrderRel` is decidable, but not computable so due to `Real.decidableLE`. -/ noncomputable instance (Ο† Ο†' : 𝓕.CrAnFieldOp) : Decidable (crAnTimeOrderRel Ο† Ο†') := inferInstanceAs (Decidable (𝓕.timeOrderRel Ο†.1 Ο†'.1)) /-- Time ordering of `CrAnFieldOp` is total. -/ instance : IsTotal 𝓕.CrAnFieldOp 𝓕.crAnTimeOrderRel where total a b := IsTotal.total (r := 𝓕.timeOrderRel) a.1 b.1 /-- Time ordering of `CrAnFieldOp` is transitive. -/ instance : IsTrans 𝓕.CrAnFieldOp 𝓕.crAnTimeOrderRel where trans a b c := IsTrans.trans (r := 𝓕.timeOrderRel) a.1 b.1 c.1 @[simp] lemma crAnTimeOrderRel_refl (Ο† : 𝓕.CrAnFieldOp) : crAnTimeOrderRel Ο† Ο† := by exact (IsTotal.to_isRefl (r := 𝓕.crAnTimeOrderRel)).refl Ο† /-- For a field specification `𝓕`, and a list `Ο†s` of `𝓕.CrAnFieldOp`, `𝓕.crAnTimeOrderSign Ο†s` is the sign corresponding to the number of `ferimionic`-`fermionic` exchanges undertaken to time-order (i.e. order with respect to `𝓕.crAnTimeOrderRel`) `Ο†s` using the insertion sort algorithm. -/ def crAnTimeOrderSign (Ο†s : List 𝓕.CrAnFieldOp) : β„‚ := Wick.koszulSign 𝓕.crAnStatistics 𝓕.crAnTimeOrderRel Ο†s @[simp] lemma crAnTimeOrderSign_nil : crAnTimeOrderSign (𝓕 := 𝓕) [] = 1 := by simp only [crAnTimeOrderSign] rfl lemma crAnTimeOrderSign_pair_ordered {Ο† ψ : 𝓕.CrAnFieldOp} (h : crAnTimeOrderRel Ο† ψ) : crAnTimeOrderSign [Ο†, ψ] = 1 := by simp only [crAnTimeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, ite_eq_left_iff, ite_eq_right_iff, and_imp] exact fun h' => False.elim (h' h) lemma crAnTimeOrderSign_pair_not_ordered {Ο† ψ : 𝓕.CrAnFieldOp} (h : Β¬ crAnTimeOrderRel Ο† ψ) : crAnTimeOrderSign [Ο†, ψ] = 𝓒(𝓕 |>β‚› Ο†, 𝓕 |>β‚› ψ) := by simp only [crAnTimeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, instCommGroup.eq_1] rw [if_neg h] simp [FieldStatistic.exchangeSign_eq_if] lemma crAnTimeOrderSign_swap_eq_time {Ο† ψ : 𝓕.CrAnFieldOp} (h1 : crAnTimeOrderRel Ο† ψ) (h2 : crAnTimeOrderRel ψ Ο†) (Ο†s Ο†s' : List 𝓕.CrAnFieldOp) : crAnTimeOrderSign (Ο†s ++ Ο† :: ψ :: Ο†s') = crAnTimeOrderSign (Ο†s ++ ψ :: Ο† :: Ο†s') := by exact Wick.koszulSign_swap_eq_rel _ _ h1 h2 _ _ /-- For a field specification `𝓕`, and a list `Ο†s` of `𝓕.CrAnFieldOp`, `𝓕.crAnTimeOrderList Ο†s` is the list `Ο†s` time-ordered using the insertion sort algorithm. -/ def crAnTimeOrderList (Ο†s : List 𝓕.CrAnFieldOp) : List 𝓕.CrAnFieldOp := List.insertionSort 𝓕.crAnTimeOrderRel Ο†s @[simp] lemma crAnTimeOrderList_nil : crAnTimeOrderList (𝓕 := 𝓕) [] = [] := by simp [crAnTimeOrderList] lemma crAnTimeOrderList_pair_ordered {Ο† ψ : 𝓕.CrAnFieldOp} (h : crAnTimeOrderRel Ο† ψ) : crAnTimeOrderList [Ο†, ψ] = [Ο†, ψ] := by simp only [crAnTimeOrderList, List.insertionSort, List.orderedInsert, ite_eq_left_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h' h) lemma crAnTimeOrderList_pair_not_ordered {Ο† ψ : 𝓕.CrAnFieldOp} (h : Β¬ crAnTimeOrderRel Ο† ψ) : crAnTimeOrderList [Ο†, ψ] = [ψ, Ο†] := by simp only [crAnTimeOrderList, List.insertionSort, List.orderedInsert, ite_eq_right_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h h') lemma orderedInsert_swap_eq_time {Ο† ψ : 𝓕.CrAnFieldOp} (h1 : crAnTimeOrderRel Ο† ψ) (h2 : crAnTimeOrderRel ψ Ο†) (Ο†s : List 𝓕.CrAnFieldOp) : List.orderedInsert crAnTimeOrderRel Ο† (List.orderedInsert crAnTimeOrderRel ψ Ο†s) = List.takeWhile (fun b => Β¬ crAnTimeOrderRel ψ b) Ο†s ++ Ο† :: ψ :: List.dropWhile (fun b => Β¬ crAnTimeOrderRel ψ b) Ο†s := by rw [List.orderedInsert_eq_take_drop crAnTimeOrderRel ψ Ο†s] simp only [decide_not] rw [List.orderedInsert_eq_take_drop] simp only [decide_not] have h1 (b : 𝓕.CrAnFieldOp) : (crAnTimeOrderRel Ο† b) ↔ (crAnTimeOrderRel ψ b) := Iff.intro (fun h => IsTrans.trans _ _ _ h2 h) (fun h => IsTrans.trans _ _ _ h1 h) simp only [h1] rw [List.takeWhile_append] rw [List.takeWhile_takeWhile] simp only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, and_self, decide_not, ↓reduceIte, crAnTimeOrderRel_refl, decide_true, Bool.false_eq_true, not_false_eq_true, List.takeWhile_cons_of_neg, List.append_nil, List.append_cancel_left_eq, List.cons.injEq, true_and] rw [List.dropWhile_append] simp only [List.isEmpty_eq_true, List.dropWhile_eq_nil_iff, Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, crAnTimeOrderRel_refl, decide_true, Bool.false_eq_true, not_false_eq_true, List.dropWhile_cons_of_neg, ite_eq_left_iff, not_forall, Classical.not_imp, Decidable.not_not, List.append_left_eq_self, forall_exists_index, and_imp] intro x hx hxψ intro y hy simpa using List.mem_takeWhile_imp hy lemma orderedInsert_in_swap_eq_time {Ο† ψ Ο†': 𝓕.CrAnFieldOp} (h1 : crAnTimeOrderRel Ο† ψ) (h2 : crAnTimeOrderRel ψ Ο†) : (Ο†s Ο†s' : List 𝓕.CrAnFieldOp) β†’ βˆƒ l1 l2, List.orderedInsert crAnTimeOrderRel Ο†' (Ο†s ++ Ο† :: ψ :: Ο†s') = l1 ++ Ο† :: ψ :: l2 ∧ List.orderedInsert crAnTimeOrderRel Ο†' (Ο†s ++ ψ :: Ο† :: Ο†s') = l1 ++ ψ :: Ο† :: l2 | [], Ο†s' => by have h1 (b : 𝓕.CrAnFieldOp) : (crAnTimeOrderRel b Ο†) ↔ (crAnTimeOrderRel b ψ) := Iff.intro (fun h => IsTrans.trans _ _ _ h h1) (fun h => IsTrans.trans _ _ _ h h2) by_cases h : crAnTimeOrderRel Ο†' Ο† Β· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 Ο†'] use [Ο†'], Ο†s' simp Β· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 Ο†'] use [], List.orderedInsert crAnTimeOrderRel Ο†' Ο†s' simp | Ο†'' :: Ο†s, Ο†s' => by obtain ⟨l1, l2, hl⟩ := orderedInsert_in_swap_eq_time (Ο†' := Ο†') h1 h2 Ο†s Ο†s' simp only [List.orderedInsert, List.append_eq] rw [hl.1, hl.2] by_cases h : crAnTimeOrderRel Ο†' Ο†'' Β· simp only [h, ↓reduceIte] use (Ο†' :: Ο†'' :: Ο†s), Ο†s' simp Β· simp only [h, ↓reduceIte] use (Ο†'' :: l1), l2 simp lemma crAnTimeOrderList_swap_eq_time {Ο† ψ : 𝓕.CrAnFieldOp} (h1 : crAnTimeOrderRel Ο† ψ) (h2 : crAnTimeOrderRel ψ Ο†) : (Ο†s Ο†s' : List 𝓕.CrAnFieldOp) β†’ βˆƒ (l1 l2 : List 𝓕.CrAnFieldOp), crAnTimeOrderList (Ο†s ++ Ο† :: ψ :: Ο†s') = l1 ++ Ο† :: ψ :: l2 ∧ crAnTimeOrderList (Ο†s ++ ψ :: Ο† :: Ο†s') = l1 ++ ψ :: Ο† :: l2 | [], Ο†s' => by simp only [crAnTimeOrderList] simp only [List.insertionSort] use List.takeWhile (fun b => Β¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel Ο†s'), List.dropWhile (fun b => Β¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel Ο†s') apply And.intro Β· exact orderedInsert_swap_eq_time h1 h2 _ Β· have h1' (b : 𝓕.CrAnFieldOp) : (crAnTimeOrderRel Ο† b) ↔ (crAnTimeOrderRel ψ b) := Iff.intro (fun h => IsTrans.trans _ _ _ h2 h) (fun h => IsTrans.trans _ _ _ h1 h) simp only [← h1', decide_not] simpa using orderedInsert_swap_eq_time h2 h1 _ | Ο†'' :: Ο†s, Ο†s' => by rw [crAnTimeOrderList, crAnTimeOrderList] simp only [List.insertionSort, List.append_eq] obtain ⟨l1, l2, hl⟩ := crAnTimeOrderList_swap_eq_time h1 h2 Ο†s Ο†s' simp only [crAnTimeOrderList] at hl rw [hl.1, hl.2] obtain ⟨l1', l2', hl'⟩ := orderedInsert_in_swap_eq_time (Ο†' := Ο†'') h1 h2 l1 l2 rw [hl'.1, hl'.2] use l1', l2' /-! ## Relationship to sections -/ lemma koszulSignInsert_crAnTimeOrderRel_crAnSection {Ο† : 𝓕.FieldOp} {ψ : 𝓕.CrAnFieldOp} (h : ψ.1 = Ο†) : {Ο†s : List 𝓕.FieldOp} β†’ (ψs : CrAnSection Ο†s) β†’ Wick.koszulSignInsert 𝓕.crAnStatistics 𝓕.crAnTimeOrderRel ψ ψs.1 = Wick.koszulSignInsert 𝓕.fieldOpStatistic 𝓕.timeOrderRel Ο† Ο†s | [], ⟨[], h⟩ => by simp [Wick.koszulSignInsert] | Ο†' :: Ο†s, ⟨ψ' :: ψs, h1⟩ => by simp only [Wick.koszulSignInsert, crAnTimeOrderRel, h] simp only [List.map_cons, List.cons.injEq] at h1 have hi := koszulSignInsert_crAnTimeOrderRel_crAnSection h (Ο†s := Ο†s) ⟨ψs, h1.2⟩ rw [hi] congr Β· exact h1.1 Β· simp only [crAnStatistics, crAnFieldOpToFieldOp, Function.comp_apply] congr Β· simp only [crAnStatistics, crAnFieldOpToFieldOp, Function.comp_apply] congr exact h1.1 @[simp] lemma crAnTimeOrderSign_crAnSection : {Ο†s : List 𝓕.FieldOp} β†’ (ψs : CrAnSection Ο†s) β†’ crAnTimeOrderSign ψs.1 = timeOrderSign Ο†s | [], ⟨[], h⟩ => by simp | Ο† :: Ο†s, ⟨ψ :: ψs, h⟩ => by simp only [crAnTimeOrderSign, Wick.koszulSign, timeOrderSign] simp only [List.map_cons, List.cons.injEq] at h congr 1 Β· rw [koszulSignInsert_crAnTimeOrderRel_crAnSection h.1 ⟨ψs, h.2⟩] Β· exact crAnTimeOrderSign_crAnSection ⟨ψs, h.2⟩ lemma orderedInsert_crAnTimeOrderRel_crAnSection {Ο† : 𝓕.FieldOp} {ψ : 𝓕.CrAnFieldOp} (h : ψ.1 = Ο†) : {Ο†s : List 𝓕.FieldOp} β†’ (ψs : CrAnSection Ο†s) β†’ (List.orderedInsert 𝓕.crAnTimeOrderRel ψ ψs.1).map 𝓕.crAnFieldOpToFieldOp = List.orderedInsert 𝓕.timeOrderRel Ο† Ο†s | [], ⟨[], _⟩ => by simp only [List.orderedInsert, List.map_cons, List.map_nil, List.cons.injEq, and_true] exact h | Ο†' :: Ο†s, ⟨ψ' :: ψs, h1⟩ => by simp only [List.orderedInsert, crAnTimeOrderRel, h] simp only [List.map_cons, List.cons.injEq] at h1 by_cases hr : timeOrderRel Ο† Ο†' Β· simp only [hr, ↓reduceIte] rw [← h1.1] at hr simp only [crAnFieldOpToFieldOp] at hr simp only [hr, ↓reduceIte, List.map_cons, List.cons.injEq] exact And.intro h (And.intro h1.1 h1.2) Β· simp only [hr, ↓reduceIte] rw [← h1.1] at hr simp only [crAnFieldOpToFieldOp] at hr simp only [hr, ↓reduceIte, List.map_cons, List.cons.injEq] apply And.intro h1.1 exact orderedInsert_crAnTimeOrderRel_crAnSection h ⟨ψs, h1.2⟩ lemma crAnTimeOrderList_crAnSection_is_crAnSection : {Ο†s : List 𝓕.FieldOp} β†’ (ψs : CrAnSection Ο†s) β†’ (crAnTimeOrderList ψs.1).map 𝓕.crAnFieldOpToFieldOp = timeOrderList Ο†s | [], ⟨[], h⟩ => by simp | Ο† :: Ο†s, ⟨ψ :: ψs, h⟩ => by simp only [crAnTimeOrderList, List.insertionSort, timeOrderList] simp only [List.map_cons, List.cons.injEq] at h exact orderedInsert_crAnTimeOrderRel_crAnSection h.1 ⟨(List.insertionSort crAnTimeOrderRel ψs), crAnTimeOrderList_crAnSection_is_crAnSection ⟨ψs, h.2⟩⟩ /-- Time ordering of sections of a list of states. -/ def crAnSectionTimeOrder (Ο†s : List 𝓕.FieldOp) (ψs : CrAnSection Ο†s) : CrAnSection (timeOrderList Ο†s) := ⟨crAnTimeOrderList ψs.1, crAnTimeOrderList_crAnSection_is_crAnSection ψs⟩ lemma orderedInsert_crAnTimeOrderRel_injective {ψ ψ' : 𝓕.CrAnFieldOp} (h : ψ.1 = ψ'.1) : {Ο†s : List 𝓕.FieldOp} β†’ (ψs ψs' : 𝓕.CrAnSection Ο†s) β†’ (ho : List.orderedInsert crAnTimeOrderRel ψ ψs.1 = List.orderedInsert crAnTimeOrderRel ψ' ψs'.1) β†’ ψ = ψ' ∧ ψs = ψs' | [], ⟨[], _⟩, ⟨[], _⟩, h => by simp only [List.orderedInsert, List.cons.injEq, and_true] at h simpa using h | Ο† :: Ο†s, ⟨ψ1 :: ψs, h1⟩, ⟨ψ1' :: ψs', h1'⟩, ho => by simp only [List.map_cons, List.cons.injEq] at h1 h1' have ih := orderedInsert_crAnTimeOrderRel_injective h ⟨ψs, h1.2⟩ ⟨ψs', h1'.2⟩ simp only [List.orderedInsert] at ho by_cases hr : crAnTimeOrderRel ψ ψ1 Β· simp_all only [ite_true] by_cases hr2 : crAnTimeOrderRel ψ' ψ1' Β· simp_all Β· simp only [crAnTimeOrderRel] at hr hr2 simp_all only rw [crAnFieldOpToFieldOp] at h1 h1' rw [h1.1] at hr rw [h1'.1] at hr2 exact False.elim (hr2 hr) Β· simp_all only [ite_false] by_cases hr2 : crAnTimeOrderRel ψ' ψ1' Β· simp only [crAnTimeOrderRel] at hr hr2 simp_all only rw [crAnFieldOpToFieldOp] at h1 h1' rw [h1.1] at hr rw [h1'.1] at hr2 exact False.elim (hr hr2) Β· simp only [hr2, ↓reduceIte, List.cons.injEq] at ho have ih' := ih ho.2 simp_all only [and_self, implies_true, not_false_eq_true, true_and] apply Subtype.ext simp only [List.cons.injEq, true_and] rw [Subtype.eq_iff] at ih' exact ih'.2 lemma crAnSectionTimeOrder_injective : {Ο†s : List 𝓕.FieldOp} β†’ Function.Injective (𝓕.crAnSectionTimeOrder Ο†s) | [], ⟨[], _⟩, ⟨[], _⟩ => by simp | Ο† :: Ο†s, ⟨ψ :: ψs, h⟩, ⟨ψ' :: ψs', h'⟩ => by intro h1 apply Subtype.ext simp only [List.cons.injEq] simp only [crAnSectionTimeOrder] at h1 rw [Subtype.eq_iff] at h1 simp only [crAnTimeOrderList, List.insertionSort] at h1 simp only [List.map_cons, List.cons.injEq] at h h' rw [crAnFieldOpToFieldOp] at h h' have hin := orderedInsert_crAnTimeOrderRel_injective (by rw [h.1, h'.1]) (𝓕.crAnSectionTimeOrder Ο†s ⟨ψs, h.2⟩) (𝓕.crAnSectionTimeOrder Ο†s ⟨ψs', h'.2⟩) h1 apply And.intro hin.1 have hl := crAnSectionTimeOrder_injective hin.2 rw [Subtype.ext_iff] at hl simpa using hl lemma crAnSectionTimeOrder_bijective (Ο†s : List 𝓕.FieldOp) : Function.Bijective (𝓕.crAnSectionTimeOrder Ο†s) := by rw [Fintype.bijective_iff_injective_and_card] apply And.intro crAnSectionTimeOrder_injective apply CrAnSection.card_perm_eq simp only [timeOrderList] exact List.Perm.symm (List.perm_insertionSort timeOrderRel Ο†s) lemma sum_crAnSections_timeOrder {Ο†s : List 𝓕.FieldOp} [AddCommMonoid M] (f : CrAnSection (timeOrderList Ο†s) β†’ M) : βˆ‘ s, f s = βˆ‘ s, f (𝓕.crAnSectionTimeOrder Ο†s s) := by erw [(Equiv.ofBijective _ (𝓕.crAnSectionTimeOrder_bijective Ο†s)).sum_comp] /-! ## normTimeOrderRel -/ /-- The time ordering relation on `CrAnFieldOp` such that if two CrAnFieldOp have the same time, we normal order them. -/ def normTimeOrderRel (a b : 𝓕.CrAnFieldOp) : Prop := crAnTimeOrderRel a b ∧ (crAnTimeOrderRel b a β†’ normalOrderRel a b) /-- The relation `normTimeOrderRel` is decidable, but not computable so due to `Real.decidableLE`. -/ noncomputable instance (Ο† Ο†' : 𝓕.CrAnFieldOp) : Decidable (normTimeOrderRel Ο† Ο†') := instDecidableAnd /-- Norm-Time ordering of `CrAnFieldOp` is total. -/ instance : IsTotal 𝓕.CrAnFieldOp 𝓕.normTimeOrderRel where total a b := by simp only [normTimeOrderRel] match IsTotal.total (r := 𝓕.crAnTimeOrderRel) a b, IsTotal.total (r := 𝓕.normalOrderRel) a b with | Or.inl h1, Or.inl h2 => simp [h1, h2] | Or.inr h1, Or.inl h2 => simp only [h1, h2, imp_self, and_true, true_and] by_cases hn : crAnTimeOrderRel a b Β· simp [hn] Β· simp [hn] | Or.inl h1, Or.inr h2 => simp only [h1, true_and, h2, imp_self, and_true] by_cases hn : crAnTimeOrderRel b a Β· simp [hn] Β· simp [hn] | Or.inr h1, Or.inr h2 => simp [h1, h2] /-- Norm-Time ordering of `CrAnFieldOp` is transitive. -/ instance : IsTrans 𝓕.CrAnFieldOp 𝓕.normTimeOrderRel where trans a b c := by intro h1 h2 simp_all only [normTimeOrderRel] apply And.intro Β· exact IsTrans.trans _ _ _ h1.1 h2.1 Β· intro hc refine IsTrans.trans _ _ _ (h1.2 ?_) (h2.2 ?_) Β· exact IsTrans.trans _ _ _ h2.1 hc Β· exact IsTrans.trans _ _ _ hc h1.1 /-- The sign associated with putting a list of `CrAnFieldOp` into normal-time order (with the state of greatest time to the left). We pick up a minus sign for every fermion paired crossed. -/ def normTimeOrderSign (Ο†s : List 𝓕.CrAnFieldOp) : β„‚ := Wick.koszulSign 𝓕.crAnStatistics 𝓕.normTimeOrderRel Ο†s /-- Sort a list of `CrAnFieldOp` based on `normTimeOrderRel`. -/ def normTimeOrderList (Ο†s : List 𝓕.CrAnFieldOp) : List 𝓕.CrAnFieldOp := List.insertionSort 𝓕.normTimeOrderRel Ο†s end end FieldSpecification