refactor: lint
This commit is contained in:
parent
490ed0380c
commit
d448c78045
8 changed files with 343 additions and 163 deletions
|
@ -7,7 +7,18 @@ import HepLean.FlavorPhysics.CKMMatrix.Basic
|
|||
import HepLean.FlavorPhysics.CKMMatrix.Rows
|
||||
import HepLean.FlavorPhysics.CKMMatrix.PhaseFreedom
|
||||
import Mathlib.Analysis.SpecialFunctions.Complex.Arg
|
||||
/-!
|
||||
# Invariants of the CKM Matrix
|
||||
|
||||
The CKM matrix is only defined up to an equivalence.
|
||||
|
||||
This file defines some invariants of the CKM matrix, which are well-defined with respect to
|
||||
this equivalence.
|
||||
|
||||
Of note, this file defines the complex jarlskog invariant.
|
||||
|
||||
|
||||
-/
|
||||
open Matrix Complex
|
||||
open ComplexConjugate
|
||||
open CKMMatrix
|
||||
|
@ -17,6 +28,7 @@ noncomputable section
|
|||
namespace Invariant
|
||||
|
||||
|
||||
/-- The complex jarlskog invariant for a CKM matrix. -/
|
||||
@[simps!]
|
||||
def jarlskogℂCKM (V : CKMMatrix) : ℂ :=
|
||||
[V]us * [V]cb * conj [V]ub * conj [V]cs
|
||||
|
@ -36,13 +48,19 @@ lemma jarlskogℂCKM_equiv (V U : CKMMatrix) (h : V ≈ U) :
|
|||
field_simp
|
||||
ring
|
||||
|
||||
/-- The complex jarlskog invariant for an equivalence class of CKM matrices. -/
|
||||
@[simp]
|
||||
def jarlskogℂ : Quotient CKMMatrixSetoid → ℂ :=
|
||||
Quotient.lift jarlskogℂCKM jarlskogℂCKM_equiv
|
||||
|
||||
/-- An invariant for CKM mtrices corresponding to the square of the absolute values
|
||||
of the `us`, `ub` and `cb` elements multipled together divied by `(VudAbs V ^ 2 + VusAbs V ^2)` .
|
||||
-/
|
||||
def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : ℝ :=
|
||||
VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2)
|
||||
|
||||
/-- An invariant for CKM matrices. The argument of this invariant is `δ₁₃` in the
|
||||
standard parameterization. -/
|
||||
def mulExpδ₁₃ (V : Quotient CKMMatrixSetoid) : ℂ :=
|
||||
jarlskogℂ V + VusVubVcdSq V
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue