refactor: Lint
This commit is contained in:
parent
fca3f02eca
commit
c8e9c285a3
8 changed files with 229 additions and 194 deletions
|
@ -8,7 +8,7 @@ import HepLean.PerturbationTheory.WickContraction.StaticContract
|
|||
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeContraction
|
||||
/-!
|
||||
|
||||
# Sub contractions
|
||||
# Sub contractions
|
||||
|
||||
-/
|
||||
|
||||
|
@ -20,7 +20,10 @@ variable {n : ℕ} {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
|||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
|
||||
def subContraction (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) : WickContraction φs.length :=
|
||||
/-- Given a Wick contraction `φsΛ`, and a subset of `φsΛ.1`, the Wick contraction
|
||||
conisting of contracted pairs within that subset. -/
|
||||
def subContraction (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
WickContraction φs.length :=
|
||||
⟨S, by
|
||||
intro i hi
|
||||
exact φsΛ.2.1 i (ha hi),
|
||||
|
@ -32,13 +35,16 @@ lemma mem_of_mem_subContraction {S : Finset (Finset (Fin φs.length))} {hs : S
|
|||
{a : Finset (Fin φs.length)} (ha : a ∈ (φsΛ.subContraction S hs).1) : a ∈ φsΛ.1 := by
|
||||
exact hs ha
|
||||
|
||||
/-- Given a Wick contraction `φsΛ`, and a subset `S` of `φsΛ.1`, the Wick contraction
|
||||
on the uncontracted list `[φsΛ.subContraction S ha]ᵘᶜ`
|
||||
consisting of the remaining contracted pairs of `φsΛ` not in `S`. -/
|
||||
def quotContraction (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
WickContraction [φsΛ.subContraction S ha]ᵘᶜ.length :=
|
||||
⟨Finset.filter (fun a => Finset.map uncontractedListEmd a ∈ φsΛ.1) Finset.univ,
|
||||
by
|
||||
intro a ha'
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and] at ha'
|
||||
simpa using φsΛ.2.1 (Finset.map uncontractedListEmd a) ha' , by
|
||||
simpa using φsΛ.2.1 (Finset.map uncontractedListEmd a) ha', by
|
||||
intro a ha b hb
|
||||
simp only [Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
|
||||
by_cases hab : a = b
|
||||
|
@ -103,13 +109,16 @@ lemma subContraction_fstFieldOfContract {S : Finset (Finset (Fin φs.length))} {
|
|||
(a : (subContraction S hs).1) :
|
||||
(subContraction S hs).fstFieldOfContract a =
|
||||
φsΛ.fstFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩:= by
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _ (φsΛ.sndFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _
|
||||
(φsΛ.sndFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
simp
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
|
@ -121,27 +130,31 @@ lemma subContraction_sndFieldOfContract {S : Finset (Finset (Fin φs.length))} {
|
|||
(a : (subContraction S hs).1) :
|
||||
(subContraction S hs).sndFieldOfContract a =
|
||||
φsΛ.sndFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩:= by
|
||||
apply eq_sndFieldOfContract_of_mem _ _ (φsΛ.fstFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
apply eq_sndFieldOfContract_of_mem _ _
|
||||
(φsΛ.fstFieldOfContract ⟨a.1, mem_of_mem_subContraction a.2⟩)
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
simp
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _ ⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
· have ha := finset_eq_fstFieldOfContract_sndFieldOfContract _
|
||||
⟨a.1, mem_of_mem_subContraction a.2⟩
|
||||
simp only at ha
|
||||
conv_lhs =>
|
||||
rw [ha]
|
||||
simp
|
||||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ ⟨↑a, mem_of_mem_subContraction a.property⟩
|
||||
|
||||
|
||||
@[simp]
|
||||
lemma quotContraction_fstFieldOfContract_uncontractedListEmd {S : Finset (Finset (Fin φs.length))}
|
||||
{hs : S ⊆ φsΛ.1} (a : (quotContraction S hs).1) :
|
||||
uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a) =
|
||||
(φsΛ.fstFieldOfContract ⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
(φsΛ.fstFieldOfContract
|
||||
⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
symm
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _ (uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a) )
|
||||
apply eq_fstFieldOfContract_of_mem _ _ _
|
||||
(uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a))
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
|
@ -151,9 +164,11 @@ lemma quotContraction_fstFieldOfContract_uncontractedListEmd {S : Finset (Finset
|
|||
lemma quotContraction_sndFieldOfContract_uncontractedListEmd {S : Finset (Finset (Fin φs.length))}
|
||||
{hs : S ⊆ φsΛ.1} (a : (quotContraction S hs).1) :
|
||||
uncontractedListEmd ((quotContraction S hs).sndFieldOfContract a) =
|
||||
(φsΛ.sndFieldOfContract ⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
(φsΛ.sndFieldOfContract
|
||||
⟨Finset.map uncontractedListEmd a.1, mem_of_mem_quotContraction a.2⟩) := by
|
||||
symm
|
||||
apply eq_sndFieldOfContract_of_mem _ _ (uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a) )
|
||||
apply eq_sndFieldOfContract_of_mem _ _
|
||||
(uncontractedListEmd ((quotContraction S hs).fstFieldOfContract a))
|
||||
· simp only [Finset.mem_map', fstFieldOfContract_mem]
|
||||
· simp
|
||||
· apply uncontractedListEmd_strictMono
|
||||
|
@ -164,7 +179,6 @@ lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {h
|
|||
GradingCompliant [φsΛ.subContraction S hs]ᵘᶜ (quotContraction S hs) := by
|
||||
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall]
|
||||
intro a ha
|
||||
have h1' := mem_of_mem_quotContraction ha
|
||||
erw [subContraction_uncontractedList_get]
|
||||
erw [subContraction_uncontractedList_get]
|
||||
simp only [quotContraction_fstFieldOfContract_uncontractedListEmd, Fin.getElem_fin,
|
||||
|
@ -173,7 +187,7 @@ lemma quotContraction_gradingCompliant {S : Finset (Finset (Fin φs.length))} {h
|
|||
|
||||
lemma mem_quotContraction_iff {S : Finset (Finset (Fin φs.length))} {hs : S ⊆ φsΛ.1}
|
||||
{a : Finset (Fin [φsΛ.subContraction S hs]ᵘᶜ.length)} :
|
||||
a ∈ (quotContraction S hs).1 ↔ a.map uncontractedListEmd ∈ φsΛ.1
|
||||
a ∈ (quotContraction S hs).1 ↔ a.map uncontractedListEmd ∈ φsΛ.1
|
||||
∧ a.map uncontractedListEmd ∉ (subContraction S hs).1 := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
|
@ -184,5 +198,4 @@ lemma mem_quotContraction_iff {S : Finset (Finset (Fin φs.length))} {hs : S ⊆
|
|||
have h2 := mem_subContraction_or_quotContraction (S := S) (hs := hs) h.1
|
||||
simp_all
|
||||
|
||||
|
||||
end WickContraction
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue