docs: CKM Relations
This commit is contained in:
parent
3e2a801c70
commit
e9ce319101
4 changed files with 25 additions and 3 deletions
|
@ -23,6 +23,7 @@ open ComplexConjugate
|
|||
|
||||
section rows
|
||||
|
||||
/-- The absolute value squared of any rwo of a CKM matrix is `1`, in terms of `Vabs`. -/
|
||||
lemma VAbs_sum_sq_row_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
||||
(VAbs i 0 V) ^ 2 + (VAbs i 1 V) ^ 2 + (VAbs i 2 V) ^ 2 = 1 := by
|
||||
obtain ⟨V, hV⟩ := Quot.exists_rep V
|
||||
|
@ -37,25 +38,31 @@ lemma VAbs_sum_sq_row_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
|||
rw [← ofReal_inj]
|
||||
simpa using ht
|
||||
|
||||
/-- The absolute value squared of the first rwo of a CKM matrix is `1`, in terms of `abs`. -/
|
||||
lemma fst_row_normalized_abs (V : CKMMatrix) : abs [V]ud ^ 2 + abs [V]us ^ 2 + abs [V]ub ^ 2 = 1 :=
|
||||
VAbs_sum_sq_row_eq_one ⟦V⟧ 0
|
||||
|
||||
/-- The absolute value squared of the second rwo of a CKM matrix is `1`, in terms of `abs`. -/
|
||||
lemma snd_row_normalized_abs (V : CKMMatrix) : abs [V]cd ^ 2 + abs [V]cs ^ 2 + abs [V]cb ^ 2 = 1 :=
|
||||
VAbs_sum_sq_row_eq_one ⟦V⟧ 1
|
||||
|
||||
/-- The absolute value squared of the thid rwo of a CKM matrix is `1`, in terms of `abs`. -/
|
||||
lemma thd_row_normalized_abs (V : CKMMatrix) : abs [V]td ^ 2 + abs [V]ts ^ 2 + abs [V]tb ^ 2 = 1 :=
|
||||
VAbs_sum_sq_row_eq_one ⟦V⟧ 2
|
||||
|
||||
/-- The absolute value squared of the first rwo of a CKM matrix is `1`, in terms of `nomSq`. -/
|
||||
lemma fst_row_normalized_normSq (V : CKMMatrix) :
|
||||
normSq [V]ud + normSq [V]us + normSq [V]ub = 1 := by
|
||||
repeat rw [← Complex.sq_abs]
|
||||
exact V.fst_row_normalized_abs
|
||||
|
||||
/-- The absolute value squared of the second rwo of a CKM matrix is `1`, in terms of `nomSq`. -/
|
||||
lemma snd_row_normalized_normSq (V : CKMMatrix) :
|
||||
normSq [V]cd + normSq [V]cs + normSq [V]cb = 1 := by
|
||||
repeat rw [← Complex.sq_abs]
|
||||
exact V.snd_row_normalized_abs
|
||||
|
||||
/-- The absolute value squared of the third rwo of a CKM matrix is `1`, in terms of `nomSq`. -/
|
||||
lemma thd_row_normalized_normSq (V : CKMMatrix) :
|
||||
normSq [V]td + normSq [V]ts + normSq [V]tb = 1 := by
|
||||
repeat rw [← Complex.sq_abs]
|
||||
|
@ -63,10 +70,10 @@ lemma thd_row_normalized_normSq (V : CKMMatrix) :
|
|||
|
||||
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)
|
||||
linear_combination fst_row_normalized_normSq V
|
||||
|
||||
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)
|
||||
linear_combination VAbs_sum_sq_row_eq_one V 0
|
||||
|
||||
lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
||||
[V]ud ≠ 0 ∨ [V]us ≠ 0 ↔ abs [V]ub ≠ 1 := by
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue