refactor: Lint
This commit is contained in:
parent
6c17a61989
commit
b0c5ed894f
8 changed files with 16 additions and 11 deletions
|
@ -123,7 +123,6 @@ instance (M : LorentzGroup d) : Invertible M.1 where
|
|||
rw [← coe_inv]
|
||||
exact (mem_iff_self_mul_dual.mp M.2)
|
||||
|
||||
@[simp]
|
||||
lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
|
||||
trans Subtype.val (Λ⁻¹ * Λ)
|
||||
· rw [← coe_inv]
|
||||
|
@ -131,7 +130,6 @@ lemma subtype_inv_mul : (Subtype.val Λ)⁻¹ * (Subtype.val Λ) = 1 := by
|
|||
· rw [inv_mul_cancel Λ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma subtype_mul_inv : (Subtype.val Λ) * (Subtype.val Λ)⁻¹ = 1 := by
|
||||
trans Subtype.val (Λ * Λ⁻¹)
|
||||
· rw [← coe_inv]
|
||||
|
|
|
@ -145,7 +145,7 @@ lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
|
|||
rw [contrContrContractField.matrix_apply_stdBasis (Λ := toMatrix u v) μ ν, toMatrix_mulVec]
|
||||
simp only [genBoost, genBoostAux₁, genBoostAux₂, smul_add, neg_smul, LinearMap.add_apply,
|
||||
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, contrContrContractField.basis_left,
|
||||
map_add, map_smul, map_neg, toField_apply, mul_eq_mul_left_iff]
|
||||
map_add, map_smul, map_neg, mul_eq_mul_left_iff]
|
||||
ring_nf
|
||||
simp only [Pi.add_apply, Action.instMonoidalCategory_tensorObj_V,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryTheory.Equivalence.symm_inverse,
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue