refactor: Remove double empty lines
This commit is contained in:
parent
ae18a2196d
commit
f03d063c86
60 changed files with 0 additions and 232 deletions
|
@ -17,12 +17,10 @@ related by phase shifts.
|
|||
The notation `[V]ud` etc can be used for the elements of a CKM matrix, and
|
||||
`[V]ud|us` etc for the ratios of elements.
|
||||
|
||||
|
||||
-/
|
||||
|
||||
open Matrix Complex
|
||||
|
||||
|
||||
noncomputable section
|
||||
|
||||
/-- Given three real numbers `a b c` the complex matrix with `exp (I * a)` etc on the
|
||||
|
@ -107,7 +105,6 @@ lemma phaseShiftRelation_trans {U V W : unitaryGroup (Fin 3) ℂ} :
|
|||
rw [phaseShiftMatrix_mul]
|
||||
rw [add_comm k e, add_comm l f, add_comm m g]
|
||||
|
||||
|
||||
lemma phaseShiftRelation_equiv : Equivalence PhaseShiftRelation where
|
||||
refl := phaseShiftRelation_refl
|
||||
symm := phaseShiftRelation_symm
|
||||
|
@ -270,7 +267,6 @@ lemma tb (V : CKMMatrix) (a b c d e f : ℝ) :
|
|||
|
||||
end phaseShiftApply
|
||||
|
||||
|
||||
/-- The aboslute value of the `(i,j)`th element of `V`. -/
|
||||
@[simp]
|
||||
def VAbs' (V : unitaryGroup (Fin 3) ℂ) (i j : Fin 3) : ℝ := Complex.abs (V i j)
|
||||
|
@ -341,11 +337,9 @@ abbrev VtsAbs := VAbs 2 1
|
|||
@[simp]
|
||||
abbrev VtbAbs := VAbs 2 2
|
||||
|
||||
|
||||
namespace CKMMatrix
|
||||
open ComplexConjugate
|
||||
|
||||
|
||||
section ratios
|
||||
|
||||
/-- The ratio of the `ub` and `ud` elements of a CKM matrix. -/
|
||||
|
@ -394,7 +388,6 @@ lemma Rcscb_mul_cb {V : CKMMatrix} (h : [V]cb ≠ 0): [V]cs = Rcscb V * [V]cb :=
|
|||
|
||||
end ratios
|
||||
|
||||
|
||||
end CKMMatrix
|
||||
|
||||
end
|
||||
|
|
|
@ -15,17 +15,14 @@ this equivalence.
|
|||
|
||||
Of note, this file defines the complex jarlskog invariant.
|
||||
|
||||
|
||||
-/
|
||||
open Matrix Complex
|
||||
open ComplexConjugate
|
||||
open CKMMatrix
|
||||
|
||||
|
||||
noncomputable section
|
||||
namespace Invariant
|
||||
|
||||
|
||||
/-- The complex jarlskog invariant for a CKM matrix. -/
|
||||
@[simps!]
|
||||
def jarlskogℂCKM (V : CKMMatrix) : ℂ :=
|
||||
|
@ -62,6 +59,5 @@ standard parameterization. -/
|
|||
def mulExpδ₁₃ (V : Quotient CKMMatrixSetoid) : ℂ :=
|
||||
jarlskogℂ V + VusVubVcdSq V
|
||||
|
||||
|
||||
end Invariant
|
||||
end
|
||||
|
|
|
@ -19,11 +19,9 @@ In this file we define two sets of conditions on the CKM matrices
|
|||
and `ubOnePhaseCond` which we show can be satisfied by any CKM matrix up to equivalence as long as
|
||||
the ub element as absolute value 1.
|
||||
|
||||
|
||||
-/
|
||||
open Matrix Complex
|
||||
|
||||
|
||||
noncomputable section
|
||||
namespace CKMMatrix
|
||||
open ComplexConjugate
|
||||
|
@ -260,7 +258,6 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
|||
apply shift_cross_product_phase_zero _ _ _ _ _ _ hτ.symm
|
||||
ring
|
||||
|
||||
|
||||
lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0))
|
||||
(hV : FstRowThdColRealCond V) :
|
||||
∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by
|
||||
|
|
|
@ -16,7 +16,6 @@ matrix.
|
|||
|
||||
open Matrix Complex
|
||||
|
||||
|
||||
noncomputable section
|
||||
|
||||
namespace CKMMatrix
|
||||
|
@ -121,9 +120,6 @@ lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
|||
rw [← ud_us_neq_zero_iff_ub_neq_one V] at hV
|
||||
simpa [← Complex.sq_abs] using (normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hV)
|
||||
|
||||
|
||||
|
||||
|
||||
lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
||||
(normSq [V]ud : ℂ) + normSq [V]us ≠ 0 := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hb
|
||||
|
@ -131,7 +127,6 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0
|
|||
rw [← ofReal_inj] at h1
|
||||
simp_all
|
||||
|
||||
|
||||
lemma Vabs_sq_add_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
||||
((VudAbs ⟦V⟧ : ℂ) * ↑(VudAbs ⟦V⟧) + ↑(VusAbs ⟦V⟧) * ↑(VusAbs ⟦V⟧)) ≠ 0 := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ hb
|
||||
|
@ -230,7 +225,6 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ}
|
|||
simp only [Fin.isValue, neg_mul]
|
||||
ring
|
||||
|
||||
|
||||
lemma cs_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
{τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) :
|
||||
[V]cs = (- conj [V]ub * [V]us * [V]cb +
|
||||
|
@ -249,10 +243,8 @@ lemma cd_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
|||
field_simp
|
||||
ring
|
||||
|
||||
|
||||
end rows
|
||||
|
||||
|
||||
section individual
|
||||
|
||||
lemma VAbs_ge_zero (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : 0 ≤ VAbs i j V := by
|
||||
|
@ -270,12 +262,10 @@ lemma VAbs_leq_one (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : VAbs i j V ≤
|
|||
change VAbs i 2 V ≤ 1
|
||||
nlinarith
|
||||
|
||||
|
||||
end individual
|
||||
|
||||
section columns
|
||||
|
||||
|
||||
lemma VAbs_sum_sq_col_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
||||
(VAbs 0 i V) ^ 2 + (VAbs 1 i V) ^ 2 + (VAbs 2 i V) ^ 2 = 1 := by
|
||||
obtain ⟨V, hV⟩ := Quot.exists_rep V
|
||||
|
@ -312,7 +302,6 @@ lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) :
|
|||
simp at h1
|
||||
exact h1.1
|
||||
|
||||
|
||||
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
|
||||
|
@ -336,8 +325,6 @@ lemma VcbAbs_sq_add_VtbAbs_sq (V : Quotient CKMMatrixSetoid) :
|
|||
VcbAbs V ^ 2 + VtbAbs V ^ 2 = 1 - VubAbs V ^2 := by
|
||||
linear_combination (VAbs_sum_sq_col_eq_one V 2)
|
||||
|
||||
|
||||
|
||||
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
|
||||
|
|
|
@ -14,7 +14,6 @@ proves some properties between them.
|
|||
|
||||
The first row can be extracted as `[V]u` for a CKM matrix `V`.
|
||||
|
||||
|
||||
-/
|
||||
|
||||
open Matrix Complex
|
||||
|
@ -24,7 +23,6 @@ noncomputable section
|
|||
|
||||
namespace CKMMatrix
|
||||
|
||||
|
||||
/-- The `u`th row of the CKM matrix. -/
|
||||
def uRow (V : CKMMatrix) : Fin 3 → ℂ := ![[V]ud, [V]us, [V]ub]
|
||||
|
||||
|
|
|
@ -16,7 +16,6 @@ four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`.
|
|||
We will show that every CKM matrix can be written within this standard parameterization
|
||||
in the file `FlavorPhysics.CKMMatrix.StandardParameters`.
|
||||
|
||||
|
||||
-/
|
||||
open Matrix Complex
|
||||
open ComplexConjugate
|
||||
|
@ -194,7 +193,5 @@ lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤
|
|||
have h1 : cexp (I * δ₁₃) ≠ 0 := exp_ne_zero _
|
||||
field_simp
|
||||
|
||||
|
||||
|
||||
end standParam
|
||||
end
|
||||
|
|
|
@ -18,7 +18,6 @@ These, when used in the standard parameterization return `V` up to equivalence.
|
|||
This leads to the theorem `standParam.exists_for_CKMatrix` which says that up to equivalence every
|
||||
CKM matrix can be written using the standard parameterization.
|
||||
|
||||
|
||||
-/
|
||||
open Matrix Complex
|
||||
open ComplexConjugate
|
||||
|
@ -333,7 +332,6 @@ lemma Vs_zero_iff_cos_sin_zero (V : CKMMatrix) :
|
|||
|
||||
end VAbs
|
||||
|
||||
|
||||
namespace standParam
|
||||
open Invariant
|
||||
|
||||
|
@ -409,7 +407,6 @@ lemma on_param_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
rfl
|
||||
rfl
|
||||
|
||||
|
||||
lemma on_param_cos_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.cos (θ₁₂ ⟦V⟧) = 0) :
|
||||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, δ₁₃, -δ₁₃, 0, - δ₁₃
|
||||
|
@ -434,7 +431,6 @@ lemma on_param_cos_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
ring_nf
|
||||
field_simp
|
||||
|
||||
|
||||
lemma on_param_cos_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.cos (θ₂₃ ⟦V⟧) = 0) :
|
||||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, 0, 0, 0, - δ₁₃
|
||||
|
@ -451,7 +447,6 @@ lemma on_param_cos_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
ring
|
||||
field_simp
|
||||
|
||||
|
||||
lemma on_param_sin_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.sin (θ₁₃ ⟦V⟧) = 0) :
|
||||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, 0, 0, 0, 0, 0
|
||||
|
@ -488,8 +483,6 @@ lemma on_param_sin_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
ring_nf
|
||||
field_simp
|
||||
|
||||
|
||||
|
||||
lemma on_param_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.sin (θ₂₃ ⟦V⟧) = 0) :
|
||||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, 0, δ₁₃, 0, 0, - δ₁₃
|
||||
|
@ -506,9 +499,6 @@ lemma on_param_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
ring
|
||||
field_simp
|
||||
|
||||
|
||||
|
||||
|
||||
lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
(hV : FstRowThdColRealCond V) : V = standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (- arg [V]ub) := by
|
||||
have hb' : VubAbs ⟦V⟧ ≠ 1 := by
|
||||
|
@ -565,7 +555,6 @@ lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0
|
|||
rw [VcbAbs_eq_S₂₃_mul_C₁₃ ⟦V⟧, S₂₃_eq_ℂsin_θ₂₃ ⟦V⟧, C₁₃]
|
||||
simp
|
||||
|
||||
|
||||
lemma eq_standParam_of_ubOnePhaseCond {V : CKMMatrix} (hV : ubOnePhaseCond V) :
|
||||
V = standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
have h1 : VubAbs ⟦V⟧ = 1 := by
|
||||
|
@ -610,7 +599,6 @@ lemma eq_standParam_of_ubOnePhaseCond {V : CKMMatrix} (hV : ubOnePhaseCond V) :
|
|||
rw [C₁₃_eq_ℂcos_θ₁₃ ⟦V⟧, C₁₃_of_Vub_eq_one h1, hV.2.2.1]
|
||||
simp
|
||||
|
||||
|
||||
theorem exists_δ₁₃ (V : CKMMatrix) :
|
||||
∃ (δ₃ : ℝ), V ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₃ := by
|
||||
obtain ⟨U, hU⟩ := fstRowThdColRealCond_holds_up_to_equiv V
|
||||
|
@ -670,7 +658,6 @@ theorem eq_standardParameterization_δ₃ (V : CKMMatrix) :
|
|||
exact on_param_sin_θ₁₃_eq_zero δ₁₃' h
|
||||
exact on_param_sin_θ₂₃_eq_zero δ₁₃' h
|
||||
|
||||
|
||||
theorem exists_for_CKMatrix (V : CKMMatrix) :
|
||||
∃ (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ), V ≈ standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by
|
||||
use θ₁₂ ⟦V⟧, θ₁₃ ⟦V⟧, θ₂₃ ⟦V⟧, δ₁₃ ⟦V⟧
|
||||
|
@ -678,9 +665,6 @@ theorem exists_for_CKMatrix (V : CKMMatrix) :
|
|||
|
||||
end standParam
|
||||
|
||||
|
||||
open CKMMatrix
|
||||
|
||||
|
||||
|
||||
end
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue