feat: Every CKM can be written with standard parameters

This commit is contained in:
jstoobysmith 2024-04-26 10:32:03 -04:00
parent e5337f021d
commit 5a5540ba78
6 changed files with 2043 additions and 271 deletions

View file

@ -14,7 +14,7 @@ noncomputable section
@[simp]
def phaseShiftMatrix (a b c : ) : Matrix (Fin 3) (Fin 3) :=
![![ Complex.exp (I * a), 0, 0], ![0, Complex.exp (I * b), 0], ![0, 0, Complex.exp (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
@ -97,8 +97,125 @@ def phaseShiftEquivRelation : Equivalence phaseShiftRelation where
def CKMMatrix : Type := unitaryGroup (Fin 3)
lemma CKMMatrix_ext {U V : CKMMatrix} (h : U.val = V.val) : U = V := by
cases U; cases V
simp at h
subst h
rfl
scoped[CKMMatrix] notation (name := ud_element) "[" V "]ud" => V.1 0 0
scoped[CKMMatrix] notation (name := us_element) "[" V "]us" => V.1 0 1
scoped[CKMMatrix] notation (name := ub_element) "[" V "]ub" => V.1 0 2
scoped[CKMMatrix] notation (name := cd_element) "[" V "]cd" => V.1 1 0
scoped[CKMMatrix] notation (name := cs_element) "[" V "]cs" => V.1 1 1
scoped[CKMMatrix] notation (name := cb_element) "[" V "]cb" => V.1 1 2
scoped[CKMMatrix] notation (name := td_element) "[" V "]td" => V.1 2 0
scoped[CKMMatrix] notation (name := ts_element) "[" V "]ts" => V.1 2 1
scoped[CKMMatrix] notation (name := tb_element) "[" V "]tb" => V.1 2 2
instance CKMMatrixSetoid : Setoid CKMMatrix := ⟨phaseShiftRelation, phaseShiftEquivRelation⟩
@[simps!]
def phaseShiftApply (V : CKMMatrix) (a b c d e f : ) : CKMMatrix :=
phaseShift a b c * ↑V * phaseShift d e f
namespace phaseShiftApply
lemma equiv (V : CKMMatrix) (a b c d e f : ) :
V ≈ phaseShiftApply V a b c d e f := by
symm
use 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
simp
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]
simp
change _ + _ * _ * 0 = _
rw [exp_add]
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
simp
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]
simp
rw [exp_add]
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
simp
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]
simp
rw [exp_add]
ring_nf
lemma cd (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 1 0= cexp (b * I + d * I) * V.1 1 0 := by
simp
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]
simp
change _ + _ * _ * 0 = _
rw [exp_add]
ring_nf
lemma cs (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 1 1 = cexp (b * I + e * I) * V.1 1 1 := by
simp
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]
simp
rw [exp_add]
ring_nf
lemma cb (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 1 2 = cexp (b * I + f * I) * V.1 1 2 := by
simp
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]
simp
rw [exp_add]
ring_nf
lemma td (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 2 0= cexp (c * I + d * I) * V.1 2 0 := by
simp
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]
simp
change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _
simp
rw [exp_add]
ring_nf
lemma ts (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 2 1 = cexp (c * I + e * I) * V.1 2 1 := by
simp
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]
simp
change (0 * _ + _) * _ = _
rw [exp_add]
ring_nf
lemma tb (V : CKMMatrix) (a b c d e f : ) :
(phaseShiftApply V a b c d e f).1 2 2 = cexp (c * I + f * I) * V.1 2 2 := by
simp
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]
simp
change (0 * _ + _) * _ = _
rw [exp_add]
ring_nf
end phaseShiftApply
@[simp]
def VAbs' (V : unitaryGroup (Fin 3) ) (i j : Fin 3) : := Complex.abs (V i j)
@ -191,6 +308,7 @@ lemma VAbs_leq_one (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : VAbs i j V ≤
change VAbs i 2 V ≤ 1
nlinarith
lemma VAbs_thd_eq_one_fst_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3} (hV : VAbs i 2 V = 1) :
VAbs i 0 V = 0 := by
have h := VAbs_sum_sq_row_eq_one V i
@ -266,6 +384,8 @@ 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,217 @@
/-
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 Mathlib.Analysis.SpecialFunctions.Complex.Arg
import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic
open Matrix Complex
noncomputable section
namespace CKMMatrix
open ComplexConjugate
variable (a b c d e f : )
lemma ud_eq_abs (V : CKMMatrix) (h1 : a + d = - arg [V]ud) :
[phaseShiftApply V a b c d e f]ud = VudAbs ⟦V⟧ := by
rw [phaseShiftApply.ud]
rw [← abs_mul_exp_arg_mul_I [V]ud]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg (V.1 0 0)) * I + (↑a * I + ↑d * I) = ↑(arg (V.1 0 0) + (a + d)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma us_eq_abs {V : CKMMatrix} (h1 : a + e = - arg [V]us) :
[phaseShiftApply V a b c d e f]us = VusAbs ⟦V⟧ := by
rw [phaseShiftApply.us]
rw [← abs_mul_exp_arg_mul_I [V]us]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]us) * I + (↑a * I + ↑e * I) = ↑(arg [V]us + (a + e)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma ub_eq_abs {V : CKMMatrix} (h1 : a + f = - arg [V]ub) :
[phaseShiftApply V a b c d e f]ub = VubAbs ⟦V⟧ := by
rw [phaseShiftApply.ub]
rw [← abs_mul_exp_arg_mul_I [V]ub]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]ub) * I + (↑a * I + ↑f * I) = ↑(arg [V]ub + (a + f)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma cs_eq_abs {V : CKMMatrix} (h1 : b + e = - arg [V]cs) :
[phaseShiftApply V a b c d e f]cs = VcsAbs ⟦V⟧ := by
rw [phaseShiftApply.cs]
rw [← abs_mul_exp_arg_mul_I [V]cs]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]cs) * I + (↑b * I + ↑e * I) = ↑(arg [V]cs + (b + e)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma cb_eq_abs {V : CKMMatrix} (h1 : b + f = - arg [V]cb) :
[phaseShiftApply V a b c d e f]cb = VcbAbs ⟦V⟧ := by
rw [phaseShiftApply.cb]
rw [← abs_mul_exp_arg_mul_I [V]cb]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]cb) * I + (↑b * I + ↑f * I) = ↑(arg [V]cb + (b + f)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma tb_eq_abs {V : CKMMatrix} (h1 : c + f = - arg [V]tb) :
[phaseShiftApply V a b c d e f]tb = VtbAbs ⟦V⟧ := by
rw [phaseShiftApply.tb]
rw [← abs_mul_exp_arg_mul_I [V]tb]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]tb) * I + (↑c * I + ↑f * I) = ↑(arg [V]tb + (c + f)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma cd_eq_neg_abs {V : CKMMatrix} (h1 : b + d = Real.pi - arg [V]cd) :
[phaseShiftApply V a b c d e f]cd = - VcdAbs ⟦V⟧ := by
rw [phaseShiftApply.cd]
rw [← abs_mul_exp_arg_mul_I [V]cd]
rw [mul_comm, mul_assoc, ← exp_add]
have h2 : ↑(arg [V]cd) * I + (↑b * I + ↑d * I) = ↑(arg [V]cd + (b + d)) * I := by
simp [add_assoc]
ring
rw [h2, h1]
simp
rfl
lemma t_eq_conj {V : CKMMatrix} {τ : } (hτ : cexp (τ * I) • (conj [V]u ×₃ conj [V]c) = [V]t)
(h1 : τ = - a - b - c - d - e - f) :
[phaseShiftApply V a b c d e f]t =
conj [phaseShiftApply V a b c d e f]u ×₃ conj [phaseShiftApply V a b c d e f]c := by
change _ = phaseShiftApply.ucCross _ _ _ _ _ _ _
funext i
fin_cases i
· simp
rw [phaseShiftApply.ucCross_fst]
simp [tRow, phaseShiftApply.td]
have hτ0 := congrFun hτ 0
simp [tRow] at hτ0
rw [← hτ0]
rw [← mul_assoc, ← exp_add, h1]
congr 2
simp
ring
· simp
rw [phaseShiftApply.ucCross_snd]
simp [tRow, phaseShiftApply.ts]
have hτ0 := congrFun hτ 1
simp [tRow] at hτ0
rw [← hτ0]
rw [← mul_assoc, ← exp_add, h1]
congr 2
simp
ring
· simp
rw [phaseShiftApply.ucCross_thd]
simp [tRow, phaseShiftApply.tb]
have hτ0 := congrFun hτ 2
simp [tRow] at hτ0
rw [← hτ0]
rw [← mul_assoc, ← exp_add, h1]
congr 2
simp
ring
-- bad name for this lemma
lemma all_cond_sol {V : CKMMatrix} (h1 : a + d = - arg [V]ud) (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 ∧
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
have hd : d = - arg [V]ud - a := by
linear_combination h1
subst hd
have he : e = - arg [V]us - a := by
linear_combination h2
subst he
simp_all
have hbf : b = - arg [V]cb - f := by
linear_combination h3
have hcf : c = - arg [V]tb - f := by
linear_combination h4
rw [hbf, hcf] at h5
simp_all
ring_nf at h5
have hf : f = τ - a - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb := by
linear_combination -(1 * h5)
rw [hf] at hbf hcf
ring_nf at hbf hcf
subst hf hbf hcf
ring_nf
simp
-- rename
def UCond₁ (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U]us = VusAbs ⟦U⟧
∧ [U]cb = VcbAbs ⟦U⟧ ∧ [U]tb = VtbAbs ⟦U⟧ ∧ [U]t = conj [U]u ×₃ conj [U]c
lemma all_eq_abs (V : CKMMatrix) :
∃ (U : CKMMatrix), V ≈ U ∧ UCond₁ U:= by
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]ud - arg [V]us - arg [V]cb - arg [V]tb)
have hUV : ⟦U⟧ = ⟦V⟧ := by
simp
symm
exact phaseShiftApply.equiv _ _ _ _ _ _ _
use U
apply And.intro
exact phaseShiftApply.equiv _ _ _ _ _ _ _
apply And.intro
rw [hUV]
apply ud_eq_abs _ _ _ _ _ _ _
ring
apply And.intro
rw [hUV]
apply us_eq_abs
ring
apply And.intro
rw [hUV]
apply cb_eq_abs
ring
apply And.intro
rw [hUV]
apply tb_eq_abs
ring
apply t_eq_conj _ _ _ _ _ _ hτ.symm
ring
end CKMMatrix
end

View file

@ -0,0 +1,655 @@
/-
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.Relations
import Mathlib.Analysis.SpecialFunctions.Complex.Arg
import HepLean.FlavorPhysics.CKMMatrix.PhaseFreedom
open Matrix Complex
noncomputable section
namespace CKMMatrix
open ComplexConjugate
-- A
def Rubud (V : CKMMatrix) : := [V]ub / [V]ud
scoped[CKMMatrix] notation (name := ub_ud_ratio) "[" V "]ub|ud" => Rubud V
-- B
def Rusud (V : CKMMatrix) : := [V]us / [V]ud
scoped[CKMMatrix] notation (name := us_ud_ratio) "[" V "]us|ud" => Rusud V
def Rudus (V : CKMMatrix) : := [V]ud / [V]us
scoped[CKMMatrix] notation (name := ud_us_ratio) "[" V "]ud|us" => Rudus V
def Rubus (V : CKMMatrix) : := [V]ub / [V]us
scoped[CKMMatrix] notation (name := ub_us_ratio) "[" V "]ub|us" => Rubus V
-- D
def Rcdcb (V : CKMMatrix) : := [V]cd / [V]cb
scoped[CKMMatrix] notation (name := cd_cb_ratio) "[" V "]cd|cb" => Rcdcb V
lemma Rcdcb_mul_cb {V : CKMMatrix} (h : [V]cb ≠ 0): [V]cd = Rcdcb V * [V]cb := by
rw [Rcdcb]
exact (div_mul_cancel₀ (V.1 1 0) h).symm
-- C'
def Rcscb (V : CKMMatrix) : := [V]cs / [V]cb
scoped[CKMMatrix] notation (name := cs_cb_ratio) "[" V "]cs|cb" => Rcscb V
lemma Rcscb_mul_cb {V : CKMMatrix} (h : [V]cb ≠ 0): [V]cs = Rcscb V * [V]cb := by
rw [Rcscb]
exact (div_mul_cancel₀ [V]cs h).symm
def Rtb!cbud (V : CKMMatrix) : := conj [V]tb / ([V]cb * [V]ud)
scoped[CKMMatrix] notation (name := tb_cb_ud_ratio) "[" V "]tb*|cb|ud" => Rtb!cbud V
def Rtb!cbus (V : CKMMatrix) : := conj [V]tb / ([V]cb * [V]us)
scoped[CKMMatrix] notation (name := tb_cb_us_ratio) "[" V "]tb*|cb|us" => Rtb!cbus V
lemma orthog_fst_snd_row_ratios {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 0) :
[V]cd|cb + [V]cs|cb * conj ([V]us|ud) + conj ([V]ub|ud) = 0 := by
have h1 : ([V]cd * conj ([V]ud) + [V]cs * conj ([V]us)
+ [V]cb * conj ([V]ub)) / ([V]cb * conj ([V]ud)) = 0 := by
rw [fst_row_snd_row V ]
simp only [Fin.isValue, zero_div]
rw [← div_add_div_same, ← div_add_div_same] at h1
rw [mul_div_mul_comm, mul_div_mul_comm, mul_div_mul_comm] at h1
rw [div_self, div_self] at h1
change Rcdcb V * 1 + _ + _ = _ at h1
have h2 : (starRingEnd ) (V.1 0 2) / (starRingEnd ) (V.1 0 0) = conj (Rubud V) := by
simp [Rubud]
have h3 : ((starRingEnd ) (V.1 0 1) / (starRingEnd ) (V.1 0 0)) = conj (Rusud V) := by
simp [Rusud]
rw [h2, h3] at h1
simp at h1
exact h1
exact ha
simpa using hb
lemma orthog_fst_snd_row_ratios_cb_us {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 0) :
[V]cd|cb * conj ([V]ud|us) + [V]cs|cb + conj ([V]ub|us) = 0 := by
have h1 : ([V]cd * conj ([V]ud) + [V]cs * conj ([V]us)
+ [V]cb * conj ([V]ub)) / ([V]cb * conj ([V]us)) = 0 := by
rw [fst_row_snd_row V ]
simp only [Fin.isValue, zero_div]
rw [← div_add_div_same, ← div_add_div_same] at h1
rw [mul_div_mul_comm, mul_div_mul_comm, mul_div_mul_comm] at h1
rw [div_self, div_self] at h1
change _ + _ * 1 + _ = _ at h1
have h2 : (starRingEnd ) (V.1 0 2) / (starRingEnd ) (V.1 0 1) = conj (Rubus V) := by
simp [Rubus]
have h3 : ((starRingEnd ) (V.1 0 0) / (starRingEnd ) (V.1 0 1)) = conj (Rudus V) := by
simp [Rudus]
rw [h2, h3] at h1
simp at h1
exact h1
exact ha
simpa using hb
def R' (V : CKMMatrix) : := [V]cs|cb * (1 + normSq [V]us|ud) + conj ([V]ub|ud) * [V]us|ud
def R's (V : CKMMatrix) : := [V]cd|cb * (normSq [V]ud|us + 1) + conj ([V]ub|us) * [V]ud|us
lemma R'_eq (V : CKMMatrix) :
R' V = [V]cs|cb * (1 + normSq ([V]us|ud)) + conj ([V]ub|ud) * [V]us|ud := by rfl
lemma R's_eq (V : CKMMatrix) :
R's V = [V]cd|cb * (normSq ([V]ud|us) + 1) + conj ([V]ub|us) * [V]ud|us := by rfl
lemma one_plus_normSq_Rusud_neq_zero_ (V : CKMMatrix):
1 + normSq ([V]us|ud) ≠ 0 := by
have h1 : 0 ≤ (normSq ([V]us|ud)) := normSq_nonneg ([V]us|ud)
have h2 : 0 < 1 + normSq ([V]us|ud) := by linarith
by_contra hn
have h3 := lt_of_lt_of_eq h2 hn
simp at h3
lemma normSq_Rudus_plus_one_neq_zero_ (V : CKMMatrix):
normSq ([V]ud|us) + 1 ≠ 0 := by
have h1 : 0 ≤ (normSq ([V]ud|us)) := normSq_nonneg ([V]ud|us)
have h2 : 0 < normSq ([V]ud|us) + 1 := by linarith
by_contra hn
have h3 := lt_of_lt_of_eq h2 hn
simp at h3
lemma one_plus_normSq_Rusud_neq_zero_ (V : CKMMatrix):
1 + (normSq ([V]us|ud) : ) ≠ 0 := by
by_contra hn
have h1 := one_plus_normSq_Rusud_neq_zero_ V
simp at h1
rw [← ofReal_inj] at h1
simp_all only [ofReal_add, ofReal_one, ofReal_zero, not_true_eq_false]
lemma normSq_Rudus_plus_one_neq_zero_ (V : CKMMatrix):
(normSq ([V]ud|us) : ) + 1 ≠ 0 := by
by_contra hn
have h1 := normSq_Rudus_plus_one_neq_zero_ V
simp at h1
rw [← ofReal_inj] at h1
simp_all only [ofReal_add, ofReal_one, ofReal_zero, not_true_eq_false]
lemma Rcscb_of_R' (V : CKMMatrix) :
[V]cs|cb = (R' V - conj [V]ub|ud * [V]us|ud) / (1 + normSq [V]us|ud) := by
have h2 : R' V - conj ([V]ub|ud) * [V]us|ud = [V]cs|cb * (1 + normSq [V]us|ud) := by
linear_combination R'_eq V
rw [h2]
rw [mul_div_cancel_right₀]
exact one_plus_normSq_Rusud_neq_zero_ V
lemma Rcdcb_of_R's (V : CKMMatrix) :
[V]cd|cb = (R's V - conj [V]ub|us * [V]ud|us) / (normSq [V]ud|us + 1) := by
have h2 : R's V - conj ([V]ub|us) * [V]ud|us = [V]cd|cb * (normSq [V]ud|us + 1) := by
linear_combination R's_eq V
rw [h2]
rw [mul_div_cancel_right₀]
exact normSq_Rudus_plus_one_neq_zero_ V
lemma Rcdcb_R'_orthog {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 0) :
[V]cd|cb + (conj [V]ub|ud + R' V * conj [V]us|ud ) / (1 + normSq [V]us|ud) = 0 := by
have h1 := orthog_fst_snd_row_ratios hb ha
rw [Rcscb_of_R'] at h1
rw [div_mul_eq_mul_div] at h1
rw [add_assoc] at h1
rw [div_add'] at h1
have h2 : (R' V - conj [V]ub|ud * [V]us|ud) * conj [V]us|ud +
conj [V]ub|ud * (1 + normSq [V]us|ud) = conj [V]ub|ud + R' V * conj [V]us|ud := by
rw [← mul_conj]
ring
rw [h2] at h1
exact h1
exact one_plus_normSq_Rusud_neq_zero_ V
lemma Rcdcb_of_R' {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 0) :
[V]cd|cb = - (conj [V]ub|ud + R' V * conj [V]us|ud ) / (1 + normSq [V]us|ud) := by
linear_combination (Rcdcb_R'_orthog hb ha)
lemma Rcscb_R's_orthog {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 0) :
[V]cs|cb + (conj [V]ub|us + R's V * conj [V]ud|us ) / (normSq [V]ud|us + 1) = 0 := by
have h1 := orthog_fst_snd_row_ratios_cb_us hb ha
rw [Rcdcb_of_R's] at h1
rw [div_mul_eq_mul_div] at h1
rw [add_assoc, add_comm [V]cs|cb, ← add_assoc] at h1
rw [div_add', add_comm] at h1
have h2 : (R's V - conj [V]ub|us * [V]ud|us) * conj [V]ud|us +
(starRingEnd ) [V]ub|us * ((normSq [V]ud|us) + 1)= conj [V]ub|us + R's V * conj [V]ud|us := by
rw [← mul_conj]
ring
rw [h2] at h1
exact h1
exact normSq_Rudus_plus_one_neq_zero_ V
lemma Rcscb_of_R's {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 0) :
[V]cs|cb = - (conj [V]ub|us + R's V * conj [V]ud|us ) / (normSq [V]ud|us + 1) := by
linear_combination (Rcscb_R's_orthog hb ha)
lemma R'_of_Rcscb_Rcdcb {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 0) :
R' V = [V]cs|cb - [V]us|ud * [V]cd|cb := by
rw [Rcdcb_of_R' hb ha, Rcscb_of_R']
have h1 := one_plus_normSq_Rusud_neq_zero_ V
have h2 : conj [V]ud ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
rw [Rubud, map_div₀, Rusud, map_div₀, map_div₀]
simp
rw [normSq_eq_conj_mul_self, normSq_eq_conj_mul_self]
field_simp
ring
lemma R's_of_Rcscb_Rcdcb {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 0) :
R's V = [V]cd|cb - [V]ud|us * [V]cs|cb := by
rw [Rcscb_of_R's hb ha, Rcdcb_of_R's]
have h1 := normSq_Rudus_plus_one_neq_zero_ V
have h2 : conj [V]us ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
rw [Rubus, map_div₀, Rudus, map_div₀, map_div₀]
simp
rw [normSq_eq_conj_mul_self, normSq_eq_conj_mul_self]
field_simp
ring
lemma R'_of_Rtb!cbud {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 0)
{τ : } (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) :
R' V = cexp (τ * I) * [V]tb*|cb|ud := by
have h1 : cexp (- τ * I) = conj (cexp (τ * I)) := by
rw [← exp_conj]
simp only [neg_mul, _root_.map_mul, conj_I, mul_neg]
rw [conj_ofReal]
have h2 : [V]tb*|cb|ud = cexp (- τ * I) * R' V := by
rw [h1, R'_of_Rcscb_Rcdcb hb ha]
have h1 := congrFun hτ 2
simp [crossProduct, tRow, uRow, cRow] at h1
apply congrArg conj at h1
simp at h1
rw [Rtb!cbud, Rcscb, Rusud, Rcdcb, h1]
field_simp
ring
rw [h2, ← mul_assoc, ← exp_add]
simp
lemma R's_of_Rtb!cbus {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 0)
{τ : } (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) :
R's V = - cexp (τ * I) * [V]tb*|cb|us := by
have h1 : cexp (- τ * I) = conj (cexp (τ * I)) := by
rw [← exp_conj]
simp only [neg_mul, _root_.map_mul, conj_I, mul_neg]
rw [conj_ofReal]
have h2 : [V]tb*|cb|us = - cexp (- τ * I) * R's V := by
rw [h1, R's_of_Rcscb_Rcdcb hb ha]
have h1 := congrFun hτ 2
simp [crossProduct, tRow, uRow, cRow] at h1
apply congrArg conj at h1
simp at h1
rw [Rtb!cbus, Rcscb, Rudus, Rcdcb, h1]
field_simp
ring
rw [h2]
simp
rw [← mul_assoc, ← exp_add]
simp
lemma over_normSq_Rusud (V : CKMMatrix) (hb : [V]ud ≠ 0) (a : ) : a / (1 + normSq [V]us|ud) =
(a * normSq [V]ud) / (normSq [V]ud + normSq [V]us) := by
rw [Rusud]
field_simp
rw [one_add_div]
field_simp
simp
exact hb
lemma over_normSq_Rudus (V : CKMMatrix) (hb : [V]us ≠ 0) (a : ) : a / (normSq [V]ud|us + 1) =
(a * normSq [V]us) / (normSq [V]ud + normSq [V]us) := by
rw [Rudus]
field_simp
rw [div_add_one]
field_simp
simp
exact hb
lemma cd_of_ud_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 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) /
(normSq [V]ud + normSq [V]us) := by
obtain hτ2 := R'_of_Rtb!cbud hb ha hτ
rw [Rcdcb_mul_cb ha]
rw [Rcdcb_of_R' hb ha, hτ2]
rw [over_normSq_Rusud V hb, div_mul_eq_mul_div]
congr 1
rw [normSq_eq_conj_mul_self, Rubud, map_div₀, Rusud, map_div₀, Rtb!cbud]
have h1 : conj [V]ud ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
ring
lemma cs_of_ud_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0) (ha : [V]cb ≠ 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
have hτ2 := R'_of_Rtb!cbud hb ha hτ
rw [Rcscb_mul_cb ha]
rw [Rcscb_of_R', hτ2]
rw [over_normSq_Rusud V hb, div_mul_eq_mul_div]
congr 1
rw [normSq_eq_conj_mul_self, Rusud, Rtb!cbud, Rubud, map_div₀]
have h1 : conj [V]ud ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
ring
lemma cd_of_us_neq_zero {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 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) /
(normSq [V]ud + normSq [V]us) := by
have hτ2 := R's_of_Rtb!cbus hb ha hτ
rw [Rcdcb_mul_cb ha]
rw [Rcdcb_of_R's, hτ2]
rw [over_normSq_Rudus V hb, div_mul_eq_mul_div]
congr 1
rw [normSq_eq_conj_mul_self, Rubus, map_div₀, Rudus, Rtb!cbus]
have h1 : conj [V]us ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
ring
lemma cs_of_us_neq_zero {V : CKMMatrix} (hb : [V]us ≠ 0) (ha : [V]cb ≠ 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
have hτ2 := R's_of_Rtb!cbus hb ha hτ
rw [Rcscb_mul_cb ha]
rw [Rcscb_of_R's hb ha, hτ2]
rw [over_normSq_Rudus V hb, div_mul_eq_mul_div]
congr 1
rw [normSq_eq_conj_mul_self, Rudus, map_div₀, Rtb!cbus, Rubus, map_div₀]
have h1 : conj [V]us ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr hb
field_simp
ring
lemma cd_of_us_or_ud_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb ≠ 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) /
(normSq [V]ud + normSq [V]us) := by
cases' hb with hb hb
exact cd_of_ud_neq_zero hb ha hτ
exact cd_of_us_neq_zero hb ha hτ
lemma cs_of_us_or_ud_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb ≠ 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
cases' hb with hb hb
exact cs_of_ud_neq_zero hb ha hτ
exact cs_of_us_neq_zero hb ha hτ
lemma cd_of_us_or_ud_neq_zero_UCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb ≠ 0)
(hV : UCond₁ V) : [V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) +
(- VubAbs ⟦V⟧ * VudAbs ⟦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
exact hV.2.2.2.2
rw [cd_of_us_or_ud_neq_zero hb ha hτ]
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
simp [sq, conj_ofReal]
have hx := Vabs_sq_add_neq_zero hb
field_simp
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I [V]ub]
rw [@RingHom.map_mul]
simp [conj_ofReal, ← exp_conj, VAbs]
rw [h1]
ring_nf
lemma cs_of_us_or_ud_neq_zero_UCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb ≠ 0)
(hV : UCond₁ 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)) * cexp (- arg [V]ub * I)
:= by
have hτ : [V]t = cexp ((0 : ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
simp
exact hV.2.2.2.2
rw [cs_of_us_or_ud_neq_zero hb ha hτ]
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
simp [sq, conj_ofReal]
have hx := Vabs_sq_add_neq_zero hb
field_simp
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I [V]ub]
rw [@RingHom.map_mul]
simp [conj_ofReal, ← exp_conj, VAbs]
rw [h1]
ring_nf
lemma cd_of_cb_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb = 0)
{κ : } (hκ : [V]u = cexp (κ * I) • (conj [V]c ×₃ conj [V]t)) :
[V]cd = - cexp (κ * I) * conj [V]tb * conj [V]us / (normSq [V]ud + normSq [V]us) := by
have h2 : [V]cd = - cexp (κ * I) * conj [V]us / [V]tb := by
have h1 := congrFun hκ 1
simp [crossProduct, tRow, uRow, cRow] at h1
rw [ha] at h1
apply congrArg conj at h1
simp at h1
rw [h1]
have hx := tb_neq_zero_of_cb_zero_ud_us_neq_zero ha hb
field_simp
have h1 : conj (cexp (↑κ * I)) = cexp (- κ * I) := by
simp [← exp_conj, conj_ofReal]
rw [h1]
rw [← mul_assoc]
rw [← exp_add]
simp
rw [div_td_of_cb_zero_ud_us_neq_zero ha hb] at h2
rw [h2]
congr 1
ring
lemma cs_of_cb_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb = 0)
{κ : } (hκ1 : [V]u = cexp (κ * I) • (conj [V]c ×₃ conj [V]t)) :
[V]cs = cexp (κ * I) * conj [V]tb * conj [V]ud / (normSq [V]ud + normSq [V]us) := by
have h2 : [V]cs = cexp (κ * I) * conj [V]ud / [V]tb := by
have h1 := congrFun hκ1 0
simp [crossProduct, tRow, uRow, cRow] at h1
rw [ha] at h1
apply congrArg conj at h1
simp at h1
rw [h1]
have hx := tb_neq_zero_of_cb_zero_ud_us_neq_zero ha hb
field_simp
have h1 : conj (cexp (↑κ * I)) = cexp (- κ * I) := by
rw [← exp_conj]
simp only [neg_mul, _root_.map_mul, conj_I, mul_neg]
rw [conj_ofReal]
rw [h1]
rw [← mul_assoc]
rw [← exp_add]
simp
rw [div_td_of_cb_zero_ud_us_neq_zero ha hb] at h2
rw [h2]
congr 1
ring
lemma cd_of_cb_zero_UCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb = 0)
(hV : UCond₁ V) {κ : } (hκ1 : [V]u = cexp (κ * I) • (conj [V]c ×₃ conj [V]t)) :
[V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) * cexp (κ * I) := by
rw [cd_of_cb_zero hb ha hκ1]
rw [hV.1, hV.2.1, hV.2.2.2.1]
simp [sq, conj_ofReal]
have hx := Vabs_sq_add_neq_zero hb
field_simp
ring_nf
simp
lemma cs_of_cb_zero_UCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb = 0)
(hV : UCond₁ V) {κ : } (hκ1 : [V]u = cexp (κ * I) • (conj [V]c ×₃ conj [V]t)) :
[V]cs = (VtbAbs ⟦V⟧ * VudAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) * cexp (κ * I) := by
rw [cs_of_cb_zero hb ha hκ1]
rw [hV.1, hV.2.1, hV.2.2.2.1]
simp [sq, conj_ofReal]
have hx := Vabs_sq_add_neq_zero hb
field_simp
ring_nf
def UCond₂ (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U]us = VusAbs ⟦U⟧
∧ [U]ub = VubAbs ⟦U⟧ ∧ [U]t = conj [U]u ×₃ conj [U]c
∧ [U]cd = (- VtbAbs ⟦U⟧ * VusAbs ⟦U⟧ / (VudAbs ⟦U⟧ ^2 + VusAbs ⟦U⟧ ^2)) ∧
[U]cs = (VtbAbs ⟦U⟧ * VudAbs ⟦U⟧ / (VudAbs ⟦U⟧ ^2 + VusAbs ⟦U⟧ ^2))
lemma UCond₂_solv {V : CKMMatrix} (h1 : a + d = 0) (h2 : a + e = 0) (h3 : a + f = - arg [V]ub)
(h4 : 0 = - a - b - c - d - e - f) (h5 : b + d = - κ) (h6 : b + e = - κ) :
b = - κ + a ∧
c = κ + arg [V]ub + a ∧
d = - a ∧
e = - a ∧
f = - arg [V]ub - a := by
have hd : d = - a := by
linear_combination h1
subst hd
have he : e = - a := by
linear_combination h2
subst he
simp_all
have hb : b = - κ + a := by
linear_combination h6
subst hb
simp_all
have hf : f = - arg [V]ub - a := by
linear_combination h3
subst hf
simp_all
ring_nf at h4
have hc : c = κ + arg [V]ub + a := by
linear_combination h4
simp_all
lemma UCond₂_exists {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) (ha : [V]cb = 0)
(hV : UCond₁ V) : ∃ (U : CKMMatrix), V ≈ U ∧ UCond₂ U:= by
obtain ⟨κ, hκ⟩ := V.cRow_cross_tRow_eq_uRow
let U : CKMMatrix := phaseShiftApply V 0 (- κ) (κ + arg [V]ub) 0 0 (- arg [V]ub)
use U
have hUV : ⟦U⟧ = ⟦V⟧ := by
simp
symm
exact phaseShiftApply.equiv _ _ _ _ _ _ _
apply And.intro
exact phaseShiftApply.equiv _ _ _ _ _ _ _
apply And.intro
· rw [hUV]
apply ud_eq_abs _ _ _ _ _ _ _
rw [hV.1, arg_ofReal_of_nonneg]
simp
exact Complex.abs.nonneg _
apply And.intro
· rw [hUV]
apply us_eq_abs _ _ _ _ _ _ _
rw [hV.2.1, arg_ofReal_of_nonneg]
simp
exact Complex.abs.nonneg _
apply And.intro
· rw [hUV]
apply ub_eq_abs _ _ _ _ _ _ _
ring
apply And.intro
· have hτ : [V]t = cexp ((0 : ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
simp
exact hV.2.2.2.2
apply t_eq_conj _ _ _ _ _ _ hτ.symm
ring
apply And.intro
· rw [phaseShiftApply.cd]
rw [cd_of_cb_zero_UCond hb ha hV hκ]
rw [mul_comm, mul_assoc, ← exp_add, hUV]
simp
· rw [phaseShiftApply.cs]
rw [cs_of_cb_zero_UCond hb ha hV hκ]
rw [mul_comm, mul_assoc, ← exp_add, hUV]
simp
section zero
def UCond₃ (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 UCond₃_solv {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
have hf : f = - arg [V]ub - a := by
linear_combination h1
subst hf
have he : e = - arg [V]cs - b := by
linear_combination h5
have hd : d = Real.pi - arg [V]cd - b := by
linear_combination h3
subst he hd
simp_all
ring_nf at h2
have hc : c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b := by
linear_combination h2
subst hc
ring
lemma cs_of_ud_us_zero {V : CKMMatrix} (ha : ¬ ([V]ud ≠ 0 [V]us ≠ 0)) :
VcsAbs ⟦V⟧ = √(1 - VcdAbs ⟦V⟧ ^ 2) := by
have h1 := snd_row_normalized_abs V
symm
rw [Real.sqrt_eq_iff_sq_eq]
simp [not_or] at ha
rw [cb_eq_zero_of_ud_us_zero ha] at h1
simp at h1
simp [VAbs]
linear_combination h1
simp
rw [@abs_le]
have h1 := VAbs_leq_one 1 0 ⟦V⟧
have h0 := VAbs_ge_zero 1 0 ⟦V⟧
simp_all
have hn : -1 ≤ (0 : ) := by simp
exact hn.trans h0
exact VAbs_ge_zero _ _ ⟦V⟧
lemma UCond₃_exists {V : CKMMatrix} (hb :¬ ([V]ud ≠ 0 [V]us ≠ 0)) (hV : UCond₁ V) :
∃ (U : CKMMatrix), V ≈ U ∧ UCond₃ 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 )
use U
have hUV : ⟦U⟧ = ⟦V⟧ := by
simp
symm
exact phaseShiftApply.equiv _ _ _ _ _ _ _
apply And.intro
exact phaseShiftApply.equiv _ _ _ _ _ _ _
apply And.intro
· simp [not_or] at hb
have h1 : VudAbs ⟦U⟧ = 0 := by
rw [hUV]
simp [VAbs, hb]
simp [VAbs] at h1
exact h1
apply And.intro
· simp [not_or] at hb
have h1 : VusAbs ⟦U⟧ = 0 := by
rw [hUV]
simp [VAbs, hb]
simp [VAbs] at h1
exact h1
apply And.intro
· simp [not_or] at hb
have h3 := cb_eq_zero_of_ud_us_zero hb
have h1 : VcbAbs ⟦U⟧ = 0 := by
rw [hUV]
simp [VAbs, h3]
simp [VAbs] at h1
exact h1
apply And.intro
· have hU1 : [U]ub = VubAbs ⟦V⟧ := by
apply ub_eq_abs _ _ _ _ _ _ _
ring
rw [hU1]
have h1:= (ud_us_neq_zero_iff_ub_neq_one V).mpr.mt hb
simpa using h1
apply And.intro
· have hτ : [V]t = cexp ((0 : ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
simp
exact hV.2.2.2.2
apply t_eq_conj _ _ _ _ _ _ hτ.symm
ring
apply And.intro
· rw [hUV]
apply cd_eq_neg_abs _ _ _ _ _ _ _
ring
have hcs : [U]cs = VcsAbs ⟦U⟧ := by
rw [hUV]
apply cs_eq_abs _ _ _ _ _ _ _
ring
rw [hcs, hUV, cs_of_ud_us_zero hb]
end zero
end CKMMatrix
end

View file

@ -0,0 +1,174 @@
/-
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 Mathlib.Analysis.SpecialFunctions.Complex.Arg
open Matrix Complex
noncomputable section
namespace CKMMatrix
open ComplexConjugate
lemma fst_row_normalized_abs (V : CKMMatrix) :
abs [V]ud ^ 2 + abs [V]us ^ 2 + abs [V]ub ^ 2 = 1 := by
have h1 := VAbs_sum_sq_row_eq_one ⟦V⟧ 0
simp [VAbs] at h1
exact h1
lemma fst_row_normalized_normSq (V : CKMMatrix) :
normSq [V]ud + normSq [V]us + normSq [V]ub = 1 := by
have h1 := V.fst_row_normalized_abs
repeat rw [Complex.sq_abs] at h1
exact h1
lemma snd_row_normalized_abs (V : CKMMatrix) :
abs [V]cd ^ 2 + abs [V]cs ^ 2 + abs [V]cb ^ 2 = 1 := by
have h1 := VAbs_sum_sq_row_eq_one ⟦V⟧ 1
simp [VAbs] at h1
exact h1
lemma snd_row_normalized_normSq (V : CKMMatrix) :
normSq [V]cd + normSq [V]cs + normSq [V]cb = 1 := by
have h1 := V.snd_row_normalized_abs
repeat rw [Complex.sq_abs] at h1
exact h1
lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
[V]ud ≠ 0 [V]us ≠ 0 ↔ abs [V]ub ≠ 1 := by
have h2 := V.fst_row_normalized_abs
apply Iff.intro
intro h
intro h1
rw [h1] at h2
simp at h2
rw [add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)] at h2
simp_all
intro h
by_contra hn
rw [not_or] at hn
simp at hn
simp_all
have h1 := Complex.abs.nonneg [V]ub
rw [h2] at h1
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
exact h2 h1
lemma thd_col_normalized_abs (V : CKMMatrix) :
abs [V]ub ^ 2 + abs [V]cb ^ 2 + abs [V]tb ^ 2 = 1 := by
have h1 := VAbs_sum_sq_col_eq_one ⟦V⟧ 2
simp [VAbs] at h1
exact h1
lemma thd_col_normalized_normSq (V : CKMMatrix) :
normSq [V]ub + normSq [V]cb + normSq [V]tb = 1 := by
have h1 := V.thd_col_normalized_abs
repeat rw [Complex.sq_abs] at h1
exact h1
lemma cb_tb_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
[V]cb ≠ 0 [V]tb ≠ 0 ↔ abs [V]ub ≠ 1 := by
have h2 := V.thd_col_normalized_abs
apply Iff.intro
intro h
intro h1
rw [h1] at h2
simp at h2
have h2 : Complex.abs (V.1 1 2) ^ 2 + Complex.abs (V.1 2 2) ^ 2 = 0 := by
linear_combination h2
rw [add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)] at h2
simp_all
intro h
by_contra hn
rw [not_or] at hn
simp at hn
simp_all
have h1 := Complex.abs.nonneg [V]ub
rw [h2] at h1
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
exact h2 h1
lemma tb_neq_zero_of_cb_zero_ud_us_neq_zero {V : CKMMatrix} (h : [V]cb = 0)
(h1 : [V]ud ≠ 0 [V]us ≠ 0) : [V]tb ≠ 0 := by
rw [ud_us_neq_zero_iff_ub_neq_one] at h1
rw [← cb_tb_neq_zero_iff_ub_neq_one] at h1
simp_all
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
simp [h] at h1
rw [add_assoc] at h1
simp at h1
rw [add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)] at h1
simp at h1
exact h1.1
lemma normSq_tb_of_cb_zero {V : CKMMatrix} (h : [V]cb = 0) : normSq [V]tb =
normSq [V]ud + normSq [V]us := by
have h1 := fst_row_normalized_normSq V
rw [← thd_col_normalized_normSq V] at h1
rw [h] at h1
simp at h1
linear_combination -(1 * h1)
lemma div_td_of_cb_zero_ud_us_neq_zero {V : CKMMatrix} (h : [V]cb = 0)
(h1 : [V]ud ≠ 0 [V]us ≠ 0) (a : ) : a / [V]tb =
(a * conj [V]tb) / (normSq [V]ud + normSq [V]us) := by
have h2 := tb_neq_zero_of_cb_zero_ud_us_neq_zero h h1
have h3 : conj [V]tb ≠ 0 := (AddEquivClass.map_ne_zero_iff starRingAut).mpr h2
have h1 : a / [V]tb = a * conj [V]tb / normSq [V]tb := by
rw [normSq_eq_conj_mul_self]
field_simp
ring
rw [normSq_tb_of_cb_zero h] at h1
simp at h1
exact h1
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 normSq_Vud_plus_normSq_Vus_neq_zero_ {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) :
normSq [V]ud + normSq [V]us ≠ 0 := by
rw [normSq_Vud_plus_normSq_Vus V]
rw [ud_us_neq_zero_iff_ub_neq_one] at hb
by_contra hn
rw [← Complex.sq_abs] at hn
have h2 : Complex.abs (V.1 0 2) ^2 = 1 := by
linear_combination -(1 * hn)
simp at h2
cases' h2 with h2 h2
exact hb h2
have h3 := Complex.abs.nonneg [V]ub
rw [h2] at h3
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
exact h2 h3
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
rw [← ofReal_inj] at h1
simp_all
lemma Vabs_sq_add_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) :
((VAbs 0 0 ⟦V⟧ : ) * ↑(VAbs 0 0 ⟦V⟧) + ↑(VAbs 0 1 ⟦V⟧) * ↑(VAbs 0 1 ⟦V⟧)) ≠ 0 := by
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ hb
rw [← Complex.sq_abs, ← Complex.sq_abs] at h1
simp [sq] at h1
exact h1
end CKMMatrix
end

View file

@ -0,0 +1,396 @@
/-
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.SpecialFunctions.Complex.Arg
import Mathlib.LinearAlgebra.CrossProduct
open Matrix Complex
open ComplexConjugate
noncomputable section
namespace CKMMatrix
def uRow (V : CKMMatrix) : Fin 3 → := ![[V]ud, [V]us, [V]ub]
scoped[CKMMatrix] notation (name := u_row) "[" V "]u" => uRow V
def cRow (V : CKMMatrix) : Fin 3 → := ![[V]cd, [V]cs, [V]cb]
scoped[CKMMatrix] notation (name := c_row) "[" V "]c" => cRow V
def tRow (V : CKMMatrix) : Fin 3 → := ![[V]td, [V]ts, [V]tb]
scoped[CKMMatrix] notation (name := t_row) "[" V "]t" => tRow V
lemma uRow_normalized (V : CKMMatrix) : conj [V]u ⬝ᵥ [V]u = 1 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 0) 0
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 0 0) _, mul_comm (V.1 0 1) _, mul_comm (V.1 0 2) _] at ht
exact ht
lemma cRow_normalized (V : CKMMatrix) : conj [V]c ⬝ᵥ [V]c = 1 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 1) 1
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 1 0) _, mul_comm (V.1 1 1) _, mul_comm (V.1 1 2) _] at ht
exact ht
lemma uRow_cRow_orthog (V : CKMMatrix) : conj [V]u ⬝ᵥ [V]c = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 1) 0
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
exact ht
lemma uRow_tRow_orthog (V : CKMMatrix) : conj [V]u ⬝ᵥ [V]t = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 2) 0
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
exact ht
lemma cRow_uRow_orthog (V : CKMMatrix) : conj [V]c ⬝ᵥ [V]u = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 0) 1
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
exact ht
lemma cRow_tRow_orthog (V : CKMMatrix) : conj [V]c ⬝ᵥ [V]t = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 2) 1
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
exact ht
lemma tRow_normalized (V : CKMMatrix) : conj [V]t ⬝ᵥ [V]t = 1 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 2) 2
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 2 0) _, mul_comm (V.1 2 1) _, mul_comm (V.1 2 2) _] at ht
exact ht
lemma tRow_uRow_orthog (V : CKMMatrix) : conj [V]t ⬝ᵥ [V]u = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 0) 2
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht
exact ht
lemma tRow_cRow_orthog (V : CKMMatrix) : conj [V]t ⬝ᵥ [V]c = 0 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 1) 2
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
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
simp [crossProduct]
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
simp [crossProduct]
funext i
fin_cases i <;> simp
lemma uRow_cross_cRow_normalized (V : CKMMatrix) :
conj (conj [V]u ×₃ conj [V]c) ⬝ᵥ (conj [V]u ×₃ conj [V]c) = 1 := by
rw [uRow_cross_cRow_conj, cross_dot_cross]
rw [dotProduct_comm, uRow_normalized, dotProduct_comm, cRow_normalized]
rw [dotProduct_comm, cRow_uRow_orthog, dotProduct_comm, uRow_cRow_orthog]
simp
lemma cRow_cross_tRow_normalized (V : CKMMatrix) :
conj (conj [V]c ×₃ conj [V]t) ⬝ᵥ (conj [V]c ×₃ conj [V]t) = 1 := by
rw [cRow_cross_tRow_conj, cross_dot_cross]
rw [dotProduct_comm, cRow_normalized, dotProduct_comm, tRow_normalized]
rw [dotProduct_comm, tRow_cRow_orthog, dotProduct_comm, cRow_tRow_orthog]
simp
@[simp]
def rows (V : CKMMatrix) : Fin 3 → Fin 3 → := fun i =>
match i with
| 0 => uRow V
| 1 => cRow V
| 2 => tRow V
def rowsLinearlyIndependent (V : CKMMatrix) : LinearIndependent (rows V) := by
apply Fintype.linearIndependent_iff.mpr
intro g hg
rw [Fin.sum_univ_three] 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 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
rw [cRow_normalized, cRow_uRow_orthog, cRow_tRow_orthog] at h1
rw [tRow_normalized, tRow_uRow_orthog, tRow_cRow_orthog] at h2
simp at h0 h1 h2
intro i
fin_cases i
exact h0
exact h1
exact h2
lemma rows_card : Fintype.card (Fin 3) = FiniteDimensional.finrank (Fin 3 → ) := by
simp only [Fintype.card_fin, FiniteDimensional.finrank_fintype_fun_eq_card]
@[simps!]
noncomputable def rowBasis (V : CKMMatrix) : Basis (Fin 3) (Fin 3 → ) :=
basisOfLinearIndependentOfCardEqFinrank (rowsLinearlyIndependent V) rows_card
lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
∃ (κ : ), [V]u = cexp (κ * I) • (conj [V]c ×₃ conj [V]t) := by
obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span (rowBasis V)
(conj [V]c ×₃ conj [V]t))
rw [Fin.sum_univ_three, rowBasis] at hg
simp at 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 at h0 h1
rw [h0, h1] at hg
simp at 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 [cRow_cross_tRow_normalized] at h2
have h3 : conj (g 0 • [V]u) = conj (g 0) • conj [V]u := by
funext i
fin_cases i <;> simp
rw [h3] at h2
simp only [Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul] at h2
rw [uRow_normalized] at h2
simp at h2
rw [mul_conj] at h2
rw [← Complex.sq_abs] at h2
simp at h2
cases' h2 with h2 h2
swap
have hx : Complex.abs (g 0) = -1 := by
rw [← ofReal_inj]
simp
exact h2
have h3 := Complex.abs.nonneg (g 0)
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
rw [← hg]
rw [@smul_smul]
rw [inv_mul_cancel]
simp
by_contra hn
rw [hn] at h2
simp at h2
have hg2 : Complex.abs (g 0)⁻¹ = 1 := by
rw [@map_inv₀, h2]
simp
have hg22 : ∃ (τ : ), (g 0)⁻¹ = Complex.exp (τ * I) := by
rw [← abs_mul_exp_arg_mul_I (g 0)⁻¹]
rw [hg2]
use arg (g 0)⁻¹
simp
obtain ⟨τ, hτ⟩ := hg22
use τ
rw [hx, hτ]
lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
∃ (τ : ), [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
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 at h0 h1
rw [h0, h1] at hg
simp at 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
funext i
fin_cases i <;> simp
rw [h3] at h2
simp only [Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul] at h2
rw [tRow_normalized] at h2
simp at h2
rw [mul_conj] at h2
rw [← Complex.sq_abs] at h2
simp at h2
cases' h2 with h2 h2
swap
have hx : Complex.abs (g 2) = -1 := by
rw [← ofReal_inj]
simp
exact h2
have h3 := Complex.abs.nonneg (g 2)
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
rw [← hg]
rw [@smul_smul]
rw [inv_mul_cancel]
simp
by_contra hn
rw [hn] at h2
simp at h2
have hg2 : Complex.abs (g 2)⁻¹ = 1 := by
rw [@map_inv₀, h2]
simp
have hg22 : ∃ (τ : ), (g 2)⁻¹ = Complex.exp (τ * I) := by
rw [← abs_mul_exp_arg_mul_I (g 2)⁻¹]
rw [hg2]
use arg (g 2)⁻¹
simp
obtain ⟨τ, hτ⟩ := hg22
use τ
rw [hx, hτ]
def uRow₁₂ (V : CKMMatrix) : Fin 2 → := ![[V]ud, [V]us]
def cRow₁₂ (V : CKMMatrix) : Fin 2 → := ![[V]cd, [V]cs]
scoped[CKMMatrix] notation (name := c₁₂_row) "[" V "]c₁₂" => cRow₁₂ V
lemma cRow₁₂_normalized_of_cb_zero {V : CKMMatrix} (hcb : [V]cb = 0) :
conj [V]c₁₂ ⬝ᵥ [V]c₁₂ = 1 := by
simp
have hV := V.prop
rw [mem_unitaryGroup_iff] at hV
have ht := congrFun (congrFun hV 1) 1
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
rw [hcb] at ht
rw [mul_comm (V.1 1 0) _, mul_comm (V.1 1 1) _] at ht
simp at ht
exact ht
lemma ext_Rows {U V : CKMMatrix} (hu : [U]u = [V]u) (hc : [U]c = [V]c) (ht : [U]t = [V]t) :
U = V := by
apply CKMMatrix_ext
funext i j
fin_cases i
have h1 := congrFun hu j
fin_cases j <;> simpa using h1
have h1 := congrFun hc j
fin_cases j <;> simpa using h1
have h1 := congrFun ht j
fin_cases j <;> simpa using h1
end CKMMatrix
namespace phaseShiftApply
open CKMMatrix
variable (V : CKMMatrix) (a b c d e f : )
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
simp [ucCross, crossProduct]
simp [uRow, us, ub, cRow, cs, cb]
rw [exp_add, exp_add]
simp [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
simp [ucCross, crossProduct]
simp [uRow, us, ub, cRow, cs, cb, ud, cd]
rw [exp_add, exp_add]
simp [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
simp [ucCross, crossProduct]
simp [uRow, us, ub, cRow, cs, cb, ud, cd]
rw [exp_add, exp_add]
simp [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
funext i
simp
fin_cases i <;>
change (phaseShiftApply V a b c 0 0 0).1 0 _ = _
rw [ud, uRow]
simp
rw [us, uRow]
simp
rw [ub, uRow]
simp
lemma cRow_mul (V : CKMMatrix) (a b c : ) :
[phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by
funext i
simp
fin_cases i <;>
change (phaseShiftApply V a b c 0 0 0).1 1 _ = _
rw [cd, cRow]
simp
rw [cs, cRow]
simp
rw [cb, cRow]
simp
lemma tRow_mul (V : CKMMatrix) (a b c : ) :
[phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by
funext i
simp
fin_cases i <;>
change (phaseShiftApply V a b c 0 0 0).1 2 _ = _
rw [td, tRow]
simp
rw [ts, tRow]
simp
rw [tb, tRow]
simp
end phaseShiftApply
end

View file

@ -4,9 +4,13 @@ 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 Mathlib.Analysis.SpecialFunctions.Complex.Arg
open Matrix Complex
open ComplexConjugate
noncomputable section
@ -16,23 +20,83 @@ def S₁₃ (V : Quotient CKMMatrixSetoid) : := VubAbs V
lemma VubAbs_eq_S₁₃ (V : Quotient CKMMatrixSetoid) : VubAbs V = S₁₃ V := rfl
def C₁₃ (V : Quotient CKMMatrixSetoid) : := √ (1 - S₁₃ V ^ 2)
def θ₁₃ (V : Quotient CKMMatrixSetoid) : := Real.arcsin (S₁₃ V)
lemma S₁₃_eq_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁₃ V) = S₁₃ V := by
refine Real.sin_arcsin ?_ ?_
have h1 := VAbs_ge_zero 0 2 V
rw [← VubAbs_eq_S₁₃]
simp
linarith
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
def C₁₃ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₁₃ V)
lemma C₁₃_eq_cos_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₁₃ V) = C₁₃ V := by
simp [C₁₃]
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)
lemma C₁₃_of_Vub_eq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : C₁₃ V = 0 := by
rw [C₁₃, ← VubAbs_eq_S₁₃, ha]
rw [C₁₃, θ₁₃, Real.cos_arcsin, ← VubAbs_eq_S₁₃, ha]
simp
lemma C₁₃_eq_add_sq (V : Quotient CKMMatrixSetoid) : C₁₃ V = √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by
rw [C₁₃, S₁₃]
rw [C₁₃, θ₁₃, Real.cos_arcsin, S₁₃]
have h1 : 1 - VubAbs V ^ 2 = VudAbs V ^ 2 + VusAbs V ^ 2 := by
linear_combination - (VAbs_sum_sq_row_eq_one V 0)
rw [h1]
/-- If VubAbs V = 1 this will give zero.-/
def S₁₂ (V : Quotient CKMMatrixSetoid) : := VusAbs V / (√ (VudAbs V ^ 2 + VusAbs V ^ 2))
lemma S₁₂_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₁₂ V = 0 := by
def θ₁₂ (V : Quotient CKMMatrixSetoid) : := Real.arcsin (S₁₂ V)
lemma S₁₂_nonneg (V : Quotient CKMMatrixSetoid) : 0 ≤ S₁₂ V := by
rw [S₁₂]
rw [@div_nonneg_iff]
apply Or.inl
apply And.intro
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]
by_cases h1 : √(VudAbs V ^ 2 + VusAbs V ^ 2) = 0
simp [h1]
have h2 := Real.sqrt_nonneg (VudAbs V ^ 2 + VusAbs V ^ 2)
rw [le_iff_eq_or_lt] at h2
have h3 : 0 < √(VudAbs V ^ 2 + VusAbs V ^ 2) := by
cases' h2 with h2 h2
simp_all
exact h2
apply Or.inl
simp_all
rw [Real.le_sqrt]
simp
exact sq_nonneg (VAbs 0 0 V)
exact VAbs_ge_zero 0 1 V
exact le_of_lt h3
lemma S₁₂_eq_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁₂ V) = S₁₂ V :=
Real.sin_arcsin (le_trans (by simp) (S₁₂_nonneg V)) (S₁₂_leq_one V)
lemma S₁₂_eq_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.sin (θ₁₂ V) = S₁₂ V := by
rw [← S₁₂_eq_sin_θ₁₂]
simp
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
linear_combination - (VAbs_sum_sq_row_eq_one V 0)
@ -40,12 +104,22 @@ lemma S₁₂_Vub_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₁
rw [ha]
simp
def C₁₂ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₁₂ V)
def C₁₂ (V : Quotient CKMMatrixSetoid) : := √ (1 - S₁₂ V ^ 2)
lemma C₁₂_eq_cos_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₁₂ V) = C₁₂ V := by
simp [C₁₂]
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
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)
lemma C₁₂_eq_Vud_div_sqrt {V : Quotient CKMMatrixSetoid} (ha : VubAbs V ≠ 1) :
C₁₂ V = VudAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by
rw [C₁₂, S₁₂, div_pow, Real.sq_sqrt]
rw [C₁₂, θ₁₂, Real.cos_arcsin, S₁₂, div_pow, Real.sq_sqrt]
rw [one_sub_div]
simp
rw [Real.sqrt_div]
@ -55,9 +129,8 @@ lemma C₁₂_eq_Vud_div_sqrt {V : Quotient CKMMatrixSetoid} (ha : VubAbs V ≠
exact VAbs_thd_neq_one_fst_snd_sq_neq_zero ha
exact (Left.add_nonneg (sq_nonneg (VAbs 0 0 V)) (sq_nonneg (VAbs 0 1 V)))
theorem VusAbs_eq_S₁₂_mul_C₁₃ (V : Quotient CKMMatrixSetoid) : VusAbs V = S₁₂ V * C₁₃ V := by
rw [C₁₃, S₁₂, S₁₃]
rw [C₁₃, θ₁₃, Real.cos_arcsin, S₁₂, S₁₃]
have h1 : 1 - VubAbs V ^ 2 = VudAbs V ^ 2 + VusAbs V ^ 2 := by
linear_combination - (VAbs_sum_sq_row_eq_one V 0)
rw [h1]
@ -75,9 +148,9 @@ theorem VudAbs_eq_C₁₂_mul_C₁₃ (V : Quotient CKMMatrixSetoid) : VudAbs V
by_cases ha : VubAbs V = 1
change VAbs 0 0 V = C₁₂ V * C₁₃ V
rw [VAbs_thd_eq_one_fst_eq_zero ha]
rw [C₁₃, ← VubAbs_eq_S₁₃, ha]
rw [C₁₃, θ₁₃, Real.cos_arcsin, ← VubAbs_eq_S₁₃, ha]
simp
rw [C₁₂_eq_Vud_div_sqrt ha, C₁₃, S₁₃]
rw [C₁₂_eq_Vud_div_sqrt ha, C₁₃, θ₁₃, Real.cos_arcsin, S₁₃]
have h1 : 1 - VubAbs V ^ 2 = VudAbs V ^ 2 + VusAbs V ^ 2 := by
linear_combination - (VAbs_sum_sq_row_eq_one V 0)
rw [h1, mul_comm]
@ -89,6 +162,50 @@ def S₂₃ (V : Quotient CKMMatrixSetoid) : :=
else
VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2)
lemma S₂₃_nonneg (V : Quotient CKMMatrixSetoid) : 0 ≤ S₂₃ V := by
by_cases ha : VubAbs V = 1
rw [S₂₃, if_pos ha]
exact VAbs_ge_zero 1 0 V
rw [S₂₃, if_neg ha]
rw [@div_nonneg_iff]
apply Or.inl
apply And.intro
exact VAbs_ge_zero 1 2 V
exact Real.sqrt_nonneg (VudAbs V ^ 2 + VusAbs V ^ 2)
lemma S₂₃_leq_one (V : Quotient CKMMatrixSetoid) : S₂₃ V ≤ 1 := by
by_cases ha : VubAbs V = 1
rw [S₂₃, if_pos ha]
exact VAbs_leq_one 1 0 V
rw [S₂₃, if_neg ha]
rw [@div_le_one_iff]
by_cases h1 : √(VudAbs V ^ 2 + VusAbs V ^ 2) = 0
simp [h1]
have h2 := Real.sqrt_nonneg (VudAbs V ^ 2 + VusAbs V ^ 2)
rw [le_iff_eq_or_lt] at h2
have h3 : 0 < √(VudAbs V ^ 2 + VusAbs V ^ 2) := by
cases' h2 with h2 h2
simp_all
exact h2
apply Or.inl
simp_all
rw [Real.le_sqrt]
rw [VudAbs_sq_add_VusAbs_sq, ← VcbAbs_sq_add_VtbAbs_sq]
simp
exact sq_nonneg (VAbs 2 2 V)
exact VAbs_ge_zero 1 2 V
exact le_of_lt h3
def θ₂₃ (V : Quotient CKMMatrixSetoid) : := Real.arcsin (S₂₃ V)
lemma S₂₃_eq_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₂₃ V) = S₂₃ V :=
Real.sin_arcsin (le_trans (by simp) (S₂₃_nonneg V)) (S₂₃_leq_one V)
lemma S₂₃_eq_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.sin (θ₂₃ V) = S₂₃ V := by
rw [← S₂₃_eq_sin_θ₂₃]
simp
lemma S₂₃_of_Vub_eq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : S₂₃ V = VcdAbs V := by
rw [S₂₃, if_pos ha]
@ -96,11 +213,18 @@ lemma S₂₃_of_Vub_neq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V ≠ 1
S₂₃ V = VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by
rw [S₂₃, if_neg ha]
def C₂₃ (V : Quotient CKMMatrixSetoid) : := √(1 - S₂₃ V ^ 2)
def C₂₃ (V : Quotient CKMMatrixSetoid) : := Real.cos (θ₂₃ V)
lemma C₂₃_eq_cos_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.cos (θ₂₃ V) = C₂₃ V := by
simp [C₂₃]
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)
lemma C₂₃_of_Vub_neq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V ≠ 1) :
C₂₃ V = VtbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by
rw [C₂₃, S₂₃_of_Vub_neq_one ha, div_pow, Real.sq_sqrt]
rw [C₂₃, θ₂₃, Real.cos_arcsin, S₂₃_of_Vub_neq_one ha, div_pow, Real.sq_sqrt]
rw [VudAbs_sq_add_VusAbs_sq, ← VcbAbs_sq_add_VtbAbs_sq]
rw [one_sub_div]
simp only [VcbAbs, Fin.isValue, VtbAbs, add_sub_cancel_left]
@ -144,281 +268,367 @@ lemma C₁₂_mul_S₂₃_mul_S₁₃_of_Vud_neq_one {V : Quotient CKMMatrixSeto
rw [Real.mul_self_sqrt (Left.add_nonneg (sq_nonneg (VAbs 0 0 V)) (sq_nonneg (VAbs 0 1 V)))]
exact (mul_div_right_comm (VudAbs V * VcbAbs V) (VubAbs V) (VudAbs V ^ 2 + VusAbs V ^ 2)).symm
namespace CKMMatrix
open ComplexConjugate
lemma Vud_eq_C₁₂_mul_C₁₃_arg (V : CKMMatrix) :
V.1 0 0 = (C₁₂ ⟦V⟧ * C₁₃ ⟦V⟧) * exp (I * arg (V.1 0 0)) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (V.1 0 0)]
rw [mul_comm _ I]
congr
have h1 := VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧
rw [← ofReal_inj] at h1
simpa using h1
lemma Vus_eq_C₁₂_mul_C₁₃_arg (V : CKMMatrix) :
V.1 0 1 = (S₁₂ ⟦V⟧ * C₁₃ ⟦V⟧ ) * exp (I * arg (V.1 0 1)) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (V.1 0 1)]
rw [mul_comm _ I]
congr
have h1 := VusAbs_eq_S₁₂_mul_C₁₃ ⟦V⟧
rw [← ofReal_inj] at h1
simpa using h1
lemma Vub_eq_C₁₃_arg (V : CKMMatrix) : V.1 0 2 = S₁₃ ⟦V⟧ * exp (I * arg (V.1 0 2)) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (V.1 0 2)]
rw [mul_comm _ I]
congr
lemma Vcb_eq_S₂₃_mul_C₁₃_arg (V : CKMMatrix) :
V.1 1 2 = (S₂₃ ⟦V⟧ * C₁₃ ⟦V⟧) * exp (I * arg (V.1 1 2)) := by
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (V.1 1 2)]
rw [mul_comm _ I]
congr
have h1 := VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧
rw [← ofReal_inj] at h1
simpa using h1
lemma Vud_neq_zero_abs_neq_zero (V : CKMMatrix) : abs (V.1 0 0) ≠ 0 ↔ V.1 0 0 ≠ 0 := by
exact AbsoluteValue.ne_zero_iff Complex.abs
lemma Vcb_neq_zero_abs_neq_zero (V : CKMMatrix) : abs (V.1 1 2) ≠ 0 ↔ V.1 1 2 ≠ 0 := by
exact AbsoluteValue.ne_zero_iff Complex.abs
lemma Vud_neq_zero_C₁₂_neq_zero {V : CKMMatrix} (ha : V.1 0 0 ≠ 0) : C₁₂ ⟦V⟧ ≠ 0 := by
rw [← Vud_neq_zero_abs_neq_zero] at ha
have h2 := VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧
simp [VAbs] at h2
rw [h2] at ha
simp at ha
rw [not_or] at ha
exact ha.1
lemma Vud_neq_zero_C₁₃_neq_zero {V : CKMMatrix} (ha : V.1 0 0 ≠ 0) : C₁₃ ⟦V⟧ ≠ 0 := by
rw [← Vud_neq_zero_abs_neq_zero] at ha
have h2 := VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧
simp [VAbs] at h2
rw [h2] at ha
simp at ha
rw [not_or] at ha
exact ha.2
lemma Vcb_neq_zero_S₂₃_neq_zero {V : CKMMatrix} (hb : V.1 1 2 ≠ 0) : S₂₃ ⟦V⟧ ≠ 0 := by
rw [← Vcb_neq_zero_abs_neq_zero] at hb
have h2 := VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧
simp [VAbs] at h2
rw [h2] at hb
simp at hb
rw [not_or] at hb
exact hb.1
def B (V : CKMMatrix) : := V.1 0 1 / V.1 0 0
lemma B_abs (V : CKMMatrix) (ha : V.1 0 0 ≠ 0) : abs (B V) = S₁₂ ⟦V⟧ / C₁₂ ⟦V⟧ := by
rw [B]
simp only [Fin.isValue, map_div₀]
have h1 := VusAbs_eq_S₁₂_mul_C₁₃ ⟦V⟧
have h2 := VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧
simp [VAbs] at h1 h2
rw [h1, h2]
rw [@mul_div_mul_comm]
rw [div_self]
simp
exact Vud_neq_zero_C₁₃_neq_zero ha
lemma B_normSq (V : CKMMatrix) (ha : V.1 0 0 ≠ 0) :
normSq (B V) = (S₁₂ ⟦V⟧ ^ 2) / (C₁₂ ⟦V⟧ ^ 2) := by
rw [← Complex.sq_abs]
rw [B_abs V ha]
simp
def A (V : CKMMatrix) : := V.1 0 2 / V.1 0 0
lemma A_abs (V : CKMMatrix) :
abs (A V) = (S₁₃ ⟦V⟧ / C₁₃ ⟦V⟧) * (1 / C₁₂ ⟦V⟧):= by
rw [A]
simp only [Fin.isValue, map_div₀]
have h1 := VubAbs_eq_S₁₃ ⟦V⟧
have h2 := VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧
simp [VAbs] at h1 h2
rw [h1, h2]
rw [← @mul_div_mul_comm]
simp
rw [mul_comm]
lemma A_normSq (V : CKMMatrix) :
normSq (A V) = (S₁₃ ⟦V⟧ ^ 2) / (C₁₃ ⟦V⟧ ^ 2) * (1 / (C₁₂ ⟦V⟧ ^ 2)) := by
rw [← Complex.sq_abs]
rw [A_abs V]
rw [mul_pow]
simp
def D (V : CKMMatrix) : := V.1 1 0 / V.1 1 2
def C' (V : CKMMatrix) : := V.1 1 1 / V.1 1 2
def C (V : CKMMatrix) : := C' V * (1 + abs (B V)^2) + conj (A V) * B V
lemma C'_of_C (V : CKMMatrix) : C' V = (C V - conj (A V) * B V) / (1 + abs (B V)^2) := by
have h1 : C V = C' V * (1 + abs (B V)^2) + conj (A V) * B V := by rfl
have h2 : C V - conj (A V) * B V = C' V * (1 + abs (B V)^2) := by
linear_combination h1
rw [h2]
rw [mul_div_cancel_right₀]
have h2 : 0 ≤ ↑(Complex.abs (B V)) ^ 2 := by simp
have h3 : 0 < 1 + ↑(Complex.abs (B V)) ^ 2 := by linarith
by_contra hn
have hx : 1 + Complex.abs (B V) ^ 2 = 0 := by
rw [← ofReal_inj]
simpa using hn
have h4 := lt_of_lt_of_eq h3 hx
simp at h4
lemma norm_D_C'_A {V : CKMMatrix} (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
D V + C' V * conj (B V) + conj (A V) = 0 := by
have h1 : (V.1 1 0 * conj (V.1 0 0) + V.1 1 1 * conj (V.1 0 1)
+ V.1 1 2 * conj (V.1 0 2))/ (V.1 1 2 * conj (V.1 0 0)) = 0 := by
rw [fst_row_snd_row V ]
simp
rw [← div_add_div_same, ← div_add_div_same] at h1
rw [mul_div_mul_comm, mul_div_mul_comm, mul_div_mul_comm] at h1
rw [div_self, div_self] at h1
change D V * 1 + _ + _ = _ at h1
have h2 : (starRingEnd ) (V.1 0 2) / (starRingEnd ) (V.1 0 0) = conj (A V) := by
simp [A]
have h3 : ((starRingEnd ) (V.1 0 1) / (starRingEnd ) (V.1 0 0)) = conj (B V) := by
simp [B]
rw [h2, h3] at h1
simp at h1
exact h1
exact ha
simpa using hb
lemma norm_D_C_A {V : CKMMatrix} (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
D V + (conj (A V) + C V * conj (B V) ) / (1 + abs (B V)^2) = 0 := by
have h1 := norm_D_C'_A ha hb
rw [C'_of_C] at h1
rw [div_mul_eq_mul_div] at h1
rw [add_assoc] at h1
rw [div_add'] at h1
have h2 : ((C V - (starRingEnd ) (A V) * B V) * (starRingEnd ) (B V) +
(starRingEnd ) (A V) * (1 + ↑(Complex.abs (B V)) ^ 2)) =
(conj (A V) + C V * conj (B V) ) := by
have h3 : ↑(Complex.abs (B V)) ^ 2 = (B V) * conj (B V) := by
rw [mul_conj]
rw [ ← Complex.sq_abs]
simp
rw [h3]
def standardParameterizationAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : 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 * δ₁₃),
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
lemma standardParameterizationAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
((standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
funext j i
simp only [standardParameterizationAsMatrix, neg_mul, Fin.isValue]
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
rw [h2] at h1
exact h1
have h2 : 0 ≤ ↑(Complex.abs (B V)) ^ 2 := by simp
have h3 : 0 < 1 + ↑(Complex.abs (B V)) ^ 2 := by linarith
by_contra hn
have hx : 1 + Complex.abs (B V) ^ 2 = 0 := by
rw [← ofReal_inj]
simpa using hn
have h4 := lt_of_lt_of_eq h3 hx
simp at h4
lemma D_of_A_C {V : CKMMatrix} (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
D V = - (conj (A V) + C V * conj (B V) ) / (1 + abs (B V)^2) := by
linear_combination (norm_D_C_A ha hb)
def sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : CKMMatrix :=
⟨standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃, by
rw [mem_unitaryGroup_iff']
exact standardParameterizationAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
lemma normSq_D_C'_add' (V : CKMMatrix) (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
normSq (D V) + normSq (C' V) + 1 = ((1 + normSq (B V)) *
(normSq (C V) + normSq (B V) + normSq (A V) + 1)) / (1 + normSq (B V)) ^ 2 := by
rw [← ofReal_inj]
lemma sP_cross (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
[sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]t = (conj [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]u ×₃ conj [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) := by
have h1 := exp_ne_zero (I * ↑δ₁₃)
funext i
fin_cases i
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg,
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
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg, 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.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
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg, 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.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
lemma eq_sP (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : } (hu : [U]u = [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]u)
(hc : [U]c = [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) :
U = sP θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by
apply ext_Rows hu hc
rw [hU, sP_cross, hu, hc]
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
have hb' : VubAbs ⟦V⟧ ≠ 1 := by
rw [ud_us_neq_zero_iff_ub_neq_one] at hb
simp [VAbs, hb]
have h1 : ofReal (√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) * ↑√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) )
= ofReal ((VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) ) := by
rw [Real.mul_self_sqrt ]
apply add_nonneg (sq_nonneg _) (sq_nonneg _)
simp at h1
have hx := Vabs_sq_add_neq_zero hb
refine eq_sP V ?_ ?_ hV.2.2.2.2
funext i
fin_cases i
simp only [uRow, Fin.isValue, Fin.zero_eta, cons_val_zero, sP, standardParameterizationAsMatrix,
ofReal_cos, ofReal_sin, ofReal_neg, mul_neg, neg_mul, neg_neg, cons_val', empty_val',
cons_val_fin_one, cons_val_one, head_cons, cons_val_two, tail_cons]
rw [hV.1, VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧]
simp [C₁₂, C₁₃]
simp [uRow, sP, standardParameterizationAsMatrix]
rw [hV.2.1, VusAbs_eq_S₁₂_mul_C₁₃ ⟦V⟧, ← S₁₂_eq_sin_θ₁₂ ⟦V⟧, C₁₃]
simp
rw [← mul_conj, ← mul_conj]
rw [D_of_A_C, C'_of_C]
simp [uRow, sP, standardParameterizationAsMatrix]
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (V.1 0 2)]
rw [show Complex.abs (V.1 0 2) = VubAbs ⟦V⟧ from rfl]
rw [VubAbs_eq_S₁₃, ← S₁₃_eq_sin_θ₁₃ ⟦V⟧]
simp
rw [conj_ofReal]
rw [div_mul_div_comm, div_mul_div_comm]
rw [div_add_div_same]
rw [div_add_one]
have h1 : ((Complex.abs (B V)) : ) ^ 2 = normSq (B V) := by
rw [← Complex.sq_abs]
simp
congr 1
rw [← mul_conj, ← mul_conj, ← mul_conj]
rw [h1, ← mul_conj]
ring
ring_nf
simp
funext i
fin_cases i
simp [cRow, sP, standardParameterizationAsMatrix]
rw [cd_of_us_or_ud_neq_zero_UCond hb ha hV]
rw [S₁₂_eq_sin_θ₁₂ ⟦V⟧, S₁₂, C₁₂_eq_cos_θ₁₂ ⟦V⟧, C₁₂_eq_Vud_div_sqrt hb']
rw [S₂₃_eq_sin_θ₂₃ ⟦V⟧, S₂₃_of_Vub_neq_one hb', C₂₃_eq_cos_θ₂₃ ⟦V⟧,
C₂₃_of_Vub_neq_one hb', S₁₃_eq_sin_θ₁₃ ⟦V⟧, S₁₃]
field_simp
rw [h1]
ring
simp [sq]
field_simp
ring_nf
simp [cRow, sP, standardParameterizationAsMatrix]
rw [C₁₂_eq_cos_θ₁₂ ⟦V⟧, C₂₃_eq_cos_θ₂₃ ⟦V⟧, S₁₂_eq_sin_θ₁₂ ⟦V⟧,
S₁₃_eq_sin_θ₁₃ ⟦V⟧, S₂₃_eq_sin_θ₂₃ ⟦V⟧]
rw [C₁₂_eq_Vud_div_sqrt hb', C₂₃_of_Vub_neq_one hb', S₁₂, S₁₃, S₂₃_of_Vub_neq_one hb']
rw [cs_of_us_or_ud_neq_zero_UCond hb ha hV]
field_simp
rw [h1]
simp [sq]
field_simp
ring_nf
simp [cRow, sP, standardParameterizationAsMatrix]
rw [hV.2.2.1]
rw [VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧, S₂₃_eq_sin_θ₂₃ ⟦V⟧, C₁₃]
simp
have h2 : 0 ≤ ↑(Complex.abs (B V)) ^ 2 := by simp
have h3 : 0 < 1 + ↑(Complex.abs (B V)) ^ 2 := by linarith
by_contra hn
have hx : 1 + Complex.abs (B V) ^ 2 = 0 := by
rw [← ofReal_inj]
simpa using hn
have h4 := lt_of_lt_of_eq h3 hx
simp at h4
exact ha
exact hb
lemma normS1_B_neq_zero (V : CKMMatrix):
1 + normSq (B V) ≠ 0 := by
have h2 : 0 ≤ normSq (B V) := by
exact normSq_nonneg (B V)
have h3 : 0 < 1 + normSq (B V) := by linarith
by_contra hn
have ht := lt_of_lt_of_eq h3 hn
simp at ht
lemma normSq_D_C'_add (V : CKMMatrix) (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
normSq (D V) + normSq (C' V) + 1 =
(normSq (C V) + normSq (B V) + normSq (A V) + 1) / (1 + normSq (B V)) := by
rw [normSq_D_C'_add' V ha hb]
rw [sq]
rw [← div_mul_div_comm]
rw [div_self]
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⟧) 0 := by
have h23 : S₂₃ ⟦V⟧ = 0 := by
have h1 := VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧
simp [VAbs] at h1
rw [ha] at h1
simp at h1
cases' h1 with h1 h1
exact h1
have h2 : abs [V]ud = 0 := by
change VudAbs ⟦V⟧ = 0
rw [VudAbs_eq_C₁₂_mul_C₁₃, h1]
simp
have h2b : abs [V]us = 0 := by
change VusAbs ⟦V⟧ = 0
rw [VusAbs_eq_S₁₂_mul_C₁₃, h1]
simp
simp_all
have h1 : ofReal (√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) * ↑√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) )
= ofReal ((VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) ) := by
rw [Real.mul_self_sqrt ]
apply add_nonneg (sq_nonneg _) (sq_nonneg _)
have hx := Vabs_sq_add_neq_zero hb
simp at h1
have hb' : VubAbs ⟦V⟧ ≠ 1 := by
rw [ud_us_neq_zero_iff_ub_neq_one] at hb
simp [VAbs, hb]
refine eq_sP V ?_ ?_ hV.2.2.2.1
funext i
fin_cases i
simp [uRow, sP, standardParameterizationAsMatrix]
rw [hV.1, VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧, C₁₂, C₁₃]
simp
exact normS1_B_neq_zero V
lemma D_C'_of_S_C (V : CKMMatrix) (ha : V.1 1 2 ≠ 0) :
normSq (D V) + normSq (C' V) + 1 = 1 / (S₂₃ ⟦V⟧ * C₁₃ ⟦V⟧) ^ 2 := by
have h1 : ((abs (V.1 1 0)) ^ 2 + (abs (V.1 1 1)) ^ 2 + (abs (V.1 1 2)) ^ 2)/ (abs (V.1 1 2))^2
= 1 /(abs (V.1 1 2))^2 := by
have h2 := VAbs_sum_sq_row_eq_one ⟦V⟧ 1
erw [h2]
rw [← div_add_div_same, ← div_add_div_same] at h1
rw [← div_pow, ← div_pow] at h1
rw [div_self] at h1
have h3 : (Complex.abs (V.1 1 0) / Complex.abs (V.1 1 2)) = abs (D V) := by
simp [D]
have h4 : (Complex.abs (V.1 1 1) / Complex.abs (V.1 1 2)) = abs (C' V) := by
simp [C']
rw [h3, h4] at h1
have h5 := VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧
simp [VAbs] at h5
rw [h5] at h1
rw [Complex.sq_abs, Complex.sq_abs] at h1
exact h1
simp [uRow, sP, standardParameterizationAsMatrix]
rw [hV.2.1, VusAbs_eq_S₁₂_mul_C₁₃ ⟦V⟧, ← S₁₂_eq_sin_θ₁₂ ⟦V⟧, C₁₃]
simp
simp [uRow, sP, standardParameterizationAsMatrix]
rw [hV.2.2.1, VubAbs_eq_S₁₃, S₁₃_eq_sin_θ₁₃ ⟦V⟧]
funext i
fin_cases i
simp [cRow, sP, standardParameterizationAsMatrix]
rw [S₂₃_eq_sin_θ₂₃ ⟦V⟧, h23]
simp
rw [C₂₃_eq_cos_θ₂₃ ⟦V⟧, C₂₃_of_Vub_neq_one hb']
rw [S₁₂_eq_sin_θ₁₂ ⟦V⟧, S₁₂]
rw [hV.2.2.2.2.1]
field_simp
rw [h1]
simp [sq]
field_simp
ring_nf
simp
simp [cRow, sP, standardParameterizationAsMatrix]
rw [S₂₃_eq_sin_θ₂₃ ⟦V⟧, h23]
simp
rw [C₂₃_eq_cos_θ₂₃ ⟦V⟧, C₂₃_of_Vub_neq_one hb']
rw [C₁₂_eq_cos_θ₁₂ ⟦V⟧, C₁₂_eq_Vud_div_sqrt hb']
rw [hV.2.2.2.2.2]
field_simp
rw [h1]
simp [sq]
field_simp
ring_nf
simp [cRow, sP, standardParameterizationAsMatrix]
rw [ha, S₂₃_eq_sin_θ₂₃ ⟦V⟧, h23]
simp
exact ha
lemma C_of_S_C (V : CKMMatrix) (ha : V.1 1 2 ≠ 0) (hb : V.1 0 0 ≠ 0) :
normSq (C V) = 0 := by
have h1 := D_C'_of_S_C V ha
rw [normSq_D_C'_add V ha hb] at h1
rw [div_eq_mul_inv] at h1
have h3 : 1 / (S₂₃ ⟦V⟧ * C₁₃ ⟦V⟧) ^ 2 ≠ 0 := by
lemma UCond₃_eq_sP {V : CKMMatrix} (hV : UCond₃ V) : V = sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
have h1 : VubAbs ⟦V⟧ = 1 := by
simp [VAbs]
rw [hV.2.2.2.1]
simp
rw [not_or]
have h2 := eq_mul_of_mul_inv_eq₀ h3 h1
have h4 : normSq (C V) = 1 / (S₂₃ ⟦V⟧ * C₁₃ ⟦V⟧) ^ 2 * (1 + normSq (B V))
- normSq (B V) - normSq (A V) - 1 := by
linear_combination h2
rw [B_normSq] at h4
rw [A_normSq] at h4
field_simp at h4
refine eq_sP V ?_ ?_ hV.2.2.2.2.1
funext i
fin_cases i
simp [uRow, sP, standardParameterizationAsMatrix]
rw [C₁₃_eq_cos_θ₁₃ ⟦V⟧, C₁₃_of_Vub_eq_one h1, hV.1]
simp
simp [uRow, sP, standardParameterizationAsMatrix]
rw [C₁₃_eq_cos_θ₁₃ ⟦V⟧, C₁₃_of_Vub_eq_one h1, hV.2.1]
simp
simp [uRow, sP, standardParameterizationAsMatrix]
rw [S₁₃_eq_sin_θ₁₃ ⟦V⟧, S₁₃]
simp [VAbs]
rw [hV.2.2.2.1]
simp
funext i
fin_cases i
simp [cRow, sP, standardParameterizationAsMatrix]
rw [S₂₃_eq_sin_θ₂₃ ⟦V⟧, S₂₃_of_Vub_eq_one h1]
rw [S₁₂_eq_sin_θ₁₂ ⟦V⟧, S₁₂_of_Vub_one h1]
rw [C₁₂_eq_cos_θ₁₂ ⟦V⟧, C₁₂_of_Vub_one h1]
rw [S₁₃_eq_sin_θ₁₃ ⟦V⟧, S₁₃_of_Vub_one h1]
rw [hV.2.2.2.2.2.1]
simp
simp [cRow, sP, standardParameterizationAsMatrix]
rw [S₂₃_eq_sin_θ₂₃ ⟦V⟧, S₂₃_of_Vub_eq_one h1]
rw [S₁₂_eq_sin_θ₁₂ ⟦V⟧, S₁₂_of_Vub_one h1]
rw [C₁₂_eq_cos_θ₁₂ ⟦V⟧, C₁₂_of_Vub_one h1]
rw [S₁₃_eq_sin_θ₁₃ ⟦V⟧, S₁₃_of_Vub_one h1]
simp
have h3 : (Real.cos (θ₂₃ ⟦V⟧) : ) = √(1 - S₂₃ ⟦V⟧ ^ 2) := by
rw [θ₂₃, Real.cos_arcsin]
simp at h3
rw [h3, S₂₃_of_Vub_eq_one h1, hV.2.2.2.2.2.2]
simp [cRow, sP, standardParameterizationAsMatrix]
rw [C₁₃_eq_cos_θ₁₃ ⟦V⟧, C₁₃_of_Vub_eq_one h1, hV.2.2.1]
simp
theorem exists_standardParameterization (V : CKMMatrix) :
∃ (δ₃ : ), V ≈ sP (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₃ := by
obtain ⟨U, hU⟩ := all_eq_abs V
have hUV : ⟦U⟧ = ⟦V⟧ := (Quotient.eq.mpr (phaseShiftEquivRelation.symm hU.1))
by_cases ha : [V]ud ≠ 0 [V]us ≠ 0
have haU : [U]ud ≠ 0 [U]us ≠ 0 := by
by_contra hn
simp [not_or] at hn
have hna : VudAbs ⟦U⟧ = 0 ∧ VusAbs ⟦U⟧ =0 := by
simp [VAbs]
exact hn
rw [hUV] at hna
simp [VAbs] at hna
simp_all
by_cases hb : [V]cb ≠ 0
have hbU : [U]cb ≠ 0 := by
by_contra hn
simp at hn
have hna : VcbAbs ⟦U⟧ = 0 := by
simp [VAbs]
exact hn
rw [hUV] at hna
simp [VAbs] at hna
simp_all
have hU' := UCond₁_eq_sP haU hbU hU.2
rw [hU'] at hU
use (- arg ([U]ub))
rw [← hUV]
exact hU.1
simp at hb
have hbU : [U]cb = 0 := by
have hna : VcbAbs ⟦U⟧ = 0 := by
rw [hUV]
simp [VAbs]
exact hb
simpa [VAbs] using hna
obtain ⟨U2, hU2⟩ := UCond₂_exists haU hbU hU.2
have hUVa2 : V ≈ U2 := phaseShiftEquivRelation.trans hU.1 hU2.1
have hUV2 : ⟦U2⟧ = ⟦V⟧ := (Quotient.eq.mpr (phaseShiftEquivRelation.symm hUVa2))
have haU2 : [U2]ud ≠ 0 [U2]us ≠ 0 := by
by_contra hn
simp [not_or] at hn
have hna : VudAbs ⟦U2⟧ = 0 ∧ VusAbs ⟦U2⟧ =0 := by
simp [VAbs]
exact hn
rw [hUV2] at hna
simp [VAbs] at hna
simp_all
have hbU2 : [U2]cb = 0 := by
have hna : VcbAbs ⟦U2⟧ = 0 := by
rw [hUV2]
simp [VAbs]
exact hb
simpa [VAbs] using hna
have hU2' := UCond₂_eq_sP haU2 hbU2 hU2.2
use 0
rw [← hUV2, ← hU2']
exact hUVa2
have haU : ¬ ([U]ud ≠ 0 [U]us ≠ 0) := by
simp [not_or] at ha
have h1 : VudAbs ⟦U⟧ = 0 ∧ VusAbs ⟦U⟧ = 0 := by
rw [hUV]
simp [VAbs]
exact ha
simpa [not_or, VAbs] using h1
have ⟨U2, hU2⟩ := UCond₃_exists haU hU.2
have hUVa2 : V ≈ U2 := phaseShiftEquivRelation.trans hU.1 hU2.1
have hUV2 : ⟦U2⟧ = ⟦V⟧ := (Quotient.eq.mpr (phaseShiftEquivRelation.symm hUVa2))
have hx := UCond₃_eq_sP hU2.2
use 0
rw [← hUV2, ← hx]
exact hUVa2
end CKMMatrix