Refactor: Names

This commit is contained in:
jstoobysmith 2024-04-29 09:22:32 -04:00
parent ff89c3f79d
commit 490ed0380c
6 changed files with 342 additions and 442 deletions

View file

@ -16,7 +16,7 @@ open CKMMatrix
noncomputable section
-- to be renamed stanParamAsMatrix ...
def standardParameterizationAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : 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 * δ₁₃),
@ -27,10 +27,10 @@ def standardParameterizationAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : )
open CKMMatrix
lemma standardParameterizationAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
((standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by
funext j i
simp only [standardParameterizationAsMatrix, neg_mul, Fin.isValue]
simp only [standParamAsMatrix, neg_mul, Fin.isValue]
rw [mul_apply]
have h1 := exp_ne_zero (I * ↑δ₁₃)
fin_cases j <;> rw [Fin.sum_univ_three]
@ -87,17 +87,19 @@ lemma standardParameterizationAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ
rw [sin_sq, sin_sq]
ring
def sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : CKMMatrix :=
⟨standardParameterizationAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃, by
def standParam (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : CKMMatrix :=
⟨standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃, by
rw [mem_unitaryGroup_iff']
exact standardParameterizationAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
exact standParamAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
lemma sP_cross (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
[sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]t = (conj [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]u ×₃ conj [sP θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) := by
namespace standParam
lemma cross_product_t (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
[standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]t =
(conj [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u ×₃ conj [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) := by
have h1 := exp_ne_zero (I * ↑δ₁₃)
funext i
fin_cases i
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg,
· simp only [tRow, standParam, standParamAsMatrix, 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,
@ -106,7 +108,7 @@ lemma sP_cross (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
ring_nf
rw [sin_sq]
ring
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg, Fin.isValue, cons_val',
· simp only [tRow, standParam, standParamAsMatrix, 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,
@ -115,7 +117,7 @@ lemma sP_cross (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
ring_nf
rw [sin_sq]
ring
· simp only [tRow, sP, standardParameterizationAsMatrix, neg_mul, exp_neg, Fin.isValue,
· simp only [tRow, standParam, standParamAsMatrix, 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,
@ -125,27 +127,26 @@ lemma sP_cross (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
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
lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : } (hu : [U]u = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u)
(hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) :
U = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by
apply ext_Rows hu hc
rw [hU, sP_cross, hu, hc]
rw [hU, cross_product_t, hu, hc]
lemma eq_phases_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
sP θ₁₂ θ₁₃ θ₂₃ δ₁₃ = sP θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
simp [sP, standardParameterizationAsMatrix]
lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
simp [standParam, standParamAsMatrix]
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]]
namespace Invariant
lemma VusVubVcdSq_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
open Invariant in
lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
VusVubVcdSq ⟦sP θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
VusVubVcdSq ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
Real.sin θ₁₂ ^ 2 * Real.cos θ₁₃ ^ 2 * Real.sin θ₁₃ ^ 2 * Real.sin θ₂₃ ^ 2 := by
simp only [VusVubVcdSq, VusAbs, VAbs, VAbs', Fin.isValue, sP, standardParameterizationAsMatrix,
simp only [VusVubVcdSq, VusAbs, VAbs, VAbs', Fin.isValue, standParam, standParamAsMatrix,
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]
@ -164,12 +165,13 @@ lemma VusVubVcdSq_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Rea
rw [hx]
simp
lemma mulExpδ₃_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
open Invariant in
lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Real.sin θ₁₂)
(h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) :
mulExpδ₃ ⟦sP θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
mulExpδ₁₃ ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by
rw [mulExpδ₃, VusVubVcdSq_sP _ _ _ _ h1 h2 h3 h4 ]
simp only [jarlskog, sP, standardParameterizationAsMatrix, neg_mul,
rw [mulExpδ₁₃, VusVubVcdSq_eq _ _ _ _ h1 h2 h3 h4 ]
simp only [jarlskog, standParam, standParamAsMatrix, neg_mul,
Quotient.lift_mk, jarlskogCKM, 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]
@ -179,6 +181,7 @@ lemma mulExpδ₃_sP (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Rea
have h1 : cexp (I * δ₁₃) ≠ 0 := exp_ne_zero _
field_simp
end Invariant
end standParam
end