refactor: Lint
This commit is contained in:
parent
c993de36f6
commit
cd63ec0716
13 changed files with 181 additions and 84 deletions
|
@ -14,6 +14,8 @@ import Mathlib.Analysis.Complex.Basic
|
|||
|
||||
namespace Wick
|
||||
|
||||
/-- Given a grading map `q : I → Fin 2` and a list `l : List I` counts the parity of the number of
|
||||
elements in `l` which are of grade `1`. -/
|
||||
def grade {I : Type} (q : I → Fin 2) : (l : List I) → Fin 2
|
||||
| [] => 0
|
||||
| a :: l => if q a = grade q l then 0 else 1
|
||||
|
|
|
@ -14,7 +14,8 @@ import HepLean.PerturbationTheory.Wick.Signs.SuperCommuteCoef
|
|||
namespace Wick
|
||||
open HepLean.List
|
||||
|
||||
|
||||
/-- The sign associated with inserting `r0` into `r` at the position `n`.
|
||||
That is the sign associated with commuting `r0` with `List.take n r`. -/
|
||||
def insertSign {I : Type} (q : I → Fin 2) (n : ℕ) (r0 : I) (r : List I) : ℂ :=
|
||||
superCommuteCoef q [r0] (List.take n r)
|
||||
|
||||
|
|
|
@ -24,7 +24,6 @@ def koszulSign {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin
|
|||
| [] => 1
|
||||
| a :: l => koszulSignInsert r q a l * koszulSign r q l
|
||||
|
||||
|
||||
lemma koszulSign_mul_self {I : Type} (r : I → I → Prop) [DecidableRel r] (q : I → Fin 2)
|
||||
(l : List I) : koszulSign r q l * koszulSign r q l = 1 := by
|
||||
induction l with
|
||||
|
@ -84,36 +83,6 @@ lemma koszulSign_erase_boson {I : Type} (q : I → Fin 2) (le1 :I → I → Prop
|
|||
rw [koszulSignInsert_erase_boson q le1 r0 r ⟨n, Nat.succ_lt_succ_iff.mp h⟩ h']
|
||||
exact h'
|
||||
|
||||
|
||||
|
||||
def koszulSignCons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (r0 r1 : I) :
|
||||
ℂ :=
|
||||
if le1 r0 r1 then 1 else
|
||||
if q r0 = 1 ∧ q r1 = 1 then -1 else 1
|
||||
|
||||
lemma koszulSignCons_eq_superComuteCoef {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
|
||||
[DecidableRel le1] (r0 r1 : I) : koszulSignCons q le1 r0 r1 =
|
||||
if le1 r0 r1 then 1 else superCommuteCoef q [r0] [r1] := by
|
||||
simp only [koszulSignCons, Fin.isValue, superCommuteCoef, grade, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false]
|
||||
congr 1
|
||||
by_cases h0 : q r0 = 1
|
||||
· by_cases h1 : q r1 = 1
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = 0 := by omega
|
||||
simp [h0, h1]
|
||||
· have h0 : q r0 = 0 := by omega
|
||||
by_cases h1 : q r1 = 1
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = 0 := by omega
|
||||
simp [h0, h1]
|
||||
|
||||
lemma koszulSignInsert_cons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
[IsTotal I le1] [IsTrans I le1] (r0 r1 : I) (r : List I) :
|
||||
koszulSignInsert le1 q r0 (r1 :: r) = (koszulSignCons q le1 r0 r1) *
|
||||
koszulSignInsert le1 q r0 r := by
|
||||
simp [koszulSignInsert, koszulSignCons]
|
||||
|
||||
lemma koszulSign_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(i : I) [IsTotal I le1] [IsTrans I le1] : (r : List I) → (n : ℕ) → (hn : n ≤ r.length) →
|
||||
koszulSign le1 q (List.insertIdx n i r) = insertSign q n i r
|
||||
|
@ -199,8 +168,6 @@ lemma koszulSign_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
|
|||
rw [← insertionSortEquiv_get]
|
||||
simp only [Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem, ni]
|
||||
simp_all only [List.length_cons, add_le_add_iff_right, List.getElem_insertIdx_self]
|
||||
have hms : (List.orderedInsert le1 r0 rs).get ⟨nro, by simp⟩ = r0 := by
|
||||
simp [nro]
|
||||
have hc1 : ni.castSucc < nro → ¬ le1 r0 i := by
|
||||
intro hninro
|
||||
rw [← hns]
|
||||
|
@ -231,5 +198,4 @@ lemma koszulSign_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
|
|||
· exact Nat.le_of_lt_succ h
|
||||
· exact Nat.le_of_lt_succ h
|
||||
|
||||
|
||||
end Wick
|
||||
|
|
|
@ -98,14 +98,13 @@ lemma koszulSignInsert_eq_filter {I : Type} (q : I → Fin 2) (le1 : I → I →
|
|||
simp
|
||||
|
||||
lemma koszulSignInsert_eq_cons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
[IsTotal I le1] [IsTrans I le1] (r0 : I) (r : List I) :
|
||||
[IsTotal I le1] (r0 : I) (r : List I) :
|
||||
koszulSignInsert le1 q r0 r = koszulSignInsert le1 q r0 (r0 :: r) := by
|
||||
simp only [koszulSignInsert, Fin.isValue, and_self]
|
||||
have h1 : le1 r0 r0 := by
|
||||
simpa using IsTotal.total (r := le1) r0 r0
|
||||
simp [h1]
|
||||
|
||||
|
||||
lemma koszulSignInsert_eq_grade {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(r0 : I) (r : List I) : koszulSignInsert le1 q r0 r = if grade q [r0] = 1 ∧
|
||||
grade q (List.filter (fun i => decide (¬ le1 r0 i)) r) = 1 then -1 else 1 := by
|
||||
|
@ -206,9 +205,39 @@ lemma koszulSignInsert_eq_insertSign {I : Type} (q : I → Fin 2) (le1 : I → I
|
|||
· simp [orderedInsertPos]
|
||||
|
||||
lemma koszulSignInsert_insertIdx {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(i j : I) [IsTotal I le1] [IsTrans I le1] (r : List I) (n : ℕ) (hn : n ≤ r.length) :
|
||||
(i j : I) (r : List I) (n : ℕ) (hn : n ≤ r.length) :
|
||||
koszulSignInsert le1 q j (List.insertIdx n i r) = koszulSignInsert le1 q j (i :: r) := by
|
||||
apply koszulSignInsert_eq_perm
|
||||
exact List.perm_insertIdx i r hn
|
||||
|
||||
/-- The difference in `koszulSignInsert` on inserting `r0` into `r` compared to
|
||||
into `r1 :: r` for any `r`. -/
|
||||
def koszulSignCons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1] (r0 r1 : I) :
|
||||
ℂ :=
|
||||
if le1 r0 r1 then 1 else
|
||||
if q r0 = 1 ∧ q r1 = 1 then -1 else 1
|
||||
|
||||
lemma koszulSignCons_eq_superComuteCoef {I : Type} (q : I → Fin 2) (le1 : I → I → Prop)
|
||||
[DecidableRel le1] (r0 r1 : I) : koszulSignCons q le1 r0 r1 =
|
||||
if le1 r0 r1 then 1 else superCommuteCoef q [r0] [r1] := by
|
||||
simp only [koszulSignCons, Fin.isValue, superCommuteCoef, grade, ite_eq_right_iff, zero_ne_one,
|
||||
imp_false]
|
||||
congr 1
|
||||
by_cases h0 : q r0 = 1
|
||||
· by_cases h1 : q r1 = 1
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = 0 := by omega
|
||||
simp [h0, h1]
|
||||
· have h0 : q r0 = 0 := by omega
|
||||
by_cases h1 : q r1 = 1
|
||||
· simp [h0, h1]
|
||||
· have h1 : q r1 = 0 := by omega
|
||||
simp [h0, h1]
|
||||
|
||||
lemma koszulSignInsert_cons {I : Type} (q : I → Fin 2) (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(r0 r1 : I) (r : List I) :
|
||||
koszulSignInsert le1 q r0 (r1 :: r) = (koszulSignCons q le1 r0 r1) *
|
||||
koszulSignInsert le1 q r0 r := by
|
||||
simp [koszulSignInsert, koszulSignCons]
|
||||
|
||||
end Wick
|
||||
|
|
|
@ -17,23 +17,25 @@ namespace Wick
|
|||
|
||||
open HepLean.List
|
||||
|
||||
def superCommuteCoefLE {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
/-- The sign that appears in the static version of Wicks theorem.
|
||||
This is actually equal to `superCommuteCoef q [r.get n] (r.take n)`, something
|
||||
which will be proved in a lemma. -/
|
||||
def staticWickCoef {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
[DecidableRel le1] (i : I) (n : Fin r.length) : ℂ :=
|
||||
koszulSign le1 q r *
|
||||
superCommuteCoef q [i] (List.take (↑((HepLean.List.insertionSortEquiv le1 r) n))
|
||||
(List.insertionSort le1 r)) *
|
||||
koszulSign le1 q (r.eraseIdx ↑n)
|
||||
|
||||
lemma superCommuteCoefLE_eq_q {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
lemma staticWickCoef_eq_q {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
[DecidableRel le1] (i : I) (n : Fin r.length)
|
||||
(hq : q i = q (r.get n)) :
|
||||
superCommuteCoefLE q le1 r i n =
|
||||
staticWickCoef q le1 r i n =
|
||||
koszulSign le1 q r *
|
||||
superCommuteCoef q [r.get n] (List.take (↑(insertionSortEquiv le1 r n))
|
||||
(List.insertionSort le1 r)) *
|
||||
koszulSign le1 q (r.eraseIdx ↑n) := by
|
||||
simp [superCommuteCoefLE, superCommuteCoef, grade, hq]
|
||||
|
||||
simp [staticWickCoef, superCommuteCoef, grade, hq]
|
||||
|
||||
lemma insertIdx_eraseIdx {I : Type} :
|
||||
(n : ℕ) → (r : List I) → (hn : n < r.length) →
|
||||
|
@ -47,11 +49,11 @@ lemma insertIdx_eraseIdx {I : Type} :
|
|||
List.eraseIdx_cons_succ, List.insertIdx_succ_cons, List.cons.injEq, true_and]
|
||||
exact insertIdx_eraseIdx n r _
|
||||
|
||||
lemma superCommuteCoefLE_eq_get {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
lemma staticWickCoef_eq_get {I : Type} (q : I → Fin 2) (le1 :I → I → Prop) (r : List I)
|
||||
[DecidableRel le1] [IsTotal I le1] [IsTrans I le1] (i : I) (n : Fin r.length)
|
||||
(heq : q i = q (r.get n)) :
|
||||
superCommuteCoefLE q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
|
||||
rw [superCommuteCoefLE_eq_q]
|
||||
staticWickCoef q le1 r i n = superCommuteCoef q [r.get n] (r.take n) := by
|
||||
rw [staticWickCoef_eq_q]
|
||||
let r' := r.eraseIdx ↑n
|
||||
have hr : List.insertIdx n (r.get n) (r.eraseIdx n) = r := by
|
||||
exact insertIdx_eraseIdx n.1 r n.prop
|
||||
|
|
|
@ -14,6 +14,10 @@ import HepLean.PerturbationTheory.Wick.Signs.Grade
|
|||
namespace Wick
|
||||
open HepLean.List
|
||||
|
||||
/-- Given two lists `la` and `lb` returns `-1` if they are both of grade `1` and
|
||||
`1` otherwise. This corresponds to the sign associated with the super commutator
|
||||
when commuting `la` and `lb` in the free algebra.
|
||||
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
||||
def superCommuteCoef {I : Type} (q : I → Fin 2) (la lb : List I) : ℂ :=
|
||||
if grade q la = 1 ∧ grade q lb = 1 then - 1 else 1
|
||||
|
||||
|
@ -23,11 +27,17 @@ lemma superCommuteCoef_comm {I : Type} (q : I → Fin 2) (la lb : List I) :
|
|||
congr 1
|
||||
exact Eq.propIntro (fun a => id (And.symm a)) fun a => id (And.symm a)
|
||||
|
||||
def superCommuteLiftCoef {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
|
||||
/-- Given a list `l : List (Σ i, f i)` and a list `r : List I` returns `-1` if the
|
||||
grade of `l` is `1` and the grade of `r` is `1` and `1` otherwise. This corresponds
|
||||
to the sign associated with the super commutator when commuting
|
||||
the lift of `l` and `r` (by summing over fibers) in the
|
||||
free algebra over `Σ i, f i`.
|
||||
In terms of physics it is `-1` if commuting two fermionic operators and `1` otherwise. -/
|
||||
def superCommuteLiftCoef {I : Type} {f : I → Type}
|
||||
(q : I → Fin 2) (l : List (Σ i, f i)) (r : List I) : ℂ :=
|
||||
(if grade (fun i => q i.fst) l = 1 ∧ grade q r = 1 then -1 else 1)
|
||||
|
||||
lemma superCommuteLiftCoef_empty {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
|
||||
lemma superCommuteLiftCoef_empty {I : Type} {f : I → Type}
|
||||
(q : I → Fin 2) (l : List (Σ i, f i)) :
|
||||
superCommuteLiftCoef q l [] = 1 := by
|
||||
simp [superCommuteLiftCoef]
|
||||
|
@ -79,5 +89,4 @@ lemma superCommuteCoef_cons {I : Type} (q : I → Fin 2) (i : I) (la lb : List I
|
|||
simp only [List.singleton_append]
|
||||
rw [superCommuteCoef_append]
|
||||
|
||||
|
||||
end Wick
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue