2024-04-29 08:13:52 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
2024-06-25 07:06:32 -04:00
|
|
|
|
import Mathlib.Analysis.Complex.Basic
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-!
|
|
|
|
|
# Invariants of the CKM Matrix
|
2024-04-29 08:13:52 -04:00
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
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.
|
|
|
|
|
|
|
|
|
|
-/
|
2024-04-29 08:13:52 -04:00
|
|
|
|
open Matrix Complex
|
|
|
|
|
open ComplexConjugate
|
|
|
|
|
open CKMMatrix
|
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
namespace Invariant
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- The complex jarlskog invariant for a CKM matrix. -/
|
2024-04-29 08:13:52 -04:00
|
|
|
|
@[simps!]
|
|
|
|
|
def jarlskogℂCKM (V : CKMMatrix) : ℂ :=
|
|
|
|
|
[V]us * [V]cb * conj [V]ub * conj [V]cs
|
|
|
|
|
|
|
|
|
|
lemma jarlskogℂCKM_equiv (V U : CKMMatrix) (h : V ≈ U) :
|
|
|
|
|
jarlskogℂCKM V = jarlskogℂCKM 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 [jarlskogℂCKM, 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
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- The complex jarlskog invariant for an equivalence class of CKM matrices. -/
|
2024-04-29 08:13:52 -04:00
|
|
|
|
@[simp]
|
|
|
|
|
def jarlskogℂ : Quotient CKMMatrixSetoid → ℂ :=
|
|
|
|
|
Quotient.lift jarlskogℂCKM jarlskogℂCKM_equiv
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- 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)` .
|
|
|
|
|
-/
|
2024-04-29 08:13:52 -04:00
|
|
|
|
def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : ℝ :=
|
|
|
|
|
VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2)
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- An invariant for CKM matrices. The argument of this invariant is `δ₁₃` in the
|
|
|
|
|
standard parameterization. -/
|
2024-04-29 09:22:32 -04:00
|
|
|
|
def mulExpδ₁₃ (V : Quotient CKMMatrixSetoid) : ℂ :=
|
2024-04-29 08:13:52 -04:00
|
|
|
|
jarlskogℂ V + VusVubVcdSq V
|
|
|
|
|
|
|
|
|
|
end Invariant
|
|
|
|
|
end
|