feat: add explicit equation for phase

This commit is contained in:
jstoobysmith 2024-04-26 14:52:56 -04:00
parent 5a5540ba78
commit fe63fc9994
3 changed files with 376 additions and 5 deletions

View file

@ -384,8 +384,6 @@ lemma fst_row_thd_row (V : CKMMatrix) : V.1 2 0 * conj (V.1 0 0) + V.1 2 1 * c
exact ht
end CKMMatrix
end

View file

@ -0,0 +1,176 @@
/-
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 HepLean.FlavorPhysics.CKMMatrix.Rows
import HepLean.FlavorPhysics.CKMMatrix.PhaseFreedom
import HepLean.FlavorPhysics.CKMMatrix.Ratios
import HepLean.FlavorPhysics.CKMMatrix.StandardParameters
import Mathlib.Analysis.SpecialFunctions.Complex.Arg
open Matrix Complex
open ComplexConjugate
open CKMMatrix
noncomputable section
@[simps!]
def jarlskogComplexCKM (V : CKMMatrix) : :=
[V]us * [V]cb * conj [V]ub * conj [V]cs
lemma jarlskogComplexCKM_equiv (V U : CKMMatrix) (h : V ≈ U) :
jarlskogComplexCKM V = jarlskogComplexCKM 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 [jarlskogComplexCKM, 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
def inv₁ (V : Quotient CKMMatrixSetoid) : :=
VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 /(VudAbs V ^ 2 + VusAbs V ^2)
lemma inv₁_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
inv₁ ⟦sP θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
Real.sin θ₁₂ ^ 2 * Real.cos θ₁₃ ^ 2 * Real.sin θ₁₃ ^ 2 * Real.sin θ₂₃ ^ 2 := by
simp only [inv₁, VusAbs, VAbs, VAbs', Fin.isValue, sP, standardParameterizationAsMatrix,
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]
simp
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
@[simp]
def jarlskogComplex : Quotient CKMMatrixSetoid → :=
Quotient.lift jarlskogComplexCKM jarlskogComplexCKM_equiv
-- bad name
def expδ₁₃ (V : Quotient CKMMatrixSetoid) : :=
jarlskogComplex V + inv₁ V
lemma expδ₁₃_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
expδ₁₃ ⟦sP θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by
rw [expδ₁₃]
rw [inv₁_sP _ _ _ _ h1 h2 h3 h4 ]
simp only [expδ₁₃, jarlskogComplex, sP, standardParameterizationAsMatrix, neg_mul,
Quotient.lift_mk, jarlskogComplexCKM, 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]
simp
ring_nf
rw [exp_neg]
have h1 : cexp (I * δ₁₃) ≠ 0 := exp_ne_zero _
field_simp
lemma expδ₁₃_sP_V (V : CKMMatrix) (δ₁₃ : ) :
expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ =
sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧)
* cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) * cexp (I * δ₁₃) := by
refine expδ₁₃_sP _ _ _ _ ?_ ?_ ?_ ?_
rw [S₁₂_eq_sin_θ₁₂]
exact S₁₂_nonneg _
exact Real.cos_arcsin_nonneg _
rw [S₂₃_eq_sin_θ₂₃]
exact S₂₃_nonneg _
exact Real.cos_arcsin_nonneg _
lemma expδ₁₃_eq_zero (V : CKMMatrix) (δ₁₃ : ) :
expδ₁₃ ⟦sP (θ₁₂ ⟦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]
simp only [ofReal_mul]
rw [← S₁₃_eq_sin_θ₁₃, ← S₁₂_eq_sin_θ₁₂, ← S₂₃_eq_sin_θ₂₃,
← C₁₃_eq_cos_θ₁₃, ← C₂₃_eq_cos_θ₂₃,← C₁₂_eq_cos_θ₁₂]
simp
rw [expδ₁₃_sP_V]
simp
have h1 := exp_ne_zero (I * δ₁₃)
simp_all
aesop
lemma inv₂_V_arg (V : CKMMatrix) (δ₁₃ : )
(h1 : expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) :
cexp (arg (expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * I) =
cexp (δ₁₃ * I) := by
have h1a := expδ₁₃_sP_V V δ₁₃
have habs : Complex.abs (expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) =
sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧)
* cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) := by
rw [h1a]
simp [abs_exp]
rw [complexAbs_sin_θ₁₃, complexAbs_cos_θ₁₃, complexAbs_sin_θ₁₂, complexAbs_cos_θ₁₂,
complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃]
have h2 : expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ =
Complex.abs (expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by
rw [habs, h1a]
ring_nf
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2
have habs_neq_zero : (Complex.abs (expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ) ≠ 0 := by
simp
exact h1
rw [← mul_right_inj' habs_neq_zero]
rw [← h2]
def δ₁₃ (V : Quotient CKMMatrixSetoid) : := arg (expδ₁₃ V)
theorem eq_standardParameterization_δ₃ (V : CKMMatrix) :
V ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (δ₁₃ ⟦V⟧) := by
obtain ⟨δ₁₃', hδ₃⟩ := exists_standardParameterization V
have hSV := (Quotient.eq.mpr (hδ₃))
by_cases h : expδ₁₃ ⟦sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0
have h1 := inv₂_V_arg V δ₁₃' h
have h2 := eq_phases_sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'
(δ₁₃ ⟦V⟧) (by rw [← h1, ← hSV, δ₁₃])
rw [h2] at hδ₃
exact hδ₃
simp at h
have h1 : δ₁₃ ⟦V⟧ = 0 := by
rw [hSV, δ₁₃, h]
simp
rw [h1]
rw [expδ₁₃_eq_zero, Vs_zero_iff_cos_sin_zero] at h
cases' h with h h
exact phaseShiftEquivRelation.trans hδ₃ (sP_cos_θ₁₂_eq_zero δ₁₃' h )
cases' h with h h
exact phaseShiftEquivRelation.trans hδ₃ (sP_cos_θ₁₃_eq_zero δ₁₃' h )
cases' h with h h
exact phaseShiftEquivRelation.trans hδ₃ (sP_cos_θ₂₃_eq_zero δ₁₃' h )
cases' h with h h
exact phaseShiftEquivRelation.trans hδ₃ (sP_sin_θ₁₂_eq_zero δ₁₃' h )
cases' h with h h
exact phaseShiftEquivRelation.trans hδ₃ (sP_sin_θ₁₃_eq_zero δ₁₃' h )
exact phaseShiftEquivRelation.trans hδ₃ (sP_sin_θ₂₃_eq_zero δ₁₃' h )
end

View file

@ -31,18 +31,49 @@ lemma S₁₃_eq_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁
rw [← VubAbs_eq_S₁₃]
exact (VAbs_leq_one 0 2 V)
lemma S₁₃_of_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₁₃ V = 1 := by
rw [← VubAbs_eq_S₁₃, ha]
lemma S₁₃_eq_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.sin (θ₁₃ V) = S₁₃ V := by
rw [← S₁₃_eq_sin_θ₁₃]
simp
lemma complexAbs_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.sin (θ₁₃ V)) =
sin (θ₁₃ V):= by
rw [S₁₃_eq_sin_θ₁₃, ← VubAbs_eq_S₁₃]
rw [Complex.abs_ofReal]
simp
exact VAbs_ge_zero 0 2 V
lemma S₁₃_of_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₁₃ V = 1 := by
rw [← VubAbs_eq_S₁₃, ha]
def C₁₃ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₁₃ V)
lemma C₁₃_eq_cos_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₁₃ V) = C₁₃ V := by
simp [C₁₃]
lemma complexAbs_cos_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₁₃ V)) =
cos (θ₁₃ V):= by
rw [C₁₃_eq_cos_θ₁₃, Complex.abs_ofReal]
simp
exact Real.cos_arcsin_nonneg _
lemma cos_θ₁₃_zero {V : Quotient CKMMatrixSetoid} (h1 : Real.cos (θ₁₃ V) = 0) :
VubAbs V = 1 := by
rw [θ₁₃, Real.cos_arcsin, ← VubAbs_eq_S₁₃, Real.sqrt_eq_zero] at h1
have h2 : VubAbs V ^ 2 = 1 := by linear_combination -(1 * h1)
simp at h2
cases' h2 with h2 h2
exact h2
have h3 := VAbs_ge_zero 0 2 V
rw [h2] at h3
simp at h3
linarith
simp
rw [_root_.abs_of_nonneg (VAbs_ge_zero 0 2 V)]
exact VAbs_leq_one 0 2 V
lemma S₁₃_sq_add_C₁₃_sq (V : Quotient CKMMatrixSetoid) : S₁₃ V ^ 2 + C₁₃ V ^ 2 = 1 := by
rw [← S₁₃_eq_sin_θ₁₃ V, C₁₃]
exact Real.sin_sq_add_cos_sq (θ₁₃ V)
@ -70,6 +101,7 @@ lemma S₁₂_nonneg (V : Quotient CKMMatrixSetoid) : 0 ≤ S₁₂ V := by
exact VAbs_ge_zero 0 1 V
exact Real.sqrt_nonneg (VudAbs V ^ 2 + VusAbs V ^ 2)
lemma S₁₂_leq_one (V : Quotient CKMMatrixSetoid) : S₁₂ V ≤ 1 := by
rw [S₁₂]
rw [@div_le_one_iff]
@ -96,6 +128,12 @@ lemma S₁₂_eq_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.sin (
rw [← S₁₂_eq_sin_θ₁₂]
simp
lemma complexAbs_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.sin (θ₁₂ V)) =
sin (θ₁₂ V):= by
rw [S₁₂_eq_sin_θ₁₂, Complex.abs_ofReal]
simp
exact S₁₂_nonneg _
lemma S₁₂_of_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₁₂ V = 0 := by
rw [S₁₂]
have h1 : 1 - VubAbs V ^ 2 = VudAbs V ^ 2 + VusAbs V ^ 2 := by
@ -109,6 +147,12 @@ def C₁₂ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₁₂ V)
lemma C₁₂_eq_cos_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₁₂ V) = C₁₂ V := by
simp [C₁₂]
lemma complexAbs_cos_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₁₂ V)) =
cos (θ₁₂ V):= by
rw [C₁₂_eq_cos_θ₁₂, Complex.abs_ofReal]
simp
exact Real.cos_arcsin_nonneg _
lemma C₁₂_of_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : C₁₂ V = 1 := by
rw [C₁₂, θ₁₂, Real.cos_arcsin, S₁₂_of_Vub_one ha]
simp
@ -205,6 +249,11 @@ lemma S₂₃_eq_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.sin (
rw [← S₂₃_eq_sin_θ₂₃]
simp
lemma complexAbs_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.sin (θ₂₃ V)) =
sin (θ₂₃ V):= by
rw [S₂₃_eq_sin_θ₂₃, Complex.abs_ofReal]
simp
exact S₂₃_nonneg _
lemma S₂₃_of_Vub_eq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₂₃ V = VcdAbs V := by
rw [S₂₃, if_pos ha]
@ -218,6 +267,12 @@ def C₂₃ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₂₃ V)
lemma C₂₃_eq_cos_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₂₃ V) = C₂₃ V := by
simp [C₂₃]
lemma complexAbs_cos_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₂₃ V)) =
cos (θ₂₃ V):= by
rw [C₂₃_eq_cos_θ₂₃, Complex.abs_ofReal]
simp
exact Real.cos_arcsin_nonneg _
lemma S₂₃_sq_add_C₂₃_sq (V : Quotient CKMMatrixSetoid) : S₂₃ V ^ 2 + C₂₃ V ^ 2 = 1 := by
rw [← S₂₃_eq_sin_θ₂₃ V, C₂₃]
exact Real.sin_sq_add_cos_sq (θ₂₃ V)
@ -385,6 +440,148 @@ lemma eq_sP (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : } (hu : [U
apply ext_Rows hu hc
rw [hU, sP_cross, hu, hc]
lemma eq_phases_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
sP θ₁₂ θ₁₃ θ₂₃ δ₁₃ = sP θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
simp [sP, standardParameterizationAsMatrix]
apply CKMMatrix_ext
simp
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]]
section zeroEntries
variable (a b c d e f : )
lemma sP_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.cos (θ₁₃ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
have hS13 := congrArg ofReal (S₁₃_of_Vub_one (cos_θ₁₃_zero h))
simp [← S₁₃_eq_sin_θ₁₃] at hS13
have hC12 := congrArg ofReal (C₁₂_of_Vub_one (cos_θ₁₃_zero h))
simp [← C₁₂_eq_cos_θ₁₂] at hC12
have hS12 := congrArg ofReal (S₁₂_of_Vub_one (cos_θ₁₃_zero h))
simp [← S₁₂_eq_sin_θ₁₂] at hS12
use 0, 0, 0, δ₁₃, 0, -δ₁₃
simp [sP, standardParameterizationAsMatrix, h, phaseShift, hS13, hC12, hS12]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
rfl
rfl
lemma sP_cos_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.cos (θ₁₂ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
use 0, δ₁₃, δ₁₃, -δ₁₃, 0, - δ₁₃
have hb := exp_ne_zero (I * δ₁₃)
simp [sP, standardParameterizationAsMatrix, h, phaseShift, exp_neg]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
apply Or.inr
rfl
change _ = _ + _ * 0
simp
field_simp
ring
ring
field_simp
ring
change _ = _ + _ * 0
simp
field_simp
ring
ring
field_simp
ring
lemma sP_cos_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.cos (θ₂₃ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
use 0, δ₁₃, 0, 0, 0, - δ₁₃
have hb := exp_ne_zero (I * δ₁₃)
simp [sP, standardParameterizationAsMatrix, h, phaseShift, exp_neg]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
apply Or.inr
rfl
ring_nf
change _ = _ + _ * 0
simp
ring
field_simp
ring
lemma sP_sin_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.sin (θ₁₃ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
use 0, 0, 0, 0, 0, 0
simp [sP, standardParameterizationAsMatrix, h, phaseShift, exp_neg]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
apply Or.inr
rfl
apply Or.inr
rfl
lemma sP_sin_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.sin (θ₁₂ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
use 0, δ₁₃, δ₁₃, 0, -δ₁₃, - δ₁₃
have hb := exp_ne_zero (I * δ₁₃)
simp [sP, standardParameterizationAsMatrix, h, phaseShift, exp_neg]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
apply Or.inr
rfl
change _ = _ + _ * 0
simp
ring
field_simp
ring
field_simp
ring
change _ = _ + _ * 0
simp
ring
field_simp
ring
field_simp
ring
lemma sP_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ) (h : Real.sin (θ₂₃ ⟦V⟧) = 0) :
sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
use 0, 0, δ₁₃, 0, 0, - δ₁₃
have hb := exp_ne_zero (I * δ₁₃)
simp [sP, standardParameterizationAsMatrix, h, phaseShift, exp_neg]
funext i j
fin_cases i <;> fin_cases j <;>
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
apply Or.inr
rfl
change _ = _ + _ * 0
simp
ring
ring
field_simp
ring
lemma Vs_zero_iff_cos_sin_zero (V : CKMMatrix) :
VudAbs ⟦V⟧ = 0 VubAbs ⟦V⟧ = 0 VusAbs ⟦V⟧ = 0 VcbAbs ⟦V⟧ = 0 VtbAbs ⟦V⟧ = 0
↔ Real.cos (θ₁₂ ⟦V⟧) = 0 Real.cos (θ₁₃ ⟦V⟧) = 0 Real.cos (θ₂₃ ⟦V⟧) = 0
Real.sin (θ₁₂ ⟦V⟧) = 0 Real.sin (θ₁₃ ⟦V⟧) = 0 Real.sin (θ₂₃ ⟦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₁₃]
rw [C₁₂, C₁₃, C₂₃, S₁₂_eq_sin_θ₁₂, S₂₃_eq_sin_θ₂₃, S₁₃_eq_sin_θ₁₃]
aesop
end zeroEntries
lemma UCond₁_eq_sP {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb ≠ 0)
(hV : UCond₁ V) : V = sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (- arg [V]ub) := by