refactor: Simp lemmas
This commit is contained in:
parent
cd1b30c069
commit
1651b265e7
16 changed files with 116 additions and 86 deletions
|
@ -116,8 +116,8 @@ lemma phaseShiftRelation_equiv : Equivalence PhaseShiftRelation where
|
|||
def CKMMatrix : Type := unitaryGroup (Fin 3) ℂ
|
||||
|
||||
lemma CKMMatrix_ext {U V : CKMMatrix} (h : U.val = V.val) : U = V := by
|
||||
cases U; cases V
|
||||
simp at h
|
||||
cases U
|
||||
cases V
|
||||
subst h
|
||||
rfl
|
||||
|
||||
|
|
|
@ -62,7 +62,8 @@ lemma shift_ub_phase_zero {V : CKMMatrix} (h1 : u + b = - arg [V]ub) :
|
|||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VubAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cs_phase_zero {V : CKMMatrix} (h1 : c + s = - arg [V]cs) :
|
||||
|
@ -71,10 +72,11 @@ lemma shift_cs_phase_zero {V : CKMMatrix} (h1 : c + s = - arg [V]cs) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]cs]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]cs) * I + (↑c * I + ↑s * I) = ↑(arg [V]cs + (c + s)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VcsAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cb_phase_zero {V : CKMMatrix} (h1 : c + b = - arg [V]cb) :
|
||||
|
@ -83,10 +85,11 @@ lemma shift_cb_phase_zero {V : CKMMatrix} (h1 : c + b = - arg [V]cb) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]cb]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]cb) * I + (↑c * I + ↑b * I) = ↑(arg [V]cb + (c + b)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VcbAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_tb_phase_zero {V : CKMMatrix} (h1 : t + b = - arg [V]tb) :
|
||||
|
@ -95,10 +98,11 @@ lemma shift_tb_phase_zero {V : CKMMatrix} (h1 : t + b = - arg [V]tb) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]tb]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]tb) * I + (↑t * I + ↑b * I) = ↑(arg [V]tb + (t + b)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VtbAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cd_phase_pi {V : CKMMatrix} (h1 : c + d = Real.pi - arg [V]cd) :
|
||||
|
@ -110,7 +114,8 @@ lemma shift_cd_phase_pi {V : CKMMatrix} (h1 : c + d = Real.pi - arg [V]cd) :
|
|||
simp [add_assoc]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_sub_cancel, exp_pi_mul_I, mul_neg, mul_one, VcdAbs, neg_inj,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ}
|
||||
|
@ -120,33 +125,35 @@ lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ}
|
|||
change _ = phaseShiftApply.ucCross _ _ _ _ _ _ _
|
||||
funext i
|
||||
fin_cases i
|
||||
· simp
|
||||
· simp only [Fin.zero_eta, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_fst]
|
||||
simp [tRow, phaseShiftApply.td]
|
||||
have hτ0 := congrFun hτ 0
|
||||
simp [tRow] at hτ0
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_zero] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
· simp
|
||||
· simp only [Fin.mk_one, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_snd]
|
||||
simp [tRow, phaseShiftApply.ts]
|
||||
simp only [tRow, Fin.isValue, phaseShiftApply_coe, phaseShiftApply.ts, cons_val_one, head_cons,
|
||||
neg_mul]
|
||||
have hτ0 := congrFun hτ 1
|
||||
simp [tRow] at hτ0
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_one, head_cons] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
· simp
|
||||
· simp only [Fin.reduceFinMk, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_thd]
|
||||
simp [tRow, phaseShiftApply.tb]
|
||||
simp only [tRow, Fin.isValue, phaseShiftApply_coe, phaseShiftApply.tb, cons_val_two,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons, neg_mul]
|
||||
have hτ0 := congrFun hτ 2
|
||||
simp [tRow] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_two, Nat.succ_eq_add_one,
|
||||
Nat.reduceAdd, tail_cons, head_cons] at hτ0
|
||||
rw [← hτ0, ← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
|
@ -271,21 +278,21 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
|||
apply And.intro
|
||||
· exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
· apply And.intro
|
||||
· simp [not_or] at hb
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at hb
|
||||
have h1 : VudAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
simp only [VudAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, map_eq_zero] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] 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
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at hb
|
||||
have h3 := cb_eq_zero_of_ud_us_zero hb
|
||||
have h1 : VcbAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
|
@ -325,7 +332,8 @@ lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠
|
|||
exact hV.2.2.2.2
|
||||
rw [cd_of_ud_us_ub_cb_tb hb hτ]
|
||||
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp [sq, conj_ofReal]
|
||||
simp only [Fin.isValue, VudAbs, VcbAbs, ofReal_zero, zero_mul, exp_zero, VtbAbs, conj_ofReal,
|
||||
one_mul, VusAbs, neg_add_rev, normSq_ofReal, ofReal_mul, neg_mul, sq, VubAbs]
|
||||
have hx := Vabs_sq_add_neq_zero hb
|
||||
field_simp
|
||||
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
|
||||
|
@ -343,15 +351,16 @@ lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠
|
|||
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
|
||||
rw [cs_of_ud_us_ub_cb_tb hb hτ]
|
||||
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp [sq, conj_ofReal]
|
||||
rw [cs_of_ud_us_ub_cb_tb hb hτ, hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp only [Fin.isValue, VusAbs, neg_mul, VcbAbs, ofReal_zero, zero_mul, exp_zero, VtbAbs,
|
||||
conj_ofReal, one_mul, VudAbs, normSq_ofReal, ofReal_mul, sq, VubAbs]
|
||||
have hx := Vabs_sq_add_neq_zero hb
|
||||
field_simp
|
||||
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
|
||||
nth_rewrite 1 [← abs_mul_exp_arg_mul_I [V]ub]
|
||||
rw [@RingHom.map_mul]
|
||||
simp [conj_ofReal, ← exp_conj, VAbs]
|
||||
simp only [Fin.isValue, conj_ofReal, ← exp_conj, _root_.map_mul, conj_I, mul_neg, VubAbs, VAbs,
|
||||
VAbs', Quotient.lift_mk, neg_mul]
|
||||
rw [h1]
|
||||
ring_nf
|
||||
|
||||
|
|
|
@ -30,7 +30,8 @@ lemma VAbs_sum_sq_row_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_conj, mul_conj, mul_conj] at ht
|
||||
repeat rw [← Complex.sq_abs] at ht
|
||||
rw [← ofReal_inj]
|
||||
|
|
|
@ -308,21 +308,26 @@ def ucCross : Fin 3 → ℂ :=
|
|||
|
||||
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
|
||||
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]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow,
|
||||
phaseShiftApply_coe, us, exp_add, ub, cRow, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply,
|
||||
cons_val_one, head_cons, _root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two,
|
||||
tail_cons, cons_val_zero, neg_mul]
|
||||
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
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add,
|
||||
exp_sub, ← exp_conj, conj_ofReal]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow, ud,
|
||||
exp_add, us, ub, cRow, cd, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_one, head_cons,
|
||||
_root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two, tail_cons,
|
||||
cons_val_zero, neg_mul]
|
||||
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
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub,
|
||||
← exp_conj, conj_ofReal]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow, ud,
|
||||
exp_add, us, ub, cRow, cd, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_one, head_cons,
|
||||
_root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two, tail_cons,
|
||||
cons_val_zero, neg_mul]
|
||||
ring
|
||||
|
||||
lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
|
|
|
@ -47,7 +47,7 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
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]
|
||||
simp only [ofReal_cos, ofReal_sin]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
|
@ -65,7 +65,7 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
· 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]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
· ring_nf
|
||||
|
@ -169,15 +169,16 @@ lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Rea
|
|||
neg_zero, Real.exp_zero, mul_one, _root_.sq_abs]
|
||||
rw [_root_.abs_of_nonneg h1, _root_.abs_of_nonneg h3, _root_.abs_of_nonneg h2,
|
||||
_root_.abs_of_nonneg h4]
|
||||
simp [sq]
|
||||
simp only [sq]
|
||||
ring_nf
|
||||
nth_rewrite 2 [Real.sin_sq θ₁₂]
|
||||
ring_nf
|
||||
field_simp
|
||||
ring
|
||||
· simp at hx
|
||||
· simp only [ne_eq, Decidable.not_not] at hx
|
||||
rw [hx]
|
||||
simp
|
||||
simp only [abs_zero, mul_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
|
||||
zero_mul, add_zero, div_zero]
|
||||
|
||||
open Invariant in
|
||||
lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂)
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue