refactor: Lint
This commit is contained in:
parent
3693dee740
commit
1c9d66ee19
12 changed files with 79 additions and 138 deletions
|
@ -12,13 +12,11 @@ import HepLean.SpaceTime.LorentzTensor.IndexNotation.Basic
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
variable (l l2 l3 : IndexList X)
|
variable (l l2 l3 : IndexList X)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Are dual indices
|
## Are dual indices
|
||||||
|
@ -74,7 +72,6 @@ lemma append_inr_inl (l l2 : IndexList X) (i : Fin l2.length) (j : Fin l.length)
|
||||||
|
|
||||||
end AreDualInSelf
|
end AreDualInSelf
|
||||||
|
|
||||||
|
|
||||||
namespace AreDualInOther
|
namespace AreDualInOther
|
||||||
|
|
||||||
variable {l l2 l3 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
|
variable {l l2 l3 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
|
||||||
|
|
|
@ -243,7 +243,6 @@ instance : Fintype l.toPosSet where
|
||||||
def toPosFinset (l : IndexList X) : Finset (Fin l.length × Index X) :=
|
def toPosFinset (l : IndexList X) : Finset (Fin l.length × Index X) :=
|
||||||
l.toPosSet.toFinset
|
l.toPosSet.toFinset
|
||||||
|
|
||||||
|
|
||||||
/-- The construction of a list of indices from a map
|
/-- The construction of a list of indices from a map
|
||||||
from `Fin n` to `Index X`. -/
|
from `Fin n` to `Index X`. -/
|
||||||
def fromFinMap {n : ℕ} (f : Fin n → Index X) : IndexList X where
|
def fromFinMap {n : ℕ} (f : Fin n → Index X) : IndexList X where
|
||||||
|
@ -368,7 +367,6 @@ lemma colorMap_sumELim (l1 l2 : IndexList X) :
|
||||||
| Sum.inl i => simp
|
| Sum.inl i => simp
|
||||||
| Sum.inr i => simp
|
| Sum.inr i => simp
|
||||||
|
|
||||||
|
|
||||||
end append
|
end append
|
||||||
|
|
||||||
end IndexList
|
end IndexList
|
||||||
|
|
|
@ -20,7 +20,6 @@ a Tensor Color structure.
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {𝓒 : TensorColor}
|
variable {𝓒 : TensorColor}
|
||||||
|
@ -158,7 +157,6 @@ lemma triple_drop_mid (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).wit
|
||||||
rw [append_withDual_eq_withUniqueDual_symm, append_assoc] at hu
|
rw [append_withDual_eq_withUniqueDual_symm, append_assoc] at hu
|
||||||
exact append_withDual_eq_withUniqueDual_inr _ _ hu
|
exact append_withDual_eq_withUniqueDual_inr _ _ hu
|
||||||
|
|
||||||
|
|
||||||
lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
|
lemma swap (hu : (l ++ l2 ++ l3).withUniqueDual = (l ++ l2 ++ l3).withDual)
|
||||||
(h : ColorCond (l ++ l2 ++ l3)) :
|
(h : ColorCond (l ++ l2 ++ l3)) :
|
||||||
ColorCond (l2 ++ l ++ l3) := by
|
ColorCond (l2 ++ l ++ l3) := by
|
||||||
|
@ -268,11 +266,8 @@ lemma iff_bool : l.ColorCond ↔ bool l := by
|
||||||
|
|
||||||
end ColorCond
|
end ColorCond
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
end IndexList
|
end IndexList
|
||||||
|
|
||||||
|
|
||||||
variable (𝓒 : TensorColor)
|
variable (𝓒 : TensorColor)
|
||||||
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||||
|
|
||||||
|
@ -317,14 +312,13 @@ lemma ext {l l' : ColorIndexList 𝓒} (h : l.val = l'.val) : l = l' := by
|
||||||
apply IndexList.ext
|
apply IndexList.ext
|
||||||
exact h
|
exact h
|
||||||
|
|
||||||
|
|
||||||
/-! TODO: `orderEmbOfFin_univ` should be replaced with a mathlib lemma if possible. -/
|
/-! TODO: `orderEmbOfFin_univ` should be replaced with a mathlib lemma if possible. -/
|
||||||
lemma orderEmbOfFin_univ (n m : ℕ) (h : n = m):
|
lemma orderEmbOfFin_univ (n m : ℕ) (h : n = m) :
|
||||||
Finset.orderEmbOfFin (Finset.univ : Finset (Fin n)) (by simp [h]: Finset.univ.card = m) =
|
Finset.orderEmbOfFin (Finset.univ : Finset (Fin n)) (by simp [h]: Finset.univ.card = m) =
|
||||||
(Fin.castOrderIso h.symm).toOrderEmbedding := by
|
(Fin.castOrderIso h.symm).toOrderEmbedding := by
|
||||||
symm
|
symm
|
||||||
have h1 : (Fin.castOrderIso h.symm).toFun =
|
have h1 : (Fin.castOrderIso h.symm).toFun =
|
||||||
( Finset.orderEmbOfFin (Finset.univ : Finset (Fin n))
|
(Finset.orderEmbOfFin (Finset.univ : Finset (Fin n))
|
||||||
(by simp [h]: Finset.univ.card = m)).toFun := by
|
(by simp [h]: Finset.univ.card = m)).toFun := by
|
||||||
apply Finset.orderEmbOfFin_unique
|
apply Finset.orderEmbOfFin_unique
|
||||||
intro x
|
intro x
|
||||||
|
@ -332,7 +326,6 @@ lemma orderEmbOfFin_univ (n m : ℕ) (h : n = m):
|
||||||
exact fun ⦃a b⦄ a => a
|
exact fun ⦃a b⦄ a => a
|
||||||
exact Eq.symm (Fin.orderEmbedding_eq (congrArg Set.range (id (Eq.symm h1))))
|
exact Eq.symm (Fin.orderEmbedding_eq (congrArg Set.range (id (Eq.symm h1))))
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Contracting an `ColorIndexList`
|
## Contracting an `ColorIndexList`
|
||||||
|
@ -388,7 +381,6 @@ lemma contr_areDualInSelf (i j : Fin l.contr.length) :
|
||||||
l.contr.AreDualInSelf i j ↔ False := by
|
l.contr.AreDualInSelf i j ↔ False := by
|
||||||
simp [contr]
|
simp [contr]
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Contract equiv
|
## Contract equiv
|
||||||
|
@ -558,7 +550,6 @@ lemma iff_bool (l l2 : ColorIndexList 𝓒) : AppendCond l l2 ↔ bool l.toIndex
|
||||||
|
|
||||||
end AppendCond
|
end AppendCond
|
||||||
|
|
||||||
|
|
||||||
end ColorIndexList
|
end ColorIndexList
|
||||||
|
|
||||||
end IndexNotation
|
end IndexNotation
|
||||||
|
|
|
@ -14,7 +14,6 @@ import Mathlib.Data.Finset.Sort
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
|
@ -131,7 +130,7 @@ lemma contrIndexList_areDualInSelf_false (i j : Fin l.contrIndexList.length) :
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma contrIndexList_of_withDual_empty (h : l.withDual = ∅) : l.contrIndexList = l := by
|
lemma contrIndexList_of_withDual_empty (h : l.withDual = ∅) : l.contrIndexList = l := by
|
||||||
have h1 := l.withDual_union_withoutDual
|
have h1 := l.withDual_union_withoutDual
|
||||||
rw [h , Finset.empty_union] at h1
|
rw [h, Finset.empty_union] at h1
|
||||||
apply ext
|
apply ext
|
||||||
rw [@List.ext_get_iff]
|
rw [@List.ext_get_iff]
|
||||||
change l.contrIndexList.length = l.length ∧ _
|
change l.contrIndexList.length = l.length ∧ _
|
||||||
|
@ -167,7 +166,6 @@ lemma contrIndexList_getDualInOther?_self (i : Fin l.contrIndexList.length) :
|
||||||
exact (contrIndexList_areDualInSelf_false l i j).mp h
|
exact (contrIndexList_areDualInSelf_false l i j).mp h
|
||||||
exact Fin.ge_of_eq (id (Eq.symm h1))
|
exact Fin.ge_of_eq (id (Eq.symm h1))
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Splitting withUniqueDual
|
## Splitting withUniqueDual
|
||||||
|
@ -276,13 +274,11 @@ def withUniqueDualLTEquivGT : l.withUniqueDualLT ≃ l.withUniqueDualGT where
|
||||||
The first `l.withUniqueDualLT` are taken to themselves, whilst the second copy are
|
The first `l.withUniqueDualLT` are taken to themselves, whilst the second copy are
|
||||||
taken to their dual. -/
|
taken to their dual. -/
|
||||||
def withUniqueLTGTEquiv : l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual := by
|
def withUniqueLTGTEquiv : l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual := by
|
||||||
apply (Equiv.sumCongr (Equiv.refl _ ) l.withUniqueDualLTEquivGT).trans
|
apply (Equiv.sumCongr (Equiv.refl _) l.withUniqueDualLTEquivGT).trans
|
||||||
apply (Equiv.Finset.union _ _ l.withUniqueDualLT_disjoint_withUniqueDualGT).trans
|
apply (Equiv.Finset.union _ _ l.withUniqueDualLT_disjoint_withUniqueDualGT).trans
|
||||||
apply (Equiv.subtypeEquivRight (by simp only [l.withUniqueDualLT_union_withUniqueDualGT,
|
apply (Equiv.subtypeEquivRight (by simp only [l.withUniqueDualLT_union_withUniqueDualGT,
|
||||||
implies_true]))
|
implies_true]))
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## withUniqueDualInOther equal univ
|
## withUniqueDualInOther equal univ
|
||||||
|
@ -316,7 +312,7 @@ lemma withUniqueDualInOther_eq_univ_symm (hl : l.length = l2.length)
|
||||||
let S' : Finset (Fin l2.length) :=
|
let S' : Finset (Fin l2.length) :=
|
||||||
Finset.map ⟨Subtype.val, Subtype.val_injective⟩
|
Finset.map ⟨Subtype.val, Subtype.val_injective⟩
|
||||||
(Equiv.finsetCongr
|
(Equiv.finsetCongr
|
||||||
(l.getDualInOtherEquiv l2) Finset.univ )
|
(l.getDualInOtherEquiv l2) Finset.univ)
|
||||||
have hSCard : S'.card = l2.length := by
|
have hSCard : S'.card = l2.length := by
|
||||||
rw [Finset.card_map]
|
rw [Finset.card_map]
|
||||||
simp only [Finset.univ_eq_attach, Equiv.finsetCongr_apply, Finset.card_map, Finset.card_attach]
|
simp only [Finset.univ_eq_attach, Equiv.finsetCongr_apply, Finset.card_map, Finset.card_attach]
|
||||||
|
|
|
@ -16,13 +16,11 @@ We define the functions `getDual?` and `getDualInOther?` which return the
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
variable (l l2 : IndexList X)
|
variable (l l2 : IndexList X)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## The getDual? Function
|
## The getDual? Function
|
||||||
|
@ -34,13 +32,11 @@ variable (l l2 : IndexList X)
|
||||||
def getDual? (i : Fin l.length) : Option (Fin l.length) :=
|
def getDual? (i : Fin l.length) : Option (Fin l.length) :=
|
||||||
Fin.find (fun j => l.AreDualInSelf i j)
|
Fin.find (fun j => l.AreDualInSelf i j)
|
||||||
|
|
||||||
|
|
||||||
/-- Given an `i`, if a dual exists in the other list,
|
/-- Given an `i`, if a dual exists in the other list,
|
||||||
outputs the first such dual, otherwise outputs `none`. -/
|
outputs the first such dual, otherwise outputs `none`. -/
|
||||||
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
|
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
|
||||||
Fin.find (fun j => l.AreDualInOther l2 i j)
|
Fin.find (fun j => l.AreDualInOther l2 i j)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Basic properties of getDual?
|
## Basic properties of getDual?
|
||||||
|
@ -124,7 +120,7 @@ lemma getDual?_get_areDualInSelf (i : Fin l.length) (h : (l.getDual? i).isSome)
|
||||||
|
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma getDual?_areDualInSelf_get (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
lemma getDual?_areDualInSelf_get (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||||
l.AreDualInSelf i ((l.getDual? i).get h):=
|
l.AreDualInSelf i ((l.getDual? i).get h) :=
|
||||||
AreDualInSelf.symm (getDual?_get_areDualInSelf l i h)
|
AreDualInSelf.symm (getDual?_get_areDualInSelf l i h)
|
||||||
|
|
||||||
@[simp]
|
@[simp]
|
||||||
|
|
|
@ -18,7 +18,6 @@ namespace IndexList
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
variable (l l2 l3 : IndexList X)
|
variable (l l2 l3 : IndexList X)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## withDual equal to withUniqueDual
|
## withDual equal to withUniqueDual
|
||||||
|
@ -179,14 +178,12 @@ lemma withUniqueDualInOther_eq_withDualInOther_of_append_symm :
|
||||||
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l2 l3
|
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l2 l3
|
||||||
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l3 l2
|
exact l.withUniqueDualInOther_eq_withDualInOther_of_append_symm' l3 l2
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## withDual equal to withUniqueDual append conditions
|
## withDual equal to withUniqueDual append conditions
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
lemma append_withDual_eq_withUniqueDual_iff :
|
lemma append_withDual_eq_withUniqueDual_iff :
|
||||||
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
|
(l ++ l2).withUniqueDual = (l ++ l2).withDual ↔
|
||||||
((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) ∪ l.withUniqueDualInOther l2)
|
((l.withUniqueDual ∩ (l.withDualInOther l2)ᶜ) ∪ l.withUniqueDualInOther l2)
|
||||||
|
@ -292,7 +289,6 @@ lemma append_withDual_eq_withUniqueDual_swap :
|
||||||
rw [withUniqueDualInOther_eq_withDualInOther_of_append_symm]
|
rw [withUniqueDualInOther_eq_withDualInOther_of_append_symm]
|
||||||
rw [withUniqueDualInOther_eq_withDualInOther_append_of_symm]
|
rw [withUniqueDualInOther_eq_withDualInOther_append_of_symm]
|
||||||
|
|
||||||
|
|
||||||
end IndexList
|
end IndexList
|
||||||
|
|
||||||
end IndexNotation
|
end IndexNotation
|
||||||
|
|
|
@ -18,7 +18,6 @@ a Tensor Color structure.
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace ColorIndexList
|
namespace ColorIndexList
|
||||||
|
|
||||||
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||||
|
@ -63,7 +62,6 @@ namespace ContrPerm
|
||||||
|
|
||||||
variable {l l' l2 l3 : ColorIndexList 𝓒}
|
variable {l l' l2 l3 : ColorIndexList 𝓒}
|
||||||
|
|
||||||
|
|
||||||
@[symm]
|
@[symm]
|
||||||
lemma symm (h : ContrPerm l l') : ContrPerm l' l := by
|
lemma symm (h : ContrPerm l l') : ContrPerm l' l := by
|
||||||
rw [ContrPerm] at h ⊢
|
rw [ContrPerm] at h ⊢
|
||||||
|
@ -148,7 +146,6 @@ lemma contrPermEquiv_colorMap_iso' {l l' : ColorIndexList 𝓒} (h : ContrPerm l
|
||||||
rw [ColorMap.MapIso.symm']
|
rw [ColorMap.MapIso.symm']
|
||||||
exact contrPermEquiv_colorMap_iso h
|
exact contrPermEquiv_colorMap_iso h
|
||||||
|
|
||||||
|
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma contrPermEquiv_refl : contrPermEquiv (@ContrPerm.refl 𝓒 _ l) = Equiv.refl _ := by
|
lemma contrPermEquiv_refl : contrPermEquiv (@ContrPerm.refl 𝓒 _ l) = Equiv.refl _ := by
|
||||||
simp [contrPermEquiv, ContrPerm.refl]
|
simp [contrPermEquiv, ContrPerm.refl]
|
||||||
|
|
|
@ -74,7 +74,6 @@ lemma ext {T₁ T₂ : 𝓣.TensorIndex} (hi : T₁.toColorIndexList = T₂.toCo
|
||||||
subst hi
|
subst hi
|
||||||
simp_all
|
simp_all
|
||||||
|
|
||||||
|
|
||||||
lemma index_eq_of_eq {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ = T₂) :
|
lemma index_eq_of_eq {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ = T₂) :
|
||||||
T₁.toColorIndexList = T₂.toColorIndexList := by
|
T₁.toColorIndexList = T₂.toColorIndexList := by
|
||||||
cases h
|
cases h
|
||||||
|
@ -197,8 +196,6 @@ lemma smul_contr (r : R) (T : 𝓣.TensorIndex) : (r • T).contr = r • T.cont
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
/-- An (equivalence) relation on two `TensorIndex`.
|
/-- An (equivalence) relation on two `TensorIndex`.
|
||||||
The point in this equivalence relation is that certain things (like the
|
The point in this equivalence relation is that certain things (like the
|
||||||
permutation of indices, the contraction of indices, or rising or lowering indices) can be placed
|
permutation of indices, the contraction of indices, or rising or lowering indices) can be placed
|
||||||
|
@ -320,10 +317,8 @@ lemma rel_left {T₁ T₁' T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h'
|
||||||
lemma rel_right {T₁ T₂ T₂' : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₂ ≈ T₂') :
|
lemma rel_right {T₁ T₂ T₂' : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₂ ≈ T₂') :
|
||||||
AddCond T₁ T₂' := h.trans h'.1
|
AddCond T₁ T₂' := h.trans h'.1
|
||||||
|
|
||||||
|
|
||||||
end AddCond
|
end AddCond
|
||||||
|
|
||||||
|
|
||||||
/-- The equivalence between indices after contraction given a `AddCond`. -/
|
/-- The equivalence between indices after contraction given a `AddCond`. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
def addCondEquiv {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
def addCondEquiv {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
||||||
|
@ -486,7 +481,6 @@ lemma add_assoc {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂} (h
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
/-- The condition on two `TensorIndex` which is true if and only if their `ColorIndexList`s
|
/-- The condition on two `TensorIndex` which is true if and only if their `ColorIndexList`s
|
||||||
are related by the condition `AppendCond`. That is, they can be appended to form a
|
are related by the condition `AppendCond`. That is, they can be appended to form a
|
||||||
`ColorIndexList`. -/
|
`ColorIndexList`. -/
|
||||||
|
@ -515,7 +509,6 @@ lemma prod_toColorIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T
|
||||||
lemma prod_toIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T₂) :
|
lemma prod_toIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T₂) :
|
||||||
(prod T₁ T₂ h).toIndexList = T₁.toIndexList ++ T₂.toIndexList := rfl
|
(prod T₁ T₂ h).toIndexList = T₁.toIndexList ++ T₂.toIndexList := rfl
|
||||||
|
|
||||||
|
|
||||||
end TensorIndex
|
end TensorIndex
|
||||||
end
|
end
|
||||||
end TensorStructure
|
end TensorStructure
|
||||||
|
|
|
@ -13,10 +13,8 @@ We define the finite sets of indices in an index list which have a dual
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
|
@ -65,7 +63,6 @@ lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma mem_withInDualOther_iff_isSome (i : Fin l.length) :
|
lemma mem_withInDualOther_iff_isSome (i : Fin l.length) :
|
||||||
i ∈ l.withDualInOther l2 ↔ (l.getDualInOther? l2 i).isSome := by
|
i ∈ l.withDualInOther l2 ↔ (l.getDualInOther? l2 i).isSome := by
|
||||||
|
@ -114,7 +111,6 @@ lemma append_withDualInOther_eq :
|
||||||
| Sum.inr k =>
|
| Sum.inr k =>
|
||||||
simp
|
simp
|
||||||
|
|
||||||
|
|
||||||
lemma withDualInOther_append_eq : l.withDualInOther (l2 ++ l3) =
|
lemma withDualInOther_append_eq : l.withDualInOther (l2 ++ l3) =
|
||||||
l.withDualInOther l2 ∪ l.withDualInOther l3 := by
|
l.withDualInOther l2 ∪ l.withDualInOther l3 := by
|
||||||
ext i
|
ext i
|
||||||
|
|
|
@ -15,13 +15,11 @@ Finite sets of indices with a unique dual.
|
||||||
|
|
||||||
namespace IndexNotation
|
namespace IndexNotation
|
||||||
|
|
||||||
|
|
||||||
namespace IndexList
|
namespace IndexList
|
||||||
|
|
||||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||||
variable (l l2 l3 : IndexList X)
|
variable (l l2 l3 : IndexList X)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Unique duals
|
## Unique duals
|
||||||
|
@ -61,8 +59,6 @@ lemma mem_withDual_of_withUniqueDual (i : l.withUniqueDual) :
|
||||||
i.1 ∈ l.withDual :=
|
i.1 ∈ l.withDual :=
|
||||||
mem_withDual_of_mem_withUniqueDual l (↑i) i.2
|
mem_withDual_of_mem_withUniqueDual l (↑i) i.2
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Basic properties of withUniqueDualInOther
|
## Basic properties of withUniqueDualInOther
|
||||||
|
@ -168,7 +164,6 @@ lemma getDual?_get_of_mem_withUnique_mem (i : Fin l.length) (h : i ∈ l.withUni
|
||||||
subst h1
|
subst h1
|
||||||
exact Eq.symm (getDual?_getDual?_get_of_withUniqueDual l i h)
|
exact Eq.symm (getDual?_getDual?_get_of_withUniqueDual l i h)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Properties of getDualInOther? and withUniqueDualInOther
|
## Properties of getDualInOther? and withUniqueDualInOther
|
||||||
|
@ -213,7 +208,7 @@ lemma getDualInOther?_get_getDualInOther?_get_of_withUniqueDualInOther
|
||||||
have h' : l.AreDualInSelf i ((l2.getDualInOther? l ((l.getDualInOther? l2 i).get
|
have h' : l.AreDualInSelf i ((l2.getDualInOther? l ((l.getDualInOther? l2 i).get
|
||||||
(mem_withUniqueDualInOther_isSome l l2 i h))).get
|
(mem_withUniqueDualInOther_isSome l l2 i h))).get
|
||||||
(l.getDualInOther?_getDualInOther?_get_isSome l2 i
|
(l.getDualInOther?_getDualInOther?_get_isSome l2 i
|
||||||
(mem_withUniqueDualInOther_isSome l l2 i h))):= by
|
(mem_withUniqueDualInOther_isSome l l2 i h))) := by
|
||||||
simp [AreDualInSelf, hn]
|
simp [AreDualInSelf, hn]
|
||||||
exact fun a => hn (id (Eq.symm a))
|
exact fun a => hn (id (Eq.symm a))
|
||||||
have h1 := l.not_mem_withDual_of_withUniqueDualInOther l2 ⟨i, h⟩
|
have h1 := l.not_mem_withDual_of_withUniqueDualInOther l2 ⟨i, h⟩
|
||||||
|
@ -285,17 +280,12 @@ lemma getDualInOther?_get_of_mem_withUniqueInOther_mem (i : Fin l.length)
|
||||||
rw [Fin.find_eq_none_iff] at h
|
rw [Fin.find_eq_none_iff] at h
|
||||||
simp_all only
|
simp_all only
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma getDualInOther?_self_of_mem_withUniqueInOther (i : Fin l.length)
|
lemma getDualInOther?_self_of_mem_withUniqueInOther (i : Fin l.length)
|
||||||
(h : i ∈ l.withUniqueDualInOther l) :
|
(h : i ∈ l.withUniqueDualInOther l) :
|
||||||
l.getDualInOther? l i = some i := by
|
l.getDualInOther? l i = some i := by
|
||||||
rw [all_dual_eq_of_withUniqueDualInOther l l i h i rfl]
|
rw [all_dual_eq_of_withUniqueDualInOther l l i h i rfl]
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## Properties of getDual?, withUniqueDual and append
|
## Properties of getDual?, withUniqueDual and append
|
||||||
|
@ -319,7 +309,6 @@ lemma append_inl_not_mem_withDual_of_withDualInOther (i : Fin l.length)
|
||||||
simp only [getDual?_isSome_append_inl_iff] at ht
|
simp only [getDual?_isSome_append_inl_iff] at ht
|
||||||
simp_all
|
simp_all
|
||||||
|
|
||||||
|
|
||||||
lemma append_inr_not_mem_withDual_of_withDualInOther (i : Fin l2.length)
|
lemma append_inr_not_mem_withDual_of_withDualInOther (i : Fin l2.length)
|
||||||
(h : appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual) :
|
(h : appendEquiv (Sum.inr i) ∈ (l ++ l2).withUniqueDual) :
|
||||||
i ∈ l2.withDual ↔ ¬ i ∈ l2.withDualInOther l := by
|
i ∈ l2.withDual ↔ ¬ i ∈ l2.withDualInOther l := by
|
||||||
|
@ -337,7 +326,6 @@ lemma append_inr_not_mem_withDual_of_withDualInOther (i : Fin l2.length)
|
||||||
simp only [getDual?_isSome_append_inr_iff] at ht
|
simp only [getDual?_isSome_append_inr_iff] at ht
|
||||||
simp_all
|
simp_all
|
||||||
|
|
||||||
|
|
||||||
lemma getDual?_append_symm_of_withUniqueDual_of_inl (i : Fin l.length) (k : Fin l2.length)
|
lemma getDual?_append_symm_of_withUniqueDual_of_inl (i : Fin l.length) (k : Fin l2.length)
|
||||||
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
|
(h : appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDual) :
|
||||||
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inr k))
|
(l ++ l2).getDual? (appendEquiv (Sum.inl i)) = some (appendEquiv (Sum.inr k))
|
||||||
|
@ -660,8 +648,6 @@ lemma append_withUniqueDual_disjSum : (l ++ l2).withUniqueDual =
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
lemma mem_append_withUniqueDualInOther_symm (i : Fin l.length) :
|
lemma mem_append_withUniqueDualInOther_symm (i : Fin l.length) :
|
||||||
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDualInOther l3 ↔
|
appendEquiv (Sum.inl i) ∈ (l ++ l2).withUniqueDualInOther l3 ↔
|
||||||
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDualInOther l3 := by
|
appendEquiv (Sum.inr i) ∈ (l2 ++ l).withUniqueDualInOther l3 := by
|
||||||
|
@ -744,9 +730,8 @@ lemma getDualInOther?_append_symm_of_withUniqueDual_of_inr (i : Fin l.length)
|
||||||
<;> by_cases ho : (l.getDualInOther? l3 i).isSome
|
<;> by_cases ho : (l.getDualInOther? l3 i).isSome
|
||||||
<;> simp_all [hs]
|
<;> simp_all [hs]
|
||||||
|
|
||||||
|
|
||||||
lemma mem_withUniqueDualInOther_symm' (i : Fin l.length)
|
lemma mem_withUniqueDualInOther_symm' (i : Fin l.length)
|
||||||
(h : i ∈ l.withUniqueDualInOther (l2 ++ l3)):
|
(h : i ∈ l.withUniqueDualInOther (l2 ++ l3)) :
|
||||||
i ∈ l.withUniqueDualInOther (l3 ++ l2) := by
|
i ∈ l.withUniqueDualInOther (l3 ++ l2) := by
|
||||||
have h' := h
|
have h' := h
|
||||||
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true,
|
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true,
|
||||||
|
@ -803,7 +788,6 @@ lemma mem_withUniqueDualInOther_symm (i : Fin l.length) :
|
||||||
i ∈ l.withUniqueDualInOther (l2 ++ l3) ↔ i ∈ l.withUniqueDualInOther (l3 ++ l2) :=
|
i ∈ l.withUniqueDualInOther (l2 ++ l3) ↔ i ∈ l.withUniqueDualInOther (l3 ++ l2) :=
|
||||||
Iff.intro (l.mem_withUniqueDualInOther_symm' l2 l3 i) (l.mem_withUniqueDualInOther_symm' l3 l2 i)
|
Iff.intro (l.mem_withUniqueDualInOther_symm' l2 l3 i) (l.mem_withUniqueDualInOther_symm' l3 l2 i)
|
||||||
|
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
||||||
## The equivalence defined by getDual?
|
## The equivalence defined by getDual?
|
||||||
|
@ -847,7 +831,6 @@ lemma getDualInOtherEquiv_self_refl : l.getDualInOtherEquiv l = Equiv.refl _ :=
|
||||||
intro x
|
intro x
|
||||||
simp [getDualInOtherEquiv]
|
simp [getDualInOtherEquiv]
|
||||||
|
|
||||||
|
|
||||||
end IndexList
|
end IndexList
|
||||||
|
|
||||||
end IndexNotation
|
end IndexNotation
|
||||||
|
|
|
@ -89,7 +89,6 @@ noncomputable def fromIndexStringColor {cn : Fin n → realTensorColor.Color}
|
||||||
IndexList.ColorCond.iff_bool.mpr hC⟩ hn
|
IndexList.ColorCond.iff_bool.mpr hC⟩ hn
|
||||||
(TensorColor.ColorMap.DualMap.boolFin_DualMap hd)
|
(TensorColor.ColorMap.DualMap.boolFin_DualMap hd)
|
||||||
|
|
||||||
|
|
||||||
/-- A tactic used to prove `boolFin` for real Lornetz tensors. -/
|
/-- A tactic used to prove `boolFin` for real Lornetz tensors. -/
|
||||||
macro "dualMapTactic" : tactic =>
|
macro "dualMapTactic" : tactic =>
|
||||||
`(tactic| {
|
`(tactic| {
|
||||||
|
@ -114,7 +113,6 @@ macro "prodTactic" : tactic =>
|
||||||
/-- The product of Real Lorentz tensors. Conditions on indices are checked automatically. -/
|
/-- The product of Real Lorentz tensors. Conditions on indices are checked automatically. -/
|
||||||
notation:10 T "⊗ᵀ" S:11 => TensorIndex.prod T S (by prodTactic)
|
notation:10 T "⊗ᵀ" S:11 => TensorIndex.prod T S (by prodTactic)
|
||||||
|
|
||||||
|
|
||||||
/-- An example showing the allowed constructions. -/
|
/-- An example showing the allowed constructions. -/
|
||||||
example (T : (realLorentzTensor d).Tensor ![ColorType.up, ColorType.down]) : True := by
|
example (T : (realLorentzTensor d).Tensor ![ColorType.up, ColorType.down]) : True := by
|
||||||
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
|
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue