Merge pull request #136 from HEPLean/simp_replace
refactor: Replace tactics with rfl
This commit is contained in:
commit
35def45dcc
22 changed files with 70 additions and 97 deletions
|
@ -97,8 +97,7 @@ def quadCoeff (T : MSSMACC.Sols) : ℚ :=
|
|||
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔ quadCoeff T = 0 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· rw [quadCoeff, h.1, h.2]
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
||||
zero_pow, add_zero, mul_zero]
|
||||
rfl
|
||||
· rw [quadCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
||||
|
@ -145,8 +144,7 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
|||
InCubeSolProp T ↔ cubicCoeff T = 0 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· rw [cubicCoeff, h.1, h.2]
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
||||
zero_pow, add_zero, mul_zero]
|
||||
rfl
|
||||
· rw [cubicCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
||||
|
|
|
@ -207,14 +207,17 @@ lemma basis_δ₂_eq_minus_δ₁ (j i : Fin n.succ) :
|
|||
simp [basisAsCharges, δ₂, δ₁]
|
||||
split <;> split
|
||||
any_goals split
|
||||
any_goals rfl
|
||||
any_goals split
|
||||
any_goals rfl
|
||||
all_goals rename_i h1 h2
|
||||
all_goals rw [Fin.ext_iff] at h1 h2
|
||||
all_goals simp_all
|
||||
all_goals rename_i h3
|
||||
all_goals rw [Fin.ext_iff] at h3
|
||||
all_goals simp_all
|
||||
all_goals
|
||||
rename_i h1 h2
|
||||
rw [Fin.ext_iff] at h1 h2
|
||||
simp_all
|
||||
all_goals
|
||||
rename_i h3
|
||||
rw [Fin.ext_iff] at h3
|
||||
simp_all
|
||||
all_goals omega
|
||||
|
||||
lemma basis!_δ!₂_eq_minus_δ!₁ (j i : Fin n) :
|
||||
|
|
|
@ -111,12 +111,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
|||
simp at h2
|
||||
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0) := by
|
||||
linear_combination -(1 * h1) - 1 * h4 - 1 * h2
|
||||
rw [h5]
|
||||
rw [δa₄_δ!₂]
|
||||
have h0 : (δa₂ (0 : Fin n.succ)) = δ!₁ 0 := by
|
||||
rw [δa₂_δ!₁]
|
||||
exact ht
|
||||
rw [h0, δa₁_δ!₃]
|
||||
rw [h5, δa₄_δ!₂, show (δa₂ (0 : Fin n.succ)) = δ!₁ 0 from rfl, δa₁_δ!₃]
|
||||
ring
|
||||
|
||||
lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
|
||||
|
|
|
@ -107,7 +107,7 @@ lemma accGrav_ext {S T : (SMCharges n).Charges}
|
|||
AddHom.coe_mk]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
repeat erw [hj]
|
||||
erw [hj, hj, hj, hj, hj]
|
||||
rfl
|
||||
|
||||
/-- The `SU(2)` anomaly equation. -/
|
||||
|
@ -166,7 +166,7 @@ lemma accSU3_ext {S T : (SMCharges n).Charges}
|
|||
AddHom.coe_mk]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
repeat erw [hj]
|
||||
erw [hj, hj, hj]
|
||||
rfl
|
||||
|
||||
/-- The `Y²` anomaly equation. -/
|
||||
|
@ -197,7 +197,7 @@ lemma accYY_ext {S T : (SMCharges n).Charges}
|
|||
AddHom.coe_mk]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
repeat erw [hj]
|
||||
erw [hj, hj, hj, hj, hj]
|
||||
rfl
|
||||
|
||||
/-- The quadratic bilinear map. -/
|
||||
|
|
|
@ -299,7 +299,7 @@ lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.
|
|||
by_contra hn
|
||||
have h : s ^ 2 < 0 := by
|
||||
rw [← hn]
|
||||
simp [discrim]
|
||||
rfl
|
||||
nlinarith
|
||||
simp_all
|
||||
exact eq_neg_of_add_eq_zero_left h'
|
||||
|
@ -317,7 +317,7 @@ lemma cube_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.v
|
|||
by_contra hn
|
||||
have h : s ^ 2 < 0 := by
|
||||
rw [← hn]
|
||||
simp [discrim]
|
||||
rfl
|
||||
nlinarith
|
||||
simp_all
|
||||
exact eq_neg_of_add_eq_zero_left h'
|
||||
|
|
|
@ -51,7 +51,7 @@ def repCharges {n : ℕ} : Representation ℚ (PermGroup n) (SMCharges n).Charge
|
|||
rw [charges_eq_toSpecies_eq]
|
||||
intro i
|
||||
simp only [chargeMap_apply, Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
|
||||
repeat erw [toSMSpecies_toSpecies_inv]
|
||||
rw [toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
map_one' := by
|
||||
apply LinearMap.ext
|
||||
|
|
|
@ -109,10 +109,8 @@ lemma sum_familyUniversal {n : ℕ} (m : ℕ) (S : (SMνCharges 1).Charges) (j :
|
|||
rw [Fin.sum_const]
|
||||
simp
|
||||
erw [h1]
|
||||
apply Finset.sum_congr
|
||||
· simp only
|
||||
· intro i _
|
||||
erw [toSpecies_familyUniversal]
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
erw [toSpecies_familyUniversal]
|
||||
|
||||
lemma sum_familyUniversal_one {n : ℕ} (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||
∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by
|
||||
|
@ -124,11 +122,9 @@ lemma sum_familyUniversal_two {n : ℕ} (S : (SMνCharges 1).Charges)
|
|||
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by
|
||||
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
|
||||
rw [Finset.mul_sum]
|
||||
apply Finset.sum_congr
|
||||
· simp only
|
||||
· intro i _
|
||||
erw [toSpecies_familyUniversal]
|
||||
rfl
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
erw [toSpecies_familyUniversal]
|
||||
rfl
|
||||
|
||||
lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges)
|
||||
(T L : (SMνCharges n).Charges) (j : Fin 6) :
|
||||
|
|
|
@ -192,50 +192,43 @@ lemma B₀_B₀_Bi_cubic {i : Fin 7} :
|
|||
cubeTriLin (B 0) (B 0) (B i) = 0 := by
|
||||
change cubeTriLin (B₀) (B₀) (B i) = 0
|
||||
rw [B₀_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₁_B₁_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 1) (B 1) (B i) = 0 := by
|
||||
change cubeTriLin (B₁) (B₁) (B i) = 0
|
||||
rw [B₁_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₂_B₂_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 2) (B 2) (B i) = 0 := by
|
||||
change cubeTriLin (B₂) (B₂) (B i) = 0
|
||||
rw [B₂_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₃_B₃_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 3) (B 3) (B i) = 0 := by
|
||||
change cubeTriLin (B₃) (B₃) (B i) = 0
|
||||
rw [B₃_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₄_B₄_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 4) (B 4) (B i) = 0 := by
|
||||
change cubeTriLin (B₄) (B₄) (B i) = 0
|
||||
rw [B₄_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₅_B₅_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 5) (B 5) (B i) = 0 := by
|
||||
change cubeTriLin (B₅) (B₅) (B i) = 0
|
||||
rw [B₅_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma B₆_B₆_Bi_cubic {i : Fin 7} :
|
||||
cubeTriLin (B 6) (B 6) (B i) = 0 := by
|
||||
change cubeTriLin (B₆) (B₆) (B i) = 0
|
||||
rw [B₆_cubic]
|
||||
fin_cases i <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
fin_cases i <;> rfl
|
||||
|
||||
lemma Bi_Bi_Bj_cubic (i j : Fin 7) :
|
||||
cubeTriLin (B i) (B i) (B j) = 0 := by
|
||||
|
|
|
@ -30,12 +30,8 @@ def familyUniversalLinear (n : ℕ) :
|
|||
(by rw [familyUniversal_accSU2, SU2Sol S, mul_zero])
|
||||
(by rw [familyUniversal_accSU3, SU3Sol S, mul_zero])
|
||||
(by rw [familyUniversal_accYY, YYsol S, mul_zero])
|
||||
map_add' S T := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
exact (familyUniversal n).map_add' _ _
|
||||
map_smul' a S := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
exact (familyUniversal n).map_smul' _ _
|
||||
map_add' S T := rfl
|
||||
map_smul' a S := rfl
|
||||
|
||||
/-- The family universal maps on `QuadSols`. -/
|
||||
def familyUniversalQuad (n : ℕ) :
|
||||
|
|
|
@ -201,7 +201,7 @@ lemma isSolution_f9 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i
|
|||
lemma isSolution_f10 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
f 10 = 0 := by
|
||||
rw [isSolution_grav f hS, isSolution_f9 f hS]
|
||||
simp
|
||||
rfl
|
||||
|
||||
lemma isSolution_f_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
|
||||
(k : Fin 11) : f k = 0 := by
|
||||
|
|
|
@ -85,7 +85,7 @@ lemma figureEight_connected : Connected figureEight := by decide
|
|||
|
||||
/-- The symmetry factor of `figureEight` is 8. We can get this from
|
||||
`#eval symmetryFactor figureEight`. -/
|
||||
lemma figureEight_symmetryFactor : symmetryFactor figureEight = 8 := by decide
|
||||
lemma figureEight_symmetryFactor : symmetryFactor figureEight = 8 := by rfl
|
||||
|
||||
end Example
|
||||
|
||||
|
|
|
@ -31,9 +31,15 @@ def phaseShiftMatrix (a b c : ℝ) : Matrix (Fin 3) (Fin 3) ℂ :=
|
|||
|
||||
lemma phaseShiftMatrix_one : phaseShiftMatrix 0 0 0 = 1 := by
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Matrix.one_apply, phaseShiftMatrix]
|
||||
rfl
|
||||
fin_cases i <;> fin_cases j
|
||||
any_goals rfl
|
||||
· simp only [phaseShiftMatrix, ofReal_zero, mul_zero, exp_zero, Fin.zero_eta, Fin.isValue,
|
||||
cons_val', cons_val_zero, empty_val', cons_val_fin_one, vecCons_const, one_apply_eq]
|
||||
· simp only [phaseShiftMatrix, ofReal_zero, mul_zero, exp_zero, Fin.mk_one, Fin.isValue,
|
||||
cons_val', cons_val_one, head_cons, empty_val', cons_val_fin_one, one_apply_eq]
|
||||
· simp only [phaseShiftMatrix, ofReal_zero, mul_zero, exp_zero, Fin.reduceFinMk, cons_val',
|
||||
Fin.isValue, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons, empty_val',
|
||||
cons_val_fin_one, head_fin_const, one_apply_eq]
|
||||
|
||||
lemma phaseShiftMatrix_star (a b c : ℝ) :
|
||||
(phaseShiftMatrix a b c)ᴴ = phaseShiftMatrix (- a) (- b) (- c) := by
|
||||
|
@ -96,14 +102,10 @@ lemma phaseShiftRelation_trans {U V W : unitaryGroup (Fin 3) ℂ} :
|
|||
obtain ⟨a, b, c, e, f, g, hUV⟩ := hUV
|
||||
obtain ⟨d, i, j, k, l, m, hVW⟩ := hVW
|
||||
use (a + d), (b + i), (c + j), (e + k), (f + l), (g + m)
|
||||
rw [Subtype.ext_iff_val]
|
||||
rw [hUV, hVW]
|
||||
rw [Subtype.ext_iff_val, hUV, hVW]
|
||||
simp only [Submonoid.coe_mul, phaseShift_coe_matrix]
|
||||
repeat rw [mul_assoc]
|
||||
rw [phaseShiftMatrix_mul]
|
||||
repeat rw [← mul_assoc]
|
||||
rw [phaseShiftMatrix_mul]
|
||||
rw [add_comm k e, add_comm l f, add_comm m g]
|
||||
rw [mul_assoc, mul_assoc, phaseShiftMatrix_mul, ← mul_assoc, ← mul_assoc, phaseShiftMatrix_mul,
|
||||
add_comm k e, add_comm l f, add_comm m g]
|
||||
|
||||
lemma phaseShiftRelation_equiv : Equivalence PhaseShiftRelation where
|
||||
refl := phaseShiftRelation_refl
|
||||
|
|
|
@ -215,12 +215,12 @@ lemma ubOnePhaseCond_shift_solution {V : CKMMatrix} (h1 : a + f = - arg [V]ub)
|
|||
have hd : d = Real.pi - arg [V]cd - b := by
|
||||
linear_combination h3
|
||||
subst he hd
|
||||
simp_all
|
||||
simp_all only [add_sub_cancel, and_self, and_true]
|
||||
ring_nf at h2
|
||||
have hc : c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b := by
|
||||
linear_combination h2
|
||||
subst hc
|
||||
ring
|
||||
rfl
|
||||
|
||||
lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
||||
∃ (U : CKMMatrix), V ≈ U ∧ FstRowThdColRealCond U:= by
|
||||
|
|
|
@ -200,7 +200,6 @@ lemma contrRightAux_comp {V1 V2 V3 V4 V5 : Type} [AddCommMonoid V1] [AddCommMono
|
|||
by simp [add_tmul, tmul_add, LinearMap.map_add, h1, h2])
|
||||
intro x3 x4
|
||||
simp [contrRightAux, contrMidAux, contrLeftAux]
|
||||
erw [TensorProduct.map_tmul]
|
||||
rfl
|
||||
|
||||
end TensorStructure
|
||||
|
@ -433,7 +432,7 @@ lemma elimPureTensor_update_left (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor c
|
|||
split_ifs
|
||||
· rename_i h
|
||||
subst h
|
||||
simp_all only
|
||||
rfl
|
||||
· rfl
|
||||
| Sum.inr y => rfl
|
||||
|
||||
|
@ -453,10 +452,10 @@ lemma inlPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Su
|
|||
funext y
|
||||
simp [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
next h =>
|
||||
· rename_i h
|
||||
subst h
|
||||
simp_all only
|
||||
rfl
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
@[simp]
|
||||
lemma inrPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (x : X)
|
||||
|
@ -473,10 +472,10 @@ lemma inrPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (S
|
|||
funext y
|
||||
simp [inrPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
next h =>
|
||||
· rename_i h
|
||||
subst h
|
||||
simp_all only
|
||||
rfl
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
@[simp]
|
||||
lemma inlPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (y : Y)
|
||||
|
|
|
@ -289,7 +289,6 @@ lemma contrAll_mapIso_right_tmul (e : X ≃ Y) (e' : Z ≃ Y)
|
|||
𝓣.contrAll e h (x ⊗ₜ[R] 𝓣.mapIso e' h' z) =
|
||||
𝓣.contrAll (e.trans e'.symm) (h.trans_mapIso h') (x ⊗ₜ[R] z) := by
|
||||
simp only [contrAll_tmul, mapIso_mapIso]
|
||||
apply congrArg
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
|
@ -429,7 +428,6 @@ lemma contr_tprod (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e) (f : (i : X)
|
|||
Function.comp_apply, LinearEquiv.trans_apply, mapIso_tprod, Equiv.symm_symm_apply,
|
||||
tensoratorEquiv_symm_tprod, congr_tmul, LinearEquiv.refl_apply, map_tmul, LinearMap.id_coe,
|
||||
id_eq, lid_tmul]
|
||||
rw [contrAll_tmul]
|
||||
rfl
|
||||
|
||||
open PiTensorProduct in
|
||||
|
|
|
@ -86,8 +86,7 @@ lemma listCharIndex_iff (l : List Char) : listCharIndex X l
|
|||
↔ (if h : l = [] then True else
|
||||
let sfst := l.head h
|
||||
if ¬ isNotationChar X sfst then False
|
||||
else listCharIndexTail sfst l.tail) := by
|
||||
rw [listCharIndex]
|
||||
else listCharIndexTail sfst l.tail) := by rfl
|
||||
|
||||
instance : Decidable (listCharIndex X l) :=
|
||||
@decidable_of_decidable_of_iff _ _
|
||||
|
|
|
@ -457,12 +457,9 @@ lemma mem_iff_dual_mem (hl : l.OnlyUniqueDuals) (hc : l.ColorCond) (I : Index
|
|||
omega
|
||||
· have hIdd : l.countId I.dual = 2 := by
|
||||
rw [← hId]
|
||||
apply countId_congr
|
||||
rfl
|
||||
have hc' := (hc I.dual h hIdd).2
|
||||
simp at hc'
|
||||
rw [← countSelf_neq_zero]
|
||||
rw [← countSelf_neq_zero] at h
|
||||
rw [← countSelf_neq_zero] at h ⊢
|
||||
rw [countDual_eq_countSelf_Dual] at hc'
|
||||
simp at hc'
|
||||
exact Ne.symm (ne_of_ne_of_eq (id (Ne.symm h)) hc')
|
||||
|
|
|
@ -45,7 +45,6 @@ def contrIndexList : IndexList X where
|
|||
|
||||
@[simp]
|
||||
lemma contrIndexList_empty : (⟨[]⟩ : IndexList X).contrIndexList = { val := [] } := by
|
||||
apply ext
|
||||
rfl
|
||||
|
||||
lemma contrIndexList_val : l.contrIndexList.val =
|
||||
|
|
|
@ -113,8 +113,7 @@ lemma countId_get_other (i : Fin l.length) : l2.countId (l.val.get i) =
|
|||
simp only [length, List.finRange_map_get]
|
||||
nth_rewrite 1 [hl2]
|
||||
rw [List.filter_map, List.length_map]
|
||||
apply congrArg
|
||||
refine List.filter_congr (fun j _ => rfl)
|
||||
rfl
|
||||
|
||||
/-! TODO: Replace with mathlib lemma. -/
|
||||
lemma filter_finRange (i : Fin l.length) :
|
||||
|
@ -269,7 +268,7 @@ lemma countId_eq_two_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withU
|
|||
((List.finRange l.length).filter fun j => l.AreDualInSelf i j)
|
||||
· exact List.filter_comm (fun j => l.AreDualInSelf i j) (fun j => j = i')
|
||||
(List.finRange l.length)
|
||||
· simp
|
||||
· simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and]
|
||||
refine List.filter_congr (fun j _ => ?_)
|
||||
simp only [Bool.and_iff_right_iff_imp, decide_eq_true_eq]
|
||||
simp [withUniqueDual] at h
|
||||
|
@ -279,7 +278,7 @@ lemma countId_eq_two_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withU
|
|||
rw [hj']
|
||||
simp [i']
|
||||
rw [List.countP_eq_length_filter, ← h1]
|
||||
simp
|
||||
rfl
|
||||
|
||||
lemma mem_withUniqueDual_of_countId_eq_two (i : Fin l.length)
|
||||
(h : l.countId (l.val.get i) = 2) : i ∈ l.withUniqueDual := by
|
||||
|
|
|
@ -141,10 +141,12 @@ lemma getDualInOtherEquiv_self_refl : l.getDualInOtherEquiv l = Equiv.refl _ :=
|
|||
apply Subtype.eq
|
||||
simp only
|
||||
have hx2 := x.2
|
||||
simp [withUniqueDualInOther] at hx2
|
||||
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, getDualInOther?_self_isSome,
|
||||
true_and, Finset.mem_filter, Finset.mem_univ] at hx2
|
||||
apply Option.some_injective
|
||||
rw [hx2.2 x.1 (by simp [AreDualInOther])]
|
||||
simp
|
||||
rw [hx2.2 x.1 (by rfl)]
|
||||
exact Option.some_get (getDualInOtherEquiv.proof_1 l l x)
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOtherEquiv_symm : (l.getDualInOtherEquiv l2).symm = l2.getDualInOtherEquiv l := by
|
||||
|
|
|
@ -215,8 +215,7 @@ lemma symm {T₁ T₂ : 𝓣.TensorIndex} (h : Rel T₁ T₂) : Rel T₂ T₁ :=
|
|||
· symm
|
||||
erw [LinearEquiv.symm_apply_eq]
|
||||
rw [h.2]
|
||||
· apply congrFun
|
||||
congr
|
||||
· rfl
|
||||
exact h'.symm
|
||||
|
||||
/-- Rel is transitive. -/
|
||||
|
@ -282,10 +281,7 @@ lemma rel_contr (T : 𝓣.TensorIndex) : T ≈ T.contr := by
|
|||
Fin.symm_castOrderIso, mapIso_mapIso, tensorIso]
|
||||
trans 𝓣.mapIso (Equiv.refl _) (by rfl) T.contr.tensor
|
||||
· simp only [contr_toColorIndexList, mapIso_refl, LinearEquiv.refl_apply]
|
||||
· apply congrFun
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -142,7 +142,7 @@ lemma as_sum :
|
|||
Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, Sum.elim_inl, one_mul,
|
||||
Finset.sum_singleton, Sum.elim_inr, neg_mul, mul_neg, Finset.sum_neg_distrib, time, space,
|
||||
Function.comp_apply, minkowskiMatrix]
|
||||
ring
|
||||
rfl
|
||||
|
||||
/-- The Minkowski metric expressed as a sum for a single vector. -/
|
||||
lemma as_sum_self : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by
|
||||
|
@ -151,7 +151,7 @@ lemma as_sum_self : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by
|
|||
|
||||
lemma eq_time_minus_inner_prod : ⟪v, w⟫ₘ = v.time * w.time - ⟪v.space, w.space⟫_ℝ := by
|
||||
rw [as_sum, @PiLp.inner_apply]
|
||||
noncomm_ring
|
||||
rfl
|
||||
|
||||
lemma self_eq_time_minus_norm : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by
|
||||
rw [← real_inner_self_eq_norm_sq, PiLp.inner_apply, as_sum]
|
||||
|
@ -192,6 +192,7 @@ lemma self_spaceReflection_eq_zero_iff : ⟪v, v.spaceReflection⟫ₘ = 0 ↔ v
|
|||
· simpa using congrFun h4 i
|
||||
· rw [h]
|
||||
exact LinearMap.map_zero₂ minkowskiMetric (spaceReflection 0)
|
||||
|
||||
/-!
|
||||
|
||||
# Non degeneracy of the Minkowski metric
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue