PhysLean/HepLean/FlavorPhysics/CKMMatrix/Invariants.lean
2024-07-03 07:56:30 -04:00

63 lines
2 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
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 Mathlib.Analysis.Complex.Basic
/-!
# 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
noncomputable section
namespace Invariant
/-- The complex jarlskog invariant for a CKM matrix. -/
@[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
/-- The complex jarlskog invariant for an equivalence class of CKM matrices. -/
@[simp]
def jarlskog : Quotient CKMMatrixSetoid → :=
Quotient.lift jarlskogCKM jarlskogCKM_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
end Invariant
end