refactor: Lint
This commit is contained in:
parent
9d98dc4854
commit
4a64acc2a2
3 changed files with 12 additions and 7 deletions
|
@ -3,7 +3,8 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.SpaceTime.LorentzTensor.Real.Basic
|
||||
import Mathlib.Data.Set.Finite
|
||||
import Mathlib.Data.Finset.Sort
|
||||
/-!
|
||||
|
||||
# Index notation for a type
|
||||
|
@ -342,7 +343,7 @@ lemma getDual_neq_self (i : l.contrSubtype) : i ≠ l.getDual i := by
|
|||
indices. -/
|
||||
def HasNoContr : Prop := ∀ i, l.NoContr i
|
||||
|
||||
instance (h : l.HasNoContr) : IsEmpty l.contrSubtype := by
|
||||
lemma hasNoContr_is_empty (h : l.HasNoContr) : IsEmpty l.contrSubtype := by
|
||||
rw [_root_.isEmpty_iff]
|
||||
intro a
|
||||
exact h a.1 a.1 (fun _ => a.2 (h a.1)) rfl
|
||||
|
|
|
@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.SpaceTime.LorentzTensor.IndexNotation.Basic
|
||||
import HepLean.SpaceTime.LorentzTensor.Basic
|
||||
/-!
|
||||
|
||||
# Index lists with color conditions
|
||||
|
@ -42,7 +43,7 @@ instance : Coe (IndexListColor 𝓒) (IndexList 𝓒.Color) := ⟨fun x => x.val
|
|||
lemma indexListColorProp_of_hasNoContr {s : IndexList 𝓒.Color} (h : s.HasNoContr) :
|
||||
IndexListColorProp 𝓒 s := by
|
||||
simp [IndexListColorProp]
|
||||
haveI : IsEmpty (s.contrSubtype) := s.instIsEmptyContrSubtypeOfHasNoContr h
|
||||
haveI : IsEmpty (s.contrSubtype) := s.hasNoContr_is_empty h
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
@ -163,6 +164,8 @@ def rel (s1 s2 : IndexListColor 𝓒) : Prop :=
|
|||
|
||||
-/
|
||||
|
||||
/-- Appending two `IndexListColor` whose correpsonding appended index list
|
||||
satisfies `IndexListColorProp`. -/
|
||||
def append (s1 s2 : IndexListColor 𝓒) (h : IndexListColorProp 𝓒 (s1.1 ++ s2.1)) :
|
||||
IndexListColor 𝓒 := ⟨s1.1 ++ s2.1, h⟩
|
||||
|
||||
|
|
|
@ -24,21 +24,22 @@ variable {d : ℕ} {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [
|
|||
|
||||
variable [IndexNotation 𝓣.Color] [Fintype 𝓣.Color] [DecidableEq 𝓣.Color]
|
||||
|
||||
/-- The structure an tensor with a index specification e.g. `ᵘ¹ᵤ₂`. -/
|
||||
structure TensorIndex (cn : Fin n → 𝓣.Color) where
|
||||
/-- The underlying tensor. -/
|
||||
tensor : 𝓣.Tensor cn
|
||||
/-- The list of indices. -/
|
||||
index : IndexListColor 𝓣.toTensorColor
|
||||
/-- The number of indices matches the number of vector spaces in the tensor. -/
|
||||
nat_eq : n = index.1.length
|
||||
/-- The equivalence classes of colors of the tensor and the index list agree. -/
|
||||
quot_eq : 𝓣.colorQuot ∘ index.1.colorMap ∘ Fin.cast nat_eq = 𝓣.colorQuot ∘ cn
|
||||
|
||||
namespace TensorIndex
|
||||
|
||||
variable {𝓣 : TensorStructure R} [IndexNotation 𝓣.Color] [Fintype 𝓣.Color] [DecidableEq 𝓣.Color]
|
||||
variable {n m : ℕ} {cn : Fin n → 𝓣.Color} {cm : Fin m → 𝓣.Color} (T : TensorIndex 𝓣 cn)
|
||||
section noncomputable
|
||||
|
||||
def smul (r : R) : TensorIndex 𝓣 cn := ⟨r • T.tensor, T.index, T.nat_eq, T.quot_eq⟩
|
||||
|
||||
end
|
||||
|
||||
end TensorIndex
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue