2024-04-29 08:13:52 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
2024-07-12 16:39:44 -04:00
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2024-04-29 08:13:52 -04:00
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
|
|
|
|
import HepLean.FlavorPhysics.CKMMatrix.Rows
|
|
|
|
|
import HepLean.FlavorPhysics.CKMMatrix.Invariants
|
|
|
|
|
import Mathlib.Analysis.SpecialFunctions.Complex.Arg
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-!
|
|
|
|
|
# Standard parameterization for the CKM Matrix
|
2024-04-29 08:13:52 -04:00
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
This file defines the standard parameterization of CKM matrices in terms of
|
|
|
|
|
four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`.
|
|
|
|
|
|
|
|
|
|
We will show that every CKM matrix can be written within this standard parameterization
|
|
|
|
|
in the file `FlavorPhysics.CKMMatrix.StandardParameters`.
|
|
|
|
|
|
|
|
|
|
-/
|
2024-04-29 08:13:52 -04:00
|
|
|
|
open Matrix Complex
|
|
|
|
|
open ComplexConjugate
|
|
|
|
|
open CKMMatrix
|
|
|
|
|
|
|
|
|
|
noncomputable section
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- Given four reals `θ₁₂ θ₁₃ θ₂₃ δ₁₃` the standard paramaterization of the CKM matrix
|
|
|
|
|
as a `3×3` complex matrix. -/
|
2024-07-12 11:23:02 -04:00
|
|
|
|
def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin 3) (Fin 3) ℂ :=
|
2024-04-29 08:13:52 -04:00
|
|
|
|
![![Real.cos θ₁₂ * Real.cos θ₁₃, Real.sin θ₁₂ * Real.cos θ₁₃, Real.sin θ₁₃ * exp (-I * δ₁₃)],
|
|
|
|
|
![(-Real.sin θ₁₂ * Real.cos θ₂₃) - (Real.cos θ₁₂ * Real.sin θ₁₃ * Real.sin θ₂₃ * exp (I * δ₁₃)),
|
|
|
|
|
Real.cos θ₁₂ * Real.cos θ₂₃ - Real.sin θ₁₂ * Real.sin θ₁₃ * Real.sin θ₂₃ * exp (I * δ₁₃),
|
|
|
|
|
Real.sin θ₂₃ * Real.cos θ₁₃],
|
|
|
|
|
![Real.sin θ₁₂ * Real.sin θ₂₃ - Real.cos θ₁₂ * Real.sin θ₁₃ * Real.cos θ₂₃ * exp (I * δ₁₃),
|
|
|
|
|
(-Real.cos θ₁₂ * Real.sin θ₂₃) - (Real.sin θ₁₂ * Real.sin θ₁₃ * Real.cos θ₂₃ * exp (I * δ₁₃)),
|
|
|
|
|
Real.cos θ₂₃ * Real.cos θ₁₃]]
|
|
|
|
|
|
|
|
|
|
open CKMMatrix
|
|
|
|
|
|
2024-07-12 11:23:02 -04:00
|
|
|
|
lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|
|
|
|
((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
|
2024-04-29 08:13:52 -04:00
|
|
|
|
funext j i
|
2024-04-29 09:22:32 -04:00
|
|
|
|
simp only [standParamAsMatrix, neg_mul, Fin.isValue]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
rw [mul_apply]
|
|
|
|
|
have h1 := exp_ne_zero (I * ↑δ₁₃)
|
|
|
|
|
fin_cases j <;> rw [Fin.sum_univ_three]
|
|
|
|
|
simp only [Fin.zero_eta, Fin.isValue, conjTranspose_apply, cons_val', cons_val_zero, empty_val',
|
|
|
|
|
cons_val_fin_one, star_mul', RCLike.star_def, conj_ofReal, cons_val_one, head_cons, star_sub,
|
|
|
|
|
star_neg, ← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const]
|
|
|
|
|
simp [conj_ofReal]
|
|
|
|
|
rw [exp_neg ]
|
|
|
|
|
fin_cases i <;> simp
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq, sin_sq, sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq, sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
simp only [Fin.mk_one, Fin.isValue, conjTranspose_apply, cons_val', cons_val_one, head_cons,
|
|
|
|
|
empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def, conj_ofReal, star_sub,
|
|
|
|
|
← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const, star_neg]
|
|
|
|
|
simp [conj_ofReal]
|
|
|
|
|
rw [exp_neg]
|
|
|
|
|
fin_cases i <;> simp
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq, sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq, sin_sq, sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
simp only [Fin.reduceFinMk, Fin.isValue, conjTranspose_apply, cons_val', cons_val_two, tail_cons,
|
|
|
|
|
head_cons, empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def, conj_ofReal,
|
|
|
|
|
← exp_conj, map_neg, _root_.map_mul, conj_I, neg_mul, neg_neg, cons_val_one, head_fin_const]
|
|
|
|
|
simp [conj_ofReal]
|
|
|
|
|
rw [exp_neg]
|
|
|
|
|
fin_cases i <;> simp
|
|
|
|
|
· ring_nf
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
· ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
rw [sin_sq, sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
|
2024-04-29 10:42:44 -04:00
|
|
|
|
/-- Given four reals `θ₁₂ θ₁₃ θ₂₃ δ₁₃` the standard paramaterization of the CKM matrix
|
|
|
|
|
as a CKM matrix. -/
|
2024-04-29 09:22:32 -04:00
|
|
|
|
def standParam (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : CKMMatrix :=
|
|
|
|
|
⟨standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃, by
|
2024-04-29 08:13:52 -04:00
|
|
|
|
rw [mem_unitaryGroup_iff']
|
2024-04-29 09:22:32 -04:00
|
|
|
|
exact standParamAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
|
2024-04-29 08:13:52 -04:00
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
namespace standParam
|
|
|
|
|
lemma cross_product_t (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|
|
|
|
[standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]t =
|
|
|
|
|
(conj [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u ×₃ conj [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) := by
|
2024-04-29 08:13:52 -04:00
|
|
|
|
have h1 := exp_ne_zero (I * ↑δ₁₃)
|
|
|
|
|
funext i
|
|
|
|
|
fin_cases i
|
2024-04-29 09:22:32 -04:00
|
|
|
|
· simp only [tRow, standParam, standParamAsMatrix, neg_mul, exp_neg,
|
2024-04-29 08:13:52 -04:00
|
|
|
|
Fin.isValue, cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_two, tail_cons,
|
|
|
|
|
head_fin_const, cons_val_one, head_cons, Fin.zero_eta, crossProduct, uRow, cRow,
|
|
|
|
|
LinearMap.mk₂_apply, Pi.conj_apply, _root_.map_mul, map_inv₀, ← exp_conj, conj_I, conj_ofReal,
|
|
|
|
|
inv_inv, map_sub, map_neg]
|
|
|
|
|
field_simp
|
|
|
|
|
ring_nf
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
2024-04-29 09:22:32 -04:00
|
|
|
|
· simp only [tRow, standParam, standParamAsMatrix, neg_mul, exp_neg, Fin.isValue, cons_val',
|
2024-04-29 08:13:52 -04:00
|
|
|
|
cons_val_zero, empty_val', cons_val_fin_one, cons_val_two, tail_cons, head_fin_const,
|
|
|
|
|
cons_val_one, head_cons, Fin.mk_one, crossProduct, uRow, cRow, LinearMap.mk₂_apply,
|
|
|
|
|
Pi.conj_apply, _root_.map_mul, conj_ofReal, map_inv₀, ← exp_conj, conj_I, inv_inv, map_sub,
|
|
|
|
|
map_neg]
|
|
|
|
|
field_simp
|
|
|
|
|
ring_nf
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
2024-04-29 09:22:32 -04:00
|
|
|
|
· simp only [tRow, standParam, standParamAsMatrix, neg_mul, exp_neg, Fin.isValue,
|
2024-04-29 08:13:52 -04:00
|
|
|
|
cons_val', cons_val_zero, empty_val', cons_val_fin_one, cons_val_two, tail_cons, head_fin_const,
|
|
|
|
|
cons_val_one, head_cons, Fin.reduceFinMk, crossProduct, uRow, cRow, LinearMap.mk₂_apply,
|
|
|
|
|
Pi.conj_apply, _root_.map_mul, conj_ofReal, map_inv₀, ← exp_conj, conj_I, inv_inv, map_sub,
|
|
|
|
|
map_neg]
|
|
|
|
|
field_simp
|
|
|
|
|
ring_nf
|
|
|
|
|
rw [sin_sq]
|
|
|
|
|
ring
|
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ} (hu : [U]u = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u)
|
2024-07-12 11:23:02 -04:00
|
|
|
|
(hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) :
|
2024-04-29 09:22:32 -04:00
|
|
|
|
U = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by
|
2024-04-29 08:13:52 -04:00
|
|
|
|
apply ext_Rows hu hc
|
2024-04-29 09:22:32 -04:00
|
|
|
|
rw [hU, cross_product_t, hu, hc]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ℝ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
|
|
|
|
|
standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
|
|
|
|
|
simp [standParam, standParamAsMatrix]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
apply CKMMatrix_ext
|
2024-04-29 10:42:44 -04:00
|
|
|
|
simp only
|
2024-07-12 11:23:02 -04:00
|
|
|
|
rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
rw [show cexp (-(I * ↑δ₁₃)) = cexp (-(I * ↑δ₁₃')) by rw [exp_neg, exp_neg, mul_comm, h, mul_comm]]
|
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
open Invariant in
|
|
|
|
|
lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂)
|
2024-04-29 10:42:44 -04:00
|
|
|
|
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
|
2024-04-29 09:22:32 -04:00
|
|
|
|
VusVubVcdSq ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
|
2024-04-29 08:13:52 -04:00
|
|
|
|
Real.sin θ₁₂ ^ 2 * Real.cos θ₁₃ ^ 2 * Real.sin θ₁₃ ^ 2 * Real.sin θ₂₃ ^ 2 := by
|
2024-04-29 09:22:32 -04:00
|
|
|
|
simp only [VusVubVcdSq, VusAbs, VAbs, VAbs', Fin.isValue, standParam, standParamAsMatrix,
|
2024-04-29 08:13:52 -04:00
|
|
|
|
neg_mul, Quotient.lift_mk, cons_val', cons_val_one, head_cons,
|
|
|
|
|
empty_val', cons_val_fin_one, cons_val_zero, _root_.map_mul, VubAbs, cons_val_two, tail_cons,
|
|
|
|
|
VcbAbs, VudAbs, Complex.abs_ofReal]
|
|
|
|
|
by_cases hx : Real.cos θ₁₃ ≠ 0
|
|
|
|
|
· rw [Complex.abs_exp]
|
2024-04-29 10:42:44 -04:00
|
|
|
|
simp only [neg_re, mul_re, I_re, ofReal_re, zero_mul, I_im, ofReal_im, mul_zero, sub_self,
|
|
|
|
|
neg_zero, Real.exp_zero, mul_one, _root_.sq_abs]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
rw [_root_.abs_of_nonneg h1, _root_.abs_of_nonneg h3, _root_.abs_of_nonneg h2,
|
|
|
|
|
_root_.abs_of_nonneg h4]
|
|
|
|
|
simp [sq]
|
|
|
|
|
ring_nf
|
|
|
|
|
nth_rewrite 2 [Real.sin_sq θ₁₂]
|
|
|
|
|
ring_nf
|
|
|
|
|
field_simp
|
|
|
|
|
ring
|
|
|
|
|
· simp at hx
|
|
|
|
|
rw [hx]
|
|
|
|
|
simp
|
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
open Invariant in
|
|
|
|
|
lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂)
|
2024-07-12 11:23:02 -04:00
|
|
|
|
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
|
2024-04-29 09:22:32 -04:00
|
|
|
|
mulExpδ₁₃ ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
|
2024-07-12 11:23:02 -04:00
|
|
|
|
sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by
|
2024-04-29 09:22:32 -04:00
|
|
|
|
rw [mulExpδ₁₃, VusVubVcdSq_eq _ _ _ _ h1 h2 h3 h4 ]
|
|
|
|
|
simp only [jarlskogℂ, standParam, standParamAsMatrix, neg_mul,
|
2024-04-29 08:13:52 -04:00
|
|
|
|
Quotient.lift_mk, jarlskogℂCKM, Fin.isValue, cons_val', cons_val_one, head_cons,
|
|
|
|
|
empty_val', cons_val_fin_one, cons_val_zero, cons_val_two, tail_cons, _root_.map_mul, ←
|
|
|
|
|
exp_conj, map_neg, conj_I, conj_ofReal, neg_neg, map_sub]
|
2024-04-29 10:42:44 -04:00
|
|
|
|
simp only [ofReal_sin, ofReal_cos, ofReal_mul, ofReal_pow]
|
2024-04-29 08:13:52 -04:00
|
|
|
|
ring_nf
|
|
|
|
|
rw [exp_neg]
|
|
|
|
|
have h1 : cexp (I * δ₁₃) ≠ 0 := exp_ne_zero _
|
|
|
|
|
field_simp
|
|
|
|
|
|
2024-04-29 09:22:32 -04:00
|
|
|
|
end standParam
|
2024-04-29 08:13:52 -04:00
|
|
|
|
end
|