reactor: Removal of double spaces
This commit is contained in:
parent
ce92e1d649
commit
13f62a50eb
64 changed files with 550 additions and 546 deletions
|
@ -27,7 +27,7 @@ noncomputable section
|
|||
leading diagonal. -/
|
||||
@[simp]
|
||||
def phaseShiftMatrix (a b c : ℝ) : Matrix (Fin 3) (Fin 3) ℂ :=
|
||||
![![cexp (I * a), 0, 0], ![0, cexp (I * b), 0], ![0, 0, cexp (I * c)]]
|
||||
![![cexp (I * a), 0, 0], ![0, cexp (I * b), 0], ![0, 0, cexp (I * c)]]
|
||||
|
||||
lemma phaseShiftMatrix_one : phaseShiftMatrix 0 0 0 = 1 := by
|
||||
ext i j
|
||||
|
@ -49,7 +49,7 @@ lemma phaseShiftMatrix_mul (a b c d e f : ℝ) :
|
|||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three]
|
||||
any_goals rw [mul_add, exp_add]
|
||||
change cexp (I * ↑c) * 0 = 0
|
||||
change cexp (I * ↑c) * 0 = 0
|
||||
simp
|
||||
|
||||
/-- Given three real numbers `a b c` the unitary matrix with `exp (I * a)` etc on the
|
||||
|
@ -161,7 +161,7 @@ lemma equiv (V : CKMMatrix) (a b c d e f : ℝ) :
|
|||
rfl
|
||||
|
||||
lemma ud (V : CKMMatrix) (a b c d e f : ℝ) :
|
||||
(phaseShiftApply V a b c d e f).1 0 0 = cexp (a * I + d * I) * V.1 0 0 := by
|
||||
(phaseShiftApply V a b c d e f).1 0 0 = cexp (a * I + d * I) * V.1 0 0 := by
|
||||
simp only [Fin.isValue, phaseShiftApply_coe]
|
||||
rw [mul_apply, Fin.sum_univ_three]
|
||||
rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three]
|
||||
|
@ -173,7 +173,7 @@ lemma ud (V : CKMMatrix) (a b c d e f : ℝ) :
|
|||
ring_nf
|
||||
|
||||
lemma us (V : CKMMatrix) (a b c d e f : ℝ) :
|
||||
(phaseShiftApply V a b c d e f).1 0 1 = cexp (a * I + e * I) * V.1 0 1 := by
|
||||
(phaseShiftApply V a b c d e f).1 0 1 = cexp (a * I + e * I) * V.1 0 1 := by
|
||||
simp only [Fin.isValue, phaseShiftApply_coe]
|
||||
rw [mul_apply, Fin.sum_univ_three]
|
||||
rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three]
|
||||
|
@ -184,7 +184,7 @@ lemma us (V : CKMMatrix) (a b c d e f : ℝ) :
|
|||
ring_nf
|
||||
|
||||
lemma ub (V : CKMMatrix) (a b c d e f : ℝ) :
|
||||
(phaseShiftApply V a b c d e f).1 0 2 = cexp (a * I + f * I) * V.1 0 2 := by
|
||||
(phaseShiftApply V a b c d e f).1 0 2 = cexp (a * I + f * I) * V.1 0 2 := by
|
||||
simp only [Fin.isValue, phaseShiftApply_coe]
|
||||
rw [mul_apply, Fin.sum_univ_three]
|
||||
rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three]
|
||||
|
@ -236,7 +236,7 @@ lemma td (V : CKMMatrix) (a b c d e f : ℝ) :
|
|||
simp only [Fin.isValue, cons_val', cons_val_zero, empty_val', cons_val_fin_one, vecCons_const,
|
||||
cons_val_two, tail_val', head_val', cons_val_one, head_cons, tail_cons, head_fin_const,
|
||||
zero_mul, add_zero, mul_zero]
|
||||
change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _
|
||||
change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _
|
||||
simp only [Fin.isValue, zero_mul, zero_add, mul_zero, add_zero]
|
||||
rw [exp_add]
|
||||
ring_nf
|
||||
|
@ -272,7 +272,7 @@ end phaseShiftApply
|
|||
def VAbs' (V : unitaryGroup (Fin 3) ℂ) (i j : Fin 3) : ℝ := Complex.abs (V i j)
|
||||
|
||||
lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) :
|
||||
VAbs' V i j = VAbs' U i j := by
|
||||
VAbs' V i j = VAbs' U i j := by
|
||||
simp only [VAbs']
|
||||
obtain ⟨a, b, c, e, f, g, h⟩ := h
|
||||
rw [h]
|
||||
|
@ -288,7 +288,7 @@ lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) :
|
|||
all_goals change Complex.abs (0 * _ + _) = _
|
||||
all_goals simp [Complex.abs_exp]
|
||||
|
||||
/-- The absolute value of the `(i,j)`th any representative of `⟦V⟧`. -/
|
||||
/-- The absolute value of the `(i,j)`th any representative of `⟦V⟧`. -/
|
||||
def VAbs (i j : Fin 3) : Quotient CKMMatrixSetoid → ℝ :=
|
||||
Quotient.lift (fun V => VAbs' V i j) (VAbs'_equiv i j)
|
||||
|
||||
|
|
|
@ -28,10 +28,10 @@ namespace Invariant
|
|||
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) :
|
||||
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
|
||||
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]
|
||||
|
@ -51,8 +51,8 @@ def jarlskogℂ : Quotient CKMMatrixSetoid → ℂ :=
|
|||
/-- 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)
|
||||
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. -/
|
||||
|
|
|
@ -126,7 +126,7 @@ lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ}
|
|||
have hτ0 := congrFun hτ 0
|
||||
simp [tRow] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
|
@ -166,15 +166,15 @@ def FstRowThdColRealCond (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U
|
|||
the cross product of the conjugates of the `u`th and `c`th rows, and the `cd`th and `cs`th
|
||||
elements are real and related in a set way.
|
||||
-/
|
||||
def ubOnePhaseCond (U : CKMMatrix) : Prop :=
|
||||
def ubOnePhaseCond (U : CKMMatrix) : Prop :=
|
||||
[U]ud = 0 ∧ [U]us = 0 ∧ [U]cb = 0 ∧ [U]ub = 1 ∧ [U]t = conj [U]u ×₃ conj [U]c
|
||||
∧ [U]cd = - VcdAbs ⟦U⟧ ∧ [U]cs = √(1 - VcdAbs ⟦U⟧ ^ 2)
|
||||
|
||||
lemma fstRowThdColRealCond_shift_solution {V : CKMMatrix} (h1 : a + d = - arg [V]ud)
|
||||
(h2 : a + e = - arg [V]us) (h3 : b + f = - arg [V]cb)
|
||||
(h2 : a + e = - arg [V]us) (h3 : b + f = - arg [V]cb)
|
||||
(h4 : c + f = - arg [V]tb) (h5 : τ = - a - b - c - d - e - f) :
|
||||
b = - τ + arg [V]ud + arg [V]us + arg [V]tb + a ∧
|
||||
c = - τ + arg [V]cb + arg [V]ud + arg [V]us + a ∧
|
||||
b = - τ + arg [V]ud + arg [V]us + arg [V]tb + a ∧
|
||||
c = - τ + arg [V]cb + arg [V]ud + arg [V]us + a ∧
|
||||
d = - arg [V]ud - a ∧
|
||||
e = - arg [V]us - a ∧
|
||||
f = τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb - a := by
|
||||
|
@ -202,11 +202,11 @@ lemma fstRowThdColRealCond_shift_solution {V : CKMMatrix} (h1 : a + d = - arg [V
|
|||
|
||||
lemma ubOnePhaseCond_shift_solution {V : CKMMatrix} (h1 : a + f = - arg [V]ub)
|
||||
(h2 : 0 = - a - b - c - d - e - f)
|
||||
(h3 : b + d = Real.pi - arg [V]cd) (h5 : b + e = - arg [V]cs) :
|
||||
c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b ∧
|
||||
d = Real.pi - arg [V]cd - b ∧
|
||||
e = - arg [V]cs - b ∧
|
||||
f = - arg [V]ub - a := by
|
||||
(h3 : b + d = Real.pi - arg [V]cd) (h5 : b + e = - arg [V]cs) :
|
||||
c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b ∧
|
||||
d = Real.pi - arg [V]cd - b ∧
|
||||
e = - arg [V]cs - b ∧
|
||||
f = - arg [V]ub - a := by
|
||||
have hf : f = - arg [V]ub - a := by
|
||||
linear_combination h1
|
||||
subst hf
|
||||
|
@ -227,21 +227,21 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
|||
obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow
|
||||
let U : CKMMatrix := phaseShiftApply V
|
||||
0
|
||||
(- τ + arg [V]ud + arg [V]us + arg [V]tb )
|
||||
(- τ + arg [V]cb + arg [V]ud + arg [V]us )
|
||||
(- τ + arg [V]ud + arg [V]us + arg [V]tb )
|
||||
(- τ + arg [V]cb + arg [V]ud + arg [V]us )
|
||||
(- arg [V]ud )
|
||||
(- arg [V]us)
|
||||
(τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb)
|
||||
have hUV : Quotient.mk CKMMatrixSetoid U = ⟦V⟧ := by
|
||||
simp only [Quotient.eq]
|
||||
symm
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
use U
|
||||
apply And.intro
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
apply shift_ud_phase_zero _ _ _ _ _ _ _
|
||||
apply shift_ud_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
|
@ -259,15 +259,15 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
|||
ring
|
||||
|
||||
lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0))
|
||||
(hV : FstRowThdColRealCond V) :
|
||||
(hV : FstRowThdColRealCond V) :
|
||||
∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by
|
||||
let U : CKMMatrix := phaseShiftApply V 0 0 (- Real.pi + arg [V]cd + arg [V]cs + arg [V]ub)
|
||||
(Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub )
|
||||
(Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub )
|
||||
use U
|
||||
have hUV : Quotient.mk CKMMatrixSetoid U= ⟦V⟧ := by
|
||||
simp only [Quotient.eq]
|
||||
symm
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
|
@ -294,7 +294,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
|||
exact h1
|
||||
apply And.intro
|
||||
· have hU1 : [U]ub = VubAbs ⟦V⟧ := by
|
||||
apply shift_ub_phase_zero _ _ _ _ _ _ _
|
||||
apply shift_ub_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
rw [hU1]
|
||||
have h1:= (ud_us_neq_zero_iff_ub_neq_one V).mpr.mt hb
|
||||
|
@ -307,7 +307,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
|||
ring
|
||||
apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_cd_phase_pi _ _ _ _ _ _ _
|
||||
apply shift_cd_phase_pi _ _ _ _ _ _ _
|
||||
ring
|
||||
have hcs : [U]cs = VcsAbs ⟦U⟧ := by
|
||||
rw [hUV]
|
||||
|
@ -338,7 +338,7 @@ lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠
|
|||
lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
(hV : FstRowThdColRealCond V) :
|
||||
[V]cs = (VtbAbs ⟦V⟧ * VudAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
|
||||
+ (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
|
||||
+ (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
|
||||
* cexp (- arg [V]ub * I) := by
|
||||
have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
simp only [ofReal_zero, zero_mul, exp_zero, one_smul]
|
||||
|
|
|
@ -64,7 +64,7 @@ lemma normSq_Vud_plus_normSq_Vus (V : CKMMatrix) :
|
|||
normSq [V]ud + normSq [V]us = 1 - normSq [V]ub := by
|
||||
linear_combination (fst_row_normalized_normSq V)
|
||||
|
||||
lemma VudAbs_sq_add_VusAbs_sq : VudAbs V ^ 2 + VusAbs V ^2 = 1 - VubAbs V ^2 := by
|
||||
lemma VudAbs_sq_add_VusAbs_sq : VudAbs V ^ 2 + VusAbs V ^2 = 1 - VubAbs V ^2 := by
|
||||
linear_combination (VAbs_sum_sq_row_eq_one V 0)
|
||||
|
||||
lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
||||
|
@ -104,7 +104,7 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℝ {V : CKMMatrix} (hb : [V]ud ≠ 0
|
|||
exact h2 h3
|
||||
|
||||
lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
||||
(hV : VAbs 0 2 V ≠ 1) :(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by
|
||||
(hV : VAbs 0 2 V ≠ 1) :(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by
|
||||
obtain ⟨V⟩ := V
|
||||
change VubAbs ⟦V⟧ ≠ 1 at hV
|
||||
simp only [VubAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk] at hV
|
||||
|
@ -112,7 +112,7 @@ lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
|||
simpa [← Complex.sq_abs] using (normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hV)
|
||||
|
||||
lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
||||
(hV : VAbs 0 2 V ≠ 1) : √(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by
|
||||
(hV : VAbs 0 2 V ≠ 1) : √(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by
|
||||
obtain ⟨V⟩ := V
|
||||
rw [Real.sqrt_ne_zero (Left.add_nonneg (sq_nonneg _) (sq_nonneg _))]
|
||||
change VubAbs ⟦V⟧ ≠ 1 at hV
|
||||
|
@ -120,7 +120,7 @@ lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
|||
rw [← ud_us_neq_zero_iff_ub_neq_one V] at hV
|
||||
simpa [← Complex.sq_abs] using (normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hV)
|
||||
|
||||
lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
||||
lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
||||
(normSq [V]ud : ℂ) + normSq [V]us ≠ 0 := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hb
|
||||
simp at h1
|
||||
|
@ -178,9 +178,9 @@ lemma VAbs_thd_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3} (hV
|
|||
|
||||
section crossProduct
|
||||
|
||||
lemma conj_Vtb_cross_product {V : CKMMatrix} {τ : ℝ}
|
||||
lemma conj_Vtb_cross_product {V : CKMMatrix} {τ : ℝ}
|
||||
(hτ : [V]t = cexp (τ * I) • (conj [V]u ×₃ conj [V]c)) :
|
||||
conj [V]tb = cexp (- τ * I) * ([V]cs * [V]ud - [V]us * [V]cd) := by
|
||||
conj [V]tb = cexp (- τ * I) * ([V]cs * [V]ud - [V]us * [V]cd) := by
|
||||
have h1 := congrFun hτ 2
|
||||
simp [crossProduct, tRow, uRow, cRow] at h1
|
||||
apply congrArg conj at h1
|
||||
|
@ -201,7 +201,7 @@ lemma conj_Vtb_mul_Vud {V : CKMMatrix} {τ : ℝ}
|
|||
simp [exp_neg]
|
||||
have h1 := exp_ne_zero (τ * I)
|
||||
field_simp
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]ud = [V]cs
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]ud = [V]cs
|
||||
* [V]ud * conj [V]ud - [V]us * ([V]cd * conj [V]ud) := by
|
||||
ring
|
||||
rw [h2, V.Vcd_mul_conj_Vud]
|
||||
|
@ -217,7 +217,7 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ}
|
|||
simp [exp_neg]
|
||||
have h1 := exp_ne_zero (τ * I)
|
||||
field_simp
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]us = ([V]cs
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]us = ([V]cs
|
||||
* conj [V]us) * [V]ud - [V]us * [V]cd * conj [V]us := by
|
||||
ring
|
||||
rw [h2, V.Vcs_mul_conj_Vus]
|
||||
|
@ -225,18 +225,18 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ}
|
|||
simp only [Fin.isValue, neg_mul]
|
||||
ring
|
||||
|
||||
lemma cs_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
lemma cs_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
{τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) :
|
||||
[V]cs = (- conj [V]ub * [V]us * [V]cb +
|
||||
cexp (τ * I) * conj [V]tb * conj [V]ud) / (normSq [V]ud + normSq [V]us) := by
|
||||
[V]cs = (- conj [V]ub * [V]us * [V]cb +
|
||||
cexp (τ * I) * conj [V]tb * conj [V]ud) / (normSq [V]ud + normSq [V]us) := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ h
|
||||
rw [conj_Vtb_mul_Vud hτ]
|
||||
field_simp
|
||||
ring
|
||||
|
||||
lemma cd_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
lemma cd_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
{τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) :
|
||||
[V]cd = - (conj [V]ub * [V]ud * [V]cb + cexp (τ * I) * conj [V]tb * conj [V]us) /
|
||||
[V]cd = - (conj [V]ub * [V]ud * [V]cb + cexp (τ * I) * conj [V]tb * conj [V]us) /
|
||||
(normSq [V]ud + normSq [V]us) := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ h
|
||||
rw [conj_Vtb_mul_Vus hτ]
|
||||
|
@ -247,7 +247,7 @@ end rows
|
|||
|
||||
section individual
|
||||
|
||||
lemma VAbs_ge_zero (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : 0 ≤ VAbs i j V := by
|
||||
lemma VAbs_ge_zero (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : 0 ≤ VAbs i j V := by
|
||||
obtain ⟨V, hV⟩ := Quot.exists_rep V
|
||||
rw [← hV]
|
||||
exact Complex.abs.nonneg _
|
||||
|
@ -291,7 +291,7 @@ lemma thd_col_normalized_normSq (V : CKMMatrix) :
|
|||
repeat rw [Complex.sq_abs] at h1
|
||||
exact h1
|
||||
|
||||
lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) :
|
||||
lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) :
|
||||
[V]cb = 0 := by
|
||||
have h1 := fst_row_normalized_abs V
|
||||
rw [← thd_col_normalized_abs V] at h1
|
||||
|
|
|
@ -130,13 +130,13 @@ lemma tRow_cRow_orthog (V : CKMMatrix) : conj [V]t ⬝ᵥ [V]c = 0 := by
|
|||
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
|
||||
exact ht
|
||||
|
||||
lemma uRow_cross_cRow_conj (V : CKMMatrix) : conj (conj [V]u ×₃ conj [V]c) = [V]u ×₃ [V]c := by
|
||||
lemma uRow_cross_cRow_conj (V : CKMMatrix) : conj (conj [V]u ×₃ conj [V]c) = [V]u ×₃ [V]c := by
|
||||
simp only [crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, LinearMap.mk₂_apply,
|
||||
Pi.conj_apply]
|
||||
funext i
|
||||
fin_cases i <;> simp
|
||||
|
||||
lemma cRow_cross_tRow_conj (V : CKMMatrix) : conj (conj [V]c ×₃ conj [V]t) = [V]c ×₃ [V]t := by
|
||||
lemma cRow_cross_tRow_conj (V : CKMMatrix) : conj (conj [V]c ×₃ conj [V]t) = [V]c ×₃ [V]t := by
|
||||
simp only [crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, LinearMap.mk₂_apply,
|
||||
Pi.conj_apply]
|
||||
funext i
|
||||
|
@ -169,9 +169,9 @@ lemma rows_linearly_independent (V : CKMMatrix) : LinearIndependent ℂ (rows V)
|
|||
intro g hg
|
||||
rw [Fin.sum_univ_three] at hg
|
||||
simp only [Fin.isValue, rows] at hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
have h2 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
have h2 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply,
|
||||
smul_eq_mul, dotProduct_zero] at h0 h1 h2
|
||||
rw [uRow_normalized, uRow_cRow_orthog, uRow_tRow_orthog] at h0
|
||||
|
@ -184,7 +184,7 @@ lemma rows_linearly_independent (V : CKMMatrix) : LinearIndependent ℂ (rows V)
|
|||
· exact h1
|
||||
· exact h2
|
||||
|
||||
lemma rows_card : Fintype.card (Fin 3) = FiniteDimensional.finrank ℂ (Fin 3 → ℂ) := by
|
||||
lemma rows_card : Fintype.card (Fin 3) = FiniteDimensional.finrank ℂ (Fin 3 → ℂ) := by
|
||||
simp
|
||||
|
||||
/-- The rows of a CKM matrix as a basis of `ℂ³`. -/
|
||||
|
@ -198,24 +198,24 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
(conj [V]c ×₃ conj [V]t))
|
||||
simp only [Fin.sum_univ_three, rowBasis, Fin.isValue,
|
||||
coe_basisOfLinearIndependentOfCardEqFinrank, rows] at hg
|
||||
have h0 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg
|
||||
have h0 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply,
|
||||
smul_eq_mul, dotProduct_zero] at h0 h1
|
||||
rw [cRow_normalized, cRow_uRow_orthog, cRow_tRow_orthog, dot_self_cross] at h0
|
||||
rw [tRow_normalized, tRow_uRow_orthog, tRow_cRow_orthog, dot_cross_self] at h1
|
||||
simp only [Fin.isValue, mul_zero, mul_one, zero_add, add_zero] at h0 h1 hg
|
||||
simp only [h0, h1, zero_smul, add_zero] at hg
|
||||
have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg
|
||||
have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_smul, Pi.conj_apply, Pi.smul_apply,
|
||||
smul_eq_mul, _root_.map_mul, cRow_cross_tRow_normalized] at h2
|
||||
have h3 : conj (g 0 • [V]u) = conj (g 0) • conj [V]u := by
|
||||
have h3 : conj (g 0 • [V]u) = conj (g 0) • conj [V]u := by
|
||||
funext i
|
||||
fin_cases i <;> simp
|
||||
simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul,
|
||||
uRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs] at h2
|
||||
simp at h2
|
||||
cases' h2 with h2 h2
|
||||
cases' h2 with h2 h2
|
||||
swap
|
||||
have hx : Complex.abs (g 0) = -1 := by
|
||||
simp [← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one, h2]
|
||||
|
@ -223,7 +223,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
simp_all
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
· exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
rw [← hg, smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
|
@ -238,30 +238,30 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
rw [hx, hτ]
|
||||
|
||||
lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
||||
∃ (τ : ℝ), [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
∃ (τ : ℝ), [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ℂ).mp (Basis.mem_span (rowBasis V)
|
||||
(conj ([V]u) ×₃ conj ([V]c)) )
|
||||
rw [Fin.sum_univ_three, rowBasis] at hg
|
||||
simp at hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply,
|
||||
smul_eq_mul, dotProduct_zero] at h0 h1
|
||||
rw [uRow_normalized, uRow_cRow_orthog, uRow_tRow_orthog, dot_self_cross] at h0
|
||||
rw [cRow_normalized, cRow_uRow_orthog, cRow_tRow_orthog, dot_cross_self] at h1
|
||||
simp only [Fin.isValue, mul_one, mul_zero, add_zero, zero_add] at h0 h1
|
||||
simp only [Fin.isValue, h0, zero_smul, h1, add_zero, zero_add] at hg
|
||||
have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg
|
||||
have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_smul, Pi.conj_apply, Pi.smul_apply,
|
||||
smul_eq_mul, _root_.map_mul] at h2
|
||||
rw [uRow_cross_cRow_normalized] at h2
|
||||
have h3 : conj (g 2 • [V]t) = conj (g 2) • conj [V]t := by
|
||||
have h3 : conj (g 2 • [V]t) = conj (g 2) • conj [V]t := by
|
||||
funext i
|
||||
fin_cases i <;> simp
|
||||
simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul, tRow_normalized,
|
||||
Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs, ofReal_pow, sq_eq_one_iff,
|
||||
ofReal_eq_one] at h2
|
||||
cases' h2 with h2 h2
|
||||
cases' h2 with h2 h2
|
||||
swap
|
||||
have hx : Complex.abs (g 2) = -1 := by
|
||||
simp [h2, ← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one]
|
||||
|
@ -269,7 +269,7 @@ lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
|||
simp_all
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
rw [← hg, @smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
|
@ -303,30 +303,30 @@ open CKMMatrix
|
|||
variable (V : CKMMatrix) (a b c d e f : ℝ)
|
||||
|
||||
/-- The cross product of the conjugate of the `u` and `c` rows of a CKM matrix. -/
|
||||
def ucCross : Fin 3 → ℂ :=
|
||||
def ucCross : Fin 3 → ℂ :=
|
||||
conj [phaseShiftApply V a b c d e f]u ×₃ conj [phaseShiftApply V a b c d e f]c
|
||||
|
||||
lemma ucCross_fst (V : CKMMatrix) : (ucCross V a b c d e f) 0 =
|
||||
cexp ((- a * I) + (- b * I) + ( - e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by
|
||||
cexp ((- a * I) + (- b * I) + ( - e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by
|
||||
simp [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue,
|
||||
LinearMap.mk₂_apply, Pi.conj_apply, cons_val_zero, neg_mul, uRow, us, ub, cRow, cs, cb,
|
||||
exp_add, exp_sub, ← exp_conj, conj_ofReal]
|
||||
ring
|
||||
|
||||
lemma ucCross_snd (V : CKMMatrix) : (ucCross V a b c d e f) 1 =
|
||||
cexp ((- a * I) + (- b * I) + ( - d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by
|
||||
cexp ((- a * I) + (- b * I) + ( - d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add,
|
||||
exp_sub, ← exp_conj, conj_ofReal]
|
||||
ring
|
||||
|
||||
lemma ucCross_thd (V : CKMMatrix) : (ucCross V a b c d e f) 2 =
|
||||
cexp ((- a * I) + (- b * I) + ( - d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by
|
||||
cexp ((- a * I) + (- b * I) + ( - d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub,
|
||||
← exp_conj, conj_ofReal]
|
||||
ring
|
||||
|
||||
lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
[phaseShiftApply V a b c 0 0 0]u = cexp (a * I) • [V]u := by
|
||||
[phaseShiftApply V a b c 0 0 0]u = cexp (a * I) • [V]u := by
|
||||
funext i
|
||||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
|
@ -336,7 +336,7 @@ lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp [ub, uRow]
|
||||
|
||||
lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
[phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by
|
||||
[phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by
|
||||
funext i
|
||||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
|
@ -346,7 +346,7 @@ lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp [cb, cRow]
|
||||
|
||||
lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
[phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by
|
||||
[phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by
|
||||
funext i
|
||||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
|
@ -355,6 +355,6 @@ lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp [ts, tRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.mk_one, cons_val_one, head_cons]
|
||||
simp [tb, tRow]
|
||||
|
||||
end phaseShiftApply
|
||||
end phaseShiftApply
|
||||
|
||||
end
|
||||
|
|
|
@ -25,7 +25,7 @@ noncomputable section
|
|||
|
||||
/-- Given four reals `θ₁₂ θ₁₃ θ₂₃ δ₁₃` the standard paramaterization of the CKM matrix
|
||||
as a `3×3` complex matrix. -/
|
||||
def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin 3) (Fin 3) ℂ :=
|
||||
def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin 3) (Fin 3) ℂ :=
|
||||
![![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 * δ₁₃),
|
||||
|
@ -36,8 +36,8 @@ def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin
|
|||
|
||||
open CKMMatrix
|
||||
|
||||
lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
||||
((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
|
||||
lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
||||
((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
|
||||
funext j i
|
||||
simp only [standParamAsMatrix, neg_mul, Fin.isValue]
|
||||
rw [mul_apply]
|
||||
|
@ -139,7 +139,7 @@ lemma cross_product_t (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
ring
|
||||
|
||||
lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ} (hu : [U]u = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u)
|
||||
(hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) :
|
||||
(hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) :
|
||||
U = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by
|
||||
apply ext_Rows hu hc
|
||||
rw [hU, cross_product_t, hu, hc]
|
||||
|
@ -149,7 +149,7 @@ lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ℝ) (h
|
|||
simp [standParam, standParamAsMatrix]
|
||||
apply CKMMatrix_ext
|
||||
simp only
|
||||
rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]]
|
||||
rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]]
|
||||
rw [show cexp (-(I * ↑δ₁₃)) = cexp (-(I * ↑δ₁₃')) by rw [exp_neg, exp_neg, mul_comm, h, mul_comm]]
|
||||
|
||||
open Invariant in
|
||||
|
@ -179,9 +179,9 @@ lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Rea
|
|||
|
||||
open Invariant in
|
||||
lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂)
|
||||
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
|
||||
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
|
||||
mulExpδ₁₃ ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
|
||||
sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by
|
||||
sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by
|
||||
rw [mulExpδ₁₃, VusVubVcdSq_eq _ _ _ _ h1 h2 h3 h4 ]
|
||||
simp only [jarlskogℂ, standParam, standParamAsMatrix, neg_mul,
|
||||
Quotient.lift_mk, jarlskogℂCKM, Fin.isValue, cons_val', cons_val_one, head_cons,
|
||||
|
|
|
@ -12,7 +12,7 @@ import Mathlib.Analysis.SpecialFunctions.Complex.Arg
|
|||
/-!
|
||||
# Standard parameters for the CKM Matrix
|
||||
|
||||
Given a CKM matrix `V` we can extract four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`.
|
||||
Given a CKM matrix `V` we can extract four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`.
|
||||
These, when used in the standard parameterization return `V` up to equivalence.
|
||||
|
||||
This leads to the theorem `standParam.exists_for_CKMatrix` which says that up to equivalence every
|
||||
|
@ -26,15 +26,15 @@ open CKMMatrix
|
|||
noncomputable section
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to `sin θ₁₂` in the
|
||||
standard parameterization. --/
|
||||
standard parameterization. --/
|
||||
def S₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := VusAbs V / (√ (VudAbs V ^ 2 + VusAbs V ^ 2))
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to `sin θ₁₃` in the
|
||||
standard parameterization. --/
|
||||
standard parameterization. --/
|
||||
def S₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := VubAbs V
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to `sin θ₂₃` in the
|
||||
standard parameterization. --/
|
||||
standard parameterization. --/
|
||||
def S₂₃ (V : Quotient CKMMatrixSetoid) : ℝ :=
|
||||
if VubAbs V = 1 then VcdAbs V
|
||||
else VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2)
|
||||
|
@ -56,7 +56,7 @@ standard parameterization. --/
|
|||
def C₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₂ V)
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to `cos θ₁₃` in the
|
||||
standard parameterization. --/
|
||||
standard parameterization. --/
|
||||
def C₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₃ V)
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to `sin θ₂₃` in the
|
||||
|
@ -64,7 +64,7 @@ standard parameterization. --/
|
|||
def C₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₂₃ V)
|
||||
|
||||
/-- Given a CKM matrix `V` the real number corresponding to the phase `δ₁₃` in the
|
||||
standard parameterization. --/
|
||||
standard parameterization. --/
|
||||
def δ₁₃ (V : Quotient CKMMatrixSetoid) : ℝ :=
|
||||
arg (Invariant.mulExpδ₁₃ V)
|
||||
|
||||
|
@ -336,7 +336,7 @@ namespace standParam
|
|||
open Invariant
|
||||
|
||||
lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ℝ) :
|
||||
mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ =
|
||||
mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ =
|
||||
sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧)
|
||||
* cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) * cexp (I * δ₁₃) := by
|
||||
refine mulExpδ₁₃_eq _ _ _ _ ?_ ?_ ?_ ?_
|
||||
|
@ -348,11 +348,11 @@ lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ℝ) :
|
|||
exact Real.cos_arcsin_nonneg _
|
||||
|
||||
lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ℝ) :
|
||||
mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = 0 ↔
|
||||
mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = 0 ↔
|
||||
VudAbs ⟦V⟧ = 0 ∨ VubAbs ⟦V⟧ = 0 ∨ VusAbs ⟦V⟧ = 0 ∨ VcbAbs ⟦V⟧ = 0 ∨ VtbAbs ⟦V⟧ = 0 := by
|
||||
rw [VudAbs_eq_C₁₂_mul_C₁₃, VubAbs_eq_S₁₃, VusAbs_eq_S₁₂_mul_C₁₃, VcbAbs_eq_S₂₃_mul_C₁₃,
|
||||
VtbAbs_eq_C₂₃_mul_C₁₃, ← ofReal_inj,
|
||||
← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj]
|
||||
← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj]
|
||||
simp only [ofReal_mul]
|
||||
rw [← S₁₃_eq_ℂsin_θ₁₃, ← S₁₂_eq_ℂsin_θ₁₂, ← S₂₃_eq_ℂsin_θ₂₃,
|
||||
← C₁₃_eq_ℂcos_θ₁₃, ← C₂₃_eq_ℂcos_θ₂₃,← C₁₂_eq_ℂcos_θ₁₂]
|
||||
|
@ -364,7 +364,7 @@ lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ℝ) :
|
|||
aesop
|
||||
|
||||
lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) :
|
||||
Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) =
|
||||
Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) =
|
||||
sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧)
|
||||
* cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) := by
|
||||
rw [mulExpδ₁₃_on_param_δ₁₃]
|
||||
|
@ -373,19 +373,19 @@ lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) :
|
|||
complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃]
|
||||
|
||||
lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ)
|
||||
(h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) :
|
||||
cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) =
|
||||
(h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) :
|
||||
cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) =
|
||||
cexp (δ₁₃ * I) := by
|
||||
have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃
|
||||
have habs := mulExpδ₁₃_on_param_abs V δ₁₃
|
||||
have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs
|
||||
(mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by
|
||||
have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs
|
||||
(mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by
|
||||
rw [habs, h1a]
|
||||
ring_nf
|
||||
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (mulExpδ₁₃
|
||||
⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2
|
||||
⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2
|
||||
have habs_neq_zero :
|
||||
(Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by
|
||||
(Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by
|
||||
simp only [ne_eq, ofReal_eq_zero, map_eq_zero]
|
||||
exact h1
|
||||
rw [← mul_right_inj' habs_neq_zero]
|
||||
|
@ -393,7 +393,7 @@ lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ)
|
|||
|
||||
lemma on_param_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.cos (θ₁₃ ⟦V⟧) = 0) :
|
||||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
have hS13 := congrArg ofReal (S₁₃_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h))
|
||||
have hS13 := congrArg ofReal (S₁₃_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h))
|
||||
simp [← S₁₃_eq_ℂsin_θ₁₃] at hS13
|
||||
have hC12 := congrArg ofReal (C₁₂_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h))
|
||||
simp [← C₁₂_eq_ℂcos_θ₁₂] at hC12
|
||||
|
@ -638,8 +638,8 @@ theorem eq_standardParameterization_δ₃ (V : CKMMatrix) :
|
|||
V ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (δ₁₃ ⟦V⟧) := by
|
||||
obtain ⟨δ₁₃', hδ₃⟩ := exists_δ₁₃ V
|
||||
have hSV := (Quotient.eq.mpr (hδ₃))
|
||||
by_cases h : Invariant.mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0
|
||||
have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'
|
||||
by_cases h : Invariant.mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0
|
||||
have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'
|
||||
(δ₁₃ ⟦V⟧) (by rw [← mulExpδ₁₃_on_param_neq_zero_arg V δ₁₃' h, ← hSV, δ₁₃, Invariant.mulExpδ₁₃])
|
||||
rw [h2] at hδ₃
|
||||
exact hδ₃
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue