chore: Bump to lean v.4.12.0
This commit is contained in:
parent
47d9649c44
commit
987bbf6013
23 changed files with 92 additions and 58 deletions
|
@ -31,7 +31,8 @@ open minkowskiMatrix
|
|||
lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by
|
||||
have h1 := A.2
|
||||
erw [mem_skewAdjointMatricesLieSubalgebra] at h1
|
||||
simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h1
|
||||
simpa only [neg_mul, mem_skewAdjointMatricesSubmodule, IsSkewAdjoint, IsAdjointPair,
|
||||
mul_neg] using h1
|
||||
|
||||
lemma mem_of_transpose_eta_eq_eta_mul_self {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ}
|
||||
(h : Aᵀ * η = - η * A) : A ∈ lorentzAlgebra := by
|
||||
|
@ -96,8 +97,7 @@ instance lorentzVectorAsLieRingModule : LieRingModule lorentzAlgebra (LorentzVec
|
|||
|
||||
@[simps!]
|
||||
instance spaceTimeAsLieModule : LieModule ℝ lorentzAlgebra (LorentzVector 3) where
|
||||
smul_lie r Λ x := by
|
||||
simp [Bracket.bracket, smul_mulVec_assoc]
|
||||
smul_lie r Λ x := smul_mulVec_assoc r Λ.1 x
|
||||
lie_smul r Λ x := by
|
||||
simp only [Bracket.bracket]
|
||||
rw [mulVec_smul]
|
||||
|
|
|
@ -98,7 +98,7 @@ lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) :
|
|||
LinearMap.id_apply, LinearMap.coe_mk, AddHom.coe_mk, basis_left, map_smul, smul_eq_mul, map_neg,
|
||||
mul_eq_mul_left_iff]
|
||||
ring_nf
|
||||
exact (true_or_iff (η μ μ = 0)).mpr trivial
|
||||
exact (true_or (η μ μ = 0)).mpr trivial
|
||||
|
||||
open minkowskiMatrix LorentzVector in
|
||||
lemma toMatrix_continuous (u : FuturePointing d) : Continuous (toMatrix u) := by
|
||||
|
|
|
@ -98,20 +98,36 @@ noncomputable def toSelfAdjointMatrix :
|
|||
· rw [show (x + y) (Sum.inl 0) = x (Sum.inl 0) + y (Sum.inl 0) from rfl]
|
||||
rw [show (x + y) (Sum.inr 2) = x (Sum.inr 2) + y (Sum.inr 2) from rfl]
|
||||
simp only [Fin.isValue, ofReal_add, Fin.zero_eta, cons_val_zero]
|
||||
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
|
||||
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
|
||||
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_zero, empty_val',
|
||||
cons_val_fin_one]
|
||||
ring
|
||||
· rw [show (x + y) (Sum.inr 0) = x (Sum.inr 0) + y (Sum.inr 0) from rfl]
|
||||
rw [show (x + y) (Sum.inr 1) = x (Sum.inr 1) + y (Sum.inr 1) from rfl]
|
||||
simp only [Fin.isValue, ofReal_add, Fin.mk_one, cons_val_one, head_cons, Fin.zero_eta,
|
||||
cons_val_zero]
|
||||
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
|
||||
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
|
||||
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_one, empty_val',
|
||||
cons_val_fin_one, cons_val_zero]
|
||||
ring
|
||||
· rw [show (x + y) (Sum.inr 0) = x (Sum.inr 0) + y (Sum.inr 0) from rfl]
|
||||
rw [show (x + y) (Sum.inr 1) = x (Sum.inr 1) + y (Sum.inr 1) from rfl]
|
||||
simp only [Fin.isValue, ofReal_add, Fin.zero_eta, cons_val_zero, Fin.mk_one, cons_val_one,
|
||||
head_fin_const]
|
||||
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
|
||||
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
|
||||
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_zero, empty_val',
|
||||
cons_val_fin_one, cons_val_one, head_fin_const]
|
||||
ring
|
||||
· rw [show (x + y) (Sum.inl 0) = x (Sum.inl 0) + y (Sum.inl 0) from rfl]
|
||||
rw [show (x + y) (Sum.inr 2) = x (Sum.inr 2) + y (Sum.inr 2) from rfl]
|
||||
simp only [Fin.isValue, ofReal_add, Fin.mk_one, cons_val_one, head_cons, head_fin_const]
|
||||
simp only [Fin.isValue, toSelfAdjointMatrix', toMatrix, LorentzVector.time,
|
||||
LorentzVector.space, Function.comp_apply, AddMemClass.mk_add_mk, of_add_of, add_cons,
|
||||
head_cons, tail_cons, empty_add_empty, of_apply, cons_val', cons_val_one, empty_val',
|
||||
cons_val_fin_one, head_fin_const]
|
||||
ring
|
||||
map_smul' r x := by
|
||||
ext i j : 2
|
||||
|
|
|
@ -97,7 +97,8 @@ noncomputable abbrev timeVec : (LorentzVector d) := e (Sum.inl 0)
|
|||
|
||||
lemma timeVec_space : (@timeVec d).space = 0 := by
|
||||
funext i
|
||||
simp only [space, Function.comp_apply, stdBasis_apply, Fin.isValue, ↓reduceIte, PiLp.zero_apply]
|
||||
simp only [space, Function.comp_apply, stdBasis_apply, Fin.isValue, reduceCtorEq, ↓reduceIte,
|
||||
PiLp.zero_apply]
|
||||
|
||||
lemma timeVec_time: (@timeVec d).time = 1 := by
|
||||
simp only [time, Fin.isValue, stdBasis_apply, ↓reduceIte]
|
||||
|
|
|
@ -156,6 +156,7 @@ lemma one_add_metric_non_zero : 1 + ⟪f, f'.1.1⟫ₘ ≠ 0 := by
|
|||
|
||||
section
|
||||
variable {v w : NormOneLorentzVector d}
|
||||
open InnerProductSpace
|
||||
|
||||
lemma metric_reflect_mem_mem (h : v ∈ FuturePointing d) (hw : w ∈ FuturePointing d) :
|
||||
0 ≤ ⟪v.1, w.1.spaceReflection⟫ₘ := by
|
||||
|
|
|
@ -17,7 +17,7 @@ of Lorentz vectors in d dimensions.
|
|||
-/
|
||||
|
||||
open Matrix
|
||||
|
||||
open InnerProductSpace
|
||||
/-!
|
||||
|
||||
# The definition of the Minkowski Matrix
|
||||
|
|
|
@ -63,6 +63,7 @@ def toLinearMapSelfAdjointMatrix (M : SL(2, ℂ)) :
|
|||
conjTranspose_mul, conjTranspose_conjTranspose,
|
||||
(star_eq_conjTranspose A.1).symm.trans $ selfAdjoint.mem_iff.mp A.2]⟩
|
||||
map_add' A B := by
|
||||
simp only [AddSubgroup.coe_add, AddMemClass.mk_add_mk, Subtype.mk.injEq]
|
||||
noncomm_ring [AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, AddSubmonoid.mk_add_mk,
|
||||
Subtype.mk.injEq]
|
||||
map_smul' r A := by
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue