diff --git a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean index d0d6586..c4f6404 100644 --- a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean +++ b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean @@ -68,7 +68,7 @@ lemma mem_of_transpose_eta_eq_eta_mul_self {A : Matrix (Fin 4) (Fin 4) ℝ} simpa using h1 -lemma mem_iff (A : Matrix (Fin 4) (Fin 4) ℝ) : A ∈ lorentzAlgebra ↔ +lemma mem_iff {A : Matrix (Fin 4) (Fin 4) ℝ} : A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A := by apply Iff.intro · intro h @@ -76,7 +76,15 @@ lemma mem_iff (A : Matrix (Fin 4) (Fin 4) ℝ) : A ∈ lorentzAlgebra ↔ · intro h exact mem_of_transpose_eta_eq_eta_mul_self h - +lemma mem_iff' (A : Matrix (Fin 4) (Fin 4) ℝ) : A ∈ lorentzAlgebra ↔ A = - η * Aᵀ * η := by + apply Iff.intro + intro h + rw [mul_assoc, mem_iff.mp h] + simp only [neg_mul, mul_neg, ← mul_assoc, η_sq, one_mul, neg_neg] + intro h + rw [mem_iff] + nth_rewrite 2 [h] + simp [← mul_assoc, η_sq] end lorentzAlgebra @@ -87,7 +95,7 @@ instance spaceTimeAsLieRingModule : LieRingModule lorentzAlgebra spaceTime where add_lie Λ1 Λ2 x := by simp [add_mulVec] lie_add Λ x1 x2 := by - simp + simp only exact mulVec_add _ _ _ leibniz_lie Λ1 Λ2 x := by simp [mulVec_add, Bracket.bracket, sub_mulVec] diff --git a/HepLean/SpaceTime/Metric.lean b/HepLean/SpaceTime/Metric.lean index 5655d61..6bdcf0c 100644 --- a/HepLean/SpaceTime/Metric.lean +++ b/HepLean/SpaceTime/Metric.lean @@ -248,9 +248,6 @@ lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ℝ) : /-- The metric as a quadratic form on `spaceTime`. -/ def quadraticForm : QuadraticForm ℝ spaceTime := ηLin.toQuadraticForm -@[simps!] -def ηTensor : (spaceTime ⊗[ℝ] spaceTime) →ₗ[ℝ] ℝ := - TensorProduct.lift ηLin