refactor: Last batch of multi-goal proofs
This commit is contained in:
parent
b9479c904d
commit
c0499483a8
43 changed files with 910 additions and 955 deletions
|
@ -40,8 +40,8 @@ lemma phaseShiftMatrix_star (a b c : ℝ) :
|
|||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [← exp_conj, conj_ofReal]
|
||||
rfl
|
||||
rfl
|
||||
· rfl
|
||||
· rfl
|
||||
|
||||
lemma phaseShiftMatrix_mul (a b c d e f : ℝ) :
|
||||
phaseShiftMatrix a b c * phaseShiftMatrix d e f = phaseShiftMatrix (a + d) (b + e) (c + f) := by
|
||||
|
|
|
@ -238,25 +238,25 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
|||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
use U
|
||||
apply And.intro
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
apply shift_ud_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
apply shift_us_phase_zero
|
||||
ring
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
apply shift_cb_phase_zero
|
||||
ring
|
||||
apply And.intro
|
||||
rw [hUV]
|
||||
apply shift_tb_phase_zero
|
||||
ring
|
||||
apply shift_cross_product_phase_zero _ _ _ _ _ _ hτ.symm
|
||||
ring
|
||||
· exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
· apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_ud_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
· apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_us_phase_zero
|
||||
ring
|
||||
· apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_cb_phase_zero
|
||||
ring
|
||||
· apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_tb_phase_zero
|
||||
ring
|
||||
· 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) :
|
||||
|
@ -269,51 +269,51 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
|||
symm
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h1 : VudAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h1 : VusAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h3 := cb_eq_zero_of_ud_us_zero hb
|
||||
have h1 : VcbAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, h3]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· have hU1 : [U]ub = VubAbs ⟦V⟧ := by
|
||||
apply shift_ub_phase_zero _ _ _ _ _ _ _
|
||||
· exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
· apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h1 : VudAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h1 : VusAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
have h3 := cb_eq_zero_of_ud_us_zero hb
|
||||
have h1 : VcbAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, h3]
|
||||
simp [VAbs] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· have hU1 : [U]ub = VubAbs ⟦V⟧ := by
|
||||
apply shift_ub_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
rw [hU1]
|
||||
have h1:= (ud_us_neq_zero_iff_ub_neq_one V).mpr.mt hb
|
||||
simpa using h1
|
||||
apply And.intro
|
||||
· have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
simp only [ofReal_zero, zero_mul, exp_zero, one_smul]
|
||||
exact hV.2.2.2.2
|
||||
apply shift_cross_product_phase_zero _ _ _ _ _ _ hτ.symm
|
||||
ring
|
||||
rw [hU1]
|
||||
have h1:= (ud_us_neq_zero_iff_ub_neq_one V).mpr.mt hb
|
||||
simpa using h1
|
||||
apply And.intro
|
||||
· have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
simp only [ofReal_zero, zero_mul, exp_zero, one_smul]
|
||||
exact hV.2.2.2.2
|
||||
apply shift_cross_product_phase_zero _ _ _ _ _ _ hτ.symm
|
||||
ring
|
||||
apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_cd_phase_pi _ _ _ _ _ _ _
|
||||
ring
|
||||
have hcs : [U]cs = VcsAbs ⟦U⟧ := by
|
||||
rw [hUV]
|
||||
apply shift_cs_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
rw [hcs, hUV, cs_of_ud_us_zero hb]
|
||||
apply And.intro
|
||||
· rw [hUV]
|
||||
apply shift_cd_phase_pi _ _ _ _ _ _ _
|
||||
ring
|
||||
have hcs : [U]cs = VcsAbs ⟦U⟧ := by
|
||||
rw [hUV]
|
||||
apply shift_cs_phase_zero _ _ _ _ _ _ _
|
||||
ring
|
||||
rw [hcs, hUV, cs_of_ud_us_zero hb]
|
||||
|
||||
lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||
(hV : FstRowThdColRealCond V) :
|
||||
|
|
|
@ -217,26 +217,25 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
uRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs] at h2
|
||||
simp at h2
|
||||
cases' h2 with h2 h2
|
||||
swap
|
||||
have hx : Complex.abs (g 0) = -1 := by
|
||||
simp [← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one, h2]
|
||||
have h3 := Complex.abs.nonneg (g 0)
|
||||
simp_all
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
· exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
rw [← hg, smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
have hg2 : Complex.abs (g 0)⁻¹ = 1 := by
|
||||
simp [@map_inv₀, h2]
|
||||
have hg22 : ∃ (τ : ℝ), (g 0)⁻¹ = Complex.exp (τ * I) := by
|
||||
rw [← abs_mul_exp_arg_mul_I (g 0)⁻¹, hg2]
|
||||
use arg (g 0)⁻¹
|
||||
simp
|
||||
obtain ⟨τ, hτ⟩ := hg22
|
||||
use τ
|
||||
rw [hx, hτ]
|
||||
· have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
rw [← hg, smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
have hg2 : Complex.abs (g 0)⁻¹ = 1 := by
|
||||
simp [@map_inv₀, h2]
|
||||
have hg22 : ∃ (τ : ℝ), (g 0)⁻¹ = Complex.exp (τ * I) := by
|
||||
rw [← abs_mul_exp_arg_mul_I (g 0)⁻¹, hg2]
|
||||
use arg (g 0)⁻¹
|
||||
simp
|
||||
obtain ⟨τ, hτ⟩ := hg22
|
||||
use τ
|
||||
rw [hx, hτ]
|
||||
· have hx : Complex.abs (g 0) = -1 := by
|
||||
simp [← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one, h2]
|
||||
have h3 := Complex.abs.nonneg (g 0)
|
||||
simp_all only [ofReal_neg, ofReal_one, Left.nonneg_neg_iff]
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
|
||||
lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
||||
∃ (τ : ℝ), [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
|
@ -264,37 +263,37 @@ lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
|||
ofReal_eq_one] at h2
|
||||
cases' h2 with h2 h2
|
||||
swap
|
||||
have hx : Complex.abs (g 2) = -1 := by
|
||||
simp [h2, ← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one]
|
||||
have h3 := Complex.abs.nonneg (g 2)
|
||||
simp_all
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
rw [← hg, @smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
have hg2 : Complex.abs (g 2)⁻¹ = 1 := by
|
||||
simp [@map_inv₀, h2]
|
||||
have hg22 : ∃ (τ : ℝ), (g 2)⁻¹ = Complex.exp (τ * I) := by
|
||||
rw [← abs_mul_exp_arg_mul_I (g 2)⁻¹]
|
||||
use arg (g 2)⁻¹
|
||||
simp [hg2]
|
||||
obtain ⟨τ, hτ⟩ := hg22
|
||||
use τ
|
||||
rw [hx, hτ]
|
||||
· have hx : Complex.abs (g 2) = -1 := by
|
||||
simp [h2, ← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one]
|
||||
have h3 := Complex.abs.nonneg (g 2)
|
||||
simp_all
|
||||
have h4 : (0 : ℝ) < 1 := by norm_num
|
||||
exact False.elim (lt_iff_not_le.mp h4 h3)
|
||||
· have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
rw [← hg, @smul_smul, inv_mul_cancel, one_smul]
|
||||
by_contra hn
|
||||
simp [hn] at h2
|
||||
have hg2 : Complex.abs (g 2)⁻¹ = 1 := by
|
||||
simp [@map_inv₀, h2]
|
||||
have hg22 : ∃ (τ : ℝ), (g 2)⁻¹ = Complex.exp (τ * I) := by
|
||||
rw [← abs_mul_exp_arg_mul_I (g 2)⁻¹]
|
||||
use arg (g 2)⁻¹
|
||||
simp [hg2]
|
||||
obtain ⟨τ, hτ⟩ := hg22
|
||||
use τ
|
||||
rw [hx, hτ]
|
||||
|
||||
lemma ext_Rows {U V : CKMMatrix} (hu : [U]u = [V]u) (hc : [U]c = [V]c) (ht : [U]t = [V]t) :
|
||||
U = V := by
|
||||
apply CKMMatrix_ext
|
||||
funext i j
|
||||
fin_cases i
|
||||
have h1 := congrFun hu j
|
||||
fin_cases j <;> simpa using h1
|
||||
have h1 := congrFun hc j
|
||||
fin_cases j <;> simpa using h1
|
||||
have h1 := congrFun ht j
|
||||
fin_cases j <;> simpa using h1
|
||||
· have h1 := congrFun hu j
|
||||
fin_cases j <;> simpa using h1
|
||||
· have h1 := congrFun hc j
|
||||
fin_cases j <;> simpa using h1
|
||||
· have h1 := congrFun ht j
|
||||
fin_cases j <;> simpa using h1
|
||||
|
||||
end CKMMatrix
|
||||
|
||||
|
@ -332,9 +331,11 @@ lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
change (phaseShiftApply V a b c 0 0 0).1 0 _ = _
|
||||
simp [ud, uRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.zero_eta, cons_val_zero]
|
||||
simp [us, uRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.mk_one, cons_val_one, head_cons]
|
||||
simp [ub, uRow]
|
||||
· simp only [Fin.isValue, ud, ofReal_zero, zero_mul, add_zero, uRow, Fin.zero_eta, cons_val_zero]
|
||||
· simp only [Fin.isValue, us, ofReal_zero, zero_mul, add_zero, uRow, Fin.mk_one, cons_val_one,
|
||||
head_cons]
|
||||
· simp only [Fin.isValue, ub, ofReal_zero, zero_mul, add_zero, uRow, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons]
|
||||
|
||||
lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
[phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by
|
||||
|
@ -342,9 +343,11 @@ lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
change (phaseShiftApply V a b c 0 0 0).1 1 _ = _
|
||||
simp [cd, cRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.zero_eta, cons_val_zero]
|
||||
simp [cs, cRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.mk_one, cons_val_one, head_cons]
|
||||
simp [cb, cRow]
|
||||
· simp only [Fin.isValue, cd, ofReal_zero, zero_mul, add_zero, cRow, Fin.zero_eta, cons_val_zero]
|
||||
· simp only [Fin.isValue, cs, ofReal_zero, zero_mul, add_zero, cRow, Fin.mk_one, cons_val_one,
|
||||
head_cons]
|
||||
· simp only [Fin.isValue, cb, ofReal_zero, zero_mul, add_zero, cRow, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons]
|
||||
|
||||
lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
[phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by
|
||||
|
@ -352,9 +355,11 @@ lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
|||
simp only [Pi.smul_apply, smul_eq_mul]
|
||||
fin_cases i <;>
|
||||
change (phaseShiftApply V a b c 0 0 0).1 2 _ = _
|
||||
simp [td, tRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.zero_eta, cons_val_zero]
|
||||
simp [ts, tRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.mk_one, cons_val_one, head_cons]
|
||||
simp [tb, tRow]
|
||||
· simp only [Fin.isValue, td, ofReal_zero, zero_mul, add_zero, tRow, Fin.zero_eta, cons_val_zero]
|
||||
· simp only [Fin.isValue, ts, ofReal_zero, zero_mul, add_zero, tRow, Fin.mk_one, cons_val_one,
|
||||
head_cons]
|
||||
· simp only [Fin.isValue, tb, ofReal_zero, zero_mul, add_zero, tRow, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons]
|
||||
|
||||
end phaseShiftApply
|
||||
|
||||
|
|
|
@ -43,58 +43,60 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
rw [mul_apply]
|
||||
have h1 := exp_ne_zero (I * ↑δ₁₃)
|
||||
fin_cases j <;> rw [Fin.sum_univ_three]
|
||||
simp only [Fin.zero_eta, Fin.isValue, conjTranspose_apply, cons_val', cons_val_zero, empty_val',
|
||||
cons_val_fin_one, star_mul', RCLike.star_def, conj_ofReal, cons_val_one, head_cons, star_sub,
|
||||
star_neg, ← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const]
|
||||
simp [conj_ofReal]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq]
|
||||
ring
|
||||
simp only [Fin.mk_one, Fin.isValue, conjTranspose_apply, cons_val', cons_val_one, head_cons,
|
||||
· simp only [Fin.zero_eta, Fin.isValue, conjTranspose_apply, cons_val', cons_val_zero,
|
||||
empty_val', cons_val_fin_one, star_mul', RCLike.star_def, conj_ofReal, cons_val_one, head_cons,
|
||||
star_sub, star_neg, ← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons,
|
||||
head_fin_const]
|
||||
simp [conj_ofReal]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· simp only [Fin.mk_one, Fin.isValue, conjTranspose_apply, cons_val', cons_val_one, head_cons,
|
||||
empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def, conj_ofReal, star_sub,
|
||||
← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const, star_neg]
|
||||
simp [conj_ofReal]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq]
|
||||
ring
|
||||
simp only [Fin.reduceFinMk, Fin.isValue, conjTranspose_apply, cons_val', cons_val_two, tail_cons,
|
||||
head_cons, empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def, conj_ofReal,
|
||||
← exp_conj, map_neg, _root_.map_mul, conj_I, neg_mul, neg_neg, cons_val_one, head_fin_const]
|
||||
simp [conj_ofReal]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
simp [conj_ofReal]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· simp only [Fin.reduceFinMk, Fin.isValue, conjTranspose_apply, cons_val', cons_val_two,
|
||||
tail_cons, head_cons, empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def,
|
||||
conj_ofReal, ← exp_conj, map_neg, _root_.map_mul, conj_I, neg_mul, neg_neg, cons_val_one,
|
||||
head_fin_const]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
ring
|
||||
|
||||
/-- A CKM Matrix from four reals `θ₁₂`, `θ₁₃`, `θ₂₃`, and `δ₁₃`. This is the standard
|
||||
parameterization of CKM matrices. -/
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue