refactor: Major refactor of CKMMatrix

This commit is contained in:
jstoobysmith 2024-04-29 08:13:52 -04:00
parent fe63fc9994
commit ff89c3f79d
9 changed files with 1096 additions and 1512 deletions

View file

@ -0,0 +1,51 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic
import HepLean.FlavorPhysics.CKMMatrix.Rows
import HepLean.FlavorPhysics.CKMMatrix.PhaseFreedom
import Mathlib.Analysis.SpecialFunctions.Complex.Arg
open Matrix Complex
open ComplexConjugate
open CKMMatrix
noncomputable section
namespace Invariant
@[simps!]
def jarlskogCKM (V : CKMMatrix) : :=
[V]us * [V]cb * conj [V]ub * conj [V]cs
lemma jarlskogCKM_equiv (V U : CKMMatrix) (h : V ≈ U) :
jarlskogCKM V = jarlskogCKM U := by
obtain ⟨a, b, c, e, f, g, h⟩ := h
change V = phaseShiftApply U a b c e f g at h
rw [h]
simp only [jarlskogCKM, Fin.isValue, phaseShiftApply.ub,
phaseShiftApply.us, phaseShiftApply.cb, phaseShiftApply.cs]
simp [← exp_conj, conj_ofReal, exp_add, exp_neg]
have ha : cexp (↑a * I) ≠ 0 := exp_ne_zero _
have hb : cexp (↑b * I) ≠ 0 := exp_ne_zero _
have hf : cexp (↑f * I) ≠ 0 := exp_ne_zero _
have hg : cexp (↑g * I) ≠ 0 := exp_ne_zero _
field_simp
ring
@[simp]
def jarlskog : Quotient CKMMatrixSetoid → :=
Quotient.lift jarlskogCKM jarlskogCKM_equiv
def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : :=
VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2)
def mulExpδ₃ (V : Quotient CKMMatrixSetoid) : :=
jarlskog V + VusVubVcdSq V
end Invariant
end