refactor: More simps
This commit is contained in:
parent
4a396783ab
commit
269f4d53a7
16 changed files with 139 additions and 62 deletions
|
@ -45,7 +45,11 @@ lemma phaseShiftMatrix_star (a b c : ℝ) :
|
|||
(phaseShiftMatrix a b c)ᴴ = phaseShiftMatrix (- a) (- b) (- c) := by
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [← exp_conj, conj_ofReal]
|
||||
simp only [phaseShiftMatrix, Fin.zero_eta, Fin.isValue, conjTranspose_apply, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, vecCons_const, star_def, ← exp_conj,
|
||||
_root_.map_mul, conj_I, conj_ofReal, neg_mul, ofReal_neg, mul_neg, Fin.mk_one, cons_val_one,
|
||||
head_fin_const, star_zero, head_cons, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
tail_val', head_val', tail_cons, Fin.reduceFinMk, map_eq_zero]
|
||||
· rfl
|
||||
· rfl
|
||||
|
||||
|
@ -53,7 +57,10 @@ 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
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three]
|
||||
simp only [phaseShiftMatrix, Fin.zero_eta, Fin.isValue, mul_apply, cons_val', empty_val',
|
||||
cons_val_fin_one, cons_val_zero, vecCons_const, Fin.sum_univ_three, cons_val_one, head_cons,
|
||||
head_fin_const, mul_zero, add_zero, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
tail_val', head_val', zero_mul, ofReal_add, Fin.mk_one, Fin.reduceFinMk, zero_add]
|
||||
any_goals rw [mul_add, exp_add]
|
||||
change cexp (I * ↑c) * 0 = 0
|
||||
simp
|
||||
|
@ -286,7 +293,11 @@ lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) :
|
|||
vecCons_const, cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons,
|
||||
head_fin_const, mul_zero]
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Complex.abs_exp]
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, zero_mul, add_zero, mul_zero,
|
||||
_root_.map_mul, abs_exp, mul_re, I_re, ofReal_re, I_im, ofReal_im, sub_self, Real.exp_zero,
|
||||
one_mul, mul_one, Fin.mk_one, cons_val_one, head_cons, zero_add, head_fin_const,
|
||||
Fin.reduceFinMk, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
tail_val', head_val']
|
||||
all_goals change Complex.abs (0 * _ + _) = _
|
||||
all_goals simp [Complex.abs_exp]
|
||||
|
||||
|
|
|
@ -273,7 +273,8 @@ lemma VAbs_sum_sq_col_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff'] at hV
|
||||
have ht := congrFun (congrFun hV i) i
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three, Fin.isValue,
|
||||
one_apply_eq] at ht
|
||||
rw [mul_comm, mul_conj, mul_comm, mul_conj, mul_comm, mul_conj] at ht
|
||||
repeat rw [← Complex.sq_abs] at ht
|
||||
rw [← ofReal_inj]
|
||||
|
@ -309,9 +310,10 @@ lemma cs_of_ud_us_zero {V : CKMMatrix} (ha : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0)) :
|
|||
have h1 := snd_row_normalized_abs V
|
||||
symm
|
||||
rw [Real.sqrt_eq_iff_eq_sq]
|
||||
· simp [not_or] at ha
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at ha
|
||||
rw [cb_eq_zero_of_ud_us_zero ha] at h1
|
||||
simp at h1
|
||||
simp only [Fin.isValue, map_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
|
||||
add_zero] at h1
|
||||
simp only [VcdAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, VcsAbs]
|
||||
rw [← h1]
|
||||
ring
|
||||
|
|
|
@ -47,7 +47,8 @@ lemma uRow_normalized (V : CKMMatrix) : conj [V]u ⬝ᵥ [V]u = 1 := by
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff] at hV
|
||||
have ht := congrFun (congrFun hV 0) 0
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [Fin.isValue, mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three,
|
||||
one_apply_eq] at ht
|
||||
rw [mul_comm (V.1 0 0) _, mul_comm (V.1 0 1) _, mul_comm (V.1 0 2) _] at ht
|
||||
exact ht
|
||||
|
||||
|
@ -215,7 +216,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
fin_cases i <;> simp
|
||||
simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul,
|
||||
uRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs] at h2
|
||||
simp at h2
|
||||
simp only [Fin.isValue, ofReal_pow, sq_eq_one_iff, ofReal_eq_one] at h2
|
||||
cases' h2 with h2 h2
|
||||
· have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
rw [← hg, smul_smul, inv_mul_cancel₀, one_smul]
|
||||
|
@ -242,7 +243,7 @@ lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
|||
obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ℂ).mp (Basis.mem_span (rowBasis V)
|
||||
(conj ([V]u) ×₃ conj ([V]c)))
|
||||
rw [Fin.sum_univ_three, rowBasis] at hg
|
||||
simp at hg
|
||||
simp only [Fin.isValue, coe_basisOfLinearIndependentOfCardEqFinrank, rows] at hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply,
|
||||
|
|
|
@ -49,7 +49,10 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
head_fin_const]
|
||||
simp only [ofReal_cos, ofReal_sin]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, one_apply_eq, Fin.mk_one, cons_val_one,
|
||||
head_cons, ne_eq, zero_ne_one, not_false_eq_true, one_apply_ne, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, Fin.reduceEq]
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
|
@ -67,7 +70,10 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const, star_neg]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, Fin.mk_one, ne_eq, one_ne_zero,
|
||||
not_false_eq_true, one_apply_ne, cons_val_one, head_cons, one_apply_eq, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, Fin.reduceEq]
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
|
@ -86,7 +92,10 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
head_fin_const]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, Fin.reduceFinMk, ne_eq, Fin.reduceEq,
|
||||
not_false_eq_true, one_apply_ne, Fin.mk_one, cons_val_one, head_cons, cons_val_two,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, one_apply_eq]
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
|
@ -148,7 +157,7 @@ lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ} (hu :
|
|||
|
||||
lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ℝ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
|
||||
standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
|
||||
simp [standParam, standParamAsMatrix]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, ofReal_sin, neg_mul]
|
||||
apply CKMMatrix_ext
|
||||
simp only
|
||||
rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]]
|
||||
|
|
|
@ -397,10 +397,16 @@ lemma on_param_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
have hS12 := congrArg ofReal (S₁₂_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h))
|
||||
simp only [ofReal_eq_coe, ← S₁₂_eq_ℂsin_θ₁₂, map_zero] at hS12
|
||||
use 0, 0, 0, δ₁₃, 0, -δ₁₃
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, hS13, hC12, hS12]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, hC12, h, ofReal_zero, mul_zero, ofReal_sin,
|
||||
hS12, hS13, neg_mul, one_mul, neg_zero, zero_mul, mul_one, zero_sub, sub_zero, phaseShift,
|
||||
phaseShiftMatrix, exp_zero, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_zero, tail_val', head_val', zero_add, Fin.mk_one, Fin.reduceFinMk,
|
||||
neg_mul, mul_one]
|
||||
· rfl
|
||||
· rfl
|
||||
|
||||
|
@ -413,7 +419,11 @@ lemma on_param_cos_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
exp_zero, mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_zero, tail_val', head_val', zero_add, zero_eq_mul, mul_neg, neg_mul,
|
||||
mul_one, neg_inj, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
field_simp
|
||||
|
@ -429,10 +439,16 @@ lemma on_param_cos_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, 0, 0, 0, - δ₁₃
|
||||
have hb := exp_ne_zero (I * δ₁₃)
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, exp_neg]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, ofReal_sin, neg_mul, exp_neg, h,
|
||||
ofReal_zero, mul_zero, zero_sub, zero_mul, sub_zero, phaseShift, phaseShiftMatrix, exp_zero,
|
||||
mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
zero_add, mul_neg, neg_inj, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· ring_nf
|
||||
change _ = _ + _ * 0
|
||||
|
@ -448,7 +464,11 @@ lemma on_param_sin_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
Submonoid.mk_mul_mk, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', zero_add, self_eq_add_right,
|
||||
mul_eq_zero, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· exact Or.inr rfl
|
||||
|
||||
|
@ -456,10 +476,16 @@ lemma on_param_sin_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, δ₁₃, 0, -δ₁₃, - δ₁₃
|
||||
have hb := exp_ne_zero (I * δ₁₃)
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, exp_neg]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, h, ofReal_zero, zero_mul, ofReal_sin,
|
||||
neg_mul, exp_neg, neg_zero, zero_sub, sub_zero, phaseShift, phaseShiftMatrix, mul_zero,
|
||||
exp_zero, mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
Fin.mk_one, zero_add, Fin.reduceFinMk, mul_neg, neg_mul, neg_inj]
|
||||
· apply Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
simp only [mul_zero, add_zero, neg_inj]
|
||||
|
@ -484,7 +510,11 @@ lemma on_param_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
Fin.mk_one, zero_add, Fin.reduceFinMk, mul_neg, neg_inj]
|
||||
· exact Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
simp only [mul_zero, add_zero, neg_inj]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue