refactor: Lint

This commit is contained in:
jstoobysmith 2024-07-19 17:00:32 -04:00
parent 52e591fa7a
commit 9f27a3a9fd
46 changed files with 176 additions and 168 deletions

View file

@ -293,47 +293,47 @@ def VAbs (i j : Fin 3) : Quotient CKMMatrixSetoid → :=
Quotient.lift (fun V => VAbs' V i j) (VAbs'_equiv i j)
/-- The absolute value of the `ud`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VudAbs := VAbs 0 0
/-- The absolute value of the `us`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VusAbs := VAbs 0 1
/-- The absolute value of the `ub`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VubAbs := VAbs 0 2
/-- The absolute value of the `cd`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VcdAbs := VAbs 1 0
/-- The absolute value of the `cs`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VcsAbs := VAbs 1 1
/-- The absolute value of the `cb`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VcbAbs := VAbs 1 2
/-- The absolute value of the `td`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VtdAbs := VAbs 2 0
/-- The absolute value of the `ts`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VtsAbs := VAbs 2 1
/-- The absolute value of the `tb`th element of a representative of an equivalence class of
CKM matrices. -/
CKM matrices. -/
@[simp]
abbrev VtbAbs := VAbs 2 2
@ -372,7 +372,7 @@ def Rcdcb (V : CKMMatrix) : := [V]cd / [V]cb
/-- The ratio of the `cd` and `cb` elements of a CKM matrix. -/
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
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
@ -382,7 +382,7 @@ def Rcscb (V : CKMMatrix) : := [V]cs / [V]cb
/-- The ratio of the `cs` and `cb` elements of a CKM matrix. -/
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
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

View file

@ -49,8 +49,8 @@ def jarlskog : Quotient CKMMatrixSetoid → :=
Quotient.lift jarlskogCKM jarlskogCKM_equiv
/-- An invariant for CKM mtrices corresponding to the square of the absolute values
of the `us`, `ub` and `cb` elements multipled together divied by `(VudAbs V ^ 2 + VusAbs V ^2)` .
-/
of the `us`, `ub` and `cb` elements multipled together divied by `(VudAbs V ^ 2 + VusAbs V ^2)` .
-/
def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : :=
VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2)

View file

@ -165,7 +165,7 @@ def FstRowThdColRealCond (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U
there is no phase difference between the `t`th-row and
the cross product of the conjugates of the `u`th and `c`th rows, and the `cd`th and `cs`th
elements are real and related in a set way.
-/
-/
def ubOnePhaseCond (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)

View file

@ -100,8 +100,8 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
parameterization of CKM matrices. -/
def standParam (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) : CKMMatrix :=
⟨standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃, by
rw [mem_unitaryGroup_iff']
exact standParamAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
rw [mem_unitaryGroup_iff']
exact standParamAsMatrix_unitary θ₁₂ θ₁₃ θ₂₃ δ₁₃⟩
namespace standParam
lemma cross_product_t (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) :
@ -158,7 +158,7 @@ lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ) (h1 : 0 ≤ Rea
VusVubVcdSq ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ =
Real.sin θ₁₂ ^ 2 * Real.cos θ₁₃ ^ 2 * Real.sin θ₁₃ ^ 2 * Real.sin θ₂₃ ^ 2 := by
simp only [VusVubVcdSq, VusAbs, VAbs, VAbs', Fin.isValue, standParam, standParamAsMatrix,
neg_mul, Quotient.lift_mk, cons_val', cons_val_one, head_cons,
neg_mul, Quotient.lift_mk, cons_val', cons_val_one, head_cons,
empty_val', cons_val_fin_one, cons_val_zero, _root_.map_mul, VubAbs, cons_val_two, tail_cons,
VcbAbs, VudAbs, Complex.abs_ofReal]
by_cases hx : Real.cos θ₁₃ ≠ 0

View file

@ -142,17 +142,17 @@ lemma S₂₃_eq_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.sin (
(ofReal_sin _).symm.trans (congrArg ofReal (S₂₃_eq_sin_θ₂₃ V))
lemma complexAbs_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) :
Complex.abs (Complex.sin (θ₁₂ V)) = sin (θ₁₂ V):= by
Complex.abs (Complex.sin (θ₁₂ V)) = sin (θ₁₂ V) := by
rw [S₁₂_eq_sin_θ₁₂, Complex.abs_ofReal, ofReal_inj, abs_eq_self]
exact S₁₂_nonneg _
lemma complexAbs_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) :
Complex.abs (Complex.sin (θ₁₃ V)) = sin (θ₁₃ V):= by
Complex.abs (Complex.sin (θ₁₃ V)) = sin (θ₁₃ V) := by
rw [S₁₃_eq_sin_θ₁₃, Complex.abs_ofReal, ofReal_inj, abs_eq_self]
exact S₁₃_nonneg _
lemma complexAbs_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) :
Complex.abs (Complex.sin (θ₂₃ V)) = sin (θ₂₃ V):= by
Complex.abs (Complex.sin (θ₂₃ V)) = sin (θ₂₃ V) := by
rw [S₂₃_eq_sin_θ₂₃, Complex.abs_ofReal, ofReal_inj, abs_eq_self]
exact S₂₃_nonneg _
@ -185,19 +185,19 @@ lemma C₂₃_eq_cos_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.cos (
simp [C₂₃]
lemma complexAbs_cos_θ₁₂ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₁₂ V)) =
cos (θ₁₂ V):= by
cos (θ₁₂ V) := by
rw [C₁₂_eq_cos_θ₁₂, Complex.abs_ofReal]
simp only [ofReal_inj, abs_eq_self]
exact Real.cos_arcsin_nonneg _
lemma complexAbs_cos_θ₁₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₁₃ V)) =
cos (θ₁₃ V):= by
cos (θ₁₃ V) := by
rw [C₁₃_eq_cos_θ₁₃, Complex.abs_ofReal]
simp only [ofReal_inj, abs_eq_self]
exact Real.cos_arcsin_nonneg _
lemma complexAbs_cos_θ₂₃ (V : Quotient CKMMatrixSetoid) : Complex.abs (Complex.cos (θ₂₃ V)) =
cos (θ₂₃ V):= by
cos (θ₂₃ V) := by
rw [C₂₃_eq_cos_θ₂₃, Complex.abs_ofReal]
simp only [ofReal_inj, abs_eq_self]
exact Real.cos_arcsin_nonneg _
@ -349,10 +349,10 @@ lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ) :
lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ) :
mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = 0 ↔
VudAbs ⟦V⟧ = 0 VubAbs ⟦V⟧ = 0 VusAbs ⟦V⟧ = 0 VcbAbs ⟦V⟧ = 0 VtbAbs ⟦V⟧ = 0 := by
VudAbs ⟦V⟧ = 0 VubAbs ⟦V⟧ = 0 VusAbs ⟦V⟧ = 0 VcbAbs ⟦V⟧ = 0 VtbAbs ⟦V⟧ = 0 := by
rw [VudAbs_eq_C₁₂_mul_C₁₃, VubAbs_eq_S₁₃, VusAbs_eq_S₁₂_mul_C₁₃, VcbAbs_eq_S₂₃_mul_C₁₃,
VtbAbs_eq_C₂₃_mul_C₁₃, ← ofReal_inj,
← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj]
VtbAbs_eq_C₂₃_mul_C₁₃, ← ofReal_inj,
← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj]
simp only [ofReal_mul]
rw [← S₁₃_eq_sin_θ₁₃, ← S₁₂_eq_sin_θ₁₂, ← S₂₃_eq_sin_θ₂₃,
← C₁₃_eq_cos_θ₁₃, ← C₂₃_eq_cos_θ₂₃,← C₁₂_eq_cos_θ₁₂]