feat: Add contracting equivalence of index sets
This commit is contained in:
parent
238233b02c
commit
17139a6cf1
3 changed files with 330 additions and 101 deletions
|
@ -28,12 +28,67 @@ under which contraction and rising and lowering etc, are invariant.
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
noncomputable section
|
|
||||||
|
|
||||||
open TensorProduct
|
open TensorProduct
|
||||||
|
|
||||||
variable {R : Type} [CommSemiring R]
|
variable {R : Type} [CommSemiring R]
|
||||||
|
|
||||||
|
structure TensorColor where
|
||||||
|
/-- The allowed colors of indices.
|
||||||
|
For example for a real Lorentz tensor these are `{up, down}`. -/
|
||||||
|
Color : Type
|
||||||
|
/-- A map taking every color to its dual color. -/
|
||||||
|
τ : Color → Color
|
||||||
|
/-- The map `τ` is an involution. -/
|
||||||
|
τ_involutive : Function.Involutive τ
|
||||||
|
|
||||||
|
namespace TensorColor
|
||||||
|
|
||||||
|
variable (𝓒 : TensorColor)
|
||||||
|
|
||||||
|
/-- A relation on colors which is true if the two colors are equal or are duals. -/
|
||||||
|
def colorRel (μ ν : 𝓒.Color) : Prop := μ = ν ∨ μ = 𝓒.τ ν
|
||||||
|
|
||||||
|
/-- An equivalence relation on colors which is true if the two colors are equal or are duals. -/
|
||||||
|
lemma colorRel_equivalence : Equivalence 𝓒.colorRel where
|
||||||
|
refl := by
|
||||||
|
intro x
|
||||||
|
left
|
||||||
|
rfl
|
||||||
|
symm := by
|
||||||
|
intro x y h
|
||||||
|
rcases h with h | h
|
||||||
|
· left
|
||||||
|
exact h.symm
|
||||||
|
· right
|
||||||
|
subst h
|
||||||
|
exact (𝓒.τ_involutive y).symm
|
||||||
|
trans := by
|
||||||
|
intro x y z hxy hyz
|
||||||
|
rcases hxy with hxy | hxy <;>
|
||||||
|
rcases hyz with hyz | hyz <;>
|
||||||
|
subst hxy hyz
|
||||||
|
· left
|
||||||
|
rfl
|
||||||
|
· right
|
||||||
|
rfl
|
||||||
|
· right
|
||||||
|
rfl
|
||||||
|
· left
|
||||||
|
exact 𝓒.τ_involutive z
|
||||||
|
|
||||||
|
/-- The structure of a setoid on colors, two colors are related if they are equal,
|
||||||
|
or dual. -/
|
||||||
|
instance colorSetoid : Setoid 𝓒.Color := ⟨𝓒.colorRel, 𝓒.colorRel_equivalence⟩
|
||||||
|
|
||||||
|
/-- A map taking a color to its equivalence class in `colorSetoid`. -/
|
||||||
|
def colorQuot (μ : 𝓒.Color) : Quotient 𝓒.colorSetoid :=
|
||||||
|
Quotient.mk 𝓒.colorSetoid μ
|
||||||
|
|
||||||
|
end TensorColor
|
||||||
|
|
||||||
|
noncomputable section
|
||||||
namespace TensorStructure
|
namespace TensorStructure
|
||||||
|
|
||||||
/-- An auxillary function to contract the vector space `V1` and `V2` in `V1 ⊗[R] V2 ⊗[R] V3`. -/
|
/-- An auxillary function to contract the vector space `V1` and `V2` in `V1 ⊗[R] V2 ⊗[R] V3`. -/
|
||||||
|
@ -86,16 +141,9 @@ end TensorStructure
|
||||||
|
|
||||||
/-- An initial structure specifying a tensor system (e.g. a system in which you can
|
/-- An initial structure specifying a tensor system (e.g. a system in which you can
|
||||||
define real Lorentz tensors or Einstein notation convention). -/
|
define real Lorentz tensors or Einstein notation convention). -/
|
||||||
structure TensorStructure (R : Type) [CommSemiring R] where
|
structure TensorStructure (R : Type) [CommSemiring R] extends TensorColor where
|
||||||
/-- The allowed colors of indices.
|
|
||||||
For example for a real Lorentz tensor these are `{up, down}`. -/
|
|
||||||
Color : Type
|
|
||||||
/-- To each color we associate a module. -/
|
/-- To each color we associate a module. -/
|
||||||
ColorModule : Color → Type
|
ColorModule : Color → Type
|
||||||
/-- A map taking every color to its dual color. -/
|
|
||||||
τ : Color → Color
|
|
||||||
/-- The map `τ` is an involution. -/
|
|
||||||
τ_involutive : Function.Involutive τ
|
|
||||||
/-- Each `ColorModule` has the structure of an additive commutative monoid. -/
|
/-- Each `ColorModule` has the structure of an additive commutative monoid. -/
|
||||||
colorModule_addCommMonoid : ∀ μ, AddCommMonoid (ColorModule μ)
|
colorModule_addCommMonoid : ∀ μ, AddCommMonoid (ColorModule μ)
|
||||||
/-- Each `ColorModule` has the structure of a module over `R`. -/
|
/-- Each `ColorModule` has the structure of a module over `R`. -/
|
||||||
|
@ -161,44 +209,6 @@ def colorModuleCast (h : μ = ν) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModu
|
||||||
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
|
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
|
||||||
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
|
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg 𝓣.ColorModule h)) x
|
||||||
|
|
||||||
/-- A relation on colors which is true if the two colors are equal or are duals. -/
|
|
||||||
def colorRel (μ ν : 𝓣.Color) : Prop := μ = ν ∨ μ = 𝓣.τ ν
|
|
||||||
|
|
||||||
/-- An equivalence relation on colors which is true if the two colors are equal or are duals. -/
|
|
||||||
lemma colorRel_equivalence : Equivalence 𝓣.colorRel where
|
|
||||||
refl := by
|
|
||||||
intro x
|
|
||||||
left
|
|
||||||
rfl
|
|
||||||
symm := by
|
|
||||||
intro x y h
|
|
||||||
rcases h with h | h
|
|
||||||
· left
|
|
||||||
exact h.symm
|
|
||||||
· right
|
|
||||||
subst h
|
|
||||||
exact (𝓣.τ_involutive y).symm
|
|
||||||
trans := by
|
|
||||||
intro x y z hxy hyz
|
|
||||||
rcases hxy with hxy | hxy <;>
|
|
||||||
rcases hyz with hyz | hyz <;>
|
|
||||||
subst hxy hyz
|
|
||||||
· left
|
|
||||||
rfl
|
|
||||||
· right
|
|
||||||
rfl
|
|
||||||
· right
|
|
||||||
rfl
|
|
||||||
· left
|
|
||||||
exact 𝓣.τ_involutive z
|
|
||||||
|
|
||||||
/-- The structure of a setoid on colors, two colors are related if they are equal,
|
|
||||||
or dual. -/
|
|
||||||
instance colorSetoid : Setoid 𝓣.Color := ⟨𝓣.colorRel, 𝓣.colorRel_equivalence⟩
|
|
||||||
|
|
||||||
/-- A map taking a color to its equivalence class in `colorSetoid`. -/
|
|
||||||
def colorQuot (μ : 𝓣.Color) : Quotient 𝓣.colorSetoid :=
|
|
||||||
Quotient.mk 𝓣.colorSetoid μ
|
|
||||||
|
|
||||||
lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M]
|
lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M]
|
||||||
{f g : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] M}
|
{f g : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] M}
|
||||||
|
|
|
@ -337,6 +337,61 @@ def id (s : Index X) : Nat := s.tailNat.foldl (fun a b => 10 * a + b) 0
|
||||||
|
|
||||||
end Index
|
end Index
|
||||||
|
|
||||||
|
def IndexList : Type := List (Index X)
|
||||||
|
|
||||||
|
namespace IndexList
|
||||||
|
|
||||||
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
|
|
||||||
|
variable (l : IndexList X)
|
||||||
|
/-- The number of indices in an index list. -/
|
||||||
|
def numIndices : Nat := l.length
|
||||||
|
|
||||||
|
/-- The map of from `Fin s.numIndices` into colors associated to an index list. -/
|
||||||
|
def colorMap : Fin l.numIndices → X :=
|
||||||
|
fun i => (l.get i).toColor
|
||||||
|
|
||||||
|
/-- The map of from `Fin s.numIndices` into the natural numbers associated to an index list. -/
|
||||||
|
def idMap : Fin l.numIndices → Nat :=
|
||||||
|
fun i => (l.get i).id
|
||||||
|
|
||||||
|
def toPosSet (l : IndexList X) : Set (Fin l.numIndices × Index X) :=
|
||||||
|
{(i, l.get i) | i : Fin l.numIndices}
|
||||||
|
|
||||||
|
def toPosSetEquiv (l : IndexList X) : l.toPosSet ≃ Fin l.numIndices where
|
||||||
|
toFun := fun x => x.1.1
|
||||||
|
invFun := fun x => ⟨(x, l.get x), by simp [toPosSet]⟩
|
||||||
|
left_inv x := by
|
||||||
|
have hx := x.prop
|
||||||
|
simp [toPosSet] at hx
|
||||||
|
simp
|
||||||
|
obtain ⟨i, hi⟩ := hx
|
||||||
|
have hi2 : i = x.1.1 := by
|
||||||
|
obtain ⟨val, property⟩ := x
|
||||||
|
obtain ⟨fst, snd⟩ := val
|
||||||
|
simp_all only [Prod.mk.injEq]
|
||||||
|
subst hi2
|
||||||
|
simp_all only [Subtype.coe_eta]
|
||||||
|
right_inv := by
|
||||||
|
intro x
|
||||||
|
rfl
|
||||||
|
|
||||||
|
lemma toPosSet_is_finite (l : IndexList X) : l.toPosSet.Finite :=
|
||||||
|
Finite.intro l.toPosSetEquiv
|
||||||
|
|
||||||
|
instance : Fintype l.toPosSet where
|
||||||
|
elems := Finset.map l.toPosSetEquiv.symm.toEmbedding Finset.univ
|
||||||
|
complete := by
|
||||||
|
intro x
|
||||||
|
simp_all only [Finset.mem_map_equiv, Equiv.symm_symm, Finset.mem_univ]
|
||||||
|
|
||||||
|
def toPosFinset (l : IndexList X) : Finset (Fin l.numIndices × Index X) :=
|
||||||
|
l.toPosSet.toFinset
|
||||||
|
|
||||||
|
|
||||||
|
end IndexList
|
||||||
|
|
||||||
|
|
||||||
/-- A string of indices to be associated with a tensor.
|
/-- A string of indices to be associated with a tensor.
|
||||||
E.g. `ᵘ⁰ᵤ₂₆₀ᵘ³`. -/
|
E.g. `ᵘ⁰ᵤ₂₆₀ᵘ³`. -/
|
||||||
def IndexString : Type := {s : String // listCharIndexStringBool X s.toList = true}
|
def IndexString : Type := {s : String // listCharIndexStringBool X s.toList = true}
|
||||||
|
@ -354,27 +409,16 @@ lemma listCharIndexString (s : IndexString X) : listCharIndexString X s.toCharLi
|
||||||
exact s.prop
|
exact s.prop
|
||||||
|
|
||||||
/-- The indices associated to an index string. -/
|
/-- The indices associated to an index string. -/
|
||||||
def toIndexList (s : IndexString X) : List (Index X) :=
|
def toIndexList (s : IndexString X) : IndexList X :=
|
||||||
(listCharIndexStringTolistCharIndex X s.toCharList (listCharIndexString s)).map
|
(listCharIndexStringTolistCharIndex X s.toCharList (listCharIndexString s)).map
|
||||||
fun x => Index.ofCharList x.1 x.2
|
fun x => Index.ofCharList x.1 x.2
|
||||||
|
|
||||||
/-- The number of indices in an index string. -/
|
|
||||||
def numIndices (s : IndexString X) : Nat := s.toIndexList.length
|
|
||||||
|
|
||||||
/-- The map of from `Fin s.numIndices` into colors associated to an index string. -/
|
|
||||||
def colorMap (s : IndexString X) : Fin s.numIndices → X :=
|
|
||||||
fun i => (s.toIndexList.get i).toColor
|
|
||||||
|
|
||||||
/-- The map of from `Fin s.numIndices` into the natural numbers associated to an index string. -/
|
|
||||||
def idMap (s : IndexString X) : Fin s.numIndices → Nat :=
|
|
||||||
fun i => (s.toIndexList.get i).id
|
|
||||||
|
|
||||||
end IndexString
|
end IndexString
|
||||||
|
|
||||||
|
|
||||||
end IndexNotation
|
end IndexNotation
|
||||||
|
|
||||||
instance : IndexNotation realTensor.ColorType where
|
instance : IndexNotation realTensorColor.Color where
|
||||||
charList := {'ᵘ', 'ᵤ'}
|
charList := {'ᵘ', 'ᵤ'}
|
||||||
notaEquiv :=
|
notaEquiv :=
|
||||||
{toFun := fun x =>
|
{toFun := fun x =>
|
||||||
|
@ -394,10 +438,7 @@ instance : IndexNotation realTensor.ColorType where
|
||||||
fin_cases x <;> rfl}
|
fin_cases x <;> rfl}
|
||||||
|
|
||||||
|
|
||||||
instance (d : ℕ) : IndexNotation (realLorentzTensor d).Color :=
|
namespace TensorColor
|
||||||
instIndexNotationColorType
|
|
||||||
|
|
||||||
namespace TensorStructure
|
|
||||||
|
|
||||||
variable {n m : ℕ}
|
variable {n m : ℕ}
|
||||||
|
|
||||||
|
@ -409,25 +450,199 @@ variable {d : ℕ} {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [
|
||||||
{cW : W → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
{cW : W → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||||
{cn : Fin n → 𝓣.Color} {cm : Fin m → 𝓣.Color}
|
{cn : Fin n → 𝓣.Color} {cm : Fin m → 𝓣.Color}
|
||||||
|
|
||||||
variable [IndexNotation 𝓣.Color] [Fintype 𝓣.Color] [DecidableEq 𝓣.Color]
|
variable (𝓒 : TensorColor)
|
||||||
|
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||||
|
|
||||||
open IndexNotation
|
open IndexNotation
|
||||||
|
|
||||||
def AllowedIndexString (s : IndexString 𝓣.Color) : Prop :=
|
/-- The proposition on an `i : Fin s.length` such the corresponding element of
|
||||||
∀ i j, s.idMap i = s.idMap j → (i = j ∨ s.colorMap i = 𝓣.τ (s.colorMap j))
|
`s` does not contract with any other element (i.e. share an index). -/
|
||||||
|
def NoContr (s : IndexList 𝓒.Color) (i : Fin s.length) : Prop :=
|
||||||
|
∀ j, i ≠ j → s.idMap i ≠ s.idMap j
|
||||||
|
|
||||||
instance (s : IndexString 𝓣.Color) : Decidable (𝓣.AllowedIndexString s) :=
|
instance (i : Fin s.length) : Decidable (NoContr 𝓒 s i) :=
|
||||||
Nat.decidableForallFin fun i =>
|
Fintype.decidableForallFintype
|
||||||
∀ (j : Fin s.numIndices), s.idMap i = s.idMap j → i = j ∨ s.colorMap i = 𝓣.τ (s.colorMap j)
|
|
||||||
|
|
||||||
def testIndex : Index (realLorentzTensor d).Color := ⟨"ᵘ¹", by
|
/-- The finset of indices of `s` corresponding to elements which do not contract. -/
|
||||||
change listCharIndex realTensor.ColorType _ ∧ _
|
def noContrFinset (s : IndexList 𝓒.Color) : Finset (Fin s.length) :=
|
||||||
decide⟩
|
Finset.univ.filter (𝓒.NoContr s)
|
||||||
|
|
||||||
def testIndexString : IndexString (realLorentzTensor 2).Color := ⟨"ᵘ⁰ᵤ₂₆₀ᵘ³", by
|
/-- An eqiuvalence between the subtype of indices of `s` which do not contract and
|
||||||
change listCharIndexStringBool realTensor.ColorType _ = _
|
`Fin _`. -/
|
||||||
rfl⟩
|
def noContrSubtypeEquiv (s : IndexList 𝓒.Color) :
|
||||||
|
{i : Fin s.length // NoContr 𝓒 s i} ≃ Fin (𝓒.noContrFinset s).card :=
|
||||||
|
(Equiv.subtypeEquivRight (by
|
||||||
|
intro x
|
||||||
|
simp only [noContrFinset, Finset.mem_filter, Finset.mem_univ, true_and])).trans
|
||||||
|
(Finset.orderIsoOfFin (𝓒.noContrFinset s) (by rfl)).toEquiv.symm
|
||||||
|
|
||||||
#eval (realLorentzTensor 2).AllowedIndexString testIndexString
|
/-- The subtype of indices `s` which do contract. -/
|
||||||
|
def contrSubtype (s : IndexList 𝓒.Color) : Type :=
|
||||||
|
{i : Fin s.length // ¬ NoContr 𝓒 s i}
|
||||||
|
|
||||||
end TensorStructure
|
instance (s : IndexList 𝓒.Color) : Fintype (𝓒.contrSubtype s) :=
|
||||||
|
Subtype.fintype fun x => ¬𝓒.NoContr s x
|
||||||
|
|
||||||
|
instance (s : IndexList 𝓒.Color) : DecidableEq (𝓒.contrSubtype s) :=
|
||||||
|
Subtype.instDecidableEq
|
||||||
|
|
||||||
|
/-- Given a `i : 𝓒.contrSubtype s` the proposition on a `j` in `Fin s.length` for
|
||||||
|
it to be an index of `s` contracting with `i`. -/
|
||||||
|
def getDualProp {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) (j : Fin s.length) : Prop :=
|
||||||
|
i.1 ≠ j ∧ s.idMap i.1 = s.idMap j
|
||||||
|
|
||||||
|
instance {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) (j : Fin s.length) :
|
||||||
|
Decidable (𝓒.getDualProp i j) :=
|
||||||
|
instDecidableAnd
|
||||||
|
|
||||||
|
/-- Given a `i : 𝓒.contrSubtype s` the index of `s` contracting with `i`. -/
|
||||||
|
def getDualFin {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) : Fin s.length :=
|
||||||
|
(Fin.find (𝓒.getDualProp i)).get (by simpa [NoContr, Fin.isSome_find_iff] using i.prop)
|
||||||
|
|
||||||
|
lemma some_getDualFin_eq_find {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) :
|
||||||
|
Fin.find (𝓒.getDualProp i) = some (getDualFin 𝓒 i) := by
|
||||||
|
simp [getDualFin]
|
||||||
|
|
||||||
|
lemma getDualFin_not_NoContr {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) :
|
||||||
|
¬ NoContr 𝓒 s (getDualFin 𝓒 i) := by
|
||||||
|
have h := 𝓒.some_getDualFin_eq_find i
|
||||||
|
rw [Fin.find_eq_some_iff] at h
|
||||||
|
simp [NoContr]
|
||||||
|
exact ⟨i.1, And.intro (fun a => h.1.1 a.symm) h.1.2.symm⟩
|
||||||
|
|
||||||
|
/-- The dual index of an element of `𝓒.contrSubtype s`, that is the index
|
||||||
|
contracting with it. -/
|
||||||
|
def getDual {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) : 𝓒.contrSubtype s :=
|
||||||
|
⟨getDualFin 𝓒 i, getDualFin_not_NoContr 𝓒 i⟩
|
||||||
|
|
||||||
|
lemma getDual_id {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) :
|
||||||
|
s.idMap i.1 = s.idMap (getDual 𝓒 i).1 := by
|
||||||
|
simp [getDual]
|
||||||
|
have h1 := 𝓒.some_getDualFin_eq_find i
|
||||||
|
rw [Fin.find_eq_some_iff] at h1
|
||||||
|
simp only [getDualProp, ne_eq, and_imp] at h1
|
||||||
|
exact h1.1.2
|
||||||
|
|
||||||
|
lemma getDual_neq_self {s : IndexList 𝓒.Color} (i : 𝓒.contrSubtype s) :
|
||||||
|
i ≠ 𝓒.getDual i := by
|
||||||
|
have h1 := 𝓒.some_getDualFin_eq_find i
|
||||||
|
rw [Fin.find_eq_some_iff] at h1
|
||||||
|
exact ne_of_apply_ne Subtype.val h1.1.1
|
||||||
|
|
||||||
|
/-- An index list is allowed if every contracting index has exactly one dual,
|
||||||
|
and the color of the dual is dual to the color of the index. -/
|
||||||
|
def AllowedIndexString (s : IndexList 𝓒.Color) : Prop :=
|
||||||
|
(∀ (i j : 𝓒.contrSubtype s), 𝓒.getDualProp i j.1 → j = 𝓒.getDual i) ∧
|
||||||
|
(∀ (i : 𝓒.contrSubtype s), s.colorMap i.1 = 𝓒.τ (s.colorMap (𝓒.getDual i).1))
|
||||||
|
|
||||||
|
@[simp]
|
||||||
|
lemma getDual_getDual {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s) (i : 𝓒.contrSubtype s) :
|
||||||
|
getDual 𝓒 (getDual 𝓒 i) = i := by
|
||||||
|
refine (h.1 (getDual 𝓒 i) i ?_).symm
|
||||||
|
simp [getDualProp]
|
||||||
|
apply And.intro
|
||||||
|
exact Subtype.coe_ne_coe.mpr (getDual_neq_self 𝓒 i).symm
|
||||||
|
exact (getDual_id 𝓒 i).symm
|
||||||
|
|
||||||
|
/-- The set of contracting ordered pairs of indices. -/
|
||||||
|
def contrPairSet (s : IndexList 𝓒.Color) : Set (𝓒.contrSubtype s × 𝓒.contrSubtype s) :=
|
||||||
|
{p | p.1.1 < p.2.1 ∧ s.idMap p.1.1 = s.idMap p.2.1}
|
||||||
|
|
||||||
|
lemma getDual_lt_self_mem_contrPairSet {s : IndexList 𝓒.Color} {i : 𝓒.contrSubtype s}
|
||||||
|
(h : (getDual 𝓒 i).1 < i.1) : (getDual 𝓒 i, i) ∈ 𝓒.contrPairSet s :=
|
||||||
|
And.intro h (𝓒.getDual_id i).symm
|
||||||
|
|
||||||
|
lemma getDual_not_lt_self_mem_contrPairSet {s : IndexList 𝓒.Color} {i : 𝓒.contrSubtype s}
|
||||||
|
(h : ¬ (getDual 𝓒 i).1 < i.1) : (i, getDual 𝓒 i) ∈ 𝓒.contrPairSet s := by
|
||||||
|
apply And.intro
|
||||||
|
have h1 := 𝓒.getDual_neq_self i
|
||||||
|
simp
|
||||||
|
simp at h
|
||||||
|
exact lt_of_le_of_ne h h1
|
||||||
|
simp
|
||||||
|
exact getDual_id 𝓒 i
|
||||||
|
|
||||||
|
lemma contrPairSet_fst_eq_dual_snd {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s)
|
||||||
|
(x : 𝓒.contrPairSet s) : x.1.1 = getDual 𝓒 x.1.2 :=
|
||||||
|
(h.1 (x.1.2) x.1.1 (And.intro (Fin.ne_of_gt x.2.1) x.2.2.symm))
|
||||||
|
|
||||||
|
lemma contrPairSet_snd_eq_dual_fst {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s)
|
||||||
|
(x : 𝓒.contrPairSet s) : x.1.2 = getDual 𝓒 x.1.1 := by
|
||||||
|
rw [contrPairSet_fst_eq_dual_snd, getDual_getDual]
|
||||||
|
exact h
|
||||||
|
exact h
|
||||||
|
|
||||||
|
lemma contrPairSet_dual_snd_lt_self {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s)
|
||||||
|
(x : 𝓒.contrPairSet s) : (getDual 𝓒 x.1.2).1 < x.1.2.1 := by
|
||||||
|
rw [← 𝓒.contrPairSet_fst_eq_dual_snd h]
|
||||||
|
exact x.2.1
|
||||||
|
|
||||||
|
/-- An equivalence between two coppies of `𝓒.contrPairSet s` and `𝓒.contrSubtype s`.
|
||||||
|
This equivalence exists due to the ordering on pairs in `𝓒.contrPairSet s`. -/
|
||||||
|
def contrPairEquiv {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s) :
|
||||||
|
𝓒.contrPairSet s ⊕ 𝓒.contrPairSet s ≃ 𝓒.contrSubtype s where
|
||||||
|
toFun x :=
|
||||||
|
match x with
|
||||||
|
| Sum.inl p => p.1.2
|
||||||
|
| Sum.inr p => p.1.1
|
||||||
|
invFun x :=
|
||||||
|
if h : (𝓒.getDual x).1 < x.1 then
|
||||||
|
Sum.inl ⟨(𝓒.getDual x, x), 𝓒.getDual_lt_self_mem_contrPairSet h⟩
|
||||||
|
else
|
||||||
|
Sum.inr ⟨(x, 𝓒.getDual x), 𝓒.getDual_not_lt_self_mem_contrPairSet h⟩
|
||||||
|
left_inv x := by
|
||||||
|
match x with
|
||||||
|
| Sum.inl x =>
|
||||||
|
simp only [Subtype.coe_lt_coe]
|
||||||
|
rw [dif_pos]
|
||||||
|
simp [← 𝓒.contrPairSet_fst_eq_dual_snd h]
|
||||||
|
exact 𝓒.contrPairSet_dual_snd_lt_self h _
|
||||||
|
| Sum.inr x =>
|
||||||
|
simp only [Subtype.coe_lt_coe]
|
||||||
|
rw [dif_neg]
|
||||||
|
simp only [← 𝓒.contrPairSet_snd_eq_dual_fst h, Prod.mk.eta, Subtype.coe_eta]
|
||||||
|
rw [← 𝓒.contrPairSet_snd_eq_dual_fst h]
|
||||||
|
have h1 := x.2.1
|
||||||
|
simp only [not_lt, ge_iff_le]
|
||||||
|
exact le_of_lt h1
|
||||||
|
right_inv x := by
|
||||||
|
by_cases h1 : (getDual 𝓒 x).1 < x.1
|
||||||
|
simp only [h1, ↓reduceDIte]
|
||||||
|
simp only [h1, ↓reduceDIte]
|
||||||
|
|
||||||
|
@[simp]
|
||||||
|
lemma contrPairEquiv_apply_inr {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s)
|
||||||
|
(x : 𝓒.contrPairSet s) : 𝓒.contrPairEquiv h (Sum.inr x) = x.1.1 := by
|
||||||
|
simp [contrPairEquiv]
|
||||||
|
|
||||||
|
@[simp]
|
||||||
|
lemma contrPairEquiv_apply_inl {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s)
|
||||||
|
(x : 𝓒.contrPairSet s) : 𝓒.contrPairEquiv h (Sum.inl x) = x.1.2 := by
|
||||||
|
simp [contrPairEquiv]
|
||||||
|
|
||||||
|
/-- An equivalence between `Fin s.length` and
|
||||||
|
`(𝓒.contrPairSet s ⊕ 𝓒.contrPairSet s) ⊕ Fin (𝓒.noContrFinset s).card`, which
|
||||||
|
can be used for contractions. -/
|
||||||
|
def splitContr {s : IndexList 𝓒.Color} (h : 𝓒.AllowedIndexString s) :
|
||||||
|
Fin s.length ≃ (𝓒.contrPairSet s ⊕ 𝓒.contrPairSet s) ⊕ Fin (𝓒.noContrFinset s).card :=
|
||||||
|
(Equiv.sumCompl (𝓒.NoContr s)).symm.trans <|
|
||||||
|
(Equiv.sumComm { i // 𝓒.NoContr s i} { i // ¬ 𝓒.NoContr s i}).trans <|
|
||||||
|
Equiv.sumCongr (𝓒.contrPairEquiv h).symm (𝓒.noContrSubtypeEquiv s)
|
||||||
|
|
||||||
|
lemma splitContr_map {s : IndexList 𝓒.Color} (hs : 𝓒.AllowedIndexString s) :
|
||||||
|
s.colorMap ∘ (𝓒.splitContr hs).symm ∘ Sum.inl ∘ Sum.inl =
|
||||||
|
𝓒.τ ∘ s.colorMap ∘ (𝓒.splitContr hs).symm ∘ Sum.inl ∘ Sum.inr := by
|
||||||
|
funext x
|
||||||
|
simp [splitContr, contrPairEquiv_apply_inr]
|
||||||
|
erw [contrPairEquiv_apply_inr, contrPairEquiv_apply_inl]
|
||||||
|
rw [contrPairSet_fst_eq_dual_snd _ hs]
|
||||||
|
exact hs.2 _
|
||||||
|
|
||||||
|
end TensorColor
|
||||||
|
/-
|
||||||
|
def testIndex : Index realTensorColor.Color := ⟨"ᵘ¹", by decide⟩
|
||||||
|
|
||||||
|
def testIndexString : IndexString realTensorColor.Color := ⟨"ᵘ⁰ᵤ₀ᵘ⁰", by rfl⟩
|
||||||
|
|
||||||
|
#eval realTensorColor.AllowedIndexString testIndexString.toIndexList
|
||||||
|
-/
|
||||||
|
|
|
@ -14,7 +14,8 @@ import HepLean.SpaceTime.LorentzTensor.MulActionTensor
|
||||||
|
|
||||||
open TensorProduct
|
open TensorProduct
|
||||||
open minkowskiMatrix
|
open minkowskiMatrix
|
||||||
namespace realTensor
|
|
||||||
|
namespace realTensorColor
|
||||||
|
|
||||||
variable {d : ℕ}
|
variable {d : ℕ}
|
||||||
/-!
|
/-!
|
||||||
|
@ -51,31 +52,25 @@ def colorTypEquivFin1Fin1 : ColorType ≃ Fin 1 ⊕ Fin 1 where
|
||||||
rename_i f
|
rename_i f
|
||||||
exact (Fin.fin_one_eq_zero f).symm
|
exact (Fin.fin_one_eq_zero f).symm
|
||||||
|
|
||||||
instance : DecidableEq realTensor.ColorType :=
|
instance : DecidableEq ColorType :=
|
||||||
Equiv.decidableEq colorTypEquivFin1Fin1
|
Equiv.decidableEq colorTypEquivFin1Fin1
|
||||||
|
|
||||||
instance : Fintype realTensor.ColorType where
|
instance : Fintype ColorType where
|
||||||
elems := {realTensor.ColorType.up, realTensor.ColorType.down}
|
elems := {ColorType.up, ColorType.down}
|
||||||
complete := by
|
complete := by
|
||||||
intro x
|
intro x
|
||||||
cases x
|
cases x
|
||||||
simp only [Finset.mem_insert, Finset.mem_singleton, or_false]
|
simp only [Finset.mem_insert, Finset.mem_singleton, or_false]
|
||||||
simp only [Finset.mem_insert, Finset.mem_singleton, or_true]
|
simp only [Finset.mem_insert, Finset.mem_singleton, or_true]
|
||||||
|
|
||||||
end realTensor
|
end realTensorColor
|
||||||
|
|
||||||
noncomputable section
|
noncomputable section
|
||||||
|
|
||||||
open realTensor
|
open realTensorColor
|
||||||
|
|
||||||
/-! TODO: Set up the notation `𝓛𝓣ℝ` or similar. -/
|
def realTensorColor : TensorColor where
|
||||||
/-- The `LorentzTensorStructure` associated with real Lorentz tensors. -/
|
|
||||||
def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
|
||||||
Color := ColorType
|
Color := ColorType
|
||||||
ColorModule μ :=
|
|
||||||
match μ with
|
|
||||||
| .up => LorentzVector d
|
|
||||||
| .down => CovariantLorentzVector d
|
|
||||||
τ μ :=
|
τ μ :=
|
||||||
match μ with
|
match μ with
|
||||||
| .up => .down
|
| .up => .down
|
||||||
|
@ -84,6 +79,19 @@ def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
||||||
match μ with
|
match μ with
|
||||||
| .up => rfl
|
| .up => rfl
|
||||||
| .down => rfl
|
| .down => rfl
|
||||||
|
|
||||||
|
instance : Fintype realTensorColor.Color := realTensorColor.instFintypeColorType
|
||||||
|
|
||||||
|
instance : DecidableEq realTensorColor.Color := realTensorColor.instDecidableEqColorType
|
||||||
|
|
||||||
|
/-! TODO: Set up the notation `𝓛𝓣ℝ` or similar. -/
|
||||||
|
/-- The `LorentzTensorStructure` associated with real Lorentz tensors. -/
|
||||||
|
def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
||||||
|
toTensorColor := realTensorColor
|
||||||
|
ColorModule μ :=
|
||||||
|
match μ with
|
||||||
|
| .up => LorentzVector d
|
||||||
|
| .down => CovariantLorentzVector d
|
||||||
colorModule_addCommMonoid μ :=
|
colorModule_addCommMonoid μ :=
|
||||||
match μ with
|
match μ with
|
||||||
| .up => instAddCommMonoidLorentzVector d
|
| .up => instAddCommMonoidLorentzVector d
|
||||||
|
@ -100,11 +108,11 @@ def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
||||||
match μ with
|
match μ with
|
||||||
| .up => by
|
| .up => by
|
||||||
intro x y
|
intro x y
|
||||||
simp only [LorentzVector.contrDownUp, Equiv.cast_refl, Equiv.refl_apply, LinearMap.coe_comp,
|
simp only [realTensorColor, LorentzVector.contrDownUp, Equiv.cast_refl, Equiv.refl_apply, LinearMap.coe_comp,
|
||||||
LinearEquiv.coe_coe, Function.comp_apply, comm_tmul]
|
LinearEquiv.coe_coe, Function.comp_apply, comm_tmul]
|
||||||
| .down => by
|
| .down => by
|
||||||
intro x y
|
intro x y
|
||||||
simp only [LorentzVector.contrDownUp, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
simp only [realTensorColor, LorentzVector.contrDownUp, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||||
Function.comp_apply, comm_tmul, Equiv.cast_refl, Equiv.refl_apply]
|
Function.comp_apply, comm_tmul, Equiv.cast_refl, Equiv.refl_apply]
|
||||||
unit μ :=
|
unit μ :=
|
||||||
match μ with
|
match μ with
|
||||||
|
@ -116,16 +124,12 @@ def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
||||||
| .down => LorentzVector.unitDown_rid
|
| .down => LorentzVector.unitDown_rid
|
||||||
metric μ :=
|
metric μ :=
|
||||||
match μ with
|
match μ with
|
||||||
| realTensor.ColorType.up => asTenProd
|
| realTensorColor.ColorType.up => asTenProd
|
||||||
| realTensor.ColorType.down => asCoTenProd
|
| realTensorColor.ColorType.down => asCoTenProd
|
||||||
metric_dual μ :=
|
metric_dual μ :=
|
||||||
match μ with
|
match μ with
|
||||||
| realTensor.ColorType.up => asTenProd_contr_asCoTenProd
|
| realTensorColor.ColorType.up => asTenProd_contr_asCoTenProd
|
||||||
| realTensor.ColorType.down => asCoTenProd_contr_asTenProd
|
| realTensorColor.ColorType.down => asCoTenProd_contr_asTenProd
|
||||||
|
|
||||||
instance : Fintype (realLorentzTensor d).Color := realTensor.instFintypeColorType
|
|
||||||
|
|
||||||
instance : DecidableEq (realLorentzTensor d).Color := realTensor.instDecidableEqColorType
|
|
||||||
|
|
||||||
/-- The action of the Lorentz group on real Lorentz tensors. -/
|
/-- The action of the Lorentz group on real Lorentz tensors. -/
|
||||||
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
|
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue