refactor: Insert Sign
This commit is contained in:
parent
9c0c499292
commit
8a3d72bd68
3 changed files with 62 additions and 42 deletions
|
@ -7,17 +7,21 @@ import HepLean.Mathematics.List
|
|||
import HepLean.PerturbationTheory.Wick.Signs.SuperCommuteCoef
|
||||
/-!
|
||||
|
||||
# Koszul signs and ordering for lists and algebras
|
||||
# Insert sign
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
/-- 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)
|
||||
section
|
||||
/-!
|
||||
|
||||
## Basic properties of lists
|
||||
|
||||
To be replaced with Mathlib or Lean definitions when/where appropraite.
|
||||
-/
|
||||
|
||||
lemma take_insert_same {I : Type} (i : I) :
|
||||
(n : ℕ) → (r : List I) →
|
||||
|
@ -28,12 +32,6 @@ lemma take_insert_same {I : Type} (i : I) :
|
|||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_insert_same i n as
|
||||
|
||||
lemma insertSign_insert {I : Type} (q : I → Fin 2) (n : ℕ)
|
||||
(r0 : I) (r : List I) : insertSign q n r0 r = insertSign q n r0 (List.insertIdx n r0 r) := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_insert_same]
|
||||
|
||||
lemma take_eraseIdx_same {I : Type} :
|
||||
(n : ℕ) → (r : List I) →
|
||||
List.take n (List.eraseIdx r n) = List.take n r
|
||||
|
@ -43,22 +41,6 @@ lemma take_eraseIdx_same {I : Type} :
|
|||
simp only [List.eraseIdx_cons_succ, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
exact take_eraseIdx_same n as
|
||||
|
||||
lemma insertSign_eraseIdx {I : Type} (q : I → Fin 2) (n : ℕ)
|
||||
(r0 : I) (r : List I) : insertSign q n r0 (r.eraseIdx n) = insertSign q n r0 r := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_eraseIdx_same]
|
||||
|
||||
lemma insertSign_zero {I : Type} (q : I → Fin 2) (r0 : I) (r : List I) :
|
||||
insertSign q 0 r0 r = 1 := by
|
||||
simp [insertSign, superCommuteCoef]
|
||||
|
||||
lemma insertSign_succ_cons {I : Type} (q : I → Fin 2) (n : ℕ)
|
||||
(r0 r1 : I) (r : List I) : insertSign q (n + 1) r0 (r1 :: r) =
|
||||
superCommuteCoef q [r0] [r1] * insertSign q n r0 r := by
|
||||
simp only [insertSign, List.take_succ_cons]
|
||||
rw [superCommuteCoef_cons]
|
||||
|
||||
lemma take_insert_gt {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : n < m) → (r : List I) →
|
||||
List.take n (List.insertIdx m i r) = List.take n r
|
||||
|
@ -69,13 +51,6 @@ lemma take_insert_gt {I : Type} (i : I) :
|
|||
simp only [List.insertIdx_succ_cons, List.take_succ_cons, List.cons.injEq, true_and]
|
||||
refine take_insert_gt i n m (Nat.succ_lt_succ_iff.mp h) as
|
||||
|
||||
lemma insertSign_insert_gt {I : Type} (q : I → Fin 2) (n m : ℕ)
|
||||
(r0 r1 : I) (r : List I) (hn : n < m) :
|
||||
insertSign q n r0 (List.insertIdx m r1 r) = insertSign q n r0 r := by
|
||||
rw [insertSign, insertSign]
|
||||
congr 1
|
||||
exact take_insert_gt r1 n m hn r
|
||||
|
||||
lemma take_insert_let {I : Type} (i : I) :
|
||||
(n m : ℕ) → (h : m ≤ n) → (r : List I) → (hm : m ≤ r.length) →
|
||||
(List.take (n + 1) (List.insertIdx m i r)).Perm (i :: List.take n r)
|
||||
|
@ -91,20 +66,65 @@ lemma take_insert_let {I : Type} (i : I) :
|
|||
refine List.Perm.cons a ?_
|
||||
exact take_insert_let i n m (Nat.le_of_succ_le_succ h) as (Nat.le_of_succ_le_succ hm)
|
||||
|
||||
lemma insertSign_insert_lt_eq_insertSort {I : Type} (q : I → Fin 2) (n m : ℕ)
|
||||
(r0 r1 : I) (r : List I) (hn : m ≤ n) (hm : m ≤ r.length) :
|
||||
end
|
||||
|
||||
/-!
|
||||
|
||||
## Insert sign
|
||||
|
||||
-/
|
||||
|
||||
section InsertSign
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
||||
/-- 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 (n : ℕ) (r0 : 𝓕) (r : List 𝓕) : ℂ :=
|
||||
superCommuteCoef q [r0] (List.take n r)
|
||||
|
||||
lemma insertSign_insert (n : ℕ) (r0 : 𝓕) (r : List 𝓕) :
|
||||
insertSign q n r0 r = insertSign q n r0 (List.insertIdx n r0 r) := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_insert_same]
|
||||
|
||||
lemma insertSign_eraseIdx (n : ℕ) (r0 : 𝓕) (r : List 𝓕) :
|
||||
insertSign q n r0 (r.eraseIdx n) = insertSign q n r0 r := by
|
||||
simp only [insertSign]
|
||||
congr 1
|
||||
rw [take_eraseIdx_same]
|
||||
|
||||
lemma insertSign_zero (r0 : 𝓕) (r : List 𝓕) : insertSign q 0 r0 r = 1 := by
|
||||
simp [insertSign, superCommuteCoef]
|
||||
|
||||
lemma insertSign_succ_cons (n : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) : insertSign q (n + 1) r0 (r1 :: r) =
|
||||
superCommuteCoef q [r0] [r1] * insertSign q n r0 r := by
|
||||
simp only [insertSign, List.take_succ_cons]
|
||||
rw [superCommuteCoef_cons]
|
||||
|
||||
|
||||
lemma insertSign_insert_gt (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : n < m) :
|
||||
insertSign q n r0 (List.insertIdx m r1 r) = insertSign q n r0 r := by
|
||||
rw [insertSign, insertSign]
|
||||
congr 1
|
||||
exact take_insert_gt r1 n m hn r
|
||||
|
||||
lemma insertSign_insert_lt_eq_insertSort (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : m ≤ n)
|
||||
(hm : m ≤ r.length) :
|
||||
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = insertSign q (n + 1) r0 (r1 :: r) := by
|
||||
rw [insertSign, insertSign]
|
||||
apply superCommuteCoef_perm_snd
|
||||
simp only [List.take_succ_cons]
|
||||
refine take_insert_let r1 n m hn r hm
|
||||
|
||||
lemma insertSign_insert_lt {I : Type} (q : I → Fin 2) (n m : ℕ)
|
||||
(r0 r1 : I) (r : List I) (hn : m ≤ n) (hm : m ≤ r.length) :
|
||||
lemma insertSign_insert_lt (n m : ℕ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : m ≤ n) (hm : m ≤ r.length) :
|
||||
insertSign q (n + 1) r0 (List.insertIdx m r1 r) = superCommuteCoef q [r0] [r1] *
|
||||
insertSign q n r0 r := by
|
||||
rw [insertSign_insert_lt_eq_insertSort, insertSign_succ_cons]
|
||||
exact hn
|
||||
exact hm
|
||||
|
||||
end InsertSign
|
||||
|
||||
end Wick
|
||||
|
|
|
@ -9,7 +9,7 @@ import Mathlib.Analysis.Complex.Basic
|
|||
import HepLean.PerturbationTheory.Wick.Signs.KoszulSign
|
||||
/-!
|
||||
|
||||
# Koszul sign insert
|
||||
# Static wick coefficent
|
||||
|
||||
-/
|
||||
|
||||
|
|
|
@ -3,16 +3,16 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Mathematics.List
|
||||
import HepLean.PerturbationTheory.FieldStatistics
|
||||
/-!
|
||||
|
||||
# Koszul signs and ordering for lists and algebras
|
||||
# Super commutation coefficent.
|
||||
|
||||
This is a complex number which is `-1` when commuting two fermionic operators and `1` otherwise.
|
||||
|
||||
-/
|
||||
|
||||
namespace Wick
|
||||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic)
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue