Skip to content

Commit

Permalink
refactor: Lint
Browse files Browse the repository at this point in the history
  • Loading branch information
jstoobysmith committed Jan 24, 2025
1 parent 2490535 commit c7bd59c
Show file tree
Hide file tree
Showing 10 changed files with 120 additions and 90 deletions.
1 change: 1 addition & 0 deletions HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -85,6 +85,7 @@ lemma ofStateAlgebra_ofState (φ : 𝓕.States) :
Roughly `[φ1, φ2]` gets sent to `(φ1ᶜ+ φ1ᵃ) * (φ2ᶜ+ φ2ᵃ)` etc. -/
def ofStateList (φs : List 𝓕.States) : CrAnAlgebra 𝓕 := (List.map ofState φs).prod

/-- Coercion from `List 𝓕.States` to `CrAnAlgebra 𝓕` through `ofStateList`. -/
instance : Coe (List 𝓕.States) (CrAnAlgebra 𝓕) := ⟨ofStateList⟩

@[simp]
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -36,6 +36,7 @@ def normalOrder : CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 :=
Basis.constr ofCrAnListBasis ℂ fun φs =>
normalOrderSign φs • ofCrAnList (normalOrderList φs)

@[inherit_doc normalOrder]
scoped[FieldSpecification.CrAnAlgebra] notation "𝓝(" a ")" => normalOrder a

lemma normalOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -406,7 +406,7 @@ lemma superCommute_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List

/--
Within the creation and annihilation algebra, we have that
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
where `sᵢ` is the exchange sign for `φᶜᵃs` and `φᶜᵃs₀ … φᶜᵃᵢ₋₁`.
-/
lemma superCommute_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -20,7 +20,7 @@ is a generalization of the operator algebra of a field theory with field specifi
It is an algebra `A` with a map `crAnF` from the creation and annihilation free algebra
satisfying a number of conditions with respect to super commutators.
The true operator algebra of a field theory with field specification `𝓕`is an
example of a proto-operator algebra.-/
example of a proto-operator algebra. -/
structure ProtoOperatorAlgebra where
/-- The algebra representing the operator algebra. -/
A : Type
Expand Down
31 changes: 19 additions & 12 deletions HepLean/PerturbationTheory/WickContraction/InsertAndContract.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,12 +29,13 @@ open HepLean.Fin
`j : Option (c.uncontracted)` of `c`.
The Wick contraction associated with `(φs.insertIdx i φ).length` formed by 'inserting' `φ`
into `φs` after the first `i` elements and contracting it optionally with j. -/
def insertAndContract {φs : List 𝓕.States} (φ : 𝓕.States) (φsΛ : WickContraction φs.length)
def insertAndContract {φs : List 𝓕.States} (φ : 𝓕.States) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : Option φsΛ.uncontracted) :
WickContraction (φs.insertIdx i φ).length :=
congr (by simp) (φsΛ.insertAndContractNat i j)

scoped[WickContraction] notation φs "↩Λ" φ:max i:max j => insertAndContract φ φs i j
@[inherit_doc insertAndContract]
scoped[WickContraction] notation φs "↩Λ" φ:max i:max j => insertAndContract φ φs i j

@[simp]
lemma insertAndContract_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
Expand All @@ -56,14 +57,16 @@ lemma insertAndContract_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.S
lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
(insertAndContract φ φsΛ i (some j)).fstFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩) =
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩) =
if i < i.succAbove j.1 then
finCongr (insertIdx_length_fin φ φs i).symm i else
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) := by
split
· rename_i h
refine (insertAndContract φ φsΛ i (some j)).eq_fstFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
· simp [congrLift]
Expand All @@ -72,7 +75,8 @@ lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
simp_all
· rename_i h
refine (insertAndContract φ φsΛ i (some j)).eq_fstFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
· simp [congrLift]
Expand Down Expand Up @@ -141,7 +145,8 @@ lemma insertAndContract_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φ
(i.succAbove j)) = Option.map (Fin.cast (insertIdx_length_fin φ φs i).symm ∘ i.succAbove)
(φsΛ.getDual? j) := by
simp only [Nat.succ_eq_add_one, insertAndContract, getDual?_congr, finCongr_apply, Fin.cast_trans,
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insertAndContractNat_some_getDual?_of_neq, Option.map_map]
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insertAndContractNat_some_getDual?_of_neq,
Option.map_map]
rfl

@[simp]
Expand All @@ -167,14 +172,16 @@ lemma insertAndContract_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕
lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
(φsΛ ↩Λ φ i (some j)).sndFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩) =
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩) =
if i < i.succAbove j.1 then
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) else
finCongr (insertIdx_length_fin φ φs i).symm i := by
split
· rename_i h
refine (φsΛ ↩Λ φ i (some j)).eq_sndFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
· simp [congrLift]
Expand All @@ -183,7 +190,8 @@ lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
simp_all
· rename_i h
refine (φsΛ ↩Λ φ i (some j)).eq_sndFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
simp [insertAndContractNat]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
· simp [congrLift]
Expand Down Expand Up @@ -240,8 +248,8 @@ lemma self_not_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List

lemma succAbove_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) (j : Fin φs.length) :
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j) ∈ insertAndContractLiftFinset φ i a ↔
j ∈ a := by
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)
∈ insertAndContractLiftFinset φ i a ↔ j ∈ a := by
simp only [insertAndContractLiftFinset, Finset.mem_map_equiv, finCongr_symm, finCongr_apply,
Fin.cast_trans, Fin.cast_eq_self]
simp only [Finset.mem_map, Fin.succAboveEmb_apply]
Expand All @@ -268,7 +276,6 @@ lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
use z
rfl


lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
∑ c, f c =
Expand Down
Loading

0 comments on commit c7bd59c

Please sign in to comment.