refactor: Linting

This commit is contained in:
jstoobysmith 2024-07-30 08:07:47 -04:00
parent a65fb06605
commit a438af453d
14 changed files with 87 additions and 76 deletions

View file

@ -23,6 +23,8 @@ variable {d : }
## Definitions and lemmas needed to define a `LorentzTensorStructure`
-/
/-- The type colors for real Lorentz tensors. -/
inductive ColorType
| up
| down
@ -34,57 +36,57 @@ open realTensor
/-! TODO: Set up the notation `𝓛𝓣` or similar. -/
/-- The `LorentzTensorStructure` associated with real Lorentz tensors. -/
def realLorentzTensor (d : ) : TensorStructure where
Color := ColorType
ColorModule μ :=
match μ with
| .up => LorentzVector d
| .down => CovariantLorentzVector d
τ μ :=
match μ with
| .up => .down
| .down => .up
τ_involutive μ :=
match μ with
| .up => rfl
| .down => rfl
colorModule_addCommMonoid μ :=
match μ with
| .up => instAddCommMonoidLorentzVector d
| .down => instAddCommMonoidCovariantLorentzVector d
colorModule_module μ :=
match μ with
| .up => instModuleRealLorentzVector d
| .down => instModuleRealCovariantLorentzVector d
contrDual μ :=
match μ with
| .up => LorentzVector.contrUpDown
| .down => LorentzVector.contrDownUp
contrDual_symm μ :=
match μ with
| .up => by
intro x y
simp only [LorentzVector.contrDownUp, Equiv.cast_refl, Equiv.refl_apply, LinearMap.coe_comp,
LinearEquiv.coe_coe, Function.comp_apply, comm_tmul]
| .down => by
intro x y
simp only [LorentzVector.contrDownUp, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
comm_tmul, Equiv.cast_refl, Equiv.refl_apply]
unit μ :=
match μ with
| .up => LorentzVector.unitUp
| .down => LorentzVector.unitDown
unit_lid μ :=
match μ with
| .up => LorentzVector.unitUp_lid
| .down => LorentzVector.unitDown_lid
metric μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector
metric_dual μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector_contr_asCovariantProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector_contr_asProdLorentzVector
Color := ColorType
ColorModule μ :=
match μ with
| .up => LorentzVector d
| .down => CovariantLorentzVector d
τ μ :=
match μ with
| .up => .down
| .down => .up
τ_involutive μ :=
match μ with
| .up => rfl
| .down => rfl
colorModule_addCommMonoid μ :=
match μ with
| .up => instAddCommMonoidLorentzVector d
| .down => instAddCommMonoidCovariantLorentzVector d
colorModule_module μ :=
match μ with
| .up => instModuleRealLorentzVector d
| .down => instModuleRealCovariantLorentzVector d
contrDual μ :=
match μ with
| .up => LorentzVector.contrUpDown
| .down => LorentzVector.contrDownUp
contrDual_symm μ :=
match μ with
| .up => by
intro x y
simp only [LorentzVector.contrDownUp, Equiv.cast_refl, Equiv.refl_apply, LinearMap.coe_comp,
LinearEquiv.coe_coe, Function.comp_apply, comm_tmul]
| .down => by
intro x y
simp only [LorentzVector.contrDownUp, LinearMap.coe_comp, LinearEquiv.coe_coe,
Function.comp_apply, comm_tmul, Equiv.cast_refl, Equiv.refl_apply]
unit μ :=
match μ with
| .up => LorentzVector.unitUp
| .down => LorentzVector.unitDown
unit_lid μ :=
match μ with
| .up => LorentzVector.unitUp_lid
| .down => LorentzVector.unitDown_lid
metric μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector
metric_dual μ :=
match μ with
| realTensor.ColorType.up => asProdLorentzVector_contr_asCovariantProdLorentzVector
| realTensor.ColorType.down => asProdCovariantLorentzVector_contr_asProdLorentzVector
/-- The action of the Lorentz group on real Lorentz tensors. -/
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where