refactor: Linting substrings

This commit is contained in:
jstoobysmith 2024-07-12 16:22:06 -04:00
parent cee38b7be8
commit ac1132c7ca
40 changed files with 133 additions and 132 deletions

View file

@ -236,7 +236,7 @@ lemma td (V : CKMMatrix) (a b c d e f : ) :
simp only [Fin.isValue, cons_val', cons_val_zero, empty_val', cons_val_fin_one, vecCons_const,
cons_val_two, tail_val', head_val', cons_val_one, head_cons, tail_cons, head_fin_const,
zero_mul, add_zero, mul_zero]
change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _
change (0 * _ + _) * _ + (0 * _ + _) * 0 = _
simp only [Fin.isValue, zero_mul, zero_add, mul_zero, add_zero]
rw [exp_add]
ring_nf

View file

@ -227,9 +227,9 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
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]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 : Quotient.mk CKMMatrixSetoid U = ⟦V⟧ := by
@ -262,7 +262,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
(hV : FstRowThdColRealCond V) :
∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond 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 )
(Real.pi - arg [V]cd) (- arg [V]cs) (- arg [V]ub)
use U
have hUV : Quotient.mk CKMMatrixSetoid U= ⟦V⟧ := by
simp only [Quotient.eq]
@ -318,7 +318,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0)
(hV : FstRowThdColRealCond 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 ))
(- 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 only [ofReal_zero, zero_mul, exp_zero, one_smul]

View file

@ -84,7 +84,7 @@ lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
simp_all
have h1 := Complex.abs.nonneg [V]ub
rw [h2] at h1
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
have h2 : ¬ 0 ≤ (-1 : ) := by simp
exact h2 h1
lemma normSq_Vud_plus_normSq_Vus_neq_zero_ {V : CKMMatrix} (hb : [V]ud ≠ 0 [V]us ≠ 0) :
@ -100,7 +100,7 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ {V : CKMMatrix} (hb : [V]ud ≠ 0
exact hb h2
have h3 := Complex.abs.nonneg [V]ub
rw [h2] at h3
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
have h2 : ¬ 0 ≤ (-1 : ) := by simp
exact h2 h3
lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
@ -154,11 +154,11 @@ lemma fst_row_orthog_thd_row (V : CKMMatrix) :
lemma Vcd_mul_conj_Vud (V : CKMMatrix) :
[V]cd * conj [V]ud = -[V]cs * conj [V]us - [V]cb * conj [V]ub := by
linear_combination (V.fst_row_orthog_snd_row )
linear_combination (V.fst_row_orthog_snd_row)
lemma Vcs_mul_conj_Vus (V : CKMMatrix) :
[V]cs * conj [V]us = - [V]cd * conj [V]ud - [V]cb * conj [V]ub := by
linear_combination (V.fst_row_orthog_snd_row )
linear_combination V.fst_row_orthog_snd_row
end orthogonal
@ -344,7 +344,7 @@ lemma cb_tb_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
simp_all
have h1 := Complex.abs.nonneg [V]ub
rw [h2] at h1
have h2 : ¬ 0 ≤ ( -1 : ) := by simp
have h2 : ¬ 0 ≤ (-1 : ) := by simp
exact h2 h1
lemma VAbs_fst_col_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3}

View file

@ -240,7 +240,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
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)) )
(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
@ -307,20 +307,20 @@ 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
cexp ((- a * I) + (- b * I) + (- e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by
simp [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue,
LinearMap.mk₂_apply, Pi.conj_apply, cons_val_zero, neg_mul, uRow, us, ub, cRow, cs, cb,
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
cexp ((- a * I) + (- b * I) + (- d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, 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
cexp ((- a * I) + (- b * I) + (- d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub,
← exp_conj, conj_ofReal]
ring

View file

@ -373,8 +373,8 @@ lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ) :
complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃]
lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : )
(h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) :
cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) =
(h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0) :
cexp (arg (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * I) =
cexp (δ₁₃ * I) := by
have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃
have habs := mulExpδ₁₃_on_param_abs V δ₁₃
@ -383,7 +383,7 @@ lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : )
rw [habs, h1a]
ring_nf
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (mulExpδ₁₃
⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2
⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧)] at h2
have habs_neq_zero :
(Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ) ≠ 0 := by
simp only [ne_eq, ofReal_eq_zero, map_eq_zero]
@ -505,7 +505,7 @@ lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0
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
↑√(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