refactor: Lint

This commit is contained in:
jstoobysmith 2024-08-06 15:56:29 -04:00
parent d02e94886d
commit cecec0c843
7 changed files with 52 additions and 43 deletions

View file

@ -15,7 +15,6 @@ to define the index notation for real Lorentz tensors.
-/
instance : IndexNotation realTensorColor.Color where
charList := {'ᵘ', 'ᵤ'}
notaEquiv :=
@ -76,7 +75,7 @@ noncomputable def fromIndexStringColor {cn : Fin n → realTensorColor.Color}
(T : (realLorentzTensor d).Tensor cn) (s : String)
(hs : listCharIndexStringBool realTensorColor.Color s.toList = true)
(hn : n = (IndexString.toIndexList (⟨s, hs⟩ : IndexString realTensorColor.Color)).length)
(hc : IndexListColor.colorPropBool (IndexString.toIndexList ⟨s, hs⟩))
(hc : IndexListColor.colorPropBool (IndexString.toIndexList ⟨s, hs⟩))
(hd : TensorColor.ColorMap.DualMap.boolFin
(IndexString.toIndexList ⟨s, hs⟩).colorMap (cn ∘ Fin.cast hn.symm)) :
(realLorentzTensor d).TensorIndex :=
@ -88,17 +87,17 @@ noncomputable def fromIndexStringColor {cn : Fin n → realTensorColor.Color}
/-- A tactics used to prove `colorPropBool` for real Lorentz tensors. -/
macro "prodTactic" : tactic =>
`(tactic| {
change @IndexListColor.colorPropBool realTensorColor _ _ _
simp only [toTensorColor_eq, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
String.toList, ↓Char.isValue, Equiv.coe_refl, Function.comp_apply, id_eq, ne_eq,
Function.comp_id, RelIso.coe_fn_toEquiv, prod_index, IndexListColor.prod]
rfl})
change @IndexListColor.colorPropBool realTensorColor _ _ _
simp only [toTensorColor_eq, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
String.toList, ↓Char.isValue, Equiv.coe_refl, Function.comp_apply, id_eq, ne_eq,
Function.comp_id, RelIso.coe_fn_toEquiv, prod_index, IndexListColor.prod]
rfl})
/-- A tactic used to prove `boolFin` for real Lornetz tensors. -/
macro "dualMapTactic" : tactic =>
`(tactic| {
simp only [String.toList, ↓Char.isValue, toTensorColor_eq]
rfl})
simp only [String.toList, ↓Char.isValue, toTensorColor_eq]
rfl})
/-- Notation for the construction of a tensor index from a tensor and a string.
Conditions are checked automatically. -/
@ -110,10 +109,10 @@ notation:10 T "⊗ᵀ" S:11 => TensorIndex.prod T S (IndexListColor.colorPropBoo
/-- An example showing the allowed constructions. -/
example (T : (realLorentzTensor d).Tensor ![ColorType.up, ColorType.down]) : True := by
let _ := T|"ᵤ₁ᵤ₂"
let _ := T|"ᵤ₁ᵤ₂"
let _ := T|"ᵘ³ᵤ₄"
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄"
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ¹ᵘ²" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ¹ᵘ²" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
exact trivial
end realLorentzTensor