feat: Time dependent Wick theorem. (#274)
feat: Proof of the time-dependent Wick's theorem
This commit is contained in:
parent
4d43698b3c
commit
17f84b7153
53 changed files with 8563 additions and 3329 deletions
|
@ -3,15 +3,14 @@ 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 Mathlib.Order.Defs.Unbundled
|
||||
import Mathlib.Data.Fintype.Basic
|
||||
import Mathlib.Algebra.BigOperators.Group.Finset
|
||||
/-!
|
||||
|
||||
# Creation and annihilation parts of fields
|
||||
# Creation and annihlation parts of fields
|
||||
|
||||
-/
|
||||
|
||||
/-- The type specifying whether an operator is a creation or annihilation operator. -/
|
||||
/-- The type specifing whether an operator is a creation or annihilation operator. -/
|
||||
inductive CreateAnnihilate where
|
||||
| create : CreateAnnihilate
|
||||
| annihilate : CreateAnnihilate
|
||||
|
@ -29,14 +28,23 @@ instance : Fintype CreateAnnihilate where
|
|||
· refine Finset.insert_eq_self.mp ?_
|
||||
exact rfl
|
||||
|
||||
/-- The normal ordering on creation and annihilation operators.
|
||||
Creation operators are put to the left. -/
|
||||
lemma eq_create_or_annihilate (φ : CreateAnnihilate) : φ = create ∨ φ = annihilate := by
|
||||
cases φ <;> simp
|
||||
|
||||
/-- The normal ordering on creation and annihlation operators.
|
||||
Under this relation, `normalOrder a b` is false only if `a` is annihlate and `b` is create. -/
|
||||
def normalOrder : CreateAnnihilate → CreateAnnihilate → Prop
|
||||
| create, create => True
|
||||
| create, _ => True
|
||||
| annihilate, annihilate => True
|
||||
| create, annihilate => True
|
||||
| annihilate, create => False
|
||||
|
||||
/-- The normal ordering on `CreateAnnihilate` is decidable. -/
|
||||
instance : (φ φ' : CreateAnnihilate) → Decidable (normalOrder φ φ')
|
||||
| create, create => isTrue True.intro
|
||||
| annihilate, annihilate => isTrue True.intro
|
||||
| create, annihilate => isTrue True.intro
|
||||
| annihilate, create => isFalse False.elim
|
||||
|
||||
/-- Normal ordering is total. -/
|
||||
instance : IsTotal CreateAnnihilate normalOrder where
|
||||
total a b := by
|
||||
|
@ -47,4 +55,16 @@ instance : IsTrans CreateAnnihilate normalOrder where
|
|||
trans a b c := by
|
||||
cases a <;> cases b <;> cases c <;> simp [normalOrder]
|
||||
|
||||
@[simp]
|
||||
lemma not_normalOrder_annihilate_iff_false (a : CreateAnnihilate) :
|
||||
(¬ normalOrder a annihilate) ↔ False := by
|
||||
cases a
|
||||
· simp [normalOrder]
|
||||
· simp [normalOrder]
|
||||
|
||||
lemma sum_eq {M : Type} [AddCommMonoid M] (f : CreateAnnihilate → M) :
|
||||
∑ i, f i = f create + f annihilate := by
|
||||
change ∑ i in {create, annihilate}, f i = f create + f annihilate
|
||||
simp
|
||||
|
||||
end CreateAnnihilate
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue