chore: Clean up index notation
This commit is contained in:
parent
ac11a510cf
commit
3cd0980f5a
34 changed files with 0 additions and 8197 deletions
|
@ -1,145 +0,0 @@
|
|||
/-
|
||||
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.LorentzVector.Contraction
|
||||
import HepLean.SpaceTime.LorentzVector.LorentzAction
|
||||
import HepLean.Tensors.MulActionTensor
|
||||
/-!
|
||||
|
||||
# Real Lorentz tensors
|
||||
|
||||
-/
|
||||
|
||||
open TensorProduct
|
||||
open minkowskiMatrix
|
||||
|
||||
namespace realTensorColor
|
||||
|
||||
variable {d : ℕ}
|
||||
/-!
|
||||
|
||||
## Definitions and lemmas needed to define a `LorentzTensorStructure`
|
||||
|
||||
-/
|
||||
|
||||
/-- The type colors for real Lorentz tensors. -/
|
||||
inductive ColorType
|
||||
| up
|
||||
| down
|
||||
|
||||
/-- An equivalence between `ColorType ≃ Fin 1 ⊕ Fin 1`. -/
|
||||
def colorTypEquivFin1Fin1 : ColorType ≃ Fin 1 ⊕ Fin 1 where
|
||||
toFun
|
||||
| ColorType.up => Sum.inl ⟨0, Nat.zero_lt_one⟩
|
||||
| ColorType.down => Sum.inr ⟨0, Nat.zero_lt_one⟩
|
||||
invFun
|
||||
| Sum.inl _ => ColorType.up
|
||||
| Sum.inr _ => ColorType.down
|
||||
left_inv := by
|
||||
intro x
|
||||
cases x
|
||||
<;> simp only
|
||||
right_inv := by
|
||||
intro x
|
||||
cases' x with f f
|
||||
<;> simpa [Fin.zero_eta, Fin.isValue, Sum.inl.injEq] using (Fin.fin_one_eq_zero f).symm
|
||||
|
||||
instance : DecidableEq ColorType :=
|
||||
Equiv.decidableEq colorTypEquivFin1Fin1
|
||||
|
||||
instance : Fintype ColorType where
|
||||
elems := {ColorType.up, ColorType.down}
|
||||
complete := by
|
||||
intro x
|
||||
cases x
|
||||
· exact Finset.mem_insert_self ColorType.up {ColorType.down}
|
||||
· apply Finset.insert_eq_self.mp
|
||||
rfl
|
||||
|
||||
end realTensorColor
|
||||
|
||||
noncomputable section
|
||||
|
||||
open realTensorColor
|
||||
|
||||
/-- The color structure for real lorentz tensors. -/
|
||||
def realTensorColor : TensorColor where
|
||||
Color := ColorType
|
||||
τ μ :=
|
||||
match μ with
|
||||
| .up => .down
|
||||
| .down => .up
|
||||
τ_involutive μ :=
|
||||
match μ with
|
||||
| .up => rfl
|
||||
| .down => rfl
|
||||
|
||||
instance : Fintype realTensorColor.Color := realTensorColor.instFintypeColorType
|
||||
|
||||
instance : DecidableEq realTensorColor.Color := realTensorColor.instDecidableEqColorType
|
||||
|
||||
/-! TODO: Set up the notation `𝓛𝓣ℝ` or similar. -/
|
||||
/-- The `TensorStructure` associated with real Lorentz tensors. -/
|
||||
def realLorentzTensor (d : ℕ) : TensorStructure ℝ where
|
||||
toTensorColor := realTensorColor
|
||||
ColorModule μ :=
|
||||
match μ with
|
||||
| .up => LorentzVector d
|
||||
| .down => CovariantLorentzVector d
|
||||
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 [realTensorColor, 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 [realTensorColor, 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_rid μ :=
|
||||
match μ with
|
||||
| .up => LorentzVector.unitUp_rid
|
||||
| .down => LorentzVector.unitDown_rid
|
||||
metric μ :=
|
||||
match μ with
|
||||
| realTensorColor.ColorType.up => asTenProd
|
||||
| realTensorColor.ColorType.down => asCoTenProd
|
||||
metric_dual μ :=
|
||||
match μ with
|
||||
| realTensorColor.ColorType.up => asTenProd_contr_asCoTenProd
|
||||
| realTensorColor.ColorType.down => asCoTenProd_contr_asTenProd
|
||||
|
||||
/-- The action of the Lorentz group on real Lorentz tensors. -/
|
||||
instance : MulActionTensor (LorentzGroup d) (realLorentzTensor d) where
|
||||
repColorModule μ :=
|
||||
match μ with
|
||||
| .up => LorentzVector.rep
|
||||
| .down => CovariantLorentzVector.rep
|
||||
contrDual_inv μ _ :=
|
||||
match μ with
|
||||
| .up => TensorProduct.ext' (fun _ _ => LorentzVector.contrUpDown_invariant_lorentzAction)
|
||||
| .down => TensorProduct.ext' (fun _ _ => LorentzVector.contrDownUp_invariant_lorentzAction)
|
||||
metric_inv μ g :=
|
||||
match μ with
|
||||
| .up => asTenProd_invariant g
|
||||
| .down => asCoTenProd_invariant g
|
||||
|
||||
end
|
|
@ -1,120 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.TensorIndex
|
||||
import HepLean.Tensors.IndexNotation.IndexString
|
||||
import HepLean.SpaceTime.LorentzTensor.Real.Basic
|
||||
/-!
|
||||
|
||||
# Index notation for real Lorentz tensors
|
||||
|
||||
This uses the general concepts of index notation in `HepLean.SpaceTime.LorentzTensor.IndexNotation`
|
||||
to define the index notation for real Lorentz tensors.
|
||||
|
||||
-/
|
||||
|
||||
instance : IndexNotation realTensorColor.Color where
|
||||
charList := {'ᵘ', 'ᵤ'}
|
||||
notaEquiv :=
|
||||
⟨fun c =>
|
||||
match c with
|
||||
| realTensorColor.ColorType.up => ⟨'ᵘ', Finset.mem_insert_self 'ᵘ' {'ᵤ'}⟩
|
||||
| realTensorColor.ColorType.down => ⟨'ᵤ', Finset.insert_eq_self.mp (by rfl)⟩,
|
||||
fun c =>
|
||||
if c = 'ᵘ' then realTensorColor.ColorType.up
|
||||
else realTensorColor.ColorType.down,
|
||||
by
|
||||
intro c
|
||||
match c with
|
||||
| realTensorColor.ColorType.up => rfl
|
||||
| realTensorColor.ColorType.down => rfl,
|
||||
by
|
||||
intro c
|
||||
by_cases hc : c = 'ᵘ'
|
||||
· simp only [↓Char.isValue, hc, ↓reduceIte]
|
||||
exact SetCoe.ext (id (Eq.symm hc))
|
||||
· have hc' : c = 'ᵤ' := by
|
||||
have hc2 := c.2
|
||||
simp only [↓Char.isValue, Finset.mem_insert, Finset.mem_singleton] at hc2
|
||||
simp_all
|
||||
simp only [↓Char.isValue, hc', Char.reduceEq, ↓reduceIte]
|
||||
exact SetCoe.ext (id (Eq.symm hc'))⟩
|
||||
|
||||
namespace realLorentzTensor
|
||||
|
||||
open realTensorColor
|
||||
open IndexNotation IndexString
|
||||
open TensorStructure TensorIndex
|
||||
|
||||
variable {d : ℕ}
|
||||
|
||||
instance : IndexNotation (realLorentzTensor d).Color := instIndexNotationColorRealTensorColor
|
||||
instance : DecidableEq (realLorentzTensor d).Color := instDecidableEqColorRealTensorColor
|
||||
|
||||
@[simp]
|
||||
lemma indexNotation_eq_color : @realLorentzTensor.instIndexNotationColor d =
|
||||
instIndexNotationColorRealTensorColor := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma decidableEq_eq_color : @realLorentzTensor.instDecidableEqColor d =
|
||||
instDecidableEqColorRealTensorColor := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma realLorentzTensor_color : (realLorentzTensor d).Color = realTensorColor.Color := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma toTensorColor_eq : (realLorentzTensor d).toTensorColor = realTensorColor := by
|
||||
rfl
|
||||
|
||||
/-- The construction of a tensor index from a tensor and a string satisfying conditions
|
||||
which can be automatically checked. This is a modified version of
|
||||
`TensorStructure.TensorIndex.mkDualMap` specific to real Lorentz tensors. -/
|
||||
noncomputable def fromIndexStringColor {cn : Fin n → realTensorColor.Color}
|
||||
(T : (realLorentzTensor d).Tensor cn) (s : String)
|
||||
(hs : listCharIsIndexString realTensorColor.Color s.toList = true)
|
||||
(hn : n = (toIndexList' s hs).length)
|
||||
(hD : (toIndexList' s hs).OnlyUniqueDuals)
|
||||
(hC : IndexList.ColorCond.bool (toIndexList' s hs))
|
||||
(hd : TensorColor.ColorMap.DualMap.boolFin
|
||||
(toIndexList' s hs).colorMap (cn ∘ Fin.cast hn.symm)) :
|
||||
(realLorentzTensor d).TensorIndex :=
|
||||
TensorStructure.TensorIndex.mkDualMap T ⟨(toIndexList' s hs), hD,
|
||||
IndexList.ColorCond.iff_bool.mpr hC⟩ hn
|
||||
(TensorColor.ColorMap.DualMap.boolFin_DualMap hd)
|
||||
|
||||
/-- A tactic used to prove `boolFin` for real Lornetz tensors. -/
|
||||
macro "dualMapTactic" : tactic =>
|
||||
`(tactic| {
|
||||
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. -/
|
||||
notation:20 T "|" S:21 => fromIndexStringColor T S
|
||||
(by rfl) (by rfl) (by rfl) (by rfl) (by dualMapTactic)
|
||||
|
||||
/-- A tactics used to prove `colorPropBool` for real Lorentz tensors. -/
|
||||
macro "prodTactic" : tactic =>
|
||||
`(tactic| {
|
||||
apply (ColorIndexList.AppendCond.iff_bool _ _).mpr
|
||||
change @ColorIndexList.AppendCond.bool realTensorColor
|
||||
instDecidableEqColorRealTensorColor _ _
|
||||
simp only [prod_toIndexList, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
|
||||
toTensorColor_eq, decidableEq_eq_color]
|
||||
rfl})
|
||||
|
||||
/-- The product of Real Lorentz tensors. Conditions on indices are checked automatically. -/
|
||||
notation:10 T "⊗ᵀ" S:11 => TensorIndex.prod T S (by prodTactic)
|
||||
|
||||
/-- An example showing the allowed constructions. -/
|
||||
example (T : (realLorentzTensor d).Tensor ![ColorType.up, ColorType.down]) : True := by
|
||||
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
|
||||
let _ := T|"ᵤ₁ᵤ₂" ⊗ᵀ T|"ᵘ³ᵤ₄" ⊗ᵀ T|"ᵘ¹ᵘ²" ⊗ᵀ T|"ᵘ⁴ᵤ₃"
|
||||
exact trivial
|
||||
|
||||
end realLorentzTensor
|
|
@ -9,7 +9,6 @@ import HepLean.SpaceTime.SL2C.Basic
|
|||
import HepLean.SpaceTime.LorentzVector.Modules
|
||||
import HepLean.Meta.Informal
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import HepLean.SpaceTime.PauliMatrices.SelfAdjoint
|
||||
/-!
|
||||
|
||||
|
|
|
@ -1,368 +0,0 @@
|
|||
/-
|
||||
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.LorentzVector.LorentzAction
|
||||
import HepLean.SpaceTime.LorentzVector.Covariant
|
||||
import HepLean.Tensors.Basic
|
||||
/-!
|
||||
|
||||
# Contractions of Lorentz vectors
|
||||
|
||||
We define the contraction between a covariant and contravariant Lorentz vector,
|
||||
as well as properties thereof.
|
||||
|
||||
The structures in this file are used in `HepLean.SpaceTime.LorentzTensor.Real.Basic`
|
||||
to define the contraction between indices of Lorentz tensors.
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open TensorProduct
|
||||
|
||||
namespace LorentzVector
|
||||
|
||||
open Matrix
|
||||
variable {d : ℕ} (v : LorentzVector d)
|
||||
|
||||
/-- The bi-linear map defining the contraction of a contravariant Lorentz vector
|
||||
and a covariant Lorentz vector. -/
|
||||
def contrUpDownBi : LorentzVector d →ₗ[ℝ] CovariantLorentzVector d →ₗ[ℝ] ℝ where
|
||||
toFun v := {
|
||||
toFun := fun w => ∑ i, v i * w i,
|
||||
map_add' := by
|
||||
intro w1 w2
|
||||
rw [← Finset.sum_add_distrib]
|
||||
refine Finset.sum_congr rfl (fun i _ => mul_add _ _ _)
|
||||
map_smul' := by
|
||||
intro r w
|
||||
simp only [RingHom.id_apply, smul_eq_mul]
|
||||
rw [Finset.mul_sum]
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
simp only [HSMul.hSMul, SMul.smul]
|
||||
ring}
|
||||
map_add' v1 v2 := by
|
||||
apply LinearMap.ext
|
||||
intro w
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.add_apply]
|
||||
rw [← Finset.sum_add_distrib]
|
||||
refine Finset.sum_congr rfl (fun i _ => add_mul _ _ _)
|
||||
map_smul' r v := by
|
||||
apply LinearMap.ext
|
||||
intro w
|
||||
simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.smul_apply]
|
||||
rw [Finset.smul_sum]
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
simp only [HSMul.hSMul, SMul.smul]
|
||||
exact mul_assoc r (v i) (w i)
|
||||
|
||||
/-- The linear map defining the contraction of a contravariant Lorentz vector
|
||||
and a covariant Lorentz vector. -/
|
||||
def contrUpDown : LorentzVector d ⊗[ℝ] CovariantLorentzVector d →ₗ[ℝ] ℝ :=
|
||||
TensorProduct.lift contrUpDownBi
|
||||
|
||||
lemma contrUpDown_tmul_eq_dotProduct {x : LorentzVector d} {y : CovariantLorentzVector d} :
|
||||
contrUpDown (x ⊗ₜ[ℝ] y) = x ⬝ᵥ y := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrUpDown_stdBasis_left (x : LorentzVector d) (i : Fin 1 ⊕ Fin d) :
|
||||
contrUpDown (x ⊗ₜ[ℝ] (CovariantLorentzVector.stdBasis i)) = x i := by
|
||||
simp only [contrUpDown, contrUpDownBi, lift.tmul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [Finset.sum_eq_single_of_mem i]
|
||||
· simp only [CovariantLorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_eq_same, mul_one]
|
||||
· exact Finset.mem_univ i
|
||||
· intro b _ hbi
|
||||
simp only [CovariantLorentzVector.stdBasis, mul_eq_zero]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_apply, ite_eq_right_iff, one_ne_zero, imp_false]
|
||||
apply Or.inr hbi
|
||||
|
||||
@[simp]
|
||||
lemma contrUpDown_stdBasis_right (x : CovariantLorentzVector d) (i : Fin 1 ⊕ Fin d) :
|
||||
contrUpDown (e i ⊗ₜ[ℝ] x) = x i := by
|
||||
simp only [contrUpDown, contrUpDownBi, lift.tmul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
rw [Finset.sum_eq_single_of_mem i]
|
||||
· erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_eq_same, one_mul]
|
||||
· exact Finset.mem_univ i
|
||||
· intro b _ hbi
|
||||
simp only [CovariantLorentzVector.stdBasis, mul_eq_zero]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_apply, ite_eq_right_iff, one_ne_zero, imp_false]
|
||||
exact Or.intro_left (x b = 0) hbi
|
||||
|
||||
/-- The linear map defining the contraction of a covariant Lorentz vector
|
||||
and a contravariant Lorentz vector. -/
|
||||
def contrDownUp : CovariantLorentzVector d ⊗[ℝ] LorentzVector d →ₗ[ℝ] ℝ :=
|
||||
contrUpDown ∘ₗ (TensorProduct.comm ℝ _ _).toLinearMap
|
||||
|
||||
lemma contrDownUp_tmul_eq_dotProduct {x : CovariantLorentzVector d} {y : LorentzVector d} :
|
||||
contrDownUp (x ⊗ₜ[ℝ] y) = x ⬝ᵥ y := by
|
||||
rw [dotProduct_comm x y]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## The unit of the contraction
|
||||
|
||||
-/
|
||||
|
||||
/-- The unit of the contraction of contravariant Lorentz vector and a
|
||||
covariant Lorentz vector. -/
|
||||
def unitUp : CovariantLorentzVector d ⊗[ℝ] LorentzVector d :=
|
||||
∑ i, CovariantLorentzVector.stdBasis i ⊗ₜ[ℝ] LorentzVector.stdBasis i
|
||||
|
||||
lemma unitUp_rid (x : LorentzVector d) :
|
||||
TensorStructure.contrLeftAux contrUpDown (x ⊗ₜ[ℝ] unitUp) = x := by
|
||||
simp only [unitUp, LinearEquiv.refl_toLinearMap]
|
||||
rw [tmul_sum]
|
||||
simp only [TensorStructure.contrLeftAux, LinearEquiv.refl_toLinearMap, Fintype.sum_sum_type,
|
||||
Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton, map_add,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul, map_tmul,
|
||||
contrUpDown_stdBasis_left, LinearMap.id_coe, id_eq, lid_tmul, map_sum, decomp_stdBasis']
|
||||
|
||||
/-- The unit of the contraction of covariant Lorentz vector and a
|
||||
contravariant Lorentz vector. -/
|
||||
def unitDown : LorentzVector d ⊗[ℝ] CovariantLorentzVector d :=
|
||||
∑ i, LorentzVector.stdBasis i ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis i
|
||||
|
||||
lemma unitDown_rid (x : CovariantLorentzVector d) :
|
||||
TensorStructure.contrLeftAux contrDownUp (x ⊗ₜ[ℝ] unitDown) = x := by
|
||||
simp only [unitDown, LinearEquiv.refl_toLinearMap]
|
||||
rw [tmul_sum]
|
||||
simp only [TensorStructure.contrLeftAux, contrDownUp, LinearEquiv.refl_toLinearMap,
|
||||
Fintype.sum_sum_type, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
Finset.sum_singleton, map_add, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
assoc_symm_tmul, map_tmul, comm_tmul, contrUpDown_stdBasis_right, LinearMap.id_coe, id_eq,
|
||||
lid_tmul, map_sum, CovariantLorentzVector.decomp_stdBasis']
|
||||
|
||||
/-!
|
||||
|
||||
# Contractions and the Lorentz actions
|
||||
|
||||
-/
|
||||
open Matrix
|
||||
|
||||
@[simp]
|
||||
lemma contrUpDown_invariant_lorentzAction : @contrUpDown d ((LorentzVector.rep g) x ⊗ₜ[ℝ]
|
||||
(CovariantLorentzVector.rep g) y) = contrUpDown (x ⊗ₜ[ℝ] y) := by
|
||||
rw [contrUpDown_tmul_eq_dotProduct, contrUpDown_tmul_eq_dotProduct]
|
||||
simp only [rep_apply, CovariantLorentzVector.rep_apply]
|
||||
rw [Matrix.dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
|
||||
simp only [LorentzGroup.subtype_inv_mul, one_mulVec]
|
||||
|
||||
@[simp]
|
||||
lemma contrDownUp_invariant_lorentzAction : @contrDownUp d ((CovariantLorentzVector.rep g) x ⊗ₜ[ℝ]
|
||||
(LorentzVector.rep g) y) = contrDownUp (x ⊗ₜ[ℝ] y) := by
|
||||
rw [contrDownUp_tmul_eq_dotProduct, contrDownUp_tmul_eq_dotProduct]
|
||||
rw [dotProduct_comm, dotProduct_comm x y]
|
||||
simp only [rep_apply, CovariantLorentzVector.rep_apply]
|
||||
rw [Matrix.dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
|
||||
simp only [LorentzGroup.subtype_inv_mul, one_mulVec]
|
||||
|
||||
end LorentzVector
|
||||
|
||||
namespace minkowskiMatrix
|
||||
open LorentzVector
|
||||
open TensorStructure
|
||||
open scoped minkowskiMatrix
|
||||
variable {d : ℕ}
|
||||
|
||||
/-- The metric tensor as an element of `LorentzVector d ⊗[ℝ] LorentzVector d`. -/
|
||||
def asTenProd : LorentzVector d ⊗[ℝ] LorentzVector d :=
|
||||
∑ μ, ∑ ν, η μ ν • (LorentzVector.stdBasis μ ⊗ₜ[ℝ] LorentzVector.stdBasis ν)
|
||||
|
||||
lemma asTenProd_diag :
|
||||
@asTenProd d = ∑ μ, η μ μ • (LorentzVector.stdBasis μ ⊗ₜ[ℝ] LorentzVector.stdBasis μ) := by
|
||||
simp only [asTenProd]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [Finset.sum_eq_single μ]
|
||||
· intro ν _ hμν
|
||||
rw [minkowskiMatrix.off_diag_zero hμν.symm]
|
||||
exact TensorProduct.zero_smul (e μ ⊗ₜ[ℝ] e ν)
|
||||
· intro a
|
||||
rename_i j
|
||||
exact False.elim (a j)
|
||||
|
||||
/-- The metric tensor as an element of `CovariantLorentzVector d ⊗[ℝ] CovariantLorentzVector d`. -/
|
||||
def asCoTenProd : CovariantLorentzVector d ⊗[ℝ] CovariantLorentzVector d :=
|
||||
∑ μ, ∑ ν, η μ ν • (CovariantLorentzVector.stdBasis μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis ν)
|
||||
|
||||
lemma asCoTenProd_diag :
|
||||
@asCoTenProd d = ∑ μ, η μ μ • (CovariantLorentzVector.stdBasis μ ⊗ₜ[ℝ]
|
||||
CovariantLorentzVector.stdBasis μ) := by
|
||||
simp only [asCoTenProd]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [Finset.sum_eq_single μ]
|
||||
· intro ν _ hμν
|
||||
rw [minkowskiMatrix.off_diag_zero hμν.symm]
|
||||
simp only [zero_smul]
|
||||
· intro a
|
||||
simp_all only
|
||||
|
||||
@[simp]
|
||||
lemma contrLeft_asTenProd (x : CovariantLorentzVector d) :
|
||||
contrLeftAux contrDownUp (x ⊗ₜ[ℝ] asTenProd) =
|
||||
∑ μ, ((η μ μ * x μ) • LorentzVector.stdBasis μ) := by
|
||||
simp only [asTenProd_diag]
|
||||
rw [tmul_sum]
|
||||
rw [map_sum]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
simp only [contrLeftAux, contrDownUp, LinearEquiv.refl_toLinearMap, tmul_smul, map_smul,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul, map_tmul,
|
||||
comm_tmul, contrUpDown_stdBasis_right, LinearMap.id_coe, id_eq, lid_tmul]
|
||||
exact smul_smul (η μ μ) (x μ) (e μ)
|
||||
|
||||
@[simp]
|
||||
lemma contrLeft_asCoTenProd (x : LorentzVector d) :
|
||||
contrLeftAux contrUpDown (x ⊗ₜ[ℝ] asCoTenProd) =
|
||||
∑ μ, ((η μ μ * x μ) • CovariantLorentzVector.stdBasis μ) := by
|
||||
simp only [asCoTenProd_diag]
|
||||
rw [tmul_sum]
|
||||
rw [map_sum]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
simp only [contrLeftAux, LinearEquiv.refl_toLinearMap, tmul_smul, LinearMapClass.map_smul,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul, map_tmul,
|
||||
contrUpDown_stdBasis_left, LinearMap.id_coe, id_eq, lid_tmul]
|
||||
exact smul_smul (η μ μ) (x μ) (CovariantLorentzVector.stdBasis μ)
|
||||
|
||||
@[simp]
|
||||
lemma asTenProd_contr_asCoTenProd :
|
||||
(contrMidAux (contrUpDown) (asTenProd ⊗ₜ[ℝ] asCoTenProd))
|
||||
= TensorProduct.comm ℝ _ _ (@unitUp d) := by
|
||||
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asTenProd_diag, LinearMap.coe_comp,
|
||||
LinearEquiv.coe_coe, Function.comp_apply]
|
||||
rw [sum_tmul, map_sum, map_sum, unitUp]
|
||||
simp only [Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
Finset.sum_singleton, map_add, comm_tmul, map_sum]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [← tmul_smul, TensorProduct.assoc_tmul]
|
||||
simp only [map_tmul, LinearMap.id_coe, id_eq, contrLeft_asCoTenProd]
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul]
|
||||
· change (η μ μ * (η μ μ * e μ μ)) • e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis μ = _
|
||||
rw [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_eq_same, mul_one, η_apply_mul_η_apply_diag, one_smul]
|
||||
· exact Finset.mem_univ μ
|
||||
· intro ν _ hμν
|
||||
rw [tmul_smul]
|
||||
change (η ν ν * (η μ μ * e μ ν)) • (e μ ⊗ₜ[ℝ] CovariantLorentzVector.stdBasis ν) = _
|
||||
rw [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_apply, mul_ite, mul_one, mul_zero, ite_smul, zero_smul,
|
||||
ite_eq_right_iff, smul_eq_zero, mul_eq_zero]
|
||||
exact fun a => False.elim (hμν a)
|
||||
|
||||
@[simp]
|
||||
lemma asCoTenProd_contr_asTenProd :
|
||||
(contrMidAux (contrDownUp) (asCoTenProd ⊗ₜ[ℝ] asTenProd))
|
||||
= TensorProduct.comm ℝ _ _ (@unitDown d) := by
|
||||
simp only [contrMidAux, LinearEquiv.refl_toLinearMap, asCoTenProd_diag,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply]
|
||||
rw [sum_tmul, map_sum, map_sum, unitDown]
|
||||
simp only [Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
Finset.sum_singleton, map_add, comm_tmul, map_sum]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [smul_tmul, TensorProduct.assoc_tmul]
|
||||
simp only [tmul_smul, LinearMapClass.map_smul, map_tmul, LinearMap.id_coe, id_eq,
|
||||
contrLeft_asTenProd]
|
||||
rw [tmul_sum, Finset.sum_eq_single_of_mem μ, tmul_smul, smul_smul, LorentzVector.stdBasis]
|
||||
· erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_eq_same, mul_one, η_apply_mul_η_apply_diag, one_smul]
|
||||
· exact Finset.mem_univ μ
|
||||
· intro ν _ hμν
|
||||
rw [tmul_smul]
|
||||
rw [LorentzVector.stdBasis]
|
||||
erw [Pi.basisFun_apply]
|
||||
simp only [Pi.single_apply, mul_ite, mul_one, mul_zero, ite_smul, zero_smul,
|
||||
ite_eq_right_iff, smul_eq_zero, mul_eq_zero]
|
||||
exact fun a => False.elim (hμν a)
|
||||
|
||||
lemma asTenProd_invariant (g : LorentzGroup d) :
|
||||
TensorProduct.map (LorentzVector.rep g) (LorentzVector.rep g) asTenProd = asTenProd := by
|
||||
simp only [asTenProd, map_sum, LinearMapClass.map_smul, map_tmul, rep_apply_stdBasis,
|
||||
Matrix.transpose_apply]
|
||||
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d,
|
||||
η μ ν • (g.1 φ μ • e φ) ⊗ₜ[ℝ] ∑ ρ : Fin 1 ⊕ Fin d, g.1 ρ ν • e ρ
|
||||
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))
|
||||
rw [sum_tmul, Finset.smul_sum]
|
||||
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
|
||||
(η μ ν • (g.1 φ μ • e φ) ⊗ₜ[ℝ] (g.1 ρ ν • e ρ))
|
||||
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ =>
|
||||
Finset.sum_congr rfl (fun φ _ => ?_)))
|
||||
rw [tmul_sum, Finset.smul_sum]
|
||||
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_comm)]
|
||||
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => Finset.sum_comm))]
|
||||
rw [Finset.sum_comm]
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_comm)]
|
||||
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d, ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d,
|
||||
((g.1 φ μ * η μ ν * g.1 ρ ν) • (e φ) ⊗ₜ[ℝ] (e ρ))
|
||||
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
|
||||
Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))))
|
||||
rw [smul_tmul, tmul_smul, tmul_smul, smul_smul, smul_smul]
|
||||
ring_nf
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
|
||||
Finset.sum_congr rfl (fun μ _ => Finset.sum_smul.symm)))]
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => Finset.sum_smul.symm))]
|
||||
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
|
||||
(((g.1 * minkowskiMatrix * g.1.transpose : Matrix _ _ _) φ ρ) • (e φ) ⊗ₜ[ℝ] (e ρ))
|
||||
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => ?_))
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
simp only [Matrix.mul_apply, Matrix.transpose_apply]
|
||||
rw [Finset.sum_comm]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [Finset.sum_mul]
|
||||
simp
|
||||
|
||||
open CovariantLorentzVector in
|
||||
lemma asCoTenProd_invariant (g : LorentzGroup d) :
|
||||
TensorProduct.map (CovariantLorentzVector.rep g) (CovariantLorentzVector.rep g)
|
||||
asCoTenProd = asCoTenProd := by
|
||||
simp only [asCoTenProd, map_sum, LinearMapClass.map_smul, map_tmul,
|
||||
CovariantLorentzVector.rep_apply_stdBasis, Matrix.transpose_apply]
|
||||
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d,
|
||||
η μ ν • (g⁻¹.1 μ φ • CovariantLorentzVector.stdBasis φ) ⊗ₜ[ℝ]
|
||||
∑ ρ : Fin 1 ⊕ Fin d, g⁻¹.1 ν ρ • CovariantLorentzVector.stdBasis ρ
|
||||
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))
|
||||
rw [sum_tmul, Finset.smul_sum]
|
||||
trans ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d, ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
|
||||
(η μ ν • (g⁻¹.1 μ φ • CovariantLorentzVector.stdBasis φ) ⊗ₜ[ℝ]
|
||||
(g⁻¹.1 ν ρ • CovariantLorentzVector.stdBasis ρ))
|
||||
· refine Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ =>
|
||||
Finset.sum_congr rfl (fun φ _ => ?_)))
|
||||
rw [tmul_sum, Finset.smul_sum]
|
||||
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_comm)]
|
||||
rw [Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => Finset.sum_comm))]
|
||||
rw [Finset.sum_comm]
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_comm)]
|
||||
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d, ∑ μ : Fin 1 ⊕ Fin d, ∑ ν : Fin 1 ⊕ Fin d,
|
||||
((g⁻¹.1 μ φ * η μ ν * g⁻¹.1 ν ρ) • (CovariantLorentzVector.stdBasis φ) ⊗ₜ[ℝ]
|
||||
(CovariantLorentzVector.stdBasis ρ))
|
||||
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
|
||||
Finset.sum_congr rfl (fun μ _ => Finset.sum_congr rfl (fun ν _ => ?_))))
|
||||
rw [smul_tmul, tmul_smul, tmul_smul, smul_smul, smul_smul]
|
||||
ring_nf
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ =>
|
||||
Finset.sum_congr rfl (fun μ _ => Finset.sum_smul.symm)))]
|
||||
rw [Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => Finset.sum_smul.symm))]
|
||||
trans ∑ φ : Fin 1 ⊕ Fin d, ∑ ρ : Fin 1 ⊕ Fin d,
|
||||
(((g⁻¹.1.transpose * minkowskiMatrix * g⁻¹.1 : Matrix _ _ _) φ ρ) •
|
||||
(CovariantLorentzVector.stdBasis φ) ⊗ₜ[ℝ] (CovariantLorentzVector.stdBasis ρ))
|
||||
· refine Finset.sum_congr rfl (fun φ _ => Finset.sum_congr rfl (fun ρ _ => ?_))
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
simp only [Matrix.mul_apply, Matrix.transpose_apply]
|
||||
rw [Finset.sum_comm]
|
||||
refine Finset.sum_congr rfl (fun μ _ => ?_)
|
||||
rw [Finset.sum_mul]
|
||||
rw [LorentzGroup.transpose_mul_minkowskiMatrix_mul_self]
|
||||
|
||||
end minkowskiMatrix
|
||||
|
||||
end
|
|
@ -6,7 +6,6 @@ Authors: Joseph Tooby-Smith
|
|||
import HepLean.Meta.Informal
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
/-!
|
||||
|
||||
|
|
|
@ -5,7 +5,6 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||
/-!
|
||||
|
|
|
@ -5,7 +5,6 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||
import HepLean.SpaceTime.PauliMatrices.Basic
|
||||
|
|
|
@ -6,7 +6,6 @@ Authors: Joseph Tooby-Smith
|
|||
import HepLean.Meta.Informal
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import HepLean.SpaceTime.WeylFermion.Modules
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
/-!
|
||||
|
|
|
@ -6,7 +6,6 @@ Authors: Joseph Tooby-Smith
|
|||
import HepLean.Meta.Informal
|
||||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import Mathlib.RepresentationTheory.Rep
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.Logic.Equiv.TransferInstance
|
||||
/-!
|
||||
|
||||
|
|
|
@ -1,771 +0,0 @@
|
|||
/-
|
||||
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 Mathlib.LinearAlgebra.PiTensorProduct
|
||||
/-!
|
||||
|
||||
# Structure of Tensors
|
||||
|
||||
This file sets up the structure `TensorStructure` which contains
|
||||
data of types (or `colors`) of indices, the dual of colors, the associated module,
|
||||
contraction between color modules, and the unit of contraction.
|
||||
|
||||
This structure is extended to `DualizeTensorStructure` which adds a metric to the tensor structure,
|
||||
allowing a vector to be taken to its dual vector by contraction with a specified metric.
|
||||
The definition of `DualizeTensorStructure` can be found in
|
||||
`HepLean.SpaceTime.LorentzTensor.RisingLowering`.
|
||||
|
||||
The structure `DualizeTensorStructure` is further extended in
|
||||
`HepLean.SpaceTime.LorentzTensor.LorentzTensorStruct` to add a group action on the tensor space,
|
||||
under which contraction and rising and lowering etc, are invariant.
|
||||
|
||||
## References
|
||||
|
||||
-- For modular operads see: [Raynor][raynor2021graphical]
|
||||
|
||||
-/
|
||||
|
||||
open TensorProduct
|
||||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
/-- The index color data associated with a tensor structure.
|
||||
This corresponds to a type with an involution. -/
|
||||
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) [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
variable {d : ℕ} {X X' Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
|
||||
/-- A relation on colors which is true if the two colors are equal or are duals. -/
|
||||
def colorRel (μ ν : 𝓒.Color) : Prop := μ = ν ∨ μ = 𝓒.τ ν
|
||||
|
||||
instance : Decidable (colorRel 𝓒 μ ν) := instDecidableOr
|
||||
|
||||
omit [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
/-- 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 μ
|
||||
|
||||
instance (μ ν : 𝓒.Color) : Decidable (μ ≈ ν) := instDecidableOr
|
||||
|
||||
instance : DecidableEq (Quotient 𝓒.colorSetoid) :=
|
||||
instDecidableEqQuotientOfDecidableEquiv
|
||||
|
||||
/-- The types of maps from an `X` to `𝓒.Color`. -/
|
||||
def ColorMap (X : Type) := X → 𝓒.Color
|
||||
|
||||
namespace ColorMap
|
||||
|
||||
variable {𝓒 : TensorColor} [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
|
||||
variable (cX : ColorMap 𝓒 X) (cY : ColorMap 𝓒 Y) (cZ : ColorMap 𝓒 Z)
|
||||
|
||||
/-- A relation, given an equivalence of types, between ColorMap which is true
|
||||
if related by composition of the equivalence. -/
|
||||
def MapIso (e : X ≃ Y) (cX : ColorMap 𝓒 X) (cY : ColorMap 𝓒 Y) : Prop := cX = cY ∘ e
|
||||
|
||||
/-- The sum of two color maps, formed by `Sum.elim`. -/
|
||||
def sum (cX : ColorMap 𝓒 X) (cY : ColorMap 𝓒 Y) : ColorMap 𝓒 (Sum X Y) :=
|
||||
Sum.elim cX cY
|
||||
|
||||
/-- The dual of a color map, formed by composition with `𝓒.τ`. -/
|
||||
def dual (cX : ColorMap 𝓒 X) : ColorMap 𝓒 X := 𝓒.τ ∘ cX
|
||||
|
||||
namespace MapIso
|
||||
|
||||
variable {e : X ≃ Y} {e' : Y ≃ Z} {cX : ColorMap 𝓒 X} {cY : ColorMap 𝓒 Y} {cZ : ColorMap 𝓒 Z}
|
||||
variable {cX' : ColorMap 𝓒 X'} {cY' : ColorMap 𝓒 Y'}
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y][Fintype Z]
|
||||
[DecidableEq Z] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
lemma symm (h : cX.MapIso e cY) : cY.MapIso e.symm cX := by
|
||||
rw [MapIso] at h
|
||||
exact (Equiv.eq_comp_symm e cY cX).mpr h.symm
|
||||
|
||||
lemma symm' : cX.MapIso e cY ↔ cY.MapIso e.symm cX := by
|
||||
refine ⟨symm, symm⟩
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z] [DecidableEq Z]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
lemma trans (h : cX.MapIso e cY) (h' : cY.MapIso e' cZ) :
|
||||
cX.MapIso (e.trans e') cZ:= by
|
||||
funext a
|
||||
subst h h'
|
||||
rfl
|
||||
|
||||
omit [Fintype Y'] [DecidableEq Y']
|
||||
|
||||
lemma sum {eX : X ≃ X'} {eY : Y ≃ Y'} (hX : cX.MapIso eX cX') (hY : cY.MapIso eY cY') :
|
||||
(cX.sum cY).MapIso (eX.sumCongr eY) (cX'.sum cY') := by
|
||||
funext x
|
||||
subst hX hY
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl
|
||||
|
||||
lemma dual {e : X ≃ Y} (h : cX.MapIso e cY) :
|
||||
cX.dual.MapIso e cY.dual := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
end MapIso
|
||||
|
||||
end ColorMap
|
||||
|
||||
end TensorColor
|
||||
|
||||
noncomputable section
|
||||
namespace TensorStructure
|
||||
|
||||
/-- An auxillary function to contract the vector space `V1` and `V2` in `V1 ⊗[R] V2 ⊗[R] V3`. -/
|
||||
def contrLeftAux {V1 V2 V3 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
|
||||
[Module R V1] [Module R V2] [Module R V3] (f : V1 ⊗[R] V2 →ₗ[R] R) :
|
||||
V1 ⊗[R] V2 ⊗[R] V3 →ₗ[R] V3 :=
|
||||
(TensorProduct.lid R _).toLinearMap ∘ₗ
|
||||
TensorProduct.map (f) (LinearEquiv.refl R V3).toLinearMap
|
||||
∘ₗ (TensorProduct.assoc R _ _ _).symm.toLinearMap
|
||||
|
||||
/-- An auxillary function to contract the vector space `V1` and `V2` in `(V3 ⊗[R] V1) ⊗[R] V2`. -/
|
||||
def contrRightAux {V1 V2 V3 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
|
||||
[Module R V1] [Module R V2] [Module R V3] (f : V1 ⊗[R] V2 →ₗ[R] R) :
|
||||
(V3 ⊗[R] V1) ⊗[R] V2 →ₗ[R] V3 :=
|
||||
(TensorProduct.rid R _).toLinearMap ∘ₗ
|
||||
TensorProduct.map (LinearEquiv.refl R V3).toLinearMap f ∘ₗ
|
||||
(TensorProduct.assoc R _ _ _).toLinearMap
|
||||
|
||||
/-- An auxillary function to contract the vector space `V1` and `V2` in
|
||||
`V4 ⊗[R] V1 ⊗[R] V2 ⊗[R] V3`. -/
|
||||
def contrMidAux {V1 V2 V3 V4 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
|
||||
[AddCommMonoid V4] [Module R V1] [Module R V2] [Module R V3] [Module R V4]
|
||||
(f : V1 ⊗[R] V2 →ₗ[R] R) : (V4 ⊗[R] V1) ⊗[R] (V2 ⊗[R] V3) →ₗ[R] V4 ⊗[R] V3 :=
|
||||
(TensorProduct.map (LinearEquiv.refl R V4).toLinearMap (contrLeftAux f)) ∘ₗ
|
||||
(TensorProduct.assoc R _ _ _).toLinearMap
|
||||
|
||||
lemma contrRightAux_comp {V1 V2 V3 V4 V5 : Type} [AddCommMonoid V1] [AddCommMonoid V2]
|
||||
[AddCommMonoid V3] [AddCommMonoid V4] [AddCommMonoid V5] [Module R V1] [Module R V3]
|
||||
[Module R V2] [Module R V4] [Module R V5] (f : V2 ⊗[R] V3 →ₗ[R] R) (g : V4 ⊗[R] V5 →ₗ[R] R) :
|
||||
(contrRightAux f ∘ₗ TensorProduct.map (LinearMap.id : V1 ⊗[R] V2 →ₗ[R] V1 ⊗[R] V2)
|
||||
(contrRightAux g)) =
|
||||
(contrRightAux g) ∘ₗ TensorProduct.map (contrMidAux f) LinearMap.id
|
||||
∘ₗ (TensorProduct.assoc R _ _ _).symm.toLinearMap := by
|
||||
apply TensorProduct.ext'
|
||||
intro x y
|
||||
refine TensorProduct.induction_on x (by simp) ?_ (fun x z h1 h2 =>
|
||||
by simp [add_tmul, LinearMap.map_add, h1, h2])
|
||||
intro x1 x2
|
||||
refine TensorProduct.induction_on y (by simp) ?_ (fun x z h1 h2 =>
|
||||
by simp [add_tmul, tmul_add, LinearMap.map_add, h1, h2])
|
||||
intro y x5
|
||||
refine TensorProduct.induction_on y (by simp) ?_ (fun x z h1 h2 =>
|
||||
by simp [add_tmul, tmul_add, LinearMap.map_add, h1, h2])
|
||||
intro x3 x4
|
||||
simp only [contrRightAux, LinearEquiv.refl_toLinearMap, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, map_tmul, LinearMap.id_coe, id_eq, assoc_tmul, rid_tmul, tmul_smul,
|
||||
map_smul, contrMidAux, contrLeftAux, assoc_symm_tmul, lid_tmul]
|
||||
rfl
|
||||
|
||||
end TensorStructure
|
||||
|
||||
/-- An initial structure specifying a tensor system (e.g. a system in which you can
|
||||
define real Lorentz tensors or Einstein notation convention). -/
|
||||
structure TensorStructure (R : Type) [CommSemiring R] extends TensorColor where
|
||||
/-- To each color we associate a module. -/
|
||||
ColorModule : Color → Type
|
||||
/-- Each `ColorModule` has the structure of an additive commutative monoid. -/
|
||||
colorModule_addCommMonoid : ∀ μ, AddCommMonoid (ColorModule μ)
|
||||
/-- Each `ColorModule` has the structure of a module over `R`. -/
|
||||
colorModule_module : ∀ μ, Module R (ColorModule μ)
|
||||
/-- The contraction of a vector with a vector with dual color. -/
|
||||
contrDual : ∀ μ, ColorModule μ ⊗[R] ColorModule (τ μ) →ₗ[R] R
|
||||
/-- The contraction is symmetric. -/
|
||||
contrDual_symm : ∀ μ x y, (contrDual μ) (x ⊗ₜ[R] y) =
|
||||
(contrDual (τ μ)) (y ⊗ₜ[R] (Equiv.cast (congrArg ColorModule (τ_involutive μ).symm) x))
|
||||
/-- The unit of the contraction. -/
|
||||
unit : (μ : Color) → ColorModule (τ μ) ⊗[R] ColorModule μ
|
||||
/-- The unit is a right identity. -/
|
||||
unit_rid : ∀ μ (x : ColorModule μ),
|
||||
TensorStructure.contrLeftAux (contrDual μ) (x ⊗ₜ[R] unit μ) = x
|
||||
/-- The metric for a given color. -/
|
||||
metric : (μ : Color) → ColorModule μ ⊗[R] ColorModule μ
|
||||
/-- The metric contracted with its dual is the unit. -/
|
||||
metric_dual : ∀ (μ : Color), (TensorStructure.contrMidAux (contrDual μ)
|
||||
(metric μ ⊗ₜ[R] metric (τ μ))) = TensorProduct.comm _ _ _ (unit μ)
|
||||
|
||||
namespace TensorStructure
|
||||
|
||||
variable (𝓣 : TensorStructure R)
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
{cX cX2 : 𝓣.ColorMap X} {cY : 𝓣.ColorMap Y} {cZ : 𝓣.ColorMap Z}
|
||||
{cW : 𝓣.ColorMap W} {cY' : 𝓣.ColorMap Y'} {μ ν η : 𝓣.Color}
|
||||
|
||||
instance : AddCommMonoid (𝓣.ColorModule μ) := 𝓣.colorModule_addCommMonoid μ
|
||||
|
||||
instance : Module R (𝓣.ColorModule μ) := 𝓣.colorModule_module μ
|
||||
|
||||
/-- The type of tensors given a map from an indexing set `X` to the type of colors,
|
||||
specifying the color of that index. -/
|
||||
def Tensor (c : 𝓣.ColorMap X) : Type := ⨂[R] x, 𝓣.ColorModule (c x)
|
||||
|
||||
instance : AddCommMonoid (𝓣.Tensor cX) :=
|
||||
PiTensorProduct.instAddCommMonoid fun i => 𝓣.ColorModule (cX i)
|
||||
|
||||
instance : Module R (𝓣.Tensor cX) := PiTensorProduct.instModule
|
||||
|
||||
/-!
|
||||
|
||||
## Color
|
||||
|
||||
Recall the `color` of an index describes the type of the index.
|
||||
|
||||
For example, in a real Lorentz tensor the colors are `{up, down}`.
|
||||
|
||||
-/
|
||||
|
||||
/-- Equivalence of `ColorModule` given an equality of colors. -/
|
||||
def colorModuleCast (h : μ = ν) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModule ν where
|
||||
toFun := Equiv.cast (congrArg 𝓣.ColorModule h)
|
||||
invFun := (Equiv.cast (congrArg 𝓣.ColorModule h)).symm
|
||||
map_add' x y := by
|
||||
subst h
|
||||
rfl
|
||||
map_smul' x y := by
|
||||
subst h
|
||||
rfl
|
||||
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
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M]
|
||||
{f g : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] M}
|
||||
(h : ∀ p q, f (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)
|
||||
= g (PiTensorProduct.tprod R p ⊗ₜ[R] PiTensorProduct.tprod R q)) : f = g := by
|
||||
apply TensorProduct.ext'
|
||||
refine fun x ↦
|
||||
PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy y
|
||||
simp [map_add, add_tmul, hx, hy])
|
||||
intro rx fx
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod] at hx hy
|
||||
simp [map_add, tmul_add, hx, hy])
|
||||
intro ry fy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, LinearMapClass.map_smul]
|
||||
apply congrArg
|
||||
simp only [smul_tmul, tmul_smul, LinearMapClass.map_smul]
|
||||
exact congrArg (HSMul.hSMul rx) (h fx fy)
|
||||
|
||||
/-!
|
||||
|
||||
## Mapping isomorphisms
|
||||
|
||||
-/
|
||||
|
||||
/-- An linear equivalence of tensor spaces given a color-preserving equivalence of indexing sets. -/
|
||||
def mapIso {c : 𝓣.ColorMap X} {d : 𝓣.ColorMap Y} (e : X ≃ Y) (h : c.MapIso e d) :
|
||||
𝓣.Tensor c ≃ₗ[R] 𝓣.Tensor d :=
|
||||
(PiTensorProduct.reindex R _ e) ≪≫ₗ
|
||||
(PiTensorProduct.congr (fun y => 𝓣.colorModuleCast (by rw [h]; simp)))
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
lemma mapIso_ext {c : 𝓣.ColorMap X} {d : 𝓣.ColorMap Y} (e e' : X ≃ Y) (h : c.MapIso e d)
|
||||
(h' : c.MapIso e' d) (he : e = e') : 𝓣.mapIso e h = 𝓣.mapIso e' h' := by
|
||||
simp [he]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma mapIso_trans (e : X ≃ Y) (e' : Y ≃ Z) (h : cX.MapIso e cY) (h' : cY.MapIso e' cZ) :
|
||||
(𝓣.mapIso e h ≪≫ₗ 𝓣.mapIso e' h') = 𝓣.mapIso (e.trans e') (h.trans h') := by
|
||||
refine LinearEquiv.toLinearMap_inj.mp ?_
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp only [mapIso, LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe,
|
||||
LinearEquiv.trans_apply, PiTensorProduct.reindex_tprod, Equiv.symm_trans_apply]
|
||||
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
|
||||
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cY x)) e')
|
||||
((PiTensorProduct.congr fun y => 𝓣.colorModuleCast _) _)) =
|
||||
(PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
|
||||
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cX x)) (e.trans e')) _)
|
||||
rw [PiTensorProduct.congr_tprod, PiTensorProduct.reindex_tprod,
|
||||
PiTensorProduct.congr_tprod, PiTensorProduct.reindex_tprod, PiTensorProduct.congr]
|
||||
simp [colorModuleCast]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma mapIso_mapIso (e : X ≃ Y) (e' : Y ≃ Z) (h : cX.MapIso e cY) (h' : cY.MapIso e' cZ)
|
||||
(T : 𝓣.Tensor cX) :
|
||||
(𝓣.mapIso e' h') (𝓣.mapIso e h T) = 𝓣.mapIso (e.trans e') (h.trans h') T := by
|
||||
rw [← LinearEquiv.trans_apply, mapIso_trans]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma mapIso_symm (e : X ≃ Y) (h : cX.MapIso e cY) :
|
||||
(𝓣.mapIso e h).symm = 𝓣.mapIso e.symm (h.symm) := by
|
||||
refine LinearEquiv.toLinearMap_inj.mp ?_
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp only [mapIso, LinearEquiv.trans_symm, LinearMap.compMultilinearMap_apply,
|
||||
LinearEquiv.coe_coe, LinearEquiv.trans_apply, Equiv.symm_symm]
|
||||
change (PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cX x)) e).symm
|
||||
((PiTensorProduct.congr fun y => 𝓣.colorModuleCast _).symm ((PiTensorProduct.tprod R) x)) =
|
||||
(PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
|
||||
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (cY x)) e.symm)
|
||||
((PiTensorProduct.tprod R) x))
|
||||
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod, PiTensorProduct.congr_symm_tprod,
|
||||
LinearEquiv.symm_apply_eq, PiTensorProduct.reindex_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
simp only [colorModuleCast, Equiv.cast_symm, LinearEquiv.coe_symm_mk,
|
||||
Equiv.symm_symm_apply, LinearEquiv.coe_mk]
|
||||
rw [← Equiv.symm_apply_eq]
|
||||
simp only [Equiv.cast_symm, Equiv.cast_apply, cast_cast]
|
||||
apply cast_eq_iff_heq.mpr
|
||||
rw [Equiv.apply_symm_apply]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma mapIso_refl : 𝓣.mapIso (Equiv.refl X) (rfl : cX = cX) = LinearEquiv.refl R _ := by
|
||||
refine LinearEquiv.toLinearMap_inj.mp ?_
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp only [mapIso, Equiv.refl_symm, Equiv.refl_apply, PiTensorProduct.reindex_refl,
|
||||
LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe, LinearEquiv.trans_apply,
|
||||
LinearEquiv.refl_apply, LinearEquiv.refl_toLinearMap, LinearMap.id, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, id_eq]
|
||||
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _) ((PiTensorProduct.tprod R) x) = _
|
||||
rw [PiTensorProduct.congr_tprod]
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma mapIso_tprod {c : 𝓣.ColorMap X} {d : 𝓣.ColorMap Y} (e : X ≃ Y) (h : c.MapIso e d)
|
||||
(f : (i : X) → 𝓣.ColorModule (c i)) : (𝓣.mapIso e h) (PiTensorProduct.tprod R f) =
|
||||
(PiTensorProduct.tprod R (fun i => 𝓣.colorModuleCast (by rw [h]; simp) (f (e.symm i)))) := by
|
||||
simp only [mapIso, LinearEquiv.trans_apply]
|
||||
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
|
||||
((PiTensorProduct.reindex R _ e) ((PiTensorProduct.tprod R) f)) = _
|
||||
rw [PiTensorProduct.reindex_tprod]
|
||||
exact PiTensorProduct.congr_tprod (fun y => 𝓣.colorModuleCast _) fun i => f (e.symm i)
|
||||
|
||||
/-!
|
||||
|
||||
## Pure tensors
|
||||
|
||||
This section is needed since: `PiTensorProduct.tmulEquiv` is not defined for dependent types.
|
||||
Hence we need to construct a version of it here.
|
||||
|
||||
-/
|
||||
|
||||
/-- The type of pure tensors, i.e. of the form `v1 ⊗ v2 ⊗ v3 ⊗ ...`. -/
|
||||
abbrev PureTensor (c : X → 𝓣.Color) := (x : X) → 𝓣.ColorModule (c x)
|
||||
|
||||
/-- A pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` constructed from a pure tensor
|
||||
in `𝓣.PureTensor cX` and a pure tensor in `𝓣.PureTensor cY`. -/
|
||||
def elimPureTensor (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY) : 𝓣.PureTensor (Sum.elim cX cY) :=
|
||||
fun x =>
|
||||
match x with
|
||||
| Sum.inl x => p x
|
||||
| Sum.inr x => q x
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma elimPureTensor_update_right (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY)
|
||||
(y : Y) (r : 𝓣.ColorModule (cY y)) : 𝓣.elimPureTensor p (Function.update q y r) =
|
||||
Function.update (𝓣.elimPureTensor p q) (Sum.inr y) r := by
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x =>
|
||||
change Function.update q y r x = _
|
||||
simp only [Function.update, Sum.inr.injEq, Sum.elim_inr]
|
||||
split_ifs
|
||||
· rename_i h
|
||||
subst h
|
||||
simp_all only
|
||||
· rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma elimPureTensor_update_left (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY)
|
||||
(x : X) (r : 𝓣.ColorModule (cX x)) : 𝓣.elimPureTensor (Function.update p x r) q =
|
||||
Function.update (𝓣.elimPureTensor p q) (Sum.inl x) r := by
|
||||
funext y
|
||||
match y with
|
||||
| Sum.inl y =>
|
||||
change (Function.update p x r) y = _
|
||||
simp only [Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split_ifs
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
| Sum.inr y => rfl
|
||||
|
||||
/-- The projection of a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` onto a pure tensor in
|
||||
`𝓣.PureTensor cX`. -/
|
||||
def inlPureTensor (p : 𝓣.PureTensor (Sum.elim cX cY)) : 𝓣.PureTensor cX := fun x => p (Sum.inl x)
|
||||
|
||||
/-- The projection of a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` onto a pure tensor in
|
||||
`𝓣.PureTensor cY`. -/
|
||||
def inrPureTensor (p : 𝓣.PureTensor (Sum.elim cX cY)) : 𝓣.PureTensor cY := fun y => p (Sum.inr y)
|
||||
|
||||
omit [Fintype X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma inlPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (x : X)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inl x))) :
|
||||
𝓣.inlPureTensor (Function.update f (Sum.inl x) v1) =
|
||||
Function.update (𝓣.inlPureTensor f) x v1 := by
|
||||
funext y
|
||||
simp only [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma inrPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (x : X)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inl x))) :
|
||||
𝓣.inrPureTensor (Function.update f (Sum.inl x) v1) = (𝓣.inrPureTensor f) := by
|
||||
funext x
|
||||
simp [inrPureTensor, Function.update]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma inrPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (y : Y)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inr y))) :
|
||||
𝓣.inrPureTensor (Function.update f (Sum.inr y) v1) =
|
||||
Function.update (𝓣.inrPureTensor f) y v1 := by
|
||||
funext y
|
||||
simp only [inrPureTensor, Function.update, Sum.inr.injEq, Sum.elim_inr]
|
||||
split
|
||||
· rename_i h
|
||||
subst h
|
||||
rfl
|
||||
· rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma inlPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim cX cY)) (y : Y)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim cX cY (Sum.inr y))) :
|
||||
𝓣.inlPureTensor (Function.update f (Sum.inr y) v1) = (𝓣.inlPureTensor f) := by
|
||||
funext x
|
||||
simp [inlPureTensor, Function.update]
|
||||
|
||||
/-- The multilinear map taking pure tensors a `𝓣.PureTensor cX` and a pure tensor in
|
||||
`𝓣.PureTensor cY`, and constructing a tensor in `𝓣.Tensor (Sum.elim cX cY))`. -/
|
||||
def elimPureTensorMulLin : MultilinearMap R (fun i => 𝓣.ColorModule (cX i))
|
||||
(MultilinearMap R (fun x => 𝓣.ColorModule (cY x)) (𝓣.Tensor (Sum.elim cX cY))) where
|
||||
toFun p := {
|
||||
toFun := fun q => PiTensorProduct.tprod R (𝓣.elimPureTensor p q)
|
||||
map_add' := fun m x v1 v2 => by
|
||||
simp only [elimPureTensor_update_right, MultilinearMap.map_add]
|
||||
map_smul' := fun m x r v => by
|
||||
simp only [elimPureTensor_update_right, MultilinearMap.map_smul]}
|
||||
map_add' p x v1 v2 := by
|
||||
apply MultilinearMap.ext
|
||||
intro y
|
||||
simp
|
||||
map_smul' p x r v := by
|
||||
apply MultilinearMap.ext
|
||||
intro y
|
||||
simp
|
||||
|
||||
/-!
|
||||
|
||||
## tensorator
|
||||
|
||||
-/
|
||||
|
||||
/-! TODO: Replace with dependent type version of `MultilinearMap.domCoprod` when in Mathlib. -/
|
||||
/-- The multi-linear map taking a pure tensor in `𝓣.PureTensor (Sum.elim cX cY)` and constructing
|
||||
a vector in `𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY`. -/
|
||||
def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim cX cY x))
|
||||
(𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY) where
|
||||
toFun f := (PiTensorProduct.tprod R (𝓣.inlPureTensor f)) ⊗ₜ
|
||||
(PiTensorProduct.tprod R (𝓣.inrPureTensor f))
|
||||
map_add' f xy v1 v2:= by
|
||||
match xy with
|
||||
| Sum.inl x => simp only [Sum.elim_inl, inlPureTensor_update_left, MultilinearMap.map_add,
|
||||
inrPureTensor_update_left, ← add_tmul]
|
||||
| Sum.inr y => simp only [Sum.elim_inr, inlPureTensor_update_right, inrPureTensor_update_right,
|
||||
MultilinearMap.map_add, ← tmul_add]
|
||||
map_smul' f xy r p := by
|
||||
match xy with
|
||||
| Sum.inl x => simp only [Sum.elim_inl, inlPureTensor_update_left, MultilinearMap.map_smul,
|
||||
inrPureTensor_update_left, smul_tmul, tmul_smul]
|
||||
| Sum.inr y => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
|
||||
|
||||
/-- The linear map combining two tensors into a single tensor
|
||||
via the tensor product i.e. `v1 v2 ↦ v1 ⊗ v2`. -/
|
||||
def tensoratorSymm : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] 𝓣.Tensor (Sum.elim cX cY) := by
|
||||
refine TensorProduct.lift {
|
||||
toFun := fun a ↦
|
||||
PiTensorProduct.lift <|
|
||||
PiTensorProduct.lift (𝓣.elimPureTensorMulLin) a,
|
||||
map_add' := fun a b ↦ by simp
|
||||
map_smul' := fun r a ↦ by simp}
|
||||
|
||||
/-! TODO: Replace with dependent type version of `PiTensorProduct.tmulEquiv` when in Mathlib. -/
|
||||
/-- Splitting a tensor in `𝓣.Tensor (Sum.elim cX cY)` into the tensor product of two tensors. -/
|
||||
def tensorator : 𝓣.Tensor (Sum.elim cX cY) →ₗ[R] 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY :=
|
||||
PiTensorProduct.lift 𝓣.domCoprod
|
||||
|
||||
/-- An equivalence formed by taking the tensor product of tensors. -/
|
||||
def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) :
|
||||
𝓣.Tensor c ⊗[R] 𝓣.Tensor d ≃ₗ[R] 𝓣.Tensor (Sum.elim c d) :=
|
||||
LinearEquiv.ofLinear (𝓣.tensoratorSymm) (𝓣.tensorator)
|
||||
(by
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro p
|
||||
simp only [tensoratorSymm, tensorator, domCoprod, LinearMap.compMultilinearMap_apply,
|
||||
LinearMap.coe_comp, Function.comp_apply, PiTensorProduct.lift.tprod, MultilinearMap.coe_mk,
|
||||
lift.tmul, LinearMap.coe_mk, AddHom.coe_mk]
|
||||
change (PiTensorProduct.lift _) ((PiTensorProduct.tprod R) _) =
|
||||
LinearMap.id ((PiTensorProduct.tprod R) p)
|
||||
rw [PiTensorProduct.lift.tprod]
|
||||
simp only [elimPureTensorMulLin, MultilinearMap.coe_mk, LinearMap.id_coe, id_eq]
|
||||
change (PiTensorProduct.tprod R) _ = _
|
||||
apply congrArg
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl)
|
||||
(by
|
||||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp only [tensorator, tensoratorSymm, LinearMap.coe_comp, Function.comp_apply, lift.tmul,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod, LinearMap.id_coe, id_eq]
|
||||
change (PiTensorProduct.lift 𝓣.domCoprod)
|
||||
((PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q)) =_
|
||||
rw [PiTensorProduct.lift.tprod]
|
||||
simp only [elimPureTensorMulLin, MultilinearMap.coe_mk, PiTensorProduct.lift.tprod]
|
||||
rfl)
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_tmul_tprod (p : 𝓣.PureTensor cX) (q : 𝓣.PureTensor cY) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((PiTensorProduct.tprod R) p ⊗ₜ[R] (PiTensorProduct.tprod R) q) =
|
||||
(PiTensorProduct.tprod R) (𝓣.elimPureTensor p q) := by
|
||||
simp only [tensoratorEquiv, tensoratorSymm, tensorator, domCoprod, LinearEquiv.ofLinear_apply,
|
||||
lift.tmul, LinearMap.coe_mk, AddHom.coe_mk, PiTensorProduct.lift.tprod]
|
||||
exact PiTensorProduct.lift.tprod q
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_symm_tprod (f : 𝓣.PureTensor (Sum.elim cX cY)) :
|
||||
(𝓣.tensoratorEquiv cX cY).symm ((PiTensorProduct.tprod R) f) =
|
||||
(PiTensorProduct.tprod R) (𝓣.inlPureTensor f) ⊗ₜ[R]
|
||||
(PiTensorProduct.tprod R) (𝓣.inrPureTensor f) := by
|
||||
simp only [tensoratorEquiv, tensorator, LinearEquiv.ofLinear_symm_apply]
|
||||
change (PiTensorProduct.lift 𝓣.domCoprod) ((PiTensorProduct.tprod R) f) = _
|
||||
simp [domCoprod]
|
||||
|
||||
omit [Fintype X] [Fintype Y] [Fintype W] [Fintype Z] in
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_mapIso (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h' : cZ.MapIso e' cY) (h'' : cW.MapIso e'' cX) :
|
||||
(TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ (𝓣.tensoratorEquiv cX cY)
|
||||
= (𝓣.tensoratorEquiv cW cZ) ≪≫ₗ (𝓣.mapIso (Equiv.sumCongr e'' e') (h''.sum h')) := by
|
||||
apply LinearEquiv.toLinearMap_inj.mp
|
||||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp only [LinearEquiv.coe_coe, LinearEquiv.trans_apply, congr_tmul, mapIso_tprod,
|
||||
tensoratorEquiv_tmul_tprod]
|
||||
erw [LinearEquiv.trans_apply]
|
||||
simp only [tensoratorEquiv_tmul_tprod, mapIso_tprod, Equiv.sumCongr_symm, Equiv.sumCongr_apply]
|
||||
apply congrArg
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] [Fintype W] [Fintype Z] in
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_mapIso_apply (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h' : cZ.MapIso e' cY) (h'' : cW.MapIso e'' cX)
|
||||
(x : 𝓣.Tensor cW ⊗[R] 𝓣.Tensor cZ) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) x) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (h''.sum h'))
|
||||
((𝓣.tensoratorEquiv cW cZ) x) := by
|
||||
trans ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv cX cY)) x
|
||||
· rfl
|
||||
· rw [tensoratorEquiv_mapIso]
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] [Fintype W] [Fintype Z] in
|
||||
lemma tensoratorEquiv_mapIso_tmul (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h' : cZ.MapIso e' cY) (h'' : cW.MapIso e'' cX)
|
||||
(x : 𝓣.Tensor cW) (y : 𝓣.Tensor cZ) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((𝓣.mapIso e'' h'' x) ⊗ₜ[R] (𝓣.mapIso e' h' y)) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (h''.sum h'))
|
||||
((𝓣.tensoratorEquiv cW cZ) (x ⊗ₜ y)) := by
|
||||
rw [← tensoratorEquiv_mapIso_apply]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## contrDual properties
|
||||
|
||||
-/
|
||||
|
||||
lemma contrDual_cast (h : μ = ν) (x : 𝓣.ColorModule μ) (y : 𝓣.ColorModule (𝓣.τ μ)) :
|
||||
𝓣.contrDual μ (x ⊗ₜ[R] y) = 𝓣.contrDual ν (𝓣.colorModuleCast h x ⊗ₜ[R]
|
||||
𝓣.colorModuleCast (congrArg 𝓣.τ h) y) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-- `𝓣.contrDual (𝓣.τ μ)` in terms of `𝓣.contrDual μ`. -/
|
||||
@[simp]
|
||||
lemma contrDual_symm' (μ : 𝓣.Color) (x : 𝓣.ColorModule (𝓣.τ μ))
|
||||
(y : 𝓣.ColorModule (𝓣.τ (𝓣.τ μ))) : 𝓣.contrDual (𝓣.τ μ) (x ⊗ₜ[R] y) =
|
||||
(𝓣.contrDual μ) ((𝓣.colorModuleCast (𝓣.τ_involutive μ) y) ⊗ₜ[R] x) := by
|
||||
rw [𝓣.contrDual_symm, 𝓣.contrDual_cast (𝓣.τ_involutive μ)]
|
||||
congr
|
||||
exact (LinearEquiv.eq_symm_apply (𝓣.colorModuleCast (congrArg 𝓣.τ (𝓣.τ_involutive μ)))).mp rfl
|
||||
|
||||
lemma contrDual_symm_contrRightAux (h : ν = η) :
|
||||
(𝓣.colorModuleCast h) ∘ₗ contrRightAux (𝓣.contrDual μ) =
|
||||
contrRightAux (𝓣.contrDual (𝓣.τ (𝓣.τ μ))) ∘ₗ
|
||||
(TensorProduct.congr (
|
||||
TensorProduct.congr (𝓣.colorModuleCast h) (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm))
|
||||
(𝓣.colorModuleCast ((𝓣.τ_involutive (𝓣.τ μ)).symm))).toLinearMap := by
|
||||
apply TensorProduct.ext'
|
||||
intro x y
|
||||
refine TensorProduct.induction_on x (by simp) ?_ ?_
|
||||
· intro x z
|
||||
simp only [contrRightAux, LinearEquiv.refl_toLinearMap, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, assoc_tmul, map_tmul, LinearMap.id_coe, id_eq, rid_tmul, map_smul,
|
||||
congr_tmul, contrDual_symm']
|
||||
congr
|
||||
· exact (LinearEquiv.symm_apply_eq (𝓣.colorModuleCast (𝓣.τ_involutive μ))).mp rfl
|
||||
· exact (LinearEquiv.symm_apply_eq (𝓣.colorModuleCast (𝓣.τ_involutive (𝓣.τ μ)))).mp rfl
|
||||
· intro x z h1 h2
|
||||
simp [add_tmul, LinearMap.map_add, h1, h2]
|
||||
|
||||
lemma contrDual_symm_contrRightAux_apply_tmul (h : ν = η)
|
||||
(m : 𝓣.ColorModule ν ⊗[R] 𝓣.ColorModule μ) (x : 𝓣.ColorModule (𝓣.τ μ)) :
|
||||
𝓣.colorModuleCast h (contrRightAux (𝓣.contrDual μ) (m ⊗ₜ[R] x)) =
|
||||
contrRightAux (𝓣.contrDual (𝓣.τ (𝓣.τ μ))) ((TensorProduct.congr
|
||||
(𝓣.colorModuleCast h) (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm) (m)) ⊗ₜ
|
||||
(𝓣.colorModuleCast (𝓣.τ_involutive (𝓣.τ μ)).symm x)) := by
|
||||
trans ((𝓣.colorModuleCast h) ∘ₗ contrRightAux (𝓣.contrDual μ)) (m ⊗ₜ[R] x)
|
||||
· rfl
|
||||
· rw [contrDual_symm_contrRightAux]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Of empty
|
||||
|
||||
-/
|
||||
|
||||
/-- The equivalence between `𝓣.Tensor cX` and `R` when the indexing set `X` is empty. -/
|
||||
def isEmptyEquiv [IsEmpty X] : 𝓣.Tensor cX ≃ₗ[R] R :=
|
||||
PiTensorProduct.isEmptyEquiv X
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma isEmptyEquiv_tprod [IsEmpty X] (f : 𝓣.PureTensor cX) :
|
||||
𝓣.isEmptyEquiv (PiTensorProduct.tprod R f) = 1 := by
|
||||
simp only [isEmptyEquiv]
|
||||
erw [PiTensorProduct.isEmptyEquiv_apply_tprod]
|
||||
/-!
|
||||
|
||||
## Splitting tensors into tensor products
|
||||
|
||||
-/
|
||||
/-! TODO: Delete the content of this section. -/
|
||||
|
||||
/-- The decomposition of a set into a direct sum based on the image of an injection. -/
|
||||
def decompEmbedSet (f : Y ↪ X) :
|
||||
X ≃ {x // x ∈ (Finset.image f Finset.univ)ᶜ} ⊕ Y :=
|
||||
(Equiv.Set.sumCompl (Set.range ⇑f)).symm.trans <|
|
||||
(Equiv.sumComm _ _).trans <|
|
||||
Equiv.sumCongr ((Equiv.subtypeEquivRight (by simp))) <|
|
||||
(Function.Embedding.toEquivRange f).symm
|
||||
|
||||
/-- The restriction of a map from an indexing set to the space to the complement of the image
|
||||
of an embedding. -/
|
||||
def decompEmbedColorLeft (c : X → 𝓣.Color) (f : Y ↪ X) :
|
||||
{x // x ∈ (Finset.image f Finset.univ)ᶜ} → 𝓣.Color :=
|
||||
(c ∘ (decompEmbedSet f).symm) ∘ Sum.inl
|
||||
|
||||
/-- The restriction of a map from an indexing set to the space to the image
|
||||
of an embedding. -/
|
||||
def decompEmbedColorRight (c : X → 𝓣.Color) (f : Y ↪ X) : Y → 𝓣.Color :=
|
||||
(c ∘ (decompEmbedSet f).symm) ∘ Sum.inr
|
||||
|
||||
omit [DecidableEq Y] in
|
||||
lemma decompEmbed_cond (c : X → 𝓣.Color) (f : Y ↪ X) : c =
|
||||
(Sum.elim (𝓣.decompEmbedColorLeft c f) (𝓣.decompEmbedColorRight c f)) ∘ decompEmbedSet f := by
|
||||
simpa [decompEmbedColorLeft, decompEmbedColorRight] using (Equiv.comp_symm_eq _ _ _).mp rfl
|
||||
|
||||
/-- Decomposes a tensor into a tensor product of two tensors
|
||||
one which has indices in the image of the given embedding, and the other has indices in
|
||||
the complement of that image. -/
|
||||
def decompEmbed (f : Y ↪ X) :
|
||||
𝓣.Tensor cX ≃ₗ[R] 𝓣.Tensor (𝓣.decompEmbedColorLeft cX f) ⊗[R] 𝓣.Tensor (cX ∘ f) :=
|
||||
(𝓣.mapIso (decompEmbedSet f) (𝓣.decompEmbed_cond cX f)) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv (𝓣.decompEmbedColorLeft cX f) (𝓣.decompEmbedColorRight cX f)).symm
|
||||
|
||||
end TensorStructure
|
||||
|
||||
end
|
|
@ -1,499 +0,0 @@
|
|||
/-
|
||||
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.Tensors.MulActionTensor
|
||||
/-!
|
||||
|
||||
# Contraction of indices
|
||||
|
||||
We define a number of ways to contract indices of tensors:
|
||||
|
||||
- `contrDualLeft`: Contracts vectors on the left as:
|
||||
`𝓣.ColorModule ν ⊗[R] 𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η →ₗ[R] 𝓣.ColorModule η`
|
||||
|
||||
- `contrDualMid`: Contracts vectors in the middle as:
|
||||
`(𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule ν) ⊗[R] (𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η) →ₗ[R]`
|
||||
`𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule η`
|
||||
|
||||
- `contrAll'`: Contracts all indices of manifestly tensors with manifestly dual colors as:
|
||||
`𝓣.Tensor cX ⊗[R] 𝓣.Tensor (𝓣.τ ∘ cX) →ₗ[R] R`
|
||||
|
||||
- `contrAll`: Contracts all indices of tensors with dual colors as:
|
||||
`𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] R`
|
||||
|
||||
- `contrAllLeft`: Contracts all indices of tensors on the left as:
|
||||
`𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY ⊗[R] 𝓣.Tensor cZ →ₗ[R] 𝓣.Tensor cZ`
|
||||
|
||||
- `contrElim`: Contracting indices of tensors indexed by `Sum.elim _ _` as:
|
||||
`𝓣.Tensor (Sum.elim cW cX) ⊗[R] 𝓣.Tensor (Sum.elim cY cZ) →ₗ[R] 𝓣.Tensor (Sum.elim cW cZ)`
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open TensorProduct
|
||||
open MulActionTensor
|
||||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
namespace TensorColor
|
||||
|
||||
variable {d : ℕ} {X X' Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
variable {d : ℕ} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
[Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P]
|
||||
|
||||
namespace ColorMap
|
||||
|
||||
variable {𝓒 : TensorColor} [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
|
||||
variable (cX : ColorMap 𝓒 X) (cY : ColorMap 𝓒 Y) (cZ : ColorMap 𝓒 Z)
|
||||
|
||||
/-- Given an equivalence `e` of types the condition that the color map `cX` is the dual to `cY`
|
||||
up to this equivalence. -/
|
||||
def ContrAll (e : X ≃ Y) (cX : ColorMap 𝓒 X) (cY : ColorMap 𝓒 Y) : Prop :=
|
||||
cX = 𝓒.τ ∘ cY ∘ e
|
||||
|
||||
namespace ContrAll
|
||||
|
||||
variable {e : X ≃ Y} {e' : Y ≃ Z} {cX : ColorMap 𝓒 X} {cY : ColorMap 𝓒 Y} {cZ : ColorMap 𝓒 Z}
|
||||
variable {cX' : ColorMap 𝓒 X'} {cY' : ColorMap 𝓒 Y'}
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color] in
|
||||
lemma toMapIso (h : cX.ContrAll e cY) : cX.MapIso e cY.dual := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color] in
|
||||
lemma symm (h : cX.ContrAll e cY) : cY.ContrAll e.symm cX := by
|
||||
subst h
|
||||
funext x
|
||||
simp only [Function.comp_apply, Equiv.apply_symm_apply]
|
||||
exact (𝓒.τ_involutive (cY x)).symm
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
lemma trans_mapIso {e : X ≃ Y} {e' : Z ≃ Y}
|
||||
(h : cX.ContrAll e cY) (h' : cZ.MapIso e' cY) : cX.ContrAll (e.trans e'.symm) cZ := by
|
||||
subst h h'
|
||||
funext x
|
||||
simp only [Function.comp_apply, Equiv.coe_trans, Equiv.apply_symm_apply]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color] [Fintype Z]
|
||||
[DecidableEq Z] in
|
||||
lemma mapIso_trans {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : cX.ContrAll e cY) (h' : cZ.MapIso e' cX) : cZ.ContrAll (e'.trans e) cY := by
|
||||
subst h h'
|
||||
rfl
|
||||
|
||||
end ContrAll
|
||||
|
||||
/-- Given an equivalence `(C ⊕ C) ⊕ P ≃ X` the restriction of a color map `cX` on to `P`. -/
|
||||
def contr (e : (C ⊕ C) ⊕ P ≃ X) (cX : ColorMap 𝓒 X) : ColorMap 𝓒 P :=
|
||||
cX ∘ e ∘ Sum.inr
|
||||
|
||||
/-- Given an equivalence `(C ⊕ C) ⊕ P ≃ X` the restriction of a color map `cX` on `X`
|
||||
to the first `C`. -/
|
||||
def contrLeft (e : (C ⊕ C) ⊕ P ≃ X) (cX : ColorMap 𝓒 X) : ColorMap 𝓒 C :=
|
||||
cX ∘ e ∘ Sum.inl ∘ Sum.inl
|
||||
|
||||
/-- Given an equivalence `(C ⊕ C) ⊕ P ≃ X` the restriction of a color map `cX` on `X`
|
||||
to the second `C`. -/
|
||||
def contrRight (e : (C ⊕ C) ⊕ P ≃ X) (cX : ColorMap 𝓒 X) : ColorMap 𝓒 C :=
|
||||
cX ∘ e ∘ Sum.inl ∘ Sum.inr
|
||||
|
||||
/-- Given an equivalence `(C ⊕ C) ⊕ P ≃ X` the condition on `cX` so that we contract
|
||||
the indices of the `C`'s under this equivalence. -/
|
||||
def ContrCond (e : (C ⊕ C) ⊕ P ≃ X) (cX : ColorMap 𝓒 X) : Prop :=
|
||||
cX ∘ e ∘ Sum.inl ∘ Sum.inl = 𝓒.τ ∘ cX ∘ e ∘ Sum.inl ∘ Sum.inr
|
||||
|
||||
namespace ContrCond
|
||||
|
||||
variable {e : (C ⊕ C) ⊕ P ≃ X} {e' : Y ≃ Z} {cX : ColorMap 𝓒 X} {cY : ColorMap 𝓒 Y}
|
||||
{cZ : ColorMap 𝓒 Z}
|
||||
|
||||
variable {cX' : ColorMap 𝓒 X'} {cY' : ColorMap 𝓒 Y'}
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype C]
|
||||
[DecidableEq C] [Fintype P] [DecidableEq P]
|
||||
[Fintype 𝓒.Color] [DecidableEq 𝓒.Color] in
|
||||
lemma to_contrAll (h : cX.ContrCond e) :
|
||||
(cX.contrLeft e).ContrAll (Equiv.refl _) (cX.contrRight e) := h
|
||||
|
||||
end ContrCond
|
||||
|
||||
end ColorMap
|
||||
|
||||
end TensorColor
|
||||
|
||||
namespace TensorStructure
|
||||
|
||||
variable (𝓣 : TensorStructure R)
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
[Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P]
|
||||
{cX cX2 : 𝓣.ColorMap X} {cY : 𝓣.ColorMap Y} {cZ : 𝓣.ColorMap Z}
|
||||
{cW : 𝓣.ColorMap W} {cY' : 𝓣.ColorMap Y'} {μ ν: 𝓣.Color}
|
||||
|
||||
variable {G H : Type} [Group G] [Group H] [MulActionTensor G 𝓣]
|
||||
local infixl:101 " • " => 𝓣.rep
|
||||
/-!
|
||||
|
||||
# Contractions of vectors
|
||||
|
||||
-/
|
||||
|
||||
/-- The contraction of a vector in `𝓣.ColorModule ν` with a vector in
|
||||
`𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η` to form a vector in `𝓣.ColorModule η`. -/
|
||||
def contrDualLeft {ν η : 𝓣.Color} :
|
||||
𝓣.ColorModule ν ⊗[R] 𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η →ₗ[R] 𝓣.ColorModule η :=
|
||||
contrLeftAux (𝓣.contrDual ν)
|
||||
|
||||
/-- The contraction of a vector in `𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule ν` with a vector in
|
||||
`𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η` to form a vector in
|
||||
`𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule η`. -/
|
||||
def contrDualMid {μ ν η : 𝓣.Color} :
|
||||
(𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule ν) ⊗[R] (𝓣.ColorModule (𝓣.τ ν) ⊗[R] 𝓣.ColorModule η) →ₗ[R]
|
||||
𝓣.ColorModule μ ⊗[R] 𝓣.ColorModule η :=
|
||||
contrMidAux (𝓣.contrDual ν)
|
||||
|
||||
/-- A linear map taking tensors mapped with the same index set to the product of paired tensors. -/
|
||||
def pairProd : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cX2 →ₗ[R]
|
||||
⨂[R] x, 𝓣.ColorModule (cX x) ⊗[R] 𝓣.ColorModule (cX2 x) :=
|
||||
TensorProduct.lift (
|
||||
PiTensorProduct.map₂ (fun x =>
|
||||
TensorProduct.mk R (𝓣.ColorModule (cX x)) (𝓣.ColorModule (cX2 x))))
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
lemma pairProd_tmul_tprod_tprod (fx : (i : X) → 𝓣.ColorModule (cX i))
|
||||
(fx2 : (i : X) → 𝓣.ColorModule (cX2 i)) :
|
||||
𝓣.pairProd (PiTensorProduct.tprod R fx ⊗ₜ[R] PiTensorProduct.tprod R fx2) =
|
||||
PiTensorProduct.tprod R (fun x => fx x ⊗ₜ[R] fx2 x) := by
|
||||
simp only [pairProd, lift.tmul]
|
||||
erw [PiTensorProduct.map₂_tprod_tprod]
|
||||
rfl
|
||||
|
||||
omit [DecidableEq X] [DecidableEq Y]
|
||||
lemma mkPiAlgebra_equiv (e : X ≃ Y) :
|
||||
(PiTensorProduct.lift (MultilinearMap.mkPiAlgebra R X R)) =
|
||||
(PiTensorProduct.lift (MultilinearMap.mkPiAlgebra R Y R)) ∘ₗ
|
||||
(PiTensorProduct.reindex R _ e).toLinearMap := by
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp only [LinearMap.compMultilinearMap_apply, PiTensorProduct.lift.tprod,
|
||||
MultilinearMap.mkPiAlgebra_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
PiTensorProduct.reindex_tprod, Equiv.prod_comp]
|
||||
|
||||
/-- Given a tensor in `𝓣.Tensor cX` and a tensor in `𝓣.Tensor (𝓣.τ ∘ cX)`, the element of
|
||||
`R` formed by contracting all of their indices. -/
|
||||
def contrAll' : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor (𝓣.τ ∘ cX) →ₗ[R] R :=
|
||||
(PiTensorProduct.lift (MultilinearMap.mkPiAlgebra R X R)) ∘ₗ
|
||||
(PiTensorProduct.map (fun x => 𝓣.contrDual (cX x))) ∘ₗ
|
||||
(𝓣.pairProd)
|
||||
|
||||
lemma contrAll'_tmul_tprod_tprod (fx : (i : X) → 𝓣.ColorModule (cX i))
|
||||
(fy : (i : X) → 𝓣.ColorModule (𝓣.τ (cX i))) :
|
||||
𝓣.contrAll' (PiTensorProduct.tprod R fx ⊗ₜ[R] PiTensorProduct.tprod R fy) =
|
||||
(PiTensorProduct.lift (MultilinearMap.mkPiAlgebra R X R))
|
||||
(PiTensorProduct.tprod R (fun x => 𝓣.contrDual (cX x) (fx x ⊗ₜ[R] fy x))) := by
|
||||
simp only [contrAll', Function.comp_apply, LinearMap.coe_comp, PiTensorProduct.lift.tprod,
|
||||
MultilinearMap.mkPiAlgebra_apply]
|
||||
erw [pairProd_tmul_tprod_tprod]
|
||||
simp only [Function.comp_apply, PiTensorProduct.map_tprod, PiTensorProduct.lift.tprod,
|
||||
MultilinearMap.mkPiAlgebra_apply]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll'_isEmpty_tmul [IsEmpty X] (x : 𝓣.Tensor cX) (y : 𝓣.Tensor (𝓣.τ ∘ cX)) :
|
||||
𝓣.contrAll' (x ⊗ₜ y) = 𝓣.isEmptyEquiv x * 𝓣.isEmptyEquiv y := by
|
||||
refine PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy
|
||||
simp [map_add, add_tmul, add_mul, hx, hy])
|
||||
intro rx fx
|
||||
refine PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, map_smul, isEmptyEquiv_tprod,
|
||||
smul_eq_mul, mul_one] at hx hy
|
||||
simp [map_add, tmul_add, mul_add, hx, hy])
|
||||
intro ry fy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, Function.comp_apply, tmul_smul, smul_tmul,
|
||||
map_smul, smul_eq_mul, isEmptyEquiv_tprod, mul_one]
|
||||
ring_nf
|
||||
rw [mul_assoc, mul_assoc]
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
simp only [contrAll', LinearMap.coe_comp, Function.comp_apply]
|
||||
erw [pairProd_tmul_tprod_tprod]
|
||||
simp only [Function.comp_apply, PiTensorProduct.map_tprod, PiTensorProduct.lift.tprod,
|
||||
MultilinearMap.mkPiAlgebra_apply, Finset.univ_eq_empty, Finset.prod_empty]
|
||||
erw [isEmptyEquiv_tprod]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll'_mapIso (e : X ≃ Y) (h : cX.MapIso e cY) :
|
||||
𝓣.contrAll' ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e h) (LinearEquiv.refl R _)).toLinearMap =
|
||||
𝓣.contrAll' ∘ₗ (TensorProduct.congr (LinearEquiv.refl R _)
|
||||
(𝓣.mapIso e.symm h.symm.dual)).toLinearMap := by
|
||||
apply TensorProduct.ext'
|
||||
refine fun x ↦
|
||||
PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy y
|
||||
simp only [add_tmul, map_add, hx, LinearMap.coe_comp, Function.comp_apply, hy])
|
||||
intro rx fx
|
||||
refine fun y ↦
|
||||
PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, congr_tmul, map_smul, mapIso_tprod, LinearEquiv.refl_apply] at hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, Function.comp_apply, tmul_add, map_add,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, congr_tmul, map_smul, mapIso_tprod,
|
||||
LinearEquiv.refl_apply, hx, hy])
|
||||
intro ry fy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, Function.comp_apply, tmul_smul,
|
||||
LinearMapClass.map_smul, LinearMap.coe_comp, LinearEquiv.coe_coe, congr_tmul, mapIso_tprod,
|
||||
LinearEquiv.refl_apply, smul_eq_mul, smul_tmul]
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
erw [contrAll'_tmul_tprod_tprod]
|
||||
erw [TensorProduct.congr_tmul]
|
||||
simp only [PiTensorProduct.lift.tprod, LinearEquiv.refl_apply]
|
||||
erw [mapIso_tprod]
|
||||
erw [contrAll'_tmul_tprod_tprod]
|
||||
rw [mkPiAlgebra_equiv e]
|
||||
simp only [Equiv.symm_symm_apply, LinearMap.coe_comp,
|
||||
LinearEquiv.coe_coe, Function.comp_apply, PiTensorProduct.reindex_tprod,
|
||||
PiTensorProduct.lift.tprod]
|
||||
apply congrArg
|
||||
funext y
|
||||
rw [𝓣.contrDual_cast (congrFun h.symm y)]
|
||||
apply congrArg
|
||||
congr 1
|
||||
· exact (LinearEquiv.eq_symm_apply
|
||||
(𝓣.colorModuleCast (congrFun (TensorColor.ColorMap.MapIso.symm h) y))).mp rfl
|
||||
· symm
|
||||
apply cast_eq_iff_heq.mpr
|
||||
simp only [Function.comp_apply, colorModuleCast, Equiv.cast_symm, LinearEquiv.coe_mk,
|
||||
Equiv.cast_apply]
|
||||
rw [Equiv.apply_symm_apply]
|
||||
exact HEq.symm (cast_heq _ _)
|
||||
|
||||
@[simp]
|
||||
lemma contrAll'_mapIso_tmul (e : X ≃ Y) (h : cX.MapIso e cY) (x : 𝓣.Tensor cX)
|
||||
(y : 𝓣.Tensor (𝓣.τ ∘ cY)) : 𝓣.contrAll' ((𝓣.mapIso e h) x ⊗ₜ[R] y) =
|
||||
𝓣.contrAll' (x ⊗ₜ[R] (𝓣.mapIso e.symm h.symm.dual y)) := by
|
||||
change (𝓣.contrAll' ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e h) (LinearEquiv.refl R _)).toLinearMap) (x ⊗ₜ[R] y) = _
|
||||
rw [contrAll'_mapIso]
|
||||
rfl
|
||||
|
||||
/-- The contraction of all the indices of two tensors with dual colors. -/
|
||||
def contrAll (e : X ≃ Y) (h : cX.ContrAll e cY) : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY →ₗ[R] R :=
|
||||
𝓣.contrAll' ∘ₗ (TensorProduct.congr (LinearEquiv.refl _ _)
|
||||
(𝓣.mapIso e.symm h.symm.toMapIso)).toLinearMap
|
||||
|
||||
omit [Fintype Y]
|
||||
|
||||
lemma contrAll_tmul (e : X ≃ Y) (h : cX.ContrAll e cY) (x : 𝓣.Tensor cX) (y : 𝓣.Tensor cY) :
|
||||
𝓣.contrAll e h (x ⊗ₜ[R] y) = 𝓣.contrAll' (x ⊗ₜ[R] ((𝓣.mapIso e.symm h.symm.toMapIso) y)) := by
|
||||
rw [contrAll]
|
||||
rfl
|
||||
|
||||
omit [Fintype Z] [DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_right_tmul (e : X ≃ Y) (e' : Z ≃ Y)
|
||||
(h : c.ContrAll e cY) (h' : cZ.MapIso e' cY) (x : 𝓣.Tensor c) (z : 𝓣.Tensor cZ) :
|
||||
𝓣.contrAll e h (x ⊗ₜ[R] 𝓣.mapIso e' h' z) =
|
||||
𝓣.contrAll (e.trans e'.symm) (h.trans_mapIso h') (x ⊗ₜ[R] z) := by
|
||||
simp only [contrAll_tmul, mapIso_mapIso]
|
||||
rfl
|
||||
|
||||
omit [Fintype Z] [DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma contrAll_comp_mapIso_right (e : X ≃ Y) (e' : Z ≃ Y)
|
||||
(h : c.ContrAll e cY) (h' : cZ.MapIso e' cY) : 𝓣.contrAll e h ∘ₗ
|
||||
(TensorProduct.congr (LinearEquiv.refl R (𝓣.Tensor c)) (𝓣.mapIso e' h')).toLinearMap
|
||||
= 𝓣.contrAll (e.trans e'.symm) (h.trans_mapIso h') := by
|
||||
apply TensorProduct.ext'
|
||||
intro x y
|
||||
exact 𝓣.contrAll_mapIso_right_tmul e e' h h' x y
|
||||
|
||||
omit [DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_left_tmul {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : cX.ContrAll e cY) (h' : cZ.MapIso e' cX) (x : 𝓣.Tensor cZ) (y : 𝓣.Tensor cY) :
|
||||
𝓣.contrAll e h (𝓣.mapIso e' h' x ⊗ₜ[R] y) =
|
||||
𝓣.contrAll (e'.trans e) (h.mapIso_trans h') (x ⊗ₜ[R] y) := by
|
||||
simp only [contrAll_tmul, contrAll'_mapIso_tmul, mapIso_mapIso]
|
||||
rfl
|
||||
|
||||
omit [DecidableEq Z] in
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_left {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : cX.ContrAll e cY) (h' : cZ.MapIso e' cX) :
|
||||
𝓣.contrAll e h ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e' h') (LinearEquiv.refl R (𝓣.Tensor cY))).toLinearMap
|
||||
= 𝓣.contrAll (e'.trans e) (h.mapIso_trans h') := by
|
||||
apply TensorProduct.ext'
|
||||
intro x y
|
||||
exact 𝓣.contrAll_mapIso_left_tmul h h' x y
|
||||
|
||||
/-- The linear map from `𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY ⊗[R] 𝓣.Tensor cZ` to
|
||||
`𝓣.Tensor cZ` obtained by contracting all indices in `𝓣.Tensor cX` and `𝓣.Tensor cY`,
|
||||
given a proof that this is possible. -/
|
||||
def contrAllLeft (e : X ≃ Y) (h : cX.ContrAll e cY) :
|
||||
𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY ⊗[R] 𝓣.Tensor cZ →ₗ[R] 𝓣.Tensor cZ :=
|
||||
(TensorProduct.lid R _).toLinearMap ∘ₗ
|
||||
TensorProduct.map (𝓣.contrAll e h) (LinearEquiv.refl R (𝓣.Tensor cZ)).toLinearMap
|
||||
∘ₗ (TensorProduct.assoc R _ _ _).symm.toLinearMap
|
||||
|
||||
/-- The linear map from `(𝓣.Tensor cW ⊗[R] 𝓣.Tensor cX) ⊗[R] (𝓣.Tensor cY ⊗[R] 𝓣.Tensor cZ)`
|
||||
to `𝓣.Tensor cW ⊗[R] 𝓣.Tensor cZ` obtained by contracting all indices of the tensors
|
||||
in the middle. -/
|
||||
def contrAllMid (e : X ≃ Y) (h : cX.ContrAll e cY) :
|
||||
(𝓣.Tensor cW ⊗[R] 𝓣.Tensor cX) ⊗[R] (𝓣.Tensor cY ⊗[R] 𝓣.Tensor cZ) →ₗ[R]
|
||||
𝓣.Tensor cW ⊗[R] 𝓣.Tensor cZ :=
|
||||
(TensorProduct.map (LinearEquiv.refl R _).toLinearMap (𝓣.contrAllLeft e h)) ∘ₗ
|
||||
(TensorProduct.assoc R _ _ _).toLinearMap
|
||||
|
||||
/-- The linear map from `𝓣.Tensor (Sum.elim cW cX) ⊗[R] 𝓣.Tensor (Sum.elim cY cZ)`
|
||||
to `𝓣.Tensor (Sum.elim cW cZ)` formed by contracting the indices specified by
|
||||
`cX` and `cY`, which are assumed to be dual. -/
|
||||
def contrElim (e : X ≃ Y) (h : cX.ContrAll e cY) :
|
||||
𝓣.Tensor (Sum.elim cW cX) ⊗[R] 𝓣.Tensor (Sum.elim cY cZ) →ₗ[R] 𝓣.Tensor (Sum.elim cW cZ) :=
|
||||
(𝓣.tensoratorEquiv cW cZ).toLinearMap ∘ₗ 𝓣.contrAllMid e h ∘ₗ
|
||||
(TensorProduct.congr (𝓣.tensoratorEquiv cW cX).symm
|
||||
(𝓣.tensoratorEquiv cY cZ).symm).toLinearMap
|
||||
|
||||
/-!
|
||||
|
||||
## Group acting on contraction
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma contrAll_rep (e : X ≃ Y) (h : cX.ContrAll e cY) (g : G) :
|
||||
𝓣.contrAll e h ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g)) = 𝓣.contrAll e h := by
|
||||
apply TensorProduct.ext'
|
||||
refine fun x ↦ PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx hy y
|
||||
simp [map_add, add_tmul, hx, hy])
|
||||
intro rx fx
|
||||
refine fun y ↦ PiTensorProduct.induction_on' y ?_ (by
|
||||
intro a b hx hy
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMap.coe_comp, Function.comp_apply,
|
||||
map_tmul, map_smul, rep_tprod] at hx hy
|
||||
simp [map_add, tmul_add, hx, hy])
|
||||
intro ry fy
|
||||
simp only [contrAll_tmul, PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, smul_tmul,
|
||||
LinearMapClass.map_smul, LinearMap.coe_comp, Function.comp_apply, map_tmul, rep_tprod,
|
||||
smul_eq_mul]
|
||||
apply congrArg
|
||||
apply congrArg
|
||||
simp only [contrAll', mapIso_tprod, Equiv.symm_symm_apply, colorModuleCast_equivariant_apply,
|
||||
LinearMap.coe_comp, Function.comp_apply]
|
||||
apply congrArg
|
||||
erw [pairProd_tmul_tprod_tprod, pairProd_tmul_tprod_tprod, PiTensorProduct.map_tprod,
|
||||
PiTensorProduct.map_tprod]
|
||||
apply congrArg
|
||||
funext x
|
||||
nth_rewrite 2 [← contrDual_inv (cX x) g]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrAll_rep_apply {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = 𝓣.τ ∘ d ∘ e)
|
||||
(g : G) (x : 𝓣.Tensor c ⊗ 𝓣.Tensor d) :
|
||||
𝓣.contrAll e h (TensorProduct.map (𝓣.rep g) (𝓣.rep g) x) = 𝓣.contrAll e h x := by
|
||||
change (𝓣.contrAll e h ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g))) x = _
|
||||
rw [contrAll_rep]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll_rep_tmul {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = 𝓣.τ ∘ d ∘ e)
|
||||
(g : G) (x : 𝓣.Tensor c) (y : 𝓣.Tensor d) :
|
||||
𝓣.contrAll e h ((g • x) ⊗ₜ[R] (g • y)) = 𝓣.contrAll e h (x ⊗ₜ[R] y) := by
|
||||
nth_rewrite 2 [← @contrAll_rep_apply R _ 𝓣 _ _ _ G]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Contraction based on specification
|
||||
|
||||
-/
|
||||
|
||||
omit [Fintype X] [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] in
|
||||
lemma contr_cond (e : (C ⊕ C) ⊕ P ≃ X) :
|
||||
cX.MapIso e.symm (Sum.elim (Sum.elim (cX.contrLeft e) (cX.contrRight e)) (cX.contr e)) := by
|
||||
rw [TensorColor.ColorMap.MapIso, Equiv.eq_comp_symm]
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl (Sum.inl x) => rfl
|
||||
| Sum.inl (Sum.inr x) => rfl
|
||||
| Sum.inr x => rfl
|
||||
|
||||
/-- Contraction of indices based on an equivalence `(C ⊕ C) ⊕ P ≃ X`. The indices
|
||||
in `C` are contracted pair-wise, whilst the indices in `P` are preserved. -/
|
||||
def contr (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e) :
|
||||
𝓣.Tensor cX →ₗ[R] 𝓣.Tensor (cX.contr e) :=
|
||||
(TensorProduct.lid R _).toLinearMap ∘ₗ
|
||||
(TensorProduct.map (𝓣.contrAll (Equiv.refl C) h.to_contrAll) LinearMap.id) ∘ₗ
|
||||
(TensorProduct.congr (𝓣.tensoratorEquiv _ _).symm (LinearEquiv.refl R _)).toLinearMap ∘ₗ
|
||||
(𝓣.tensoratorEquiv _ _).symm.toLinearMap ∘ₗ
|
||||
(𝓣.mapIso e.symm (𝓣.contr_cond e)).toLinearMap
|
||||
|
||||
open PiTensorProduct in
|
||||
|
||||
omit [Fintype X] [Fintype P] in
|
||||
lemma contr_tprod (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e) (f : (i : X) → 𝓣.ColorModule (cX i)) :
|
||||
𝓣.contr e h (tprod R f) = (𝓣.contrAll (Equiv.refl C) h.to_contrAll
|
||||
(tprod R (fun i => f (e (Sum.inl (Sum.inl i)))) ⊗ₜ[R]
|
||||
tprod R (fun i => f (e (Sum.inl (Sum.inr i)))))) •
|
||||
tprod R (fun (p : P) => f (e (Sum.inr p))) := by
|
||||
simp only [contr, LinearEquiv.comp_coe, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, LinearEquiv.trans_apply, mapIso_tprod, Equiv.symm_symm_apply,
|
||||
tensoratorEquiv_symm_tprod, congr_tmul, LinearEquiv.refl_apply, map_tmul, LinearMap.id_coe,
|
||||
id_eq, lid_tmul]
|
||||
rfl
|
||||
|
||||
open PiTensorProduct in
|
||||
omit [Fintype X] [Fintype P] in
|
||||
@[simp]
|
||||
lemma contr_tprod_isEmpty [IsEmpty C] (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e)
|
||||
(f : (i : X) → 𝓣.ColorModule (cX i)) :
|
||||
𝓣.contr e h (tprod R f) = tprod R (fun (p : P) => f (e (Sum.inr p))) := by
|
||||
rw [contr_tprod]
|
||||
rw [contrAll_tmul, contrAll'_isEmpty_tmul]
|
||||
simp only [isEmptyEquiv_tprod, Equiv.refl_symm, mapIso_tprod, Equiv.refl_apply, one_mul]
|
||||
erw [isEmptyEquiv_tprod]
|
||||
exact MulAction.one_smul ((tprod R) fun p => f (e (Sum.inr p)))
|
||||
|
||||
omit [Fintype X] [Fintype P] in
|
||||
/-- The contraction of indices via `contr` is equivariant. -/
|
||||
@[simp]
|
||||
lemma contr_equivariant (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e)
|
||||
(g : G) (x : 𝓣.Tensor cX) : 𝓣.contr e h (g • x) = g • 𝓣.contr e h x := by
|
||||
simp only [TensorColor.ColorMap.contr, contr, TensorProduct.congr, LinearEquiv.refl_toLinearMap,
|
||||
LinearEquiv.symm_symm, LinearEquiv.refl_symm, LinearEquiv.ofLinear_toLinearMap,
|
||||
LinearEquiv.comp_coe, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
LinearEquiv.trans_apply, rep_mapIso_apply, rep_tensoratorEquiv_symm_apply]
|
||||
rw [← LinearMap.comp_apply (TensorProduct.map _ _), ← TensorProduct.map_comp]
|
||||
rw [← LinearMap.comp_apply (TensorProduct.map _ _), ← TensorProduct.map_comp]
|
||||
rw [LinearMap.comp_assoc, rep_tensoratorEquiv_symm, ← LinearMap.comp_assoc]
|
||||
simp only [contrAll_rep, LinearMap.comp_id, LinearMap.id_comp]
|
||||
have h1 {M N A B : Type} [AddCommMonoid M] [AddCommMonoid N]
|
||||
[AddCommMonoid A] [AddCommMonoid B] [Module R M] [Module R N] [Module R A] [Module R B]
|
||||
(f : M →ₗ[R] N) (g : A →ₗ[R] B) : TensorProduct.map f g
|
||||
= TensorProduct.map (LinearMap.id) g ∘ₗ TensorProduct.map f (LinearMap.id) :=
|
||||
ext rfl
|
||||
rw [h1]
|
||||
simp only [LinearMap.coe_comp, Function.comp_apply, rep_lid_apply]
|
||||
rw [← LinearMap.comp_apply (TensorProduct.map _ _), ← TensorProduct.map_comp]
|
||||
rfl
|
||||
|
||||
end TensorStructure
|
|
@ -1,111 +0,0 @@
|
|||
/-
|
||||
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 Mathlib.LinearAlgebra.StdBasis
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.LinearAlgebra.DirectSum.Finsupp
|
||||
import Mathlib.LinearAlgebra.Finsupp
|
||||
/-!
|
||||
|
||||
# Einstein notation for real tensors
|
||||
|
||||
Einstein notation is a specific example of index notation, with only one color.
|
||||
|
||||
In this file we define Einstein notation for generic ring `R`.
|
||||
|
||||
-/
|
||||
|
||||
open TensorProduct
|
||||
|
||||
/-- Einstein tensors have only one color, corresponding to a `down` index. . -/
|
||||
def einsteinTensorColor : TensorColor where
|
||||
Color := Unit
|
||||
τ a := a
|
||||
τ_involutive μ := by rfl
|
||||
|
||||
instance : Fintype einsteinTensorColor.Color := Unit.fintype
|
||||
|
||||
instance : DecidableEq einsteinTensorColor.Color := instDecidableEqPUnit
|
||||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
/-- The `TensorStructure` associated with `n`-dimensional tensors. -/
|
||||
noncomputable def einsteinTensor (R : Type) [CommSemiring R] (n : ℕ) : TensorStructure R where
|
||||
toTensorColor := einsteinTensorColor
|
||||
ColorModule _ := Fin n → R
|
||||
colorModule_addCommMonoid _ := Pi.addCommMonoid
|
||||
colorModule_module _ := Pi.Function.module (Fin n) R R
|
||||
contrDual _ := TensorProduct.lift (Fintype.linearCombination R R)
|
||||
contrDual_symm a x y := by
|
||||
simp only [lift.tmul, Fintype.linearCombination_apply, smul_eq_mul, mul_comm, Equiv.cast_refl,
|
||||
Equiv.refl_apply]
|
||||
unit a := ∑ i, Pi.basisFun R (Fin n) i ⊗ₜ[R] Pi.basisFun R (Fin n) i
|
||||
unit_rid a x:= by
|
||||
simp only [Pi.basisFun_apply]
|
||||
rw [tmul_sum, map_sum]
|
||||
trans ∑ i, x i • Pi.basisFun R (Fin n) i
|
||||
· refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
simp only [TensorStructure.contrLeftAux, LinearEquiv.refl_toLinearMap, LinearMap.coe_comp,
|
||||
LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul, map_tmul, lift.tmul,
|
||||
Fintype.linearCombination_apply, Pi.single_apply, smul_eq_mul, ite_mul, one_mul, zero_mul,
|
||||
Finset.sum_ite_eq', Finset.mem_univ, ↓reduceIte, LinearMap.id_coe, id_eq, lid_tmul,
|
||||
Pi.basisFun_apply]
|
||||
· funext a
|
||||
simp only [Pi.basisFun_apply, Finset.sum_apply, Pi.smul_apply, Pi.single_apply, smul_eq_mul,
|
||||
mul_ite, mul_one, mul_zero, Finset.sum_ite_eq, Finset.mem_univ, ↓reduceIte]
|
||||
metric a := ∑ i, Pi.basisFun R (Fin n) i ⊗ₜ[R] Pi.basisFun R (Fin n) i
|
||||
metric_dual a := by
|
||||
simp only [Pi.basisFun_apply, map_sum, comm_tmul]
|
||||
rw [tmul_sum, map_sum]
|
||||
refine Finset.sum_congr rfl (fun i _ => ?_)
|
||||
rw [sum_tmul, map_sum, Fintype.sum_eq_single i]
|
||||
· simp only [TensorStructure.contrMidAux, LinearEquiv.refl_toLinearMap,
|
||||
TensorStructure.contrLeftAux, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
assoc_tmul, map_tmul, LinearMap.id_coe, id_eq, assoc_symm_tmul, lift.tmul,
|
||||
Fintype.linearCombination_apply, Pi.single_apply, smul_eq_mul, mul_ite, mul_one, mul_zero,
|
||||
Finset.sum_ite_eq', Finset.mem_univ, ↓reduceIte, lid_tmul, one_smul]
|
||||
· intro x hi
|
||||
simp only [TensorStructure.contrMidAux, LinearEquiv.refl_toLinearMap,
|
||||
TensorStructure.contrLeftAux, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
assoc_tmul, map_tmul, LinearMap.id_coe, id_eq, assoc_symm_tmul, lift.tmul,
|
||||
Fintype.linearCombination_apply, Pi.single_apply, smul_eq_mul, mul_ite, mul_one, mul_zero,
|
||||
Finset.sum_ite_eq', Finset.mem_univ, ↓reduceIte, lid_tmul, ite_smul, one_smul, zero_smul]
|
||||
rw [if_neg hi]
|
||||
exact tmul_zero (Fin n → R) (Pi.single x 1)
|
||||
|
||||
namespace einsteinTensor
|
||||
|
||||
open TensorStructure
|
||||
|
||||
noncomputable section
|
||||
|
||||
instance : OfNat einsteinTensorColor.Color 0 := ⟨PUnit.unit⟩
|
||||
instance : OfNat (einsteinTensor R n).Color 0 := ⟨PUnit.unit⟩
|
||||
|
||||
@[simp]
|
||||
lemma ofNat_inst_eq : @einsteinTensor.instOfNatColorOfNatNat R _ n =
|
||||
einsteinTensor.instOfNatColorEinsteinTensorColorOfNatNat := rfl
|
||||
|
||||
/-- A vector from an Einstein tensor with one index. -/
|
||||
def toVec : (einsteinTensor R n).Tensor ![Unit.unit] ≃ₗ[R] Fin n → R :=
|
||||
PiTensorProduct.subsingletonEquiv 0
|
||||
|
||||
/-- A matrix from an Einstein tensor with two indices. -/
|
||||
def toMatrix : (einsteinTensor R n).Tensor ![Unit.unit, Unit.unit] ≃ₗ[R] Matrix (Fin n) (Fin n) R :=
|
||||
((einsteinTensor R n).mapIso ((Fin.castOrderIso
|
||||
(by rfl : (Nat.succ 0).succ = Nat.succ 0 + Nat.succ 0)).toEquiv.trans
|
||||
finSumFinEquiv.symm) (by funext x; fin_cases x <;> rfl)).trans <|
|
||||
((einsteinTensor R n).tensoratorEquiv ![0] ![0]).symm.trans <|
|
||||
(TensorProduct.congr ((PiTensorProduct.subsingletonEquiv 0))
|
||||
((PiTensorProduct.subsingletonEquiv 0))).trans <|
|
||||
(TensorProduct.congr (Finsupp.linearEquivFunOnFinite R R (Fin n)).symm
|
||||
(Finsupp.linearEquivFunOnFinite R R (Fin n)).symm).trans <|
|
||||
(finsuppTensorFinsupp' R (Fin n) (Fin n)).trans <|
|
||||
(Finsupp.linearEquivFunOnFinite R R (Fin n × Fin n)).trans <|
|
||||
(LinearEquiv.curry R R (Fin n) (Fin n))
|
||||
|
||||
end
|
||||
|
||||
end einsteinTensor
|
|
@ -1,126 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.TensorIndex
|
||||
import HepLean.Tensors.IndexNotation.IndexString
|
||||
import HepLean.Tensors.EinsteinNotation.Basic
|
||||
/-!
|
||||
|
||||
# Index notation for Einstein tensors
|
||||
|
||||
-/
|
||||
|
||||
instance : IndexNotation einsteinTensorColor.Color where
|
||||
charList := {'ᵢ'}
|
||||
notaEquiv :=
|
||||
⟨fun _ => ⟨'ᵢ', Finset.mem_singleton.mpr rfl⟩,
|
||||
fun _ => Unit.unit,
|
||||
fun _ => rfl,
|
||||
by
|
||||
intro c
|
||||
have hc2 := c.2
|
||||
simp only [↓Char.isValue, Finset.mem_singleton] at hc2
|
||||
exact SetCoe.ext (id (Eq.symm hc2))⟩
|
||||
|
||||
namespace einsteinTensor
|
||||
|
||||
open einsteinTensorColor
|
||||
open IndexNotation IndexString
|
||||
open TensorStructure TensorIndex
|
||||
|
||||
variable {R : Type} [CommSemiring R] {n m : ℕ}
|
||||
|
||||
instance : IndexNotation (einsteinTensor R n).Color := instIndexNotationColorEinsteinTensorColor
|
||||
instance : DecidableEq (einsteinTensor R n).Color := instDecidableEqColorEinsteinTensorColor
|
||||
|
||||
@[simp]
|
||||
lemma indexNotation_eq_color : @einsteinTensor.instIndexNotationColor R _ n =
|
||||
instIndexNotationColorEinsteinTensorColor := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma decidableEq_eq_color : @einsteinTensor.instDecidableEqColor R _ n =
|
||||
instDecidableEqColorEinsteinTensorColor := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma einsteinTensor_color : (einsteinTensor R n).Color = einsteinTensorColor.Color := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma toTensorColor_eq : (einsteinTensor R n).toTensorColor = einsteinTensorColor := by
|
||||
rfl
|
||||
|
||||
/-- The construction of a tensor index from a tensor and a string satisfying conditions
|
||||
which can be automatically checked. This is a modified version of
|
||||
`TensorStructure.TensorIndex.mkDualMap` specific to real Lorentz tensors. -/
|
||||
noncomputable def fromIndexStringColor {R : Type} [CommSemiring R]
|
||||
{cn : Fin n → einsteinTensorColor.Color}
|
||||
(T : (einsteinTensor R m).Tensor cn) (s : String)
|
||||
(hs : listCharIsIndexString einsteinTensorColor.Color s.toList = true)
|
||||
(hn : n = (toIndexList' s hs).length)
|
||||
(hD : (toIndexList' s hs).OnlyUniqueDuals)
|
||||
(hC : IndexList.ColorCond.bool (toIndexList' s hs))
|
||||
(hd : TensorColor.ColorMap.DualMap.boolFin'
|
||||
(toIndexList' s hs).colorMap (cn ∘ Fin.cast hn.symm)) :
|
||||
(einsteinTensor R m).TensorIndex :=
|
||||
TensorStructure.TensorIndex.mkDualMap T ⟨(toIndexList' s hs), hD,
|
||||
IndexList.ColorCond.iff_bool.mpr hC⟩ hn
|
||||
(TensorColor.ColorMap.DualMap.boolFin'_DualMap hd)
|
||||
|
||||
@[simp]
|
||||
lemma fromIndexStringColor_indexList {R : Type} [CommSemiring R]
|
||||
{cn : Fin n → einsteinTensorColor.Color}
|
||||
(T : (einsteinTensor R m).Tensor cn) (s : String)
|
||||
(hs : listCharIsIndexString einsteinTensorColor.Color s.toList = true)
|
||||
(hn : n = (toIndexList' s hs).length)
|
||||
(hD : (toIndexList' s hs).OnlyUniqueDuals)
|
||||
(hC : IndexList.ColorCond.bool (toIndexList' s hs))
|
||||
(hd : TensorColor.ColorMap.DualMap.boolFin'
|
||||
(toIndexList' s hs).colorMap (cn ∘ Fin.cast hn.symm)) :
|
||||
(fromIndexStringColor T s hs hn hD hC hd).toIndexList = toIndexList' s hs := by
|
||||
rfl
|
||||
|
||||
/-- A tactic used to prove `boolFin` for real Lornetz tensors. -/
|
||||
macro "dualMapTactic" : tactic =>
|
||||
`(tactic| {
|
||||
simp only [toTensorColor_eq]
|
||||
decide })
|
||||
|
||||
/-- Notation for the construction of a tensor index from a tensor and a string.
|
||||
Conditions are checked automatically. -/
|
||||
notation:20 T "|" S:21 => fromIndexStringColor T S
|
||||
(by decide)
|
||||
(by decide) (by rfl)
|
||||
(by decide)
|
||||
(by dualMapTactic)
|
||||
|
||||
/-- A tactics used to prove `colorPropBool` for real Lorentz tensors. -/
|
||||
macro "prodTactic" : tactic =>
|
||||
`(tactic| {
|
||||
apply (ColorIndexList.AppendCond.iff_bool _ _).mpr
|
||||
change @ColorIndexList.AppendCond.bool einsteinTensorColor
|
||||
instDecidableEqColorEinsteinTensorColor _ _
|
||||
simp only [prod_toIndexList, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
|
||||
toTensorColor_eq, decidableEq_eq_color]
|
||||
decide})
|
||||
|
||||
lemma mem_fin_list (n : ℕ) (i : Fin n) : i ∈ Fin.list n := by
|
||||
have h1' : (Fin.list n)[i] = i := Fin.getElem_list _ _
|
||||
exact h1' ▸ List.getElem_mem _ _ _
|
||||
|
||||
instance (n : ℕ) (i : Fin n) : Decidable (i ∈ Fin.list n) :=
|
||||
isTrue (mem_fin_list n i)
|
||||
|
||||
/-- The product of Real Lorentz tensors. Conditions on indices are checked automatically. -/
|
||||
notation:10 T "⊗ᵀ" S:11 => TensorIndex.prod T S (by prodTactic)
|
||||
/-- An example showing the allowed constructions. -/
|
||||
example (T : (einsteinTensor R n).Tensor ![Unit.unit, Unit.unit]) : True := by
|
||||
let _ := T|"ᵢ₂ᵢ₃"
|
||||
let _ := T|"ᵢ₁ᵢ₂" ⊗ᵀ T|"ᵢ₂ᵢ₁"
|
||||
let _ := T|"ᵢ₁ᵢ₂" ⊗ᵀ T|"ᵢ₂ᵢ₁" ⊗ᵀ T|"ᵢ₃ᵢ₄"
|
||||
exact trivial
|
||||
|
||||
end einsteinTensor
|
|
@ -1,41 +0,0 @@
|
|||
/-
|
||||
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.Tensors.EinsteinNotation.IndexNotation
|
||||
/-!
|
||||
|
||||
# Lemmas regarding Einstein tensors
|
||||
|
||||
This file is currently a stub.
|
||||
|
||||
-/
|
||||
namespace einsteinTensor
|
||||
|
||||
open einsteinTensorColor
|
||||
open IndexNotation IndexString
|
||||
open TensorStructure TensorIndex
|
||||
|
||||
variable {R : Type} [CommSemiring R] {n m : ℕ}
|
||||
/-
|
||||
lemma swap_eq_transpose (T : (einsteinTensor R n).Tensor ![Unit.unit, Unit.unit]) :
|
||||
(T|"ᵢ₁ᵢ₂") ≈ ((toMatrix.symm (toMatrix T).transpose)|"ᵢ₂ᵢ₁") := by
|
||||
refine Rel.of_withDual_empty ?_ ?_ ?_ ?_
|
||||
· apply And.intro
|
||||
simp only [toTensorColor_eq, indexNotation_eq_color, ColorIndexList.contr, fromIndexStringColor,
|
||||
mkDualMap, decidableEq_eq_color]
|
||||
decide
|
||||
simp only [toTensorColor_eq, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
|
||||
ColorIndexList.colorMap', decidableEq_eq_color, ColorIndexList.contr]
|
||||
decide
|
||||
· simp only [toTensorColor_eq, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
|
||||
decidableEq_eq_color]
|
||||
decide
|
||||
· simp only [toTensorColor_eq, indexNotation_eq_color, fromIndexStringColor, mkDualMap,
|
||||
ColorIndexList.colorMap', decidableEq_eq_color]
|
||||
decide
|
||||
simp [fromIndexStringColor, mkDualMap]-/
|
||||
|
||||
end einsteinTensor
|
|
@ -1,21 +0,0 @@
|
|||
/-
|
||||
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 Mathlib.LinearAlgebra.StdBasis
|
||||
import HepLean.Tensors.Basic
|
||||
import Mathlib.LinearAlgebra.DirectSum.Finsupp
|
||||
import Mathlib.LinearAlgebra.Finsupp
|
||||
/-!
|
||||
|
||||
# Dualizing indices of Einstein tensors
|
||||
|
||||
Dualizing indices (the more general notion of rising and lowering indices) does
|
||||
nothing for Einstein tensors, since they only have one color of index.
|
||||
|
||||
The aim of this file is show this result.
|
||||
|
||||
This file is currently a stub.
|
||||
-/
|
||||
/-! TODO: Prove dualizing indices for Einstein tensors does nothing. -/
|
|
@ -1,209 +0,0 @@
|
|||
/-
|
||||
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 Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Index notation for a type
|
||||
|
||||
In this file we will define an index of a Lorentz tensor as a
|
||||
string satisfying certain properties.
|
||||
|
||||
For example, the string `ᵘ¹²` is an index of a real Lorentz tensors.
|
||||
The first character `ᵘ` specifies the color of the index, and the subsequent characters
|
||||
`¹²` specify the id of the index.
|
||||
|
||||
Strings of indices e.g. `ᵘ¹²ᵤ₄₃`` are defined elsewhere.
|
||||
|
||||
## Note
|
||||
|
||||
Index notation is currently being refactored. Much of the content here will likely be replaced
|
||||
or removed. See the HepLean.Tensors.Tree.Basic for the current approach of the index notation.
|
||||
|
||||
-/
|
||||
|
||||
open Lean
|
||||
open String
|
||||
|
||||
/-- The class defining index notation on a type `X`.
|
||||
Normally `X` will be taken as the type of colors of a `TensorStructure`. -/
|
||||
class IndexNotation (X : Type) where
|
||||
/-- The list of characters describing the index notation e.g.
|
||||
`{'ᵘ', 'ᵤ'}` for real tensors. -/
|
||||
charList : Finset Char
|
||||
/-- An equivalence between `X` (colors of indices) and `charList`.
|
||||
This takes every color of index to its notation character. -/
|
||||
notaEquiv : X ≃ charList
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
variable (X : Type) [IndexNotation X]
|
||||
variable [Fintype X] [DecidableEq X]
|
||||
|
||||
/-!
|
||||
|
||||
## Lists of characters forming an index
|
||||
|
||||
Here we define `listCharIndex` and properties thereof.
|
||||
|
||||
-/
|
||||
|
||||
/-- The map taking a color to its notation character. -/
|
||||
def nota {X : Type} [IndexNotation X] (x : X) : Char :=
|
||||
(IndexNotation.notaEquiv).toFun x
|
||||
|
||||
/-- A character is a `notation character` if it is in `charList`. -/
|
||||
def isNotationChar (c : Char) : Bool :=
|
||||
if c ∈ charList X then true else false
|
||||
|
||||
/-- A character is a numeric superscript if it is e.g. `⁰`, `¹`, etc. -/
|
||||
def isNumericSupscript (c : Char) : Bool :=
|
||||
c = '¹' ∨ c = '²' ∨ c = '³' ∨ c = '⁴' ∨ c = '⁵' ∨ c = '⁶' ∨ c = '⁷' ∨ c = '⁸' ∨ c = '⁹' ∨ c = '⁰'
|
||||
|
||||
/-- Given a character `f` which is a notation character, this is true if `c`
|
||||
is a subscript when `f` is a subscript or `c` is a superscript when `f` is a
|
||||
superscript. -/
|
||||
def IsIndexId (f : Char) (c : Char) : Bool :=
|
||||
(isSubScriptAlnum f ∧ isNumericSubscript c) ∨
|
||||
(¬ isSubScriptAlnum f ∧ isNumericSupscript c)
|
||||
|
||||
/-- The proposition for a list of characters to be the tail of an index
|
||||
e.g. `['¹', '⁷', ...]` -/
|
||||
def listCharIndexTail (f : Char) (l : List Char) : Prop :=
|
||||
l ≠ [] ∧ List.all l (fun c => IsIndexId f c)
|
||||
|
||||
instance : Decidable (listCharIndexTail f l) := instDecidableAnd
|
||||
|
||||
/-- The proposition for a list of characters to be the characters of an index
|
||||
e.g. `['ᵘ', '¹', '⁷', ...]` -/
|
||||
def listCharIndex (l : List Char) : Prop :=
|
||||
if h : l = [] then True
|
||||
else
|
||||
let sfst := l.head h
|
||||
if ¬ isNotationChar X sfst then False
|
||||
else
|
||||
listCharIndexTail sfst l.tail
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
/-- An auxillary rewrite lemma to prove that `listCharIndex` is decidable. -/
|
||||
lemma listCharIndex_iff (l : List Char) : listCharIndex X l
|
||||
↔ (if h : l = [] then True else
|
||||
let sfst := l.head h
|
||||
if ¬ isNotationChar X sfst then False
|
||||
else listCharIndexTail sfst l.tail) := by rfl
|
||||
|
||||
instance : Decidable (listCharIndex X l) :=
|
||||
@decidable_of_decidable_of_iff _ _
|
||||
(@instDecidableDite _ _ _ _ _ <|
|
||||
fun _ => @instDecidableDite _ _ _ _ _ <|
|
||||
fun _ => instDecidableListCharIndexTail)
|
||||
(listCharIndex_iff X l).symm
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of an index
|
||||
|
||||
-/
|
||||
|
||||
/-- An index for `X` is an pair of an element of `X` (the color of the index) and a natural
|
||||
number (the id of the index). -/
|
||||
def Index : Type := X × ℕ
|
||||
|
||||
instance : DecidableEq (Index X) := instDecidableEqProd
|
||||
|
||||
namespace Index
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- The color associated to an index. -/
|
||||
def toColor (I : Index X) : X := I.1
|
||||
|
||||
/-- The natural number representating the id of an index. -/
|
||||
def id (I : Index X) : ℕ := I.2
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma eq_iff_color_eq_and_id_eq (I J : Index X) : I = J ↔ I.toColor = J.toColor ∧ I.id = J.id := by
|
||||
refine Iff.intro (fun h => Prod.mk.inj_iff.mp h) (fun h => ?_)
|
||||
· cases I
|
||||
cases J
|
||||
simp only [toColor, id] at h
|
||||
simp [h]
|
||||
|
||||
end Index
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of an index and its properties
|
||||
|
||||
-/
|
||||
|
||||
/-- An index rep is a non-empty string satisfying the condtion `listCharIndex`,
|
||||
e.g. `ᵘ¹²` or `ᵤ₄₃` etc. -/
|
||||
def IndexRep : Type := {s : String // listCharIndex X s.toList ∧ s.toList ≠ []}
|
||||
|
||||
instance : DecidableEq (IndexRep X) := Subtype.instDecidableEq
|
||||
|
||||
namespace IndexRep
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- Creats an index from a non-empty list of characters satisfying `listCharIndex`. -/
|
||||
def ofCharList (l : List Char) (h : listCharIndex X l ∧ l ≠ []) : IndexRep X := ⟨l.asString, h⟩
|
||||
|
||||
instance : ToString (IndexRep X) := ⟨fun i => i.val⟩
|
||||
|
||||
/-- Gets the first character in an index e.g. `ᵘ` as an element of `charList X`. -/
|
||||
def head (s : IndexRep X) : charList X :=
|
||||
⟨s.val.toList.head (s.prop.2), by
|
||||
have h := s.prop.1
|
||||
have h2 := s.prop.2
|
||||
simp only [listCharIndex, toList, Bool.not_eq_true, ne_eq, if_false_left,
|
||||
Bool.not_eq_false] at h
|
||||
simp_all only [toList, ne_eq, Bool.not_eq_true, ↓reduceDIte]
|
||||
simpa [isNotationChar] using h.1⟩
|
||||
|
||||
/-- The color associated to an index. -/
|
||||
def toColor (s : IndexRep X) : X := (IndexNotation.notaEquiv).invFun s.head
|
||||
|
||||
/-- A map from super and subscript numerical characters to the natural numbers,
|
||||
returning `0` on all other characters. -/
|
||||
def charToNat (c : Char) : Nat :=
|
||||
match c with
|
||||
| '₀' => 0
|
||||
| '₁' => 1
|
||||
| '₂' => 2
|
||||
| '₃' => 3
|
||||
| '₄' => 4
|
||||
| '₅' => 5
|
||||
| '₆' => 6
|
||||
| '₇' => 7
|
||||
| '₈' => 8
|
||||
| '₉' => 9
|
||||
| '⁰' => 0
|
||||
| '¹' => 1
|
||||
| '²' => 2
|
||||
| '³' => 3
|
||||
| '⁴' => 4
|
||||
| '⁵' => 5
|
||||
| '⁶' => 6
|
||||
| '⁷' => 7
|
||||
| '⁸' => 8
|
||||
| '⁹' => 9
|
||||
| _ => 0
|
||||
|
||||
/-- The numerical characters associated with an index. -/
|
||||
def tail (s : IndexRep X) : List Char := s.val.toList.tail
|
||||
|
||||
/-- The natural numbers assocaited with an index. -/
|
||||
def tailNat (s : IndexRep X) : List ℕ := s.tail.map charToNat
|
||||
|
||||
/-- The id of an index, as a natural number. -/
|
||||
def id (s : IndexRep X) : ℕ := s.tailNat.foldl (fun a b => 10 * a + b) 0
|
||||
|
||||
/-- The index associated with a `IndexRep`. -/
|
||||
def toIndex (s : IndexRep X) : Index X := (s.toColor, s.id)
|
||||
|
||||
end IndexRep
|
||||
end IndexNotation
|
|
@ -1,175 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.ColorIndexList.Basic
|
||||
import HepLean.Tensors.IndexNotation.ColorIndexList.Contraction
|
||||
import HepLean.Tensors.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Appending two ColorIndexLists
|
||||
|
||||
We define conditional appending of `ColorIndexList`'s.
|
||||
|
||||
It is conditional on `AppendCond` being true, which we define.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
namespace ColorIndexList
|
||||
|
||||
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
(l l2 : ColorIndexList 𝓒)
|
||||
|
||||
open IndexList TensorColor
|
||||
|
||||
/-!
|
||||
|
||||
## Append
|
||||
|
||||
-/
|
||||
|
||||
/-- The condition on the `ColorIndexList`s `l` and `l2` so that on appending they form
|
||||
a `ColorIndexList`.
|
||||
|
||||
Note: `AppendCond` does not form an equivalence relation as it is not reflexive or
|
||||
transitive. -/
|
||||
def AppendCond : Prop :=
|
||||
(l.toIndexList ++ l2.toIndexList).OnlyUniqueDuals ∧ ColorCond (l.toIndexList ++ l2.toIndexList)
|
||||
|
||||
/-- Given two `ColorIndexList`s satisfying `AppendCond`. The correponding combined
|
||||
`ColorIndexList`. -/
|
||||
def append (h : AppendCond l l2) : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList ++ l2.toIndexList
|
||||
unique_duals := h.1
|
||||
dual_color := h.2
|
||||
|
||||
/-- The join of two `ColorIndexList` satisfying the condition `AppendCond` that they
|
||||
can be appended to form a `ColorIndexList`. -/
|
||||
scoped[IndexNotation.ColorIndexList] notation l " ++["h"] " l2 => append l l2 h
|
||||
|
||||
omit [DecidableEq 𝓒.Color] in
|
||||
@[simp]
|
||||
lemma append_toIndexList (h : AppendCond l l2) :
|
||||
(l ++[h] l2).toIndexList = l.toIndexList ++ l2.toIndexList := rfl
|
||||
|
||||
namespace AppendCond
|
||||
|
||||
variable {l l2 l3 : ColorIndexList 𝓒}
|
||||
|
||||
@[symm]
|
||||
lemma symm (h : AppendCond l l2) : AppendCond l2 l := by
|
||||
apply And.intro _ (h.2.symm h.1)
|
||||
rw [OnlyUniqueDuals.symm]
|
||||
exact h.1
|
||||
|
||||
lemma inr (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
||||
AppendCond l2 l3 := by
|
||||
apply And.intro
|
||||
· have h1 := h'.1
|
||||
rw [append_toIndexList, append_assoc] at h1
|
||||
exact OnlyUniqueDuals.inr h1
|
||||
· have h1 := h'.2
|
||||
simp only [append_toIndexList] at h1
|
||||
rw [append_assoc] at h1
|
||||
refine ColorCond.inr ?_ h1
|
||||
rw [← append_assoc]
|
||||
exact h'.1
|
||||
|
||||
lemma assoc (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
||||
AppendCond l (l2 ++[h.inr h'] l3) := by
|
||||
apply And.intro
|
||||
· simp only [append_toIndexList]
|
||||
rw [← append_assoc]
|
||||
simpa using h'.1
|
||||
· simp only [append_toIndexList]
|
||||
rw [← append_assoc]
|
||||
exact h'.2
|
||||
|
||||
lemma swap (h : AppendCond l l2) (h' : AppendCond (l ++[h] l2) l3) :
|
||||
AppendCond (l2 ++[h.symm] l) l3:= by
|
||||
apply And.intro
|
||||
· simp only [append_toIndexList]
|
||||
apply OnlyUniqueDuals.swap
|
||||
simpa using h'.1
|
||||
· exact ColorCond.swap h'.1 h'.2
|
||||
|
||||
/-- If `AppendCond l l2` then `AppendCond l.contr l2`. Note that the inverse
|
||||
is generally not true. -/
|
||||
lemma contr_left (h : AppendCond l l2) : AppendCond l.contr l2 :=
|
||||
And.intro (OnlyUniqueDuals.contrIndexList_left h.1) (ColorCond.contrIndexList_left h.1 h.2)
|
||||
|
||||
lemma contr_right (h : AppendCond l l2) : AppendCond l l2.contr := (contr_left h.symm).symm
|
||||
|
||||
lemma contr (h : AppendCond l l2) : AppendCond l.contr l2.contr := contr_left (contr_right h)
|
||||
|
||||
/-- A boolean which is true for two index lists `l` and `l2` if on appending
|
||||
they can form a `ColorIndexList`. -/
|
||||
def bool (l l2 : IndexList 𝓒.Color) : Bool :=
|
||||
if ¬ (l ++ l2).withUniqueDual = (l ++ l2).withDual then
|
||||
false
|
||||
else
|
||||
ColorCond.bool (l ++ l2)
|
||||
|
||||
omit [IndexNotation 𝓒.Color] in
|
||||
lemma bool_iff (l l2 : IndexList 𝓒.Color) :
|
||||
bool l l2 ↔ (l ++ l2).withUniqueDual = (l ++ l2).withDual ∧ ColorCond.bool (l ++ l2) := by
|
||||
simp [bool]
|
||||
|
||||
lemma iff_bool (l l2 : ColorIndexList 𝓒) : AppendCond l l2 ↔ bool l.toIndexList l2.toIndexList := by
|
||||
rw [AppendCond]
|
||||
simp only [bool, ite_not, Bool.if_false_right, Bool.and_eq_true, decide_eq_true_eq]
|
||||
rw [ColorCond.iff_bool]
|
||||
exact Eq.to_iff rfl
|
||||
|
||||
lemma countId_contr_fst_eq_zero_mem_snd (h : AppendCond l l2) {I : Index 𝓒.Color}
|
||||
(hI : I ∈ l2.val) : l.contr.countId I = 0 ↔ l.countId I = 0 := by
|
||||
rw [countId_contr_eq_zero_iff]
|
||||
have h1 := countId_mem l2.toIndexList I hI
|
||||
have h1I := h.1
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at h1I
|
||||
have h1I' := h1I I
|
||||
simp only [countId_append] at h1I'
|
||||
omega
|
||||
|
||||
lemma countId_contr_snd_eq_zero_mem_fst (h : AppendCond l l2) {I : Index 𝓒.Color}
|
||||
(hI : I ∈ l.val) : l2.contr.countId I = 0 ↔ l2.countId I = 0 := by
|
||||
exact countId_contr_fst_eq_zero_mem_snd h.symm hI
|
||||
|
||||
end AppendCond
|
||||
|
||||
lemma append_contr_left (h : AppendCond l l2) :
|
||||
(l.contr ++[h.contr_left] l2).contr = (l ++[h] l2).contr := by
|
||||
apply ext
|
||||
simp only [contr, append_toIndexList]
|
||||
rw [contrIndexList_append_eq_filter, contrIndexList_append_eq_filter,
|
||||
contrIndexList_contrIndexList]
|
||||
apply congrArg
|
||||
apply List.filter_congr
|
||||
intro I hI
|
||||
simp only [decide_eq_decide]
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at hI
|
||||
exact AppendCond.countId_contr_fst_eq_zero_mem_snd h hI.1
|
||||
|
||||
lemma append_contr_right (h : AppendCond l l2) :
|
||||
(l ++[h.contr_right] l2.contr).contr = (l ++[h] l2).contr := by
|
||||
apply ext
|
||||
simp only [contr, append_toIndexList]
|
||||
rw [contrIndexList_append_eq_filter, contrIndexList_append_eq_filter,
|
||||
contrIndexList_contrIndexList]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
apply List.filter_congr
|
||||
intro I hI
|
||||
simp only [decide_eq_decide]
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at hI
|
||||
exact AppendCond.countId_contr_snd_eq_zero_mem_fst h hI.1
|
||||
|
||||
lemma append_contr (h : AppendCond l l2) :
|
||||
(l.contr ++[h.contr] l2.contr).contr = (l ++[h] l2).contr := by
|
||||
rw [append_contr_left, append_contr_right]
|
||||
|
||||
end ColorIndexList
|
||||
end IndexNotation
|
|
@ -1,103 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Color
|
||||
import HepLean.Tensors.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Color index lists
|
||||
|
||||
A color index list is defined as a list of indices with two constraints. The first is that
|
||||
if an index has a dual, that dual is unique. The second is that if an index has a dual, the
|
||||
color of that dual is dual to the color of the index.
|
||||
|
||||
The indices of a tensor are required to be of this type.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
variable (𝓒 : TensorColor)
|
||||
variable [IndexNotation 𝓒.Color] [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
|
||||
/-- A list of indices with the additional constraint that if a index has a dual,
|
||||
that dual is unique, and the dual of an index has dual color to that index.
|
||||
|
||||
This is the permissible type of indices which can be used for a tensor. For example,
|
||||
the index list `['ᵘ¹', 'ᵤ₁']` can be extended to a `ColorIndexList` but the index list
|
||||
`['ᵘ¹', 'ᵤ₁', 'ᵤ₁']` cannot. -/
|
||||
structure ColorIndexList (𝓒 : TensorColor) [IndexNotation 𝓒.Color] extends IndexList 𝓒.Color where
|
||||
/-- The condition that for index with a dual, that dual is the unique other index with
|
||||
an identical `id`. -/
|
||||
unique_duals : toIndexList.OnlyUniqueDuals
|
||||
/-- The condition that for an index with a dual, that dual has dual color to the index. -/
|
||||
dual_color : IndexList.ColorCond toIndexList
|
||||
|
||||
namespace ColorIndexList
|
||||
|
||||
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color]
|
||||
|
||||
variable (l l2 : ColorIndexList 𝓒)
|
||||
open IndexList TensorColor
|
||||
|
||||
instance : Coe (ColorIndexList 𝓒) (IndexList 𝓒.Color) := ⟨fun l => l.toIndexList⟩
|
||||
|
||||
/-- The colorMap of a `ColorIndexList` as a `𝓒.ColorMap`.
|
||||
This is to be compared with `colorMap` which is a map `Fin l.length → 𝓒.Color`. -/
|
||||
def colorMap' : 𝓒.ColorMap (Fin l.length) :=
|
||||
l.colorMap
|
||||
|
||||
@[ext]
|
||||
lemma ext {l l' : ColorIndexList 𝓒} (h : l.val = l'.val) : l = l' := by
|
||||
cases l
|
||||
cases l'
|
||||
simp_all
|
||||
apply IndexList.ext
|
||||
exact h
|
||||
|
||||
lemma ext' {l l' : ColorIndexList 𝓒} (h : l.toIndexList = l'.toIndexList) : l = l' := by
|
||||
cases l
|
||||
cases l'
|
||||
simp_all
|
||||
|
||||
/-! TODO: `orderEmbOfFin_univ` should be replaced with a mathlib lemma if possible. -/
|
||||
lemma orderEmbOfFin_univ (n m : ℕ) (h : n = m) :
|
||||
Finset.orderEmbOfFin (Finset.univ : Finset (Fin n)) (by simp [h]: Finset.univ.card = m) =
|
||||
(Fin.castOrderIso h.symm).toOrderEmbedding := by
|
||||
symm
|
||||
have h1 : (Fin.castOrderIso h.symm).toFun =
|
||||
(Finset.orderEmbOfFin (Finset.univ : Finset (Fin n))
|
||||
(by simp [h]: Finset.univ.card = m)).toFun := by
|
||||
apply Finset.orderEmbOfFin_unique
|
||||
intro x
|
||||
exact Finset.mem_univ ((Fin.castOrderIso (Eq.symm h)).toFun x)
|
||||
exact fun ⦃a b⦄ a => a
|
||||
exact Eq.symm (OrderEmbedding.range_inj.mp (congrArg Set.range (id (Eq.symm h1))))
|
||||
|
||||
/-!
|
||||
|
||||
## Cons for `ColorIndexList`
|
||||
|
||||
-/
|
||||
|
||||
/-- The `ColorIndexList` whose underlying list of indices is empty. -/
|
||||
def empty : ColorIndexList 𝓒 where
|
||||
val := []
|
||||
unique_duals := rfl
|
||||
dual_color := rfl
|
||||
|
||||
/-!
|
||||
|
||||
## CountId for `ColorIndexList`
|
||||
|
||||
-/
|
||||
|
||||
lemma countId_le_two [DecidableEq 𝓒.Color] (l : ColorIndexList 𝓒) (I : Index 𝓒.Color) :
|
||||
l.countId I ≤ 2 :=
|
||||
(OnlyUniqueDuals.iff_countId_leq_two').mp l.unique_duals I
|
||||
|
||||
end ColorIndexList
|
||||
end IndexNotation
|
|
@ -1,370 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.ColorIndexList.Contraction
|
||||
import HepLean.Tensors.IndexNotation.IndexList.Subperm
|
||||
import HepLean.Tensors.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
## Permutation
|
||||
|
||||
Test whether two `ColorIndexList`s are permutations of each other.
|
||||
To prevent choice problems, this has to be done after contraction.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace ColorIndexList
|
||||
|
||||
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color] [DecidableEq 𝓒.Color]
|
||||
|
||||
variable (l l' : ColorIndexList 𝓒)
|
||||
open IndexList TensorColor
|
||||
|
||||
/--
|
||||
Two `ColorIndexList`s are said to be related by contracted permutations, `ContrPerm`,
|
||||
if and only if:
|
||||
|
||||
1) Their contractions are the same length.
|
||||
2) Every index in the contracted list of one has a unqiue dual in the contracted
|
||||
list of the other and that dual has a the same color.
|
||||
-/
|
||||
def ContrPerm : Prop :=
|
||||
l.contr.length = l'.contr.length ∧
|
||||
IndexList.Subperm l.contr l'.contr.toIndexList ∧
|
||||
l'.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l'.contr)
|
||||
= l.contr.colorMap' ∘ Subtype.val
|
||||
|
||||
namespace ContrPerm
|
||||
|
||||
variable {l l' l2 l3 : ColorIndexList 𝓒}
|
||||
|
||||
lemma getDualInOtherEquiv_eq (h : l.ContrPerm l2) (i : Fin l.contr.length) :
|
||||
l2.contr.val.get (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩).1 = l.contr.val.get i := by
|
||||
rw [Index.eq_iff_color_eq_and_id_eq]
|
||||
apply And.intro
|
||||
· trans (l2.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l2.contr)) ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩
|
||||
· rfl
|
||||
· simp only [h.2.2]
|
||||
rfl
|
||||
· trans l2.contr.idMap (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [h.2.1]
|
||||
exact Finset.mem_univ i⟩).1
|
||||
· rfl
|
||||
· simp only [getDualInOtherEquiv, Equiv.coe_fn_mk, getDualInOther?_id, List.get_eq_getElem]
|
||||
rfl
|
||||
|
||||
lemma mem_snd_of_mem_snd (h : l.ContrPerm l2) {I : Index 𝓒.Color} (hI : I ∈ l.contr.val) :
|
||||
I ∈ l2.contr.val := by
|
||||
have h1 : l.contr.val.indexOf I < l.contr.val.length := List.indexOf_lt_length.mpr hI
|
||||
have h2 : I = l.contr.val.get ⟨l.contr.val.indexOf I, h1⟩ := (List.indexOf_get h1).symm
|
||||
rw [h2]
|
||||
rw [← getDualInOtherEquiv_eq h ⟨l.contr.val.indexOf I, h1⟩]
|
||||
simp only [List.get_eq_getElem]
|
||||
exact List.getElem_mem _ _ _
|
||||
|
||||
lemma countSelf_eq_one_snd_of_countSelf_eq_one_fst (h : l.ContrPerm l2) {I : Index 𝓒.Color}
|
||||
(h1 : l.contr.countSelf I = 1) : l2.contr.countSelf I = 1 := by
|
||||
refine countSelf_eq_one_of_countId_eq_one l2.contr I ?_ (mem_snd_of_mem_snd h ?_)
|
||||
· have h2 := h.2.1
|
||||
rw [Subperm.iff_countId] at h2
|
||||
refine (h2 I).2 ?_
|
||||
have h1 := h2 I
|
||||
have h2' := countSelf_le_countId l.contr.toIndexList I
|
||||
omega
|
||||
· rw [← countSelf_neq_zero, h1]
|
||||
exact Nat.one_ne_zero
|
||||
|
||||
lemma getDualInOtherEquiv_eq_of_countSelf
|
||||
(hn : IndexList.Subperm l.contr l2.contr.toIndexList) (i : Fin l.contr.length)
|
||||
(h1 : l2.contr.countId (l.contr.val.get i) = 1)
|
||||
(h2 : l2.contr.countSelf (l.contr.val.get i) = 1) :
|
||||
l2.contr.val.get (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [hn]
|
||||
exact Finset.mem_univ i⟩).1 = l.contr.val.get i := by
|
||||
have h1' : (l.contr.val.get i) ∈ l2.contr.val := by
|
||||
rw [← countSelf_neq_zero, h2]
|
||||
exact Nat.one_ne_zero
|
||||
rw [← List.mem_singleton, ← filter_id_of_countId_eq_one_mem l2.contr.toIndexList h1' h1]
|
||||
simp only [List.get_eq_getElem, List.mem_filter, decide_eq_true_eq]
|
||||
apply And.intro (List.getElem_mem _ _ _)
|
||||
simp only [getDualInOtherEquiv, Equiv.coe_fn_mk]
|
||||
change _ = l2.contr.idMap (l.contr.getDualInOtherEquiv l2.contr ⟨i, by
|
||||
rw [hn]
|
||||
exact Finset.mem_univ i⟩).1
|
||||
simp only [getDualInOtherEquiv, Equiv.coe_fn_mk, getDualInOther?_id]
|
||||
rfl
|
||||
|
||||
lemma colorMap_eq_of_countSelf (hn : IndexList.Subperm l.contr l2.contr.toIndexList)
|
||||
(h2 : ∀ i, l.contr.countSelf (l.contr.val.get i) = 1
|
||||
→ l2.contr.countSelf (l.contr.val.get i) = 1) :
|
||||
l2.contr.colorMap' ∘ Subtype.val ∘ (l.contr.getDualInOtherEquiv l2.contr)
|
||||
= l.contr.colorMap' ∘ Subtype.val := by
|
||||
funext a
|
||||
simp only [colorMap', Function.comp_apply, colorMap, List.get_eq_getElem]
|
||||
change _ = (l.contr.val.get a.1).toColor
|
||||
rw [← getDualInOtherEquiv_eq_of_countSelf hn a.1]
|
||||
· rfl
|
||||
· rw [@Subperm.iff_countId_fin] at hn
|
||||
exact (hn a.1).2
|
||||
· refine h2 a.1
|
||||
(countSelf_eq_one_of_countId_eq_one l.contr.toIndexList (l.contr.val.get a.1) ?h1 ?hme)
|
||||
· rw [Subperm.iff_countId_fin] at hn
|
||||
exact (hn a.1).1
|
||||
· simp only [List.get_eq_getElem]
|
||||
exact List.getElem_mem l.contr.val (↑↑a) a.1.isLt
|
||||
|
||||
lemma iff_count_fin : l.ContrPerm l2 ↔
|
||||
l.contr.length = l2.contr.length ∧ IndexList.Subperm l.contr l2.contr.toIndexList ∧
|
||||
∀ i, l.contr.countSelf (l.contr.val.get i) = 1 →
|
||||
l2.contr.countSelf (l.contr.val.get i) = 1 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· refine And.intro h.1 (And.intro h.2.1 ?_)
|
||||
exact fun i a => countSelf_eq_one_snd_of_countSelf_eq_one_fst h a
|
||||
· refine And.intro h.1 (And.intro h.2.1 ?_)
|
||||
apply colorMap_eq_of_countSelf h.2.1 h.2.2
|
||||
|
||||
lemma iff_count' : l.ContrPerm l2 ↔
|
||||
l.contr.length = l2.contr.length ∧ IndexList.Subperm l.contr l2.contr.toIndexList ∧
|
||||
∀ I, l.contr.countSelf I = 1 → l2.contr.countSelf I = 1 := by
|
||||
rw [iff_count_fin]
|
||||
simp_all only [List.get_eq_getElem, and_congr_right_iff]
|
||||
intro _ _
|
||||
apply Iff.intro
|
||||
· intro ha I hI1
|
||||
have hI : I ∈ l.contr.val := by
|
||||
rw [← countSelf_neq_zero, hI1]
|
||||
exact Nat.one_ne_zero
|
||||
have hId : l.contr.val.indexOf I < l.contr.val.length := List.indexOf_lt_length.mpr hI
|
||||
have hI' : I = l.contr.val.get ⟨l.contr.val.indexOf I, hId⟩ := (List.indexOf_get hId).symm
|
||||
rw [hI'] at hI1 ⊢
|
||||
exact ha ⟨l.contr.val.indexOf I, hId⟩ hI1
|
||||
· exact fun a i a_1 => a l.contr.val[↑i] a_1
|
||||
|
||||
lemma iff_count : l.ContrPerm l2 ↔ l.contr.length = l2.contr.length ∧
|
||||
∀ I, l.contr.countSelf I = 1 → l2.contr.countSelf I = 1 := by
|
||||
rw [iff_count']
|
||||
refine Iff.intro (fun h => And.intro h.1 h.2.2) (fun h => And.intro h.1 (And.intro ?_ h.2))
|
||||
rw [Subperm.iff_countId]
|
||||
intro I
|
||||
apply And.intro (countId_contr_le_one l I)
|
||||
intro h'
|
||||
obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contr I (ne_zero_of_eq_one h')
|
||||
rw [countId_congr _ hI2] at h' ⊢
|
||||
have hi := h.2 I' (countSelf_eq_one_of_countId_eq_one l.contr.toIndexList I' h' hI1)
|
||||
have h1 := countSelf_le_countId l2.contr.toIndexList I'
|
||||
have h2 := countId_contr_le_one l2 I'
|
||||
omega
|
||||
|
||||
/-- The relation `ContrPerm` is symmetric. -/
|
||||
@[symm]
|
||||
lemma symm (h : ContrPerm l l') : ContrPerm l' l := by
|
||||
rw [ContrPerm] at h ⊢
|
||||
apply And.intro h.1.symm
|
||||
apply And.intro (Subperm.symm h.2.1 h.1)
|
||||
rw [← Function.comp_assoc, ← h.2.2, Function.comp_assoc, Function.comp_assoc]
|
||||
rw [show (l.contr.getDualInOtherEquiv l'.contr) =
|
||||
(l'.contr.getDualInOtherEquiv l.contr).symm from rfl]
|
||||
simp only [Equiv.symm_comp_self, CompTriple.comp_eq]
|
||||
|
||||
lemma iff_countSelf : l.ContrPerm l2 ↔ ∀ I, l.contr.countSelf I = l2.contr.countSelf I := by
|
||||
refine Iff.intro (fun h I => ?_) (fun h => ?_)
|
||||
· have hs := h.symm
|
||||
rw [iff_count] at hs h
|
||||
have ht := Iff.intro (fun h1 => h.2 I h1) (fun h2 => hs.2 I h2)
|
||||
have h1 : l.contr.countSelf I ≤ 1 := countSelf_contrIndexList_le_one l.toIndexList I
|
||||
have h2 : l2.contr.countSelf I ≤ 1 := countSelf_contrIndexList_le_one l2.toIndexList I
|
||||
omega
|
||||
· rw [iff_count]
|
||||
apply And.intro
|
||||
· have h1 : l.contr.val.Perm l2.contr.val := by
|
||||
rw [List.perm_iff_count]
|
||||
intro I
|
||||
rw [← countSelf_count, ← countSelf_count]
|
||||
exact h I
|
||||
exact List.Perm.length_eq h1
|
||||
· intro I
|
||||
rw [h I]
|
||||
exact fun a => a
|
||||
|
||||
lemma contr_perm (h : l.ContrPerm l2) : l.contr.val.Perm l2.contr.val := by
|
||||
rw [List.perm_iff_count]
|
||||
intro I
|
||||
rw [← countSelf_count, ← countSelf_count]
|
||||
exact iff_countSelf.mp h I
|
||||
|
||||
/-- The relation `ContrPerm` is reflexive. -/
|
||||
@[simp]
|
||||
lemma refl (l : ColorIndexList 𝓒) : ContrPerm l l :=
|
||||
iff_countSelf.mpr (congrFun rfl)
|
||||
|
||||
/-- The relation `ContrPerm` is transitive. -/
|
||||
@[trans]
|
||||
lemma trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) : ContrPerm l l3 := by
|
||||
rw [iff_countSelf] at h1 h2 ⊢
|
||||
exact fun I => (h1 I).trans (h2 I)
|
||||
|
||||
/-- `ContrPerm` forms an equivalence relation. -/
|
||||
lemma equivalence : Equivalence (@ContrPerm 𝓒 _ _) where
|
||||
refl := refl
|
||||
symm := symm
|
||||
trans := trans
|
||||
|
||||
lemma symm_trans (h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) :
|
||||
(h1.trans h2).symm = h2.symm.trans h1.symm := rfl
|
||||
|
||||
@[simp]
|
||||
lemma contr_self : ContrPerm l l.contr := by
|
||||
rw [iff_countSelf]
|
||||
intro I
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma self_contr : ContrPerm l.contr l := symm contr_self
|
||||
|
||||
lemma length_of_no_contr (h : l.ContrPerm l') (h1 : l.withDual = ∅) (h2 : l'.withDual = ∅) :
|
||||
l.length = l'.length := by
|
||||
simp only [ContrPerm] at h
|
||||
rw [contr_of_withDual_empty l h1, contr_of_withDual_empty l' h2] at h
|
||||
exact h.1
|
||||
|
||||
lemma mem_withUniqueDualInOther_of_no_contr (h : l.ContrPerm l') (h1 : l.withDual = ∅)
|
||||
(h2 : l'.withDual = ∅) : ∀ (x : Fin l.length), x ∈ l.withUniqueDualInOther l'.toIndexList := by
|
||||
simp only [ContrPerm] at h
|
||||
rw [contr_of_withDual_empty l h1, contr_of_withDual_empty l' h2] at h
|
||||
rw [h.2.1]
|
||||
exact fun x => Finset.mem_univ x
|
||||
|
||||
end ContrPerm
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences from `ContrPerm`
|
||||
|
||||
-/
|
||||
|
||||
open ContrPerm
|
||||
|
||||
/-- Given two `ColorIndexList` related by contracted permutations, the equivalence between
|
||||
the indices of the corresponding contracted index lists. This equivalence is the
|
||||
permutation between the contracted indices. -/
|
||||
def contrPermEquiv {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
Fin l.contr.length ≃ Fin l'.contr.length :=
|
||||
(Equiv.subtypeUnivEquiv (by rw [h.2.1]; exact fun x => Finset.mem_univ x)).symm.trans <|
|
||||
(l.contr.getDualInOtherEquiv l'.contr.toIndexList).trans <|
|
||||
Equiv.subtypeUnivEquiv (by rw [h.symm.2.1]; exact fun x => Finset.mem_univ x)
|
||||
|
||||
lemma contrPermEquiv_colorMap_iso {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
ColorMap.MapIso (contrPermEquiv h).symm l'.contr.colorMap' l.contr.colorMap' := by
|
||||
simp only [ColorMap.MapIso]
|
||||
funext i
|
||||
simp only [contrPermEquiv, getDualInOtherEquiv, Function.comp_apply, Equiv.symm_trans_apply,
|
||||
Equiv.symm_symm, Equiv.subtypeUnivEquiv_symm_apply, Equiv.coe_fn_symm_mk,
|
||||
Equiv.subtypeUnivEquiv_apply]
|
||||
have h' := h.symm.2.2
|
||||
have hi : i ∈ (l'.contr.withUniqueDualInOther l.contr.toIndexList) := by
|
||||
rw [h.symm.2.1]
|
||||
exact Finset.mem_univ i
|
||||
exact (congrFun h' ⟨i, hi⟩).symm
|
||||
|
||||
lemma contrPermEquiv_colorMap_iso' {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
ColorMap.MapIso (contrPermEquiv h) l.contr.colorMap' l'.contr.colorMap' := by
|
||||
rw [ColorMap.MapIso.symm']
|
||||
exact contrPermEquiv_colorMap_iso h
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_refl : contrPermEquiv (ContrPerm.refl l) = Equiv.refl _ := by
|
||||
simp [contrPermEquiv, ContrPerm.refl]
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_symm {l l' : ColorIndexList 𝓒} (h : ContrPerm l l') :
|
||||
(contrPermEquiv h).symm = contrPermEquiv h.symm := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_trans {l l2 l3 : ColorIndexList 𝓒}
|
||||
(h1 : ContrPerm l l2) (h2 : ContrPerm l2 l3) :
|
||||
(contrPermEquiv h1).trans (contrPermEquiv h2) = contrPermEquiv (h1.trans h2) := by
|
||||
simp only [contrPermEquiv]
|
||||
ext x
|
||||
simp only [getDualInOtherEquiv, Equiv.trans_apply, Equiv.subtypeUnivEquiv_symm_apply,
|
||||
Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply]
|
||||
apply congrArg
|
||||
have h1' : l.contr.countSelf (l.contr.val.get x) = 1 := by simp [contr]
|
||||
rw [iff_countSelf.mp h1, iff_countSelf.mp h2] at h1'
|
||||
have h3 : l3.contr.countId (l.contr.val.get x) = 1 := by
|
||||
have h' := countSelf_le_countId l3.contr.toIndexList (l.contr.val.get x)
|
||||
have h'' := countId_contr_le_one l3 (l.contr.val.get x)
|
||||
omega
|
||||
rw [countId_get_other, List.countP_eq_length_filter, List.length_eq_one] at h3
|
||||
obtain ⟨a, ha⟩ := h3
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_self_contr {l : ColorIndexList 𝓒} :
|
||||
contrPermEquiv (contr_self : ContrPerm l l.contr) =
|
||||
(Fin.castOrderIso (by simp)).toEquiv := by
|
||||
simp only [contrPermEquiv]
|
||||
ext1 x
|
||||
simp only [getDualInOtherEquiv, Equiv.trans_apply, Equiv.subtypeUnivEquiv_symm_apply,
|
||||
Equiv.coe_fn_mk, Equiv.subtypeUnivEquiv_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Fin.coe_cast]
|
||||
symm
|
||||
have h1' : l.contr.countSelf (l.contr.val.get x) = 1 := by simp [contr]
|
||||
rw [iff_countSelf.mp contr_self] at h1'
|
||||
have h3 : l.contr.contr.countId (l.contr.val.get x) = 1 := by
|
||||
have h' := countSelf_le_countId l.contr.contr.toIndexList (l.contr.val.get x)
|
||||
have h'' := countId_contr_le_one l.contr (l.contr.val.get x)
|
||||
omega
|
||||
rw [countId_get_other, List.countP_eq_length_filter, List.length_eq_one] at h3
|
||||
obtain ⟨a, ha⟩ := h3
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [AreDualInOther]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp only [AreDualInOther, List.mem_filter, List.mem_finRange,
|
||||
decide_eq_true_eq, true_and, getDualInOther?_id]
|
||||
|
||||
@[simp]
|
||||
lemma contrPermEquiv_contr_self {l : ColorIndexList 𝓒} :
|
||||
contrPermEquiv (self_contr : ContrPerm l.contr l) =
|
||||
(Fin.castOrderIso (by simp)).toEquiv := by
|
||||
rw [← contrPermEquiv_symm, contrPermEquiv_self_contr]
|
||||
rfl
|
||||
|
||||
/-- Given two `ColorIndexList` related by permutations and without duals, the equivalence between
|
||||
the indices of the corresponding index lists. This equivalence is the
|
||||
permutation between the indices. -/
|
||||
def permEquiv {l l' : ColorIndexList 𝓒} (h : ContrPerm l l')
|
||||
(h1 : l.withDual = ∅) (h2 : l'.withDual = ∅) : Fin l.length ≃ Fin l'.length :=
|
||||
(Equiv.subtypeUnivEquiv (mem_withUniqueDualInOther_of_no_contr h h1 h2)).symm.trans <|
|
||||
(l.getDualInOtherEquiv l'.toIndexList).trans <|
|
||||
Equiv.subtypeUnivEquiv (mem_withUniqueDualInOther_of_no_contr h.symm h2 h1)
|
||||
|
||||
lemma permEquiv_colorMap_iso {l l' : ColorIndexList 𝓒} (h : ContrPerm l l')
|
||||
(h1 : l.withDual = ∅) (h2 : l'.withDual = ∅) :
|
||||
ColorMap.MapIso (permEquiv h h1 h2).symm l'.colorMap' l.colorMap' := by
|
||||
funext i
|
||||
rw [Function.comp_apply]
|
||||
have h' := h.symm
|
||||
simp only [ContrPerm] at h'
|
||||
rw [contr_of_withDual_empty l h1, contr_of_withDual_empty l' h2] at h'
|
||||
exact (congrFun h'.2.2 ⟨i, _⟩).symm
|
||||
|
||||
end ColorIndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,204 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.ColorIndexList.Basic
|
||||
import HepLean.Tensors.IndexNotation.IndexList.Contraction
|
||||
import HepLean.Tensors.Contraction
|
||||
import HepLean.Tensors.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Contraction of ColorIndexLists
|
||||
|
||||
We define the contraction of ColorIndexLists.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
namespace ColorIndexList
|
||||
|
||||
variable {𝓒 : TensorColor} [IndexNotation 𝓒.Color]
|
||||
(l l2 : ColorIndexList 𝓒)
|
||||
|
||||
open IndexList TensorColor
|
||||
|
||||
/-!
|
||||
|
||||
## Contracting an `ColorIndexList`
|
||||
|
||||
-/
|
||||
|
||||
/-- The contraction of a `ColorIndexList`, `l`.
|
||||
That is, the `ColorIndexList` obtained by taking only those indices in `l` which do not
|
||||
have a dual. This can be thought of as contracting all of those indices with a dual. -/
|
||||
def contr : ColorIndexList 𝓒 where
|
||||
toIndexList := l.toIndexList.contrIndexList
|
||||
unique_duals := by
|
||||
simp [OnlyUniqueDuals]
|
||||
dual_color := ColorCond.contrIndexList
|
||||
|
||||
@[simp]
|
||||
lemma contr_contr : l.contr.contr = l.contr := by
|
||||
apply ext
|
||||
simp [contr]
|
||||
|
||||
@[simp]
|
||||
lemma contr_contr_idMap (i : Fin l.contr.contr.length) :
|
||||
l.contr.contr.idMap i = l.contr.idMap (Fin.cast (by simp) i) := by
|
||||
simp only [contr, contrIndexList_idMap, Fin.cast_trans]
|
||||
apply congrArg
|
||||
simp only [withoutDualEquiv, RelIso.coe_fn_toEquiv, Finset.coe_orderIsoOfFin_apply,
|
||||
EmbeddingLike.apply_eq_iff_eq]
|
||||
have h1 : l.contrIndexList.withoutDual = Finset.univ := by
|
||||
have hx := l.contrIndexList.withDual_union_withoutDual
|
||||
have hx2 := l.contrIndexList_withDual
|
||||
simp_all
|
||||
simp only [h1]
|
||||
rw [orderEmbOfFin_univ]
|
||||
· rfl
|
||||
· rw [h1]
|
||||
exact (Finset.card_fin l.contrIndexList.length).symm
|
||||
|
||||
@[simp]
|
||||
lemma contr_contr_colorMap (i : Fin l.contr.contr.length) :
|
||||
l.contr.contr.colorMap i = l.contr.colorMap (Fin.cast (by simp) i) := by
|
||||
simp only [contr, contrIndexList_colorMap, Fin.cast_trans]
|
||||
apply congrArg
|
||||
simp only [withoutDualEquiv, RelIso.coe_fn_toEquiv, Finset.coe_orderIsoOfFin_apply,
|
||||
EmbeddingLike.apply_eq_iff_eq]
|
||||
have h1 : l.contrIndexList.withoutDual = Finset.univ := by
|
||||
have hx := l.contrIndexList.withDual_union_withoutDual
|
||||
have hx2 := l.contrIndexList_withDual
|
||||
simp_all
|
||||
simp only [h1]
|
||||
rw [orderEmbOfFin_univ]
|
||||
· rfl
|
||||
· rw [h1]
|
||||
exact (Finset.card_fin l.contrIndexList.length).symm
|
||||
|
||||
@[simp]
|
||||
lemma contr_of_withDual_empty (h : l.withDual = ∅) :
|
||||
l.contr = l := by
|
||||
simp only [contr]
|
||||
apply ext
|
||||
simp [l.contrIndexList_of_withDual_empty h]
|
||||
|
||||
@[simp]
|
||||
lemma contr_getDual?_eq_none (i : Fin l.contr.length) :
|
||||
l.contr.getDual? i = none := by
|
||||
simp only [contr, contrIndexList_getDual?]
|
||||
|
||||
@[simp]
|
||||
lemma contr_areDualInSelf (i j : Fin l.contr.length) :
|
||||
l.contr.AreDualInSelf i j ↔ False := by
|
||||
simp [contr]
|
||||
|
||||
lemma contr_countId_eq_zero_of_countId_zero (I : Index 𝓒.Color)
|
||||
(h : l.countId I = 0) : l.contr.countId I = 0 := by
|
||||
simp only [contr]
|
||||
exact countId_contrIndexList_zero_of_countId l.toIndexList I h
|
||||
|
||||
lemma contr_countId_eq_filter (I : Index 𝓒.Color) :
|
||||
l.contr.countId I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun i => l.val.countP (fun j => i.id = j.id) = 1) := by
|
||||
simp only [countId, contr, contrIndexList]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
congr
|
||||
funext J
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
exact Bool.and_comm (decide (I.id = J.id))
|
||||
(decide (List.countP (fun j => decide (J.id = j.id)) l.val = 1))
|
||||
|
||||
lemma countId_contr_le_one (I : Index 𝓒.Color) :
|
||||
l.contr.countId I ≤ 1 := by
|
||||
exact l.countId_contrIndexList_le_one I
|
||||
|
||||
lemma countId_contr_eq_zero_iff [DecidableEq 𝓒.Color] (I : Index 𝓒.Color) :
|
||||
l.contr.countId I = 0 ↔ l.countId I = 0 ∨ l.countId I = 2 := by
|
||||
by_cases hI : l.contr.countId I = 1
|
||||
· have hI' := hI
|
||||
erw [countId_contrIndexList_eq_one_iff_countId_eq_one] at hI'
|
||||
omega
|
||||
· have hI' := hI
|
||||
erw [countId_contrIndexList_eq_one_iff_countId_eq_one] at hI'
|
||||
have hI2 := l.countId_le_two I
|
||||
have hI3 := l.countId_contrIndexList_le_one I
|
||||
have hI3 : l.contr.countId I = 0 := by
|
||||
simp_all only [contr]
|
||||
omega
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## Contract equiv
|
||||
|
||||
-/
|
||||
|
||||
/-- An equivalence splitting the indices of `l` into
|
||||
a sum type of those indices and their duals (with choice determined by the ordering on `Fin`),
|
||||
and those indices without duals.
|
||||
|
||||
This equivalence is used to contract the indices of tensors. -/
|
||||
def contrEquiv : (l.withUniqueDualLT ⊕ l.withUniqueDualLT) ⊕ Fin l.contr.length ≃ Fin l.length :=
|
||||
(Equiv.sumCongr (l.withUniqueLTGTEquiv) (Equiv.refl _)).trans <|
|
||||
(Equiv.sumCongr (Equiv.subtypeEquivRight (by
|
||||
rw [l.unique_duals]
|
||||
exact fun x => Eq.to_iff rfl))
|
||||
(Fin.castOrderIso l.contrIndexList_length).toEquiv).trans <|
|
||||
l.dualEquiv
|
||||
|
||||
lemma contrEquiv_inl_inl_isSome (i : l.withUniqueDualLT) :
|
||||
(l.getDual? (l.contrEquiv (Sum.inl (Sum.inl i)))).isSome :=
|
||||
mem_withUniqueDual_isSome l.toIndexList (↑i)
|
||||
(mem_withUniqueDual_of_mem_withUniqueDualLt l.toIndexList (↑i) i.2)
|
||||
|
||||
@[simp]
|
||||
lemma contrEquiv_inl_inr_eq (i : l.withUniqueDualLT) :
|
||||
(l.contrEquiv (Sum.inl (Sum.inr i))) =
|
||||
(l.getDual? i.1).get (l.contrEquiv_inl_inl_isSome i) := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrEquiv_inl_inl_eq (i : l.withUniqueDualLT) :
|
||||
(l.contrEquiv (Sum.inl (Sum.inl i))) = i := by
|
||||
rfl
|
||||
|
||||
lemma contrEquiv_colorMapIso :
|
||||
ColorMap.MapIso (Equiv.refl (Fin l.contr.length))
|
||||
(ColorMap.contr l.contrEquiv l.colorMap) l.contr.colorMap := by
|
||||
simp only [ColorMap.MapIso, ColorMap.contr, Equiv.coe_refl, CompTriple.comp_eq]
|
||||
funext i
|
||||
simp only [contr, Function.comp_apply, contrIndexList_colorMap]
|
||||
rfl
|
||||
|
||||
lemma contrEquiv_contrCond : ColorMap.ContrCond l.contrEquiv l.colorMap := by
|
||||
simp only [ColorMap.ContrCond]
|
||||
funext i
|
||||
simp only [Function.comp_apply, contrEquiv_inl_inl_eq, contrEquiv_inl_inr_eq]
|
||||
have h1 := l.dual_color
|
||||
rw [ColorCond.iff_on_isSome] at h1
|
||||
exact (h1 i.1 _).symm
|
||||
|
||||
@[simp]
|
||||
lemma contrEquiv_on_withDual_empty (i : Fin l.contr.length) (h : l.withDual = ∅) :
|
||||
l.contrEquiv (Sum.inr i) = Fin.cast (by simp [h]) i := by
|
||||
simp only [contrEquiv, Equiv.trans_apply, Equiv.sumCongr_apply, Equiv.coe_refl, Sum.map_inr,
|
||||
id_eq]
|
||||
change l.dualEquiv (Sum.inr ((Fin.castOrderIso _).toEquiv i)) = _
|
||||
simp only [dualEquiv, withoutDualEquiv, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
|
||||
Equiv.trans_apply, Equiv.sumCongr_apply, Equiv.coe_refl, Sum.map_inr,
|
||||
Equiv.Finset.union_symm_inr, Finset.coe_orderIsoOfFin_apply, Equiv.subtypeUnivEquiv_apply]
|
||||
have h : l.withoutDual = Finset.univ := by
|
||||
have hx := l.withDual_union_withoutDual
|
||||
simp_all
|
||||
simp only [h]
|
||||
rw [orderEmbOfFin_univ]
|
||||
· rfl
|
||||
· rw [h]
|
||||
exact (Finset.card_fin l.length).symm
|
||||
|
||||
end ColorIndexList
|
||||
end IndexNotation
|
|
@ -1,353 +0,0 @@
|
|||
/-
|
||||
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 Mathlib.Data.Set.Finite
|
||||
import Mathlib.Logic.Equiv.Fin
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import HepLean.Tensors.IndexNotation.Basic
|
||||
/-!
|
||||
|
||||
# Index lists
|
||||
|
||||
i.e. lists of indices.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
variable (X : Type) [IndexNotation X]
|
||||
variable [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- The type of lists of indices. -/
|
||||
structure IndexList where
|
||||
/-- The list of index values. For example `['ᵘ¹','ᵘ²','ᵤ₁']`. -/
|
||||
val : 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 length : ℕ := l.val.length
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma ext (h : l.val = l2.val) : l = l2 := by
|
||||
cases l
|
||||
cases l2
|
||||
simp_all
|
||||
|
||||
/-- The index list constructed by prepending an index to the list. -/
|
||||
def cons (i : Index X) : IndexList X := {val := i :: l.val}
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma cons_val (i : Index X) : (l.cons i).val = i :: l.val := by
|
||||
rfl
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma cons_length (i : Index X) : (l.cons i).length = l.length + 1 := by
|
||||
rfl
|
||||
|
||||
/-- The tail of an index list. That is, the index list with the first index dropped. -/
|
||||
def tail : IndexList X := {val := l.val.tail}
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma tail_val : l.tail.val = l.val.tail := by
|
||||
rfl
|
||||
|
||||
/-- The first index in a non-empty index list. -/
|
||||
def head (h : l ≠ {val := ∅}) : Index X := l.val.head (by cases' l; simpa using h)
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma head_cons_tail (h : l ≠ {val := ∅}) : l = (l.tail.cons (l.head h)) := by
|
||||
apply ext
|
||||
simp only [cons_val, tail_val]
|
||||
simp only [head, List.head_cons_tail]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma induction {P : IndexList X → Prop } (h_nil : P {val := ∅})
|
||||
(h_cons : ∀ (x : Index X) (xs : IndexList X), P xs → P (xs.cons x)) (l : IndexList X) : P l := by
|
||||
cases' l with val
|
||||
induction val with
|
||||
| nil => exact h_nil
|
||||
| cons x xs ih =>
|
||||
exact h_cons x ⟨xs⟩ ih
|
||||
|
||||
/-- The map of from `Fin s.numIndices` into colors associated to an index list. -/
|
||||
def colorMap : Fin l.length → X :=
|
||||
fun i => (l.val.get i).toColor
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma colorMap_cast {l1 l2 : IndexList X} (h : l1 = l2) :
|
||||
l1.colorMap = l2.colorMap ∘ Fin.cast (congrArg length h) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
/-- The map of from `Fin s.numIndices` into the natural numbers associated to an index list. -/
|
||||
def idMap : Fin l.length → Nat :=
|
||||
fun i => (l.val.get i).id
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
lemma idMap_cast {l1 l2 : IndexList X} (h : l1 = l2) (i : Fin l1.length) :
|
||||
l1.idMap i = l2.idMap (Fin.cast (by rw [h]) i) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
lemma ext_colorMap_idMap {l l2 : IndexList X} (hl : l.length = l2.length)
|
||||
(hi : l.idMap = l2.idMap ∘ Fin.cast hl) (hc : l.colorMap = l2.colorMap ∘ Fin.cast hl) :
|
||||
l = l2 := by
|
||||
apply ext
|
||||
refine List.ext_get hl ?h.h
|
||||
intro n h1 h2
|
||||
rw [Index.eq_iff_color_eq_and_id_eq]
|
||||
apply And.intro
|
||||
· trans l.colorMap ⟨n, h1⟩
|
||||
· rfl
|
||||
· rw [hc]
|
||||
rfl
|
||||
· trans l.idMap ⟨n, h1⟩
|
||||
· rfl
|
||||
· rw [hi]
|
||||
rfl
|
||||
|
||||
/-- Given a list of indices a subset of `Fin l.numIndices × Index X`
|
||||
of pairs of positions in `l` and the corresponding item in `l`. -/
|
||||
def toPosSet (l : IndexList X) : Set (Fin l.length × Index X) :=
|
||||
{(i, l.val.get i) | i : Fin l.length}
|
||||
|
||||
/-- Equivalence between `toPosSet` and `Fin l.numIndices`. -/
|
||||
def toPosSetEquiv (l : IndexList X) : l.toPosSet ≃ Fin l.length where
|
||||
toFun := fun x => x.1.1
|
||||
invFun := fun x => ⟨(x, l.val.get x), by simp [toPosSet]⟩
|
||||
left_inv x := by
|
||||
have hx := x.prop
|
||||
simp only [toPosSet, List.get_eq_getElem, Set.mem_setOf_eq] at hx
|
||||
simp only [List.get_eq_getElem]
|
||||
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
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
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]
|
||||
|
||||
/-- Given a list of indices a finite set of `Fin l.length × Index X`
|
||||
of pairs of positions in `l` and the corresponding item in `l`. -/
|
||||
def toPosFinset (l : IndexList X) : Finset (Fin l.length × Index X) :=
|
||||
l.toPosSet.toFinset
|
||||
|
||||
/-- The construction of a list of indices from a map
|
||||
from `Fin n` to `Index X`. -/
|
||||
def fromFinMap {n : ℕ} (f : Fin n → Index X) : IndexList X where
|
||||
val := (Fin.list n).map f
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma fromFinMap_numIndices {n : ℕ} (f : Fin n → Index X) :
|
||||
(fromFinMap f).length = n := by
|
||||
simp [fromFinMap, length]
|
||||
|
||||
/-!
|
||||
|
||||
## Appending index lists.
|
||||
|
||||
-/
|
||||
|
||||
section append
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
instance : HAppend (IndexList X) (IndexList X) (IndexList X) where
|
||||
hAppend := fun l l2 => {val := l.val ++ l2.val}
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma cons_append (i : Index X) : (l.cons i) ++ l2 = (l ++ l2).cons i := by
|
||||
rfl
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
@[simp]
|
||||
lemma append_length : (l ++ l2).length = l.length + l2.length := by
|
||||
simpa only [length] using List.length_append l.val l2.val
|
||||
|
||||
lemma append_assoc : l ++ l2 ++ l3 = l ++ (l2 ++ l3) := by
|
||||
apply ext
|
||||
change l.val ++ l2.val ++ l3.val = l.val ++ (l2.val ++ l3.val)
|
||||
exact List.append_assoc l.val l2.val l3.val
|
||||
|
||||
/-- An equivalence between the sum of the types of indices of `l` an `l2` and the type
|
||||
of indices of the joined index list `l ++ l2`. -/
|
||||
def appendEquiv {l l2 : IndexList X} : Fin l.length ⊕ Fin l2.length ≃ Fin (l ++ l2).length :=
|
||||
finSumFinEquiv.trans (Fin.castOrderIso (List.length_append _ _).symm).toEquiv
|
||||
|
||||
/-- The inclusion of the indices of `l` into the indices of `l ++ l2`. -/
|
||||
def appendInl : Fin l.length ↪ Fin (l ++ l2).length where
|
||||
toFun := appendEquiv ∘ Sum.inl
|
||||
inj' := by
|
||||
intro i j h
|
||||
simp only [Function.comp, EmbeddingLike.apply_eq_iff_eq, Sum.inl.injEq] at h
|
||||
exact h
|
||||
|
||||
/-- The inclusion of the indices of `l2` into the indices of `l ++ l2`. -/
|
||||
def appendInr : Fin l2.length ↪ Fin (l ++ l2).length where
|
||||
toFun := appendEquiv ∘ Sum.inr
|
||||
inj' i j h := by
|
||||
simpa only [Function.comp, EmbeddingLike.apply_eq_iff_eq, Sum.inr.injEq] using h
|
||||
|
||||
@[simp]
|
||||
lemma appendInl_appendEquiv :
|
||||
(l.appendInl l2).trans appendEquiv.symm.toEmbedding =
|
||||
{toFun := Sum.inl, inj' := Sum.inl_injective} := by
|
||||
ext i
|
||||
simp [appendInl]
|
||||
|
||||
@[simp]
|
||||
lemma appendInr_appendEquiv :
|
||||
(l.appendInr l2).trans appendEquiv.symm.toEmbedding =
|
||||
{toFun := Sum.inr, inj' := Sum.inr_injective} := by
|
||||
ext i
|
||||
simp [appendInr]
|
||||
|
||||
@[simp]
|
||||
lemma append_val {l l2 : IndexList X} : (l ++ l2).val = l.val ++ l2.val := by
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma idMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
|
||||
(l ++ l2).idMap (appendEquiv (Sum.inl i)) = l.idMap i := by
|
||||
simp only [idMap, append_val, appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_left,
|
||||
List.get_eq_getElem]
|
||||
rw [List.getElem_append_left]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma idMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
|
||||
(l ++ l2).idMap (appendEquiv (Sum.inr i)) = l2.idMap i := by
|
||||
simp only [idMap, append_val, length, appendEquiv, Equiv.trans_apply, finSumFinEquiv_apply_right,
|
||||
RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, List.get_eq_getElem, Fin.coe_cast,
|
||||
Fin.coe_natAdd]
|
||||
rw [List.getElem_append_right]
|
||||
· simp only [Nat.add_sub_cancel_left]
|
||||
· omega
|
||||
· omega
|
||||
|
||||
@[simp]
|
||||
lemma colorMap_append_inl {l l2 : IndexList X} (i : Fin l.length) :
|
||||
(l ++ l2).colorMap (appendEquiv (Sum.inl i)) = l.colorMap i := by
|
||||
simp only [colorMap, append_val, length, appendEquiv, Equiv.trans_apply,
|
||||
finSumFinEquiv_apply_left, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, List.get_eq_getElem,
|
||||
Fin.coe_cast, Fin.coe_castAdd]
|
||||
rw [List.getElem_append_left]
|
||||
|
||||
@[simp]
|
||||
lemma colorMap_append_inl' :
|
||||
(l ++ l2).colorMap ∘ appendEquiv ∘ Sum.inl = l.colorMap := by
|
||||
funext i
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma colorMap_append_inr {l l2 : IndexList X} (i : Fin l2.length) :
|
||||
(l ++ l2).colorMap (appendEquiv (Sum.inr i)) = l2.colorMap i := by
|
||||
simp only [colorMap, append_val, length, appendEquiv, Equiv.trans_apply,
|
||||
finSumFinEquiv_apply_right, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, List.get_eq_getElem,
|
||||
Fin.coe_cast, Fin.coe_natAdd]
|
||||
rw [List.getElem_append_right]
|
||||
· simp only [Nat.add_sub_cancel_left]
|
||||
· omega
|
||||
· omega
|
||||
|
||||
@[simp]
|
||||
lemma colorMap_append_inr' :
|
||||
(l ++ l2).colorMap ∘ appendEquiv ∘ Sum.inr = l2.colorMap := by
|
||||
funext i
|
||||
simp
|
||||
|
||||
lemma colorMap_sumELim (l1 l2 : IndexList X) :
|
||||
Sum.elim l1.colorMap l2.colorMap =
|
||||
(l1 ++ l2).colorMap ∘ appendEquiv := by
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl i => simp
|
||||
| Sum.inr i => simp
|
||||
|
||||
end append
|
||||
|
||||
/-!
|
||||
|
||||
## Filter id
|
||||
|
||||
-/
|
||||
|
||||
/-! TODO: Replace with Mathlib lemma. -/
|
||||
lemma filter_sort_comm {n : ℕ} (s : Finset (Fin n)) (p : Fin n → Prop) [DecidablePred p] :
|
||||
List.filter p (Finset.sort (fun i j => i ≤ j) s) =
|
||||
Finset.sort (fun i j => i ≤ j) (Finset.filter p s) := by
|
||||
simp only [Finset.sort, Finset.filter]
|
||||
have : ∀ (m : Multiset (Fin n)), List.filter p (Multiset.sort (fun i j => i ≤ j) m) =
|
||||
Multiset.sort (fun i j => i ≤ j) (Multiset.filter p m) := by
|
||||
apply Quot.ind
|
||||
intro m
|
||||
simp only [Multiset.quot_mk_to_coe'', Multiset.coe_sort, Multiset.filter_coe]
|
||||
have h1 : List.Sorted (fun i j => i ≤ j) (List.filter (fun b => decide (p b))
|
||||
(List.mergeSort' (fun i j => i ≤ j) m)) := by
|
||||
simp only [List.Sorted]
|
||||
rw [List.pairwise_filter, List.pairwise_iff_get]
|
||||
intro i j h1 _ _
|
||||
have hs : List.Sorted (fun i j => i ≤ j) (List.mergeSort' (fun i j => i ≤ j) m) := by
|
||||
exact List.sorted_mergeSort' (fun i j => i ≤ j) m
|
||||
simp only [List.Sorted] at hs
|
||||
rw [List.pairwise_iff_get] at hs
|
||||
exact hs i j h1
|
||||
have hp1 : (List.mergeSort' (fun i j => i ≤ j) m).Perm m := by
|
||||
exact List.perm_mergeSort' (fun i j => i ≤ j) m
|
||||
have hp2 : (List.filter (fun b => decide (p b)) ((List.mergeSort' (fun i j => i ≤ j) m))).Perm
|
||||
(List.filter (fun b => decide (p b)) m) := by
|
||||
exact List.Perm.filter (fun b => decide (p b)) hp1
|
||||
have hp3 : (List.filter (fun b => decide (p b)) m).Perm
|
||||
(List.mergeSort' (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)) := by
|
||||
exact List.Perm.symm (List.perm_mergeSort' (fun i j => i ≤ j)
|
||||
(List.filter (fun b => decide (p b)) m))
|
||||
have hp4 := hp2.trans hp3
|
||||
refine List.eq_of_perm_of_sorted hp4 h1 ?_
|
||||
exact List.sorted_mergeSort' (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)
|
||||
exact this s.val
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma filter_id_eq_sort (i : Fin l.length) : l.val.filter (fun J => (l.val.get i).id = J.id) =
|
||||
List.map l.val.get (Finset.sort (fun i j => i ≤ j)
|
||||
(Finset.filter (fun j => l.idMap i = l.idMap j) Finset.univ)) := by
|
||||
have h1 := (List.finRange_map_get l.val).symm
|
||||
have h2 : l.val = List.map l.val.get (Finset.sort (fun i j => i ≤ j) Finset.univ) := by
|
||||
nth_rewrite 1 [h1, (Fin.sort_univ l.val.length).symm]
|
||||
rfl
|
||||
nth_rewrite 3 [h2]
|
||||
rw [List.filter_map]
|
||||
apply congrArg
|
||||
rw [← filter_sort_comm]
|
||||
apply List.filter_congr
|
||||
intro x _
|
||||
rfl
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,583 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Contraction
|
||||
import HepLean.Tensors.IndexNotation.IndexList.OnlyUniqueDuals
|
||||
import HepLean.Tensors.Basic
|
||||
import Init.Data.List.Lemmas
|
||||
/-!
|
||||
|
||||
# Index lists and color
|
||||
|
||||
In this file we look at the interaction of index lists and color.
|
||||
|
||||
The main definition of this file is `ColorCond`.
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace Index
|
||||
|
||||
variable {𝓒 : TensorColor}
|
||||
variable (I : Index 𝓒.Color)
|
||||
|
||||
/-- The dual of an index is the index with the same id, but opposite color. -/
|
||||
def dual : Index 𝓒.Color := ⟨𝓒.τ I.toColor, I.id⟩
|
||||
|
||||
@[simp]
|
||||
lemma dual_dual : I.dual.dual = I := by
|
||||
simp only [dual, toColor, id]
|
||||
rw [𝓒.τ_involutive]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma dual_id : I.dual.id = I.id := rfl
|
||||
|
||||
@[simp]
|
||||
lemma dual_color : I.dual.toColor = 𝓒.τ I.toColor := rfl
|
||||
|
||||
end Index
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {𝓒 : TensorColor}
|
||||
variable [DecidableEq 𝓒.Color]
|
||||
variable (l l2 l3 : IndexList 𝓒.Color)
|
||||
|
||||
/-- The number of times `I` or its dual appears in an `IndexList`. -/
|
||||
def countColorQuot (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I = J ∨ I.dual = J)
|
||||
|
||||
lemma countColorQuot_eq_filter_id_countP (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun J => I.toColor = J.toColor ∨ I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
simp only [countColorQuot, Bool.decide_or]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp only [Index.eq_iff_color_eq_and_id_eq, Bool.decide_and, Index.dual_color, Index.dual_id,
|
||||
Bool.or_eq_true, Bool.and_eq_true, decide_eq_true_eq, Bool.decide_or]
|
||||
apply Iff.intro
|
||||
· intro a_1
|
||||
cases a_1 with
|
||||
| inl h => simp_all only [true_or, and_self]
|
||||
| inr h_1 =>
|
||||
simp_all only [and_true]
|
||||
obtain ⟨left, _⟩ := h_1
|
||||
rw [← left]
|
||||
rw [𝓒.τ_involutive]
|
||||
exact Or.inr rfl
|
||||
· intro a_1
|
||||
simp_all only [and_true]
|
||||
obtain ⟨left, _⟩ := a_1
|
||||
cases left with
|
||||
| inl h => simp_all only [true_or]
|
||||
| inr h_1 =>
|
||||
simp_all only
|
||||
rw [𝓒.τ_involutive]
|
||||
exact Or.inr rfl
|
||||
|
||||
lemma countColorQuot_eq_filter_color_countP (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I =
|
||||
(l.val.filter (fun J => I.toColor = J.toColor ∨ I.toColor = 𝓒.τ (J.toColor))).countP
|
||||
(fun J => I.id = J.id) := by
|
||||
rw [countColorQuot_eq_filter_id_countP]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simpa using And.comm
|
||||
|
||||
@[simp]
|
||||
lemma countColorQuot_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countColorQuot I = countColorQuot l I + countColorQuot l2 I := by
|
||||
simp [countColorQuot]
|
||||
|
||||
lemma countColorQuot_eq_countId_iff_of_isSome (hl : l.OnlyUniqueDuals) (i : Fin l.length)
|
||||
(hi : (l.getDual? i).isSome) :
|
||||
l.countColorQuot (l.val.get i) = l.countId (l.val.get i) ↔
|
||||
(l.colorMap i = l.colorMap ((l.getDual? i).get hi) ∨
|
||||
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))) := by
|
||||
rw [countColorQuot_eq_filter_id_countP, countId_eq_length_filter]
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp only [List.countP, List.countP.go, List.get_eq_getElem, true_or, decide_True,
|
||||
Bool.decide_or, zero_add, Nat.reduceAdd, cond_true, List.length_cons, List.length_singleton]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· by_contra hn
|
||||
have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = false := by
|
||||
simpa using hn
|
||||
erw [hn'] at h
|
||||
simp at h
|
||||
· have hn' : (decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) ||
|
||||
decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)) = true := by
|
||||
simpa using h
|
||||
erw [hn']
|
||||
rfl
|
||||
|
||||
lemma countColorQuot_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countColorQuot I = 0 := by
|
||||
rw [countColorQuot_eq_filter_id_countP]
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countColorQuot]
|
||||
|
||||
lemma countColorQuot_le_countId (I : Index 𝓒.Color) :
|
||||
l.countColorQuot I ≤ l.countId I := by
|
||||
rw [countColorQuot_eq_filter_color_countP, countId]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.filter_sublist l.val
|
||||
|
||||
lemma countColorQuot_contrIndexList_le_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I ≤ 1 :=
|
||||
(l.contrIndexList.countColorQuot_le_countId I).trans
|
||||
(countId_contrIndexList_le_one l I)
|
||||
|
||||
lemma countColorQuot_contrIndexList_eq_zero_or_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I = 0 ∨ l.contrIndexList.countColorQuot I = 1 := by
|
||||
have h1 := countColorQuot_contrIndexList_le_one l I
|
||||
exact Nat.le_one_iff_eq_zero_or_eq_one.mp h1
|
||||
|
||||
lemma countColorQuot_contrIndexList_le_countColorQuot (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countColorQuot I ≤ l.countColorQuot I := by
|
||||
rw [countColorQuot_eq_filter_color_countP, countColorQuot_eq_filter_color_countP]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.Sublist.filter _ (List.filter_sublist l.val)
|
||||
|
||||
lemma countColorQuot_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countColorQuot I = l.countColorQuot I := by
|
||||
rw [countColorQuot_eq_filter_id_countP,
|
||||
l.filter_id_contrIndexList_eq_of_countId_contrIndexList I h1,
|
||||
countColorQuot_eq_filter_id_countP]
|
||||
|
||||
/-- The number of times an index `I` appears in an index list. -/
|
||||
def countSelf (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I = J)
|
||||
|
||||
lemma countSelf_eq_filter_id_countP : l.countSelf I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = J.toColor) := by
|
||||
rw [countSelf]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp [Index.eq_iff_color_eq_and_id_eq]
|
||||
|
||||
lemma countSelf_eq_filter_color_countP :
|
||||
l.countSelf I =
|
||||
(l.val.filter (fun J => I.toColor = J.toColor)).countP (fun J => I.id = J.id) := by
|
||||
simp only [countSelf]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simpa [Index.eq_iff_color_eq_and_id_eq] using And.comm
|
||||
|
||||
lemma countSelf_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countSelf I = 0 := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countSelf_eq_filter_id_countP]
|
||||
|
||||
lemma countSelf_count (I : Index 𝓒.Color) : l.countSelf I = l.val.count I := by
|
||||
rw [countSelf, List.count]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp only [decide_eq_true_eq, beq_iff_eq]
|
||||
exact eq_comm
|
||||
|
||||
lemma countSelf_eq_zero (I : Index 𝓒.Color) : l.countSelf I = 0 ↔ I ∉ l.val := by
|
||||
rw [countSelf_count]
|
||||
exact List.count_eq_zero
|
||||
|
||||
lemma countSelf_neq_zero (I : Index 𝓒.Color) : l.countSelf I ≠ 0 ↔ I ∈ l.val := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· simpa using (l.countSelf_eq_zero I).mpr.mt h
|
||||
· exact (l.countSelf_eq_zero I).mp.mt (by simpa using h)
|
||||
|
||||
@[simp]
|
||||
lemma countSelf_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countSelf I = countSelf l I + countSelf l2 I := by
|
||||
simp [countSelf]
|
||||
|
||||
lemma countSelf_le_countId (I : Index 𝓒.Color) :
|
||||
l.countSelf I ≤ l.countId I := by
|
||||
rw [countSelf_eq_filter_color_countP, countId]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.filter_sublist l.val
|
||||
|
||||
lemma countSelf_eq_one_of_countId_eq_one (I : Index 𝓒.Color) (h1 : l.countId I = 1)
|
||||
(hme : I ∈ l.val) : l.countSelf I = 1 := by
|
||||
rw [countSelf_eq_filter_id_countP]
|
||||
rw [filter_id_of_countId_eq_one_mem l hme h1]
|
||||
simp
|
||||
|
||||
lemma countSelf_contrIndexList_le_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I ≤ 1 :=
|
||||
(l.contrIndexList.countSelf_le_countId I).trans (countId_contrIndexList_le_one l I)
|
||||
|
||||
lemma countSelf_contrIndexList_eq_zero_or_one (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I = 0 ∨ l.contrIndexList.countSelf I = 1 := by
|
||||
exact Nat.le_one_iff_eq_zero_or_eq_one.mp (countSelf_contrIndexList_le_one l I)
|
||||
|
||||
lemma countSelf_contrIndexList_eq_zero_of_zero (I : Index 𝓒.Color) (h : l.countSelf I = 0) :
|
||||
l.contrIndexList.countSelf I = 0 := by
|
||||
rw [countSelf_eq_zero] at h ⊢
|
||||
simp_all [contrIndexList]
|
||||
|
||||
lemma countSelf_contrIndexList_le_countSelf (I : Index 𝓒.Color) :
|
||||
l.contrIndexList.countSelf I ≤ l.countSelf I := by
|
||||
rw [countSelf_eq_filter_color_countP, countSelf_eq_filter_color_countP]
|
||||
apply List.Sublist.countP_le
|
||||
exact List.Sublist.filter _ (List.filter_sublist l.val)
|
||||
|
||||
lemma countSelf_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countSelf I = l.countSelf I := by
|
||||
rw [countSelf_eq_filter_id_countP,
|
||||
l.filter_id_contrIndexList_eq_of_countId_contrIndexList I h1,
|
||||
countSelf_eq_filter_id_countP]
|
||||
|
||||
@[simp]
|
||||
lemma countSelf_contrIndexList_get (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.countSelf l.contrIndexList.val[Fin.val i] = 1 := by
|
||||
refine countSelf_eq_one_of_countId_eq_one _ _ ?h1 ?hme
|
||||
· refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
· exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
|
||||
/-- The number of times the dual of an index `I` appears in an index list. -/
|
||||
def countDual (I : Index 𝓒.Color) : ℕ := l.val.countP (fun J => I.dual = J)
|
||||
|
||||
lemma countDual_eq_countSelf_Dual (I : Index 𝓒.Color) : l.countDual I = l.countSelf I.dual := by
|
||||
rw [countDual, countSelf]
|
||||
|
||||
lemma countDual_eq_filter_id_countP : l.countDual I =
|
||||
(l.val.filter (fun J => I.id = J.id)).countP (fun J => I.toColor = 𝓒.τ (J.toColor)) := by
|
||||
simp only [countDual]
|
||||
rw [List.countP_filter]
|
||||
apply List.countP_congr
|
||||
intro I' _
|
||||
simp only [Index.dual, Index.toColor, Index.id, Index.eq_iff_color_eq_and_id_eq, Bool.decide_and,
|
||||
Bool.and_eq_true, decide_eq_true_eq, and_congr_left_iff]
|
||||
intro _
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· rw [← h]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
· rw [h]
|
||||
exact (𝓒.τ_involutive _)
|
||||
|
||||
lemma countDual_of_countId_zero {I : Index 𝓒.Color} (h : l.countId I = 0) :
|
||||
l.countDual I = 0 := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
simp [h, countDual_eq_filter_id_countP]
|
||||
|
||||
@[simp]
|
||||
lemma countDual_append (I : Index 𝓒.Color) :
|
||||
(l ++ l2).countDual I = countDual l I + countDual l2 I := by
|
||||
simp [countDual]
|
||||
|
||||
lemma countDual_contrIndexList_eq_of_countId_eq
|
||||
(h1 : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.countDual I = l.countDual I := by
|
||||
rw [countDual_eq_countSelf_Dual, countDual_eq_countSelf_Dual]
|
||||
refine countSelf_contrIndexList_eq_of_countId_eq l h1
|
||||
|
||||
lemma countSelf_eq_countDual_iff_of_isSome (hl : l.OnlyUniqueDuals)
|
||||
(i : Fin l.length) (hi : (l.getDual? i).isSome) :
|
||||
l.countSelf (l.val.get i) = l.countDual (l.val.get i) ↔
|
||||
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
∨ (l.colorMap i) = 𝓒.τ (l.colorMap i) := by
|
||||
rw [countSelf_eq_filter_id_countP, countDual_eq_filter_id_countP]
|
||||
have hi1 := hi
|
||||
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
|
||||
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
|
||||
all_goals
|
||||
erw [hf]
|
||||
simp only [List.countP, List.countP.go, List.get_eq_getElem, decide_True, zero_add,
|
||||
Nat.reduceAdd, cond_true]
|
||||
by_cases hn : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
|
||||
· simp only [hn, true_or, iff_true]
|
||||
have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= true := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp only [cond_true]
|
||||
have hτ : l.colorMap ((l.getDual? i).get hi) = 𝓒.τ (l.colorMap i) := by
|
||||
rw [hn]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
simp only [colorMap, List.get_eq_getElem] at hτ
|
||||
erw [hτ]
|
||||
· have hn' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑((l.getDual? i).get hi)].toColor) =
|
||||
false := by simpa [colorMap] using hn
|
||||
erw [hn']
|
||||
simp only [cond_false, hn, false_or]
|
||||
by_cases hm : l.colorMap i = 𝓒.τ (l.colorMap i)
|
||||
· trans True
|
||||
· simp only [iff_true]
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = true := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_true]
|
||||
have hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simp only [Fin.getElem_fin, decide_eq_false_iff_not]
|
||||
simp only [colorMap, List.get_eq_getElem] at hm
|
||||
erw [hm]
|
||||
by_contra hn'
|
||||
have hn'' : l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi)) := by
|
||||
simp only [colorMap, List.get_eq_getElem]
|
||||
rw [← hn']
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
exact hn hn''
|
||||
erw [hm'']
|
||||
rfl
|
||||
· rw [true_iff]
|
||||
exact hm
|
||||
· simp only [hm, iff_false, ne_eq]
|
||||
simp only [colorMap, List.get_eq_getElem] at hm
|
||||
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = false := by simpa using hm
|
||||
erw [hm']
|
||||
simp only [cond_false, ne_eq]
|
||||
by_cases hm'' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor) = true
|
||||
· erw [hm'']
|
||||
exact Nat.add_one_ne_zero 1
|
||||
· have hm''' : decide (l.val[↑i].toColor = l.val[↑((l.getDual? i).get hi)].toColor)
|
||||
= false := by
|
||||
simpa using hm''
|
||||
erw [hm''']
|
||||
exact Nat.add_one_ne_zero 0
|
||||
|
||||
/-- The condition an index and its' dual, when it exists, have dual colors. -/
|
||||
def ColorCond : Prop := Option.map l.colorMap ∘
|
||||
l.getDual? = Option.map (𝓒.τ ∘ l.colorMap) ∘
|
||||
Option.guard fun i => (l.getDual? i).isSome
|
||||
|
||||
namespace ColorCond
|
||||
|
||||
variable {l l2 l3 : IndexList 𝓒.Color}
|
||||
|
||||
omit [DecidableEq 𝓒.Color] in
|
||||
lemma iff_withDual :
|
||||
l.ColorCond ↔ ∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· have h' := congrFun h i
|
||||
simp only [Function.comp_apply] at h'
|
||||
rw [show l.getDual? i = some ((l.getDual? i).get (l.withDual_isSome i)) by simp] at h'
|
||||
have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp [l.withDual_isSome i]
|
||||
rw [h'', Option.map_some', Option.map_some'] at h'
|
||||
simp only [Function.comp_apply, Option.some.injEq] at h'
|
||||
rw [h']
|
||||
exact 𝓒.τ_involutive (l.colorMap i)
|
||||
· funext i
|
||||
by_cases hi : (l.getDual? i).isSome
|
||||
· have h'' : (Option.guard (fun i => (l.getDual? i).isSome = true) ↑i) = i := by
|
||||
apply Option.guard_eq_some.mpr
|
||||
simp only [true_and]
|
||||
exact hi
|
||||
simp only [Function.comp_apply, h'', Option.map_some']
|
||||
rw [Option.eq_some_of_isSome hi, Option.map_some']
|
||||
simp only [Option.some.injEq]
|
||||
have hii := h ⟨i, (mem_withDual_iff_isSome l i).mpr hi⟩
|
||||
simp only at hii
|
||||
rw [← hii]
|
||||
exact (𝓒.τ_involutive _).symm
|
||||
· simp only [Function.comp_apply, Option.guard, hi, Bool.false_eq_true, ↓reduceIte,
|
||||
Option.map_none', Option.map_eq_none']
|
||||
exact Option.not_isSome_iff_eq_none.mp hi
|
||||
|
||||
omit [DecidableEq 𝓒.Color] in
|
||||
lemma iff_on_isSome : l.ColorCond ↔ ∀ (i : Fin l.length) (h : (l.getDual? i).isSome), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get h)) = l.colorMap i := by
|
||||
rw [iff_withDual]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome]
|
||||
|
||||
/-- A condition on an index list `l` and and index `I`. If the id of `I` appears
|
||||
twice in `l` (and `I` at least once) then this condition is equivalent to the dual of `I` having
|
||||
dual color to `I`, but written totally in terms of lists. -/
|
||||
@[simp]
|
||||
abbrev countColorCond (l : IndexList 𝓒.Color) (I : Index 𝓒.Color) : Prop :=
|
||||
l.countColorQuot I = l.countId I ∧
|
||||
l.countSelf I = l.countDual I
|
||||
|
||||
lemma countColorCond_of_filter_eq (l l2 : IndexList 𝓒.Color) {I : Index 𝓒.Color}
|
||||
(hf : l.val.filter (fun J => I.id = J.id) = l2.val.filter (fun J => I.id = J.id))
|
||||
(h1 : countColorCond l I) : countColorCond l2 I := by
|
||||
rw [countColorCond, countColorQuot_eq_filter_id_countP, countId_eq_length_filter,
|
||||
countSelf_eq_filter_id_countP, countDual_eq_filter_id_countP, ← hf]
|
||||
rw [countColorCond, countColorQuot_eq_filter_id_countP, countId_eq_length_filter,
|
||||
countSelf_eq_filter_id_countP, countDual_eq_filter_id_countP] at h1
|
||||
exact h1
|
||||
|
||||
lemma iff_countColorCond_isSome (hl : l.OnlyUniqueDuals) : l.ColorCond ↔
|
||||
∀ (i : Fin l.length) (_ : (l.getDual? i).isSome), countColorCond l (l.val.get i) := by
|
||||
rw [iff_on_isSome]
|
||||
simp only [countColorCond]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [l.countColorQuot_eq_countId_iff_of_isSome hl i hi,
|
||||
l.countSelf_eq_countDual_iff_of_isSome hl i hi]
|
||||
have hi' := h i hi
|
||||
exact And.intro (Or.inr hi'.symm) (Or.inl hi'.symm)
|
||||
· have hi' := h i hi
|
||||
rw [l.countColorQuot_eq_countId_iff_of_isSome hl i hi,
|
||||
l.countSelf_eq_countDual_iff_of_isSome hl i hi] at hi'
|
||||
rcases hi'.1 with hi1 | hi1
|
||||
<;> rcases hi'.2 with hi2 | hi2
|
||||
· exact hi2.symm
|
||||
· rw [← hi1]
|
||||
exact hi2.symm
|
||||
· exact hi1.symm
|
||||
· exact hi1.symm
|
||||
|
||||
lemma iff_countColorCond_index (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ (i : Fin l.length), l.countId (l.val.get i) = 2
|
||||
→ countColorCond l (l.val.get i) := by
|
||||
rw [iff_countColorCond_isSome hl]
|
||||
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
|
||||
· rw [← mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i (mem_withUniqueDual_isSome l i hi)
|
||||
· rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
exact h i hi
|
||||
|
||||
lemma iff_countColorCond_mem (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ (I : Index 𝓒.Color) (_ : I ∈ l.val),
|
||||
l.countId I = 2 → countColorCond l I := by
|
||||
rw [iff_countColorCond_index hl]
|
||||
refine Iff.intro (fun h I hI hi => ?_) (fun h i hi => ?_)
|
||||
· let i := l.val.indexOf I
|
||||
have hi' : i < l.length := List.indexOf_lt_length.mpr hI
|
||||
have hIi : I = l.val.get ⟨i, hi'⟩ := (List.indexOf_get hi').symm
|
||||
rw [hIi] at hi ⊢
|
||||
exact h ⟨i, hi'⟩ hi
|
||||
· exact h (l.val.get i) (List.getElem_mem l.val (↑i) i.isLt) hi
|
||||
|
||||
lemma mem_iff_dual_mem (hl : l.OnlyUniqueDuals) (hc : l.ColorCond) (I : Index 𝓒.Color)
|
||||
(hId : l.countId I = 2) : I ∈ l.val ↔ I.dual ∈ l.val := by
|
||||
rw [iff_countColorCond_mem hl] at hc
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· have hc' := hc I h hId
|
||||
simp only [countColorCond] at hc'
|
||||
rw [← countSelf_neq_zero] at h
|
||||
rw [← countSelf_neq_zero, ← countDual_eq_countSelf_Dual]
|
||||
omega
|
||||
· have hIdd : l.countId I.dual = 2 := by
|
||||
rw [← hId]
|
||||
rfl
|
||||
have hc' := (hc I.dual h hIdd).2
|
||||
rw [← countSelf_neq_zero] at h ⊢
|
||||
rw [countDual_eq_countSelf_Dual] at hc'
|
||||
simp only [Index.dual_dual] at hc'
|
||||
exact Ne.symm (ne_of_ne_of_eq (id (Ne.symm h)) hc')
|
||||
|
||||
lemma iff_countColorCond (hl : l.OnlyUniqueDuals) :
|
||||
l.ColorCond ↔ ∀ I, l.countSelf I ≠ 0 → l.countId I = 2 → countColorCond l I := by
|
||||
refine Iff.intro (fun h I hIs hi => ?_) (fun h => ?_)
|
||||
· rw [countSelf_neq_zero] at hIs
|
||||
rw [iff_countColorCond_mem hl] at h
|
||||
exact h I hIs hi
|
||||
· rw [iff_countColorCond_mem hl]
|
||||
intro I hmem hi
|
||||
refine h I ?_ hi
|
||||
rw [countSelf_neq_zero]
|
||||
exact hmem
|
||||
|
||||
omit [DecidableEq 𝓒.Color] in
|
||||
lemma assoc (h : ColorCond (l ++ l2 ++ l3)) : ColorCond (l ++ (l2 ++ l3)) := by
|
||||
rw [← append_assoc]
|
||||
exact h
|
||||
|
||||
lemma symm (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) :
|
||||
ColorCond (l2 ++ l) := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.symm' hl)]
|
||||
intro I hI1 hI2
|
||||
have hI' := h I (by simp_all; omega) (by simp_all; omega)
|
||||
simp_all
|
||||
omega
|
||||
|
||||
lemma inl (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) : ColorCond l := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.inl hl)]
|
||||
intro I hI1 hI2
|
||||
have hI2' : l2.countId I = 0 := by
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at hl
|
||||
have hlI := hl I
|
||||
simp only [countId_append] at hlI
|
||||
omega
|
||||
have hI' := h I (by
|
||||
simp only [countSelf_append, ne_eq, add_eq_zero, not_and, hI1, false_implies])
|
||||
(by simp_all)
|
||||
simp only [countColorCond, countColorQuot_append, countId_append, countSelf_append,
|
||||
countDual_append] at hI'
|
||||
rw [l2.countColorQuot_of_countId_zero hI2', l2.countSelf_of_countId_zero hI2',
|
||||
l2.countDual_of_countId_zero hI2', hI2'] at hI'
|
||||
exact hI'
|
||||
|
||||
lemma inr (hl : (l ++ l2).OnlyUniqueDuals) (h : ColorCond (l ++ l2)) : ColorCond l2 :=
|
||||
inl (OnlyUniqueDuals.symm.mp hl) (symm hl h)
|
||||
|
||||
lemma swap (hl : (l ++ l2 ++ l3).OnlyUniqueDuals) (h : ColorCond (l ++ l2 ++ l3)) :
|
||||
ColorCond (l2 ++ l ++ l3) := by
|
||||
rw [iff_countColorCond hl] at h
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.swap hl)]
|
||||
intro I hI1 hI2
|
||||
have hI' := h I (by simp_all) (by simp_all; omega)
|
||||
simp_all only [countSelf_append, ne_eq, add_eq_zero, not_and, and_imp, countId_append,
|
||||
countColorCond, countColorQuot_append, countDual_append, not_false_eq_true, implies_true]
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## Contractions
|
||||
|
||||
-/
|
||||
|
||||
omit [DecidableEq 𝓒.Color] in
|
||||
lemma contrIndexList : ColorCond l.contrIndexList := by
|
||||
funext i
|
||||
simp [Option.guard]
|
||||
|
||||
lemma contrIndexList_left (hl : (l ++ l2).OnlyUniqueDuals) (h1 : (l ++ l2).ColorCond) :
|
||||
ColorCond (l.contrIndexList ++ l2) := by
|
||||
rw [iff_countColorCond hl] at h1
|
||||
rw [iff_countColorCond (OnlyUniqueDuals.contrIndexList_left hl)]
|
||||
intro I hI1 hI2
|
||||
simp only [countSelf_append, ne_eq] at hI1
|
||||
have hc := countSelf_contrIndexList_le_countSelf l I
|
||||
have h2 := (countId_eq_two_ofcontrIndexList_left_of_OnlyUniqueDuals l l2 hl I hI2)
|
||||
have hI1' := h1 I (by simp_all; omega) h2
|
||||
have hIdEq : l.contrIndexList.countId I = l.countId I := by
|
||||
simp only [countId_append] at h2 hI2
|
||||
omega
|
||||
simp only [countColorCond, countColorQuot_append, countId_append, countSelf_append,
|
||||
countDual_append]
|
||||
rw [l.countColorQuot_contrIndexList_eq_of_countId_eq hIdEq,
|
||||
l.countSelf_contrIndexList_eq_of_countId_eq hIdEq,
|
||||
l.countDual_contrIndexList_eq_of_countId_eq hIdEq, hIdEq]
|
||||
simpa using hI1'
|
||||
|
||||
/-!
|
||||
|
||||
## Bool
|
||||
|
||||
-/
|
||||
/-- A bool which is true for an index list if and only if every index and its' dual, when it exists,
|
||||
have dual colors. -/
|
||||
def bool (l : IndexList 𝓒.Color) : Bool :=
|
||||
if (∀ (i : l.withDual), 𝓒.τ
|
||||
(l.colorMap ((l.getDual? i).get (l.withDual_isSome i))) = l.colorMap i) then
|
||||
true
|
||||
else false
|
||||
|
||||
lemma iff_bool : l.ColorCond ↔ bool l := by
|
||||
rw [iff_withDual, bool]
|
||||
simp only [Subtype.forall, mem_withDual_iff_isSome, Bool.if_false_right, Bool.and_true,
|
||||
decide_eq_true_eq]
|
||||
|
||||
end ColorCond
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,343 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Equivs
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Contraction of an index list.
|
||||
|
||||
In this file we define the contraction of an index list `l` to be the index list formed by
|
||||
by the subset of indices of `l` which do not have a dual in `l`.
|
||||
|
||||
For example, the contraction of the index list `['ᵘ¹', 'ᵘ²', 'ᵤ₁', 'ᵘ¹']` is the index list
|
||||
`['ᵘ²']`.
|
||||
|
||||
We also define the following finite sets
|
||||
- `l.withoutDual` the finite set of indices of `l` which do not have a dual in `l`.
|
||||
- `l.withUniqueDualLT` the finite set of those indices of `l` which have a unique dual, and
|
||||
for which that dual is greater-then (determined by the ordering on `Fin`) then the index itself.
|
||||
- `l.withUniqueDualGT` the finite set of those indices of `l` which have a unique dual, and
|
||||
for which that dual is less-then (determined by the ordering on `Fin`) then the index itself.
|
||||
|
||||
We define an equivalence `l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual`.
|
||||
|
||||
We prove properties related to when `l.withUniqueDualInOther l2 = Finset.univ` for two
|
||||
index lists `l` and `l2`.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type}
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-- The index list formed from `l` by selecting only those indices in `l` which
|
||||
do not have a dual. -/
|
||||
def contrIndexList : IndexList X where
|
||||
val := l.val.filter (fun I => l.countId I = 1)
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_empty : (⟨[]⟩ : IndexList X).contrIndexList = { val := [] } := by
|
||||
rfl
|
||||
|
||||
lemma contrIndexList_val : l.contrIndexList.val =
|
||||
l.val.filter (fun I => l.countId I = 1) := by
|
||||
rfl
|
||||
|
||||
/-- An alternative form of the contracted index list. -/
|
||||
@[simp]
|
||||
def contrIndexList' : IndexList X where
|
||||
val := List.ofFn (l.val.get ∘ Subtype.val ∘ l.withoutDualEquiv)
|
||||
|
||||
lemma withoutDual_sort_eq_filter : l.withoutDual.sort (fun i j => i ≤ j) =
|
||||
(List.finRange l.length).filter (fun i => i ∈ l.withoutDual) := by
|
||||
rw [withoutDual]
|
||||
rw [← filter_sort_comm]
|
||||
simp only [Option.isNone_iff_eq_none, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply congrArg
|
||||
exact Fin.sort_univ l.length
|
||||
|
||||
lemma contrIndexList_eq_contrIndexList' : l.contrIndexList = l.contrIndexList' := by
|
||||
apply ext
|
||||
simp only [contrIndexList']
|
||||
trans List.map l.val.get (List.ofFn (Subtype.val ∘ ⇑l.withoutDualEquiv))
|
||||
· rw [list_ofFn_withoutDualEquiv_eq_sort, withoutDual_sort_eq_filter]
|
||||
simp only [contrIndexList]
|
||||
let f1 : Index X → Bool := fun (I : Index X) => l.countId I = 1
|
||||
let f2 : (Fin l.length) → Bool := fun i => i ∈ l.withoutDual
|
||||
change List.filter f1 l.val = List.map l.val.get (List.filter f2 (List.finRange l.length))
|
||||
have hf : f2 = f1 ∘ l.val.get := by
|
||||
funext i
|
||||
simp only [mem_withoutDual_iff_countId_eq_one l, List.get_eq_getElem, Function.comp_apply, f2,
|
||||
f1]
|
||||
rw [hf, ← List.filter_map]
|
||||
apply congrArg
|
||||
simp [length]
|
||||
· exact List.map_ofFn (Subtype.val ∘ ⇑l.withoutDualEquiv) l.val.get
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_length : l.contrIndexList.length = l.withoutDual.card := by
|
||||
simp [contrIndexList_eq_contrIndexList', withoutDual, length]
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_idMap (i : Fin l.contrIndexList.length) : l.contrIndexList.idMap i
|
||||
= l.idMap (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1 := by
|
||||
simp only [idMap, List.get_eq_getElem, contrIndexList_eq_contrIndexList', contrIndexList',
|
||||
List.getElem_ofFn, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_colorMap (i : Fin l.contrIndexList.length) : l.contrIndexList.colorMap i
|
||||
= l.colorMap (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1 := by
|
||||
simp only [colorMap, List.get_eq_getElem, contrIndexList_eq_contrIndexList', contrIndexList',
|
||||
List.getElem_ofFn, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
lemma contrIndexList_areDualInSelf (i j : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.AreDualInSelf i j ↔
|
||||
l.AreDualInSelf (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).1
|
||||
(l.withoutDualEquiv (Fin.cast l.contrIndexList_length j)).1 := by
|
||||
simp only [AreDualInSelf, ne_eq, contrIndexList_idMap, and_congr_left_iff]
|
||||
intro _
|
||||
trans ¬ (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)) =
|
||||
(l.withoutDualEquiv (Fin.cast l.contrIndexList_length j))
|
||||
· rw [l.withoutDualEquiv.apply_eq_iff_eq]
|
||||
simp [Fin.ext_iff]
|
||||
· exact Iff.symm Subtype.coe_ne_coe
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_getDual? (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.getDual? i = none := by
|
||||
rw [← Option.not_isSome_iff_eq_none, ← mem_withDual_iff_isSome, mem_withDual_iff_exists]
|
||||
simp only [contrIndexList_areDualInSelf, not_exists]
|
||||
have h1 := (l.withoutDualEquiv (Fin.cast l.contrIndexList_length i)).2
|
||||
have h1' := Finset.disjoint_right.mp l.withDual_disjoint_withoutDual h1
|
||||
rw [mem_withDual_iff_exists] at h1'
|
||||
exact fun x => forall_not_of_not_exists h1'
|
||||
↑(l.withoutDualEquiv (Fin.cast (contrIndexList_length l) x))
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_withDual : l.contrIndexList.withDual = ∅ := by
|
||||
rw [Finset.eq_empty_iff_forall_not_mem]
|
||||
intro x
|
||||
simp [withDual]
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_withUniqueDual : l.contrIndexList.withUniqueDual = ∅ := by
|
||||
rw [withUniqueDual]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_areDualInSelf_false (i j : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.AreDualInSelf i j ↔ False := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => False.elim h)
|
||||
have h1 : i ∈ l.contrIndexList.withDual := by
|
||||
rw [@mem_withDual_iff_exists]
|
||||
exact Exists.intro j h
|
||||
simp_all
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_of_withDual_empty (h : l.withDual = ∅) : l.contrIndexList = l := by
|
||||
have h1 := l.withDual_union_withoutDual
|
||||
rw [h, Finset.empty_union] at h1
|
||||
apply ext
|
||||
rw [@List.ext_get_iff]
|
||||
change l.contrIndexList.length = l.length ∧ _
|
||||
rw [contrIndexList_length, h1]
|
||||
simp only [Finset.card_univ, Fintype.card_fin, List.get_eq_getElem, true_and]
|
||||
intro n h1' h2
|
||||
simp only [contrIndexList_eq_contrIndexList', contrIndexList', List.getElem_ofFn,
|
||||
Function.comp_apply, List.get_eq_getElem]
|
||||
congr
|
||||
simp only [withoutDualEquiv, RelIso.coe_fn_toEquiv, Finset.coe_orderIsoOfFin_apply]
|
||||
simp only [h1]
|
||||
rw [(Finset.orderEmbOfFin_unique' _
|
||||
(fun x => Finset.mem_univ ((Fin.castOrderIso _).toOrderEmbedding x))).symm]
|
||||
· exact Eq.symm (Nat.add_zero n)
|
||||
· rw [h1]
|
||||
exact Finset.card_fin l.length
|
||||
|
||||
lemma contrIndexList_contrIndexList : l.contrIndexList.contrIndexList = l.contrIndexList := by
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma contrIndexList_getDualInOther?_self (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.getDualInOther? l.contrIndexList i = some i := by
|
||||
simp only [getDualInOther?]
|
||||
rw [@Fin.find_eq_some_iff]
|
||||
simp only [AreDualInOther, contrIndexList_idMap, true_and]
|
||||
intro j hj
|
||||
have h1 : i = j := by
|
||||
by_contra hn
|
||||
have h : l.contrIndexList.AreDualInSelf i j := by
|
||||
simp only [AreDualInSelf]
|
||||
simp only [ne_eq, contrIndexList_idMap, hj, and_true]
|
||||
exact hn
|
||||
exact (contrIndexList_areDualInSelf_false l i j).mp h
|
||||
exact Fin.ge_of_eq (id (Eq.symm h1))
|
||||
|
||||
lemma cons_contrIndexList_of_countId_eq_zero {I : Index X}
|
||||
(h : l.countId I = 0) :
|
||||
(l.cons I).contrIndexList = l.contrIndexList.cons I := by
|
||||
apply ext
|
||||
simp only [contrIndexList, countId, cons_val]
|
||||
rw [List.filter_cons_of_pos]
|
||||
· apply congrArg
|
||||
apply List.filter_congr
|
||||
intro J hJ
|
||||
simp only [decide_eq_decide]
|
||||
rw [countId, List.countP_eq_zero] at h
|
||||
simp only [decide_eq_true_eq] at h
|
||||
rw [List.countP_cons_of_neg]
|
||||
simp only [decide_eq_true_eq]
|
||||
exact fun a => h J hJ (id (Eq.symm a))
|
||||
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
|
||||
exact h
|
||||
|
||||
lemma cons_contrIndexList_of_countId_neq_zero {I : Index X} (h : l.countId I ≠ 0) :
|
||||
(l.cons I).contrIndexList.val = l.contrIndexList.val.filter (fun J => ¬ I.id = J.id) := by
|
||||
simp only [contrIndexList, countId, cons_val, decide_not, List.filter_filter, Bool.not_eq_true',
|
||||
decide_eq_false_iff_not, decide_eq_true_eq, Bool.decide_and]
|
||||
rw [List.filter_cons_of_neg]
|
||||
· apply List.filter_congr
|
||||
intro J hJ
|
||||
by_cases hI : I.id = J.id
|
||||
· simp only [hI, decide_True, List.countP_cons_of_pos, add_left_eq_self, Bool.not_true,
|
||||
Bool.false_and, decide_eq_false_iff_not, countId]
|
||||
rw [countId, hI] at h
|
||||
exact h
|
||||
· simp only [hI, decide_False, Bool.not_false, Bool.true_and, decide_eq_decide]
|
||||
rw [List.countP_cons_of_neg]
|
||||
simp only [decide_eq_true_eq]
|
||||
exact fun a => hI (id (Eq.symm a))
|
||||
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
|
||||
exact h
|
||||
|
||||
lemma mem_contrIndexList_countId {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.countId I = 1 := by
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at h
|
||||
exact h.2
|
||||
|
||||
lemma mem_contrIndexList_filter {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.val.filter (fun J => I.id = J.id) = [I] := by
|
||||
have h1 : (l.val.filter (fun J => I.id = J.id)).length = 1 := by
|
||||
rw [← List.countP_eq_length_filter]
|
||||
exact l.mem_contrIndexList_countId h
|
||||
have h2 : I ∈ l.val.filter (fun J => I.id = J.id) := by
|
||||
simp only [List.mem_filter, decide_True, and_true]
|
||||
exact List.mem_of_mem_filter h
|
||||
rw [List.length_eq_one] at h1
|
||||
obtain ⟨J, hJ⟩ := h1
|
||||
rw [hJ] at h2
|
||||
simp only [List.mem_singleton] at h2
|
||||
subst h2
|
||||
exact hJ
|
||||
|
||||
lemma mem_contrIndexList_countId_contrIndexList {I : Index X} (h : I ∈ l.contrIndexList.val) :
|
||||
l.contrIndexList.countId I = 1 := by
|
||||
trans (l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
|
||||
· rw [contrIndexList]
|
||||
simp only [countId]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
congr
|
||||
funext a
|
||||
rw [Bool.and_comm]
|
||||
· rw [l.mem_contrIndexList_filter h, List.countP_cons_of_pos]
|
||||
· rfl
|
||||
· simp only [decide_eq_true_eq]
|
||||
exact mem_contrIndexList_countId l h
|
||||
|
||||
lemma countId_contrIndexList_zero_of_countId (I : Index X)
|
||||
(h : l.countId I = 0) : l.contrIndexList.countId I = 0 := by
|
||||
trans (l.val.filter (fun J => I.id = J.id)).countP
|
||||
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
|
||||
· rw [contrIndexList]
|
||||
simp only [countId]
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
simp only [decide_eq_true_eq, Bool.decide_and]
|
||||
congr
|
||||
funext a
|
||||
rw [Bool.and_comm]
|
||||
· rw [countId_eq_length_filter, List.length_eq_zero] at h
|
||||
rw [h]
|
||||
rfl
|
||||
|
||||
lemma countId_contrIndexList_le_one (I : Index X) :
|
||||
l.contrIndexList.countId I ≤ 1 := by
|
||||
by_cases h : l.contrIndexList.countId I = 0
|
||||
· exact StrictMono.minimal_preimage_bot (fun ⦃a b⦄ a => a) h 1
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contrIndexList I h
|
||||
rw [countId_congr l.contrIndexList hI2, mem_contrIndexList_countId_contrIndexList l hI1]
|
||||
|
||||
lemma countId_contrIndexList_eq_one_iff_countId_eq_one (I : Index X) :
|
||||
l.contrIndexList.countId I = 1 ↔ l.countId I = 1 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l.contrIndexList I (ne_zero_of_eq_one h)
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at hI1
|
||||
rw [countId_congr l hI2]
|
||||
exact hI1.2
|
||||
· obtain ⟨I', hI1, hI2⟩ := countId_neq_zero_mem l I (ne_zero_of_eq_one h)
|
||||
rw [countId_congr l hI2] at h
|
||||
rw [countId_congr _ hI2]
|
||||
refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq]
|
||||
exact ⟨hI1, h⟩
|
||||
|
||||
lemma countId_contrIndexList_le_countId (I : Index X) :
|
||||
l.contrIndexList.countId I ≤ l.countId I := by
|
||||
by_cases h : l.contrIndexList.countId I = 0
|
||||
· exact StrictMono.minimal_preimage_bot (fun ⦃a b⦄ a => a) h (l.countId I)
|
||||
· have ho : l.contrIndexList.countId I = 1 := by
|
||||
have h1 := l.countId_contrIndexList_le_one I
|
||||
omega
|
||||
rw [ho]
|
||||
rw [countId_contrIndexList_eq_one_iff_countId_eq_one] at ho
|
||||
rw [ho]
|
||||
|
||||
@[simp]
|
||||
lemma countId_contrIndexList_get (i : Fin l.contrIndexList.length) :
|
||||
l.contrIndexList.countId l.contrIndexList.val[Fin.val i] = 1 := by
|
||||
refine mem_contrIndexList_countId_contrIndexList l ?_
|
||||
exact List.getElem_mem l.contrIndexList.val (↑i) _
|
||||
|
||||
lemma filter_id_contrIndexList_eq_of_countId_contrIndexList (I : Index X)
|
||||
(h : l.contrIndexList.countId I = l.countId I) :
|
||||
l.contrIndexList.val.filter (fun J => I.id = J.id) =
|
||||
l.val.filter (fun J => I.id = J.id) := by
|
||||
apply List.Sublist.eq_of_length
|
||||
· rw [contrIndexList, List.filter_comm]
|
||||
exact List.filter_sublist (List.filter (fun J => decide (I.id = J.id)) l.val)
|
||||
· rw [← countId_eq_length_filter, h, countId_eq_length_filter]
|
||||
|
||||
lemma contrIndexList_append_eq_filter : (l ++ l2).contrIndexList.val =
|
||||
l.contrIndexList.val.filter (fun I => l2.countId I = 0) ++
|
||||
l2.contrIndexList.val.filter (fun I => l.countId I = 0) := by
|
||||
simp only [contrIndexList, countId_append, append_val, List.filter_append, List.filter_filter,
|
||||
decide_eq_true_eq, Bool.decide_and]
|
||||
congr 1
|
||||
· apply List.filter_congr
|
||||
intro I hI
|
||||
have hIz : l.countId I ≠ 0 := countId_mem l I hI
|
||||
have hx : l.countId I + l2.countId I = 1 ↔ (l2.countId I = 0 ∧ l.countId I = 1) := by
|
||||
omega
|
||||
simp only [hx, Bool.decide_and]
|
||||
· apply List.filter_congr
|
||||
intro I hI
|
||||
have hIz : l2.countId I ≠ 0 := countId_mem l2 I hI
|
||||
have hx : l.countId I + l2.countId I = 1 ↔ (l2.countId I = 1 ∧ l.countId I = 0) := by
|
||||
omega
|
||||
simp only [hx, Bool.decide_and]
|
||||
exact Bool.and_comm (decide (l2.countId I = 1)) (decide (l.countId I = 0))
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,569 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Duals
|
||||
/-!
|
||||
|
||||
# Counting ids
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-!
|
||||
|
||||
## countId
|
||||
|
||||
-/
|
||||
|
||||
/-- The number of times the id of an index `I` appears in a list of indices `l`. -/
|
||||
def countId (I : Index X) : ℕ :=
|
||||
l.val.countP (fun J => I.id = J.id)
|
||||
|
||||
/-!
|
||||
|
||||
## Basic properties
|
||||
|
||||
-/
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma countId_append (I : Index X) : (l ++ l2).countId I = l.countId I + l2.countId I := by
|
||||
simp [countId]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_eq_length_filter (I : Index X) :
|
||||
l.countId I = (l.val.filter (fun J => I.id = J.id)).length := by
|
||||
rw [countId, List.countP_eq_length_filter]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_index_neq_zero (i : Fin l.length) : l.countId (l.val.get i) ≠ 0 := by
|
||||
by_contra hn
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at hn
|
||||
refine (List.mem_nil_iff (l.val.get i)).mp ?_
|
||||
simpa [← hn] using List.getElem_mem l.val i.1 i.isLt
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_append_symm (I : Index X) : (l ++ l2).countId I = (l2 ++ l).countId I := by
|
||||
simp only [countId_append]
|
||||
exact Nat.add_comm (l.countId I) (l2.countId I)
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_eq_one_append_mem_right_self_eq_one {I : Index X} (hI : I ∈ l2.val)
|
||||
(h : (l ++ l2).countId I = 1) : l2.countId I = 1 := by
|
||||
simp only [countId_append] at h
|
||||
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
|
||||
simp [List.mem_filter, decide_True, and_true, hI]
|
||||
have h1 : l2.countId I ≠ 0 := by
|
||||
rw [countId_eq_length_filter]
|
||||
by_contra hn
|
||||
rw [@List.length_eq_zero] at hn
|
||||
rw [hn] at hmem
|
||||
exact (List.mem_nil_iff I).mp hmem
|
||||
omega
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_eq_one_append_mem_right_other_eq_zero {I : Index X} (hI : I ∈ l2.val)
|
||||
(h : (l ++ l2).countId I = 1) : l.countId I = 0 := by
|
||||
simp only [countId_append] at h
|
||||
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
|
||||
simp [List.mem_filter, decide_True, and_true, hI]
|
||||
have h1 : l2.countId I ≠ 0 := by
|
||||
rw [countId_eq_length_filter]
|
||||
by_contra hn
|
||||
rw [@List.length_eq_zero] at hn
|
||||
rw [hn] at hmem
|
||||
exact (List.mem_nil_iff I).mp hmem
|
||||
omega
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma countId_cons_eq_two {I : Index X} :
|
||||
(l.cons I).countId I = 2 ↔ l.countId I = 1 := by
|
||||
simp [countId]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_congr {I J : Index X} (h : I.id = J.id) : l.countId I = l.countId J := by
|
||||
simp [countId, h]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_neq_zero_mem (I : Index X) (h : l.countId I ≠ 0) :
|
||||
∃ I', I' ∈ l.val ∧ I.id = I'.id := by
|
||||
rw [countId_eq_length_filter] at h
|
||||
have h' := List.isEmpty_iff_length_eq_zero.mp.mt h
|
||||
simp only at h'
|
||||
have h'' := eq_false_of_ne_true h'
|
||||
rw [List.isEmpty_false_iff_exists_mem] at h''
|
||||
obtain ⟨I', hI'⟩ := h''
|
||||
simp only [List.mem_filter, decide_eq_true_eq] at hI'
|
||||
exact ⟨I', hI'⟩
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_mem (I : Index X) (hI : I ∈ l.val) : l.countId I ≠ 0 := by
|
||||
rw [countId_eq_length_filter]
|
||||
by_contra hn
|
||||
rw [List.length_eq_zero] at hn
|
||||
have hIme : I ∈ List.filter (fun J => decide (I.id = J.id)) l.val := by
|
||||
simp [hI]
|
||||
rw [hn] at hIme
|
||||
exact (List.mem_nil_iff I).mp hIme
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_get_other (i : Fin l.length) : l2.countId (l.val.get i) =
|
||||
(List.finRange l2.length).countP (fun j => l.AreDualInOther l2 i j) := by
|
||||
rw [countId_eq_length_filter]
|
||||
rw [List.countP_eq_length_filter]
|
||||
have hl2 : l2.val = List.map l2.val.get (List.finRange l2.length) := by
|
||||
simp only [length, List.finRange_map_get]
|
||||
nth_rewrite 1 [hl2]
|
||||
rw [List.filter_map, List.length_map]
|
||||
rfl
|
||||
|
||||
/-! TODO: Replace with mathlib lemma. -/
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma filter_finRange (i : Fin l.length) :
|
||||
List.filter (fun j => i = j) (List.finRange l.length) = [i] := by
|
||||
have h3 : (List.filter (fun j => i = j) (List.finRange l.length)).length = 1 := by
|
||||
rw [← List.countP_eq_length_filter]
|
||||
trans List.count i (List.finRange l.length)
|
||||
· rw [List.count]
|
||||
apply List.countP_congr (fun j _ => ?_)
|
||||
simp only [decide_eq_true_eq, beq_iff_eq]
|
||||
exact eq_comm
|
||||
· exact List.nodup_iff_count_eq_one.mp (List.nodup_finRange l.length) _ (List.mem_finRange i)
|
||||
have h4 : i ∈ List.filter (fun j => i = j) (List.finRange l.length) := by
|
||||
simp
|
||||
rw [@List.length_eq_one] at h3
|
||||
obtain ⟨a, ha⟩ := h3
|
||||
rw [ha] at h4
|
||||
simp only [List.mem_singleton] at h4
|
||||
subst h4
|
||||
exact ha
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_get (i : Fin l.length) : l.countId (l.val.get i) =
|
||||
(List.finRange l.length).countP (fun j => l.AreDualInSelf i j) + 1 := by
|
||||
rw [countId_get_other l l]
|
||||
have h1 : (List.finRange l.length).countP (fun j => l.AreDualInSelf i j)
|
||||
= ((List.finRange l.length).filter (fun j => l.AreDualInOther l i j)).countP
|
||||
(fun j => ¬ i = j) := by
|
||||
rw [List.countP_filter]
|
||||
refine List.countP_congr ?_
|
||||
intro j _
|
||||
simp [AreDualInSelf, AreDualInOther]
|
||||
rw [h1]
|
||||
have h1 := List.length_eq_countP_add_countP (fun j => i = j)
|
||||
((List.finRange l.length).filter (fun j => l.AreDualInOther l i j))
|
||||
have h2 : List.countP (fun j => i = j)
|
||||
(List.filter (fun j => l.AreDualInOther l i j) (List.finRange l.length)) =
|
||||
List.countP (fun j => l.AreDualInOther l i j)
|
||||
(List.filter (fun j => i = j) (List.finRange l.length)) := by
|
||||
rw [List.countP_filter, List.countP_filter]
|
||||
refine List.countP_congr (fun j _ => ?_)
|
||||
simpa using And.comm
|
||||
have ha := l.filter_finRange
|
||||
rw [ha] at h2
|
||||
rw [h2] at h1
|
||||
rw [List.countP_eq_length_filter, h1, add_comm]
|
||||
simp only [decide_eq_true_eq, decide_not, add_right_inj]
|
||||
simp [List.countP, List.countP.go, AreDualInOther]
|
||||
|
||||
/-!
|
||||
|
||||
## Duals and countId
|
||||
|
||||
-/
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_gt_zero_of_mem_withDual (i : Fin l.length) (h : i ∈ l.withDual) :
|
||||
1 < l.countId (l.val.get i) := by
|
||||
rw [countId_get]
|
||||
by_contra hn
|
||||
simp only [Nat.lt_add_left_iff_pos, not_lt, Nat.le_zero_eq] at hn
|
||||
rw [List.countP_eq_length_filter, List.length_eq_zero] at hn
|
||||
rw [mem_withDual_iff_exists] at h
|
||||
obtain ⟨j, hj⟩ := h
|
||||
have hjmem : j ∈ (List.finRange l.length).filter (fun j => decide (l.AreDualInSelf i j)) := by
|
||||
simpa using hj
|
||||
rw [hn] at hjmem
|
||||
exact (List.mem_nil_iff j).mp hjmem
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_of_not_mem_withDual (i : Fin l.length)(h : i ∉ l.withDual) :
|
||||
l.countId (l.val.get i) = 1 := by
|
||||
rw [countId_get]
|
||||
simp only [add_left_eq_self]
|
||||
rw [List.countP_eq_length_filter]
|
||||
simp only [List.length_eq_zero]
|
||||
rw [List.filter_eq_nil]
|
||||
simp only [List.mem_finRange, decide_eq_true_eq, true_implies]
|
||||
rw [mem_withDual_iff_exists] at h
|
||||
simpa using h
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma mem_withDual_iff_countId_gt_one (i : Fin l.length) :
|
||||
i ∈ l.withDual ↔ 1 < l.countId (l.val.get i) := by
|
||||
refine Iff.intro (fun h => countId_gt_zero_of_mem_withDual l i h) (fun h => ?_)
|
||||
by_contra hn
|
||||
have hn' := countId_of_not_mem_withDual l i hn
|
||||
omega
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_neq_zero_of_mem_withDualInOther (i : Fin l.length) (h : i ∈ l.withDualInOther l2) :
|
||||
l2.countId (l.val.get i) ≠ 0 := by
|
||||
rw [mem_withInDualOther_iff_exists] at h
|
||||
rw [countId_eq_length_filter]
|
||||
by_contra hn
|
||||
rw [List.length_eq_zero] at hn
|
||||
obtain ⟨j, hj⟩ := h
|
||||
have hjmem : l2.val.get j ∈ List.filter (fun J => decide ((l.val.get i).id = J.id)) l2.val := by
|
||||
simp only [List.get_eq_getElem, List.mem_filter, decide_eq_true_eq]
|
||||
apply And.intro
|
||||
· exact List.getElem_mem l2.val (↑j) j.isLt
|
||||
· simpa [AreDualInOther] using hj
|
||||
rw [hn] at hjmem
|
||||
exact (List.mem_nil_iff (l2.val.get j)).mp hjmem
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma countId_of_not_mem_withDualInOther (i : Fin l.length) (h : i ∉ l.withDualInOther l2) :
|
||||
l2.countId (l.val.get i) = 0 := by
|
||||
by_contra hn
|
||||
rw [countId_eq_length_filter] at hn
|
||||
rw [← List.isEmpty_iff_length_eq_zero] at hn
|
||||
have hx := eq_false_of_ne_true hn
|
||||
rw [List.isEmpty_false_iff_exists_mem] at hx
|
||||
obtain ⟨j, hj⟩ := hx
|
||||
have hjmem : j ∈ l2.val := List.mem_of_mem_filter hj
|
||||
have hj' : l2.val.indexOf j < l2.length := List.indexOf_lt_length.mpr hjmem
|
||||
rw [mem_withInDualOther_iff_exists] at h
|
||||
simp only [not_exists] at h
|
||||
have hj' := h ⟨l2.val.indexOf j, hj'⟩
|
||||
simp only [AreDualInOther, idMap, List.get_eq_getElem, List.getElem_indexOf] at hj'
|
||||
simp only [List.get_eq_getElem, List.mem_filter, decide_eq_true_eq] at hj
|
||||
simp_all only [List.get_eq_getElem, List.isEmpty_eq_true, List.getElem_indexOf, not_true_eq_false]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma mem_withDualInOther_iff_countId_neq_zero (i : Fin l.length) :
|
||||
i ∈ l.withDualInOther l2 ↔ l2.countId (l.val.get i) ≠ 0 := by
|
||||
refine Iff.intro (fun h => countId_neq_zero_of_mem_withDualInOther l l2 i h)
|
||||
(fun h => ?_)
|
||||
by_contra hn
|
||||
have hn' := countId_of_not_mem_withDualInOther l l2 i hn
|
||||
exact h hn'
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma mem_withoutDual_iff_countId_eq_one (i : Fin l.length) :
|
||||
i ∈ l.withoutDual ↔ l.countId (l.val.get i) = 1 := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· exact countId_of_not_mem_withDual l i (l.not_mem_withDual_of_mem_withoutDual i h)
|
||||
· by_contra hn
|
||||
have h : i ∈ l.withDual := by
|
||||
simp only [withoutDual, Option.isNone_iff_eq_none, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and] at hn
|
||||
simpa using Option.ne_none_iff_isSome.mp hn
|
||||
rw [mem_withDual_iff_countId_gt_one] at h
|
||||
omega
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_eq_two_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
|
||||
l.countId (l.val.get i) = 2 := by
|
||||
rw [countId_get]
|
||||
simp only [Nat.reduceEqDiff]
|
||||
let i' := (l.getDual? i).get (mem_withUniqueDual_isSome l i h)
|
||||
have h1 : [i'] = (List.finRange l.length).filter (fun j => (l.AreDualInSelf i j)) := by
|
||||
trans List.filter (fun j => (l.AreDualInSelf i j)) [i']
|
||||
· simp [List.filter, i']
|
||||
trans List.filter (fun j => (l.AreDualInSelf i j))
|
||||
((List.finRange l.length).filter (fun j => j = i'))
|
||||
· apply congrArg
|
||||
rw [← filter_finRange l i']
|
||||
apply List.filter_congr (fun j _ => ?_)
|
||||
simpa using eq_comm
|
||||
trans List.filter (fun j => j = i')
|
||||
((List.finRange l.length).filter fun j => l.AreDualInSelf i j)
|
||||
· exact List.filter_comm (fun j => l.AreDualInSelf i j) (fun j => j = i')
|
||||
(List.finRange l.length)
|
||||
· simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and]
|
||||
refine List.filter_congr (fun j _ => ?_)
|
||||
simp only [Bool.and_iff_right_iff_imp, decide_eq_true_eq]
|
||||
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and] at h
|
||||
intro hj
|
||||
have hj' := h.2 j hj
|
||||
apply Option.some_injective
|
||||
rw [hj']
|
||||
simp [i']
|
||||
rw [List.countP_eq_length_filter, ← h1]
|
||||
rfl
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma mem_withUniqueDual_of_countId_eq_two (i : Fin l.length)
|
||||
(h : l.countId (l.val.get i) = 2) : i ∈ l.withUniqueDual := by
|
||||
have hw : i ∈ l.withDual := by
|
||||
rw [mem_withDual_iff_countId_gt_one, h]
|
||||
exact Nat.one_lt_two
|
||||
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply And.intro ((mem_withDual_iff_isSome l i).mp hw)
|
||||
intro j hj
|
||||
rw [@countId_get] at h
|
||||
simp only [List.countP_eq_length_filter, Nat.reduceEqDiff] at h
|
||||
rw [List.length_eq_one] at h
|
||||
obtain ⟨a, ha⟩ := h
|
||||
have hj : j ∈ List.filter (fun j => decide (l.AreDualInSelf i j)) (List.finRange l.length) := by
|
||||
simpa using hj
|
||||
rw [ha] at hj
|
||||
simp only [List.mem_singleton] at hj
|
||||
subst hj
|
||||
have ht : (l.getDual? i).get ((mem_withDual_iff_isSome l i).mp hw) ∈
|
||||
(List.finRange l.length).filter (fun j => decide (l.AreDualInSelf i j)) := by
|
||||
simp
|
||||
rw [ha] at ht
|
||||
simp only [List.mem_singleton] at ht
|
||||
subst ht
|
||||
simp
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma mem_withUniqueDual_iff_countId_eq_two (i : Fin l.length) :
|
||||
i ∈ l.withUniqueDual ↔ l.countId (l.val.get i) = 2 :=
|
||||
Iff.intro (fun h => l.countId_eq_two_of_mem_withUniqueDual i h)
|
||||
(fun h => l.mem_withUniqueDual_of_countId_eq_two i h)
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma countId_eq_one_of_mem_withUniqueDualInOther (i : Fin l.length)
|
||||
(h : i ∈ l.withUniqueDualInOther l2) :
|
||||
l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 := by
|
||||
let i' := (l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h)
|
||||
have h1 : [i'] = (List.finRange l2.length).filter (fun j => (l.AreDualInOther l2 i j)) := by
|
||||
trans List.filter (fun j => (l.AreDualInOther l2 i j)) [i']
|
||||
· simp [List.filter, i']
|
||||
trans List.filter (fun j => (l.AreDualInOther l2 i j))
|
||||
((List.finRange l2.length).filter (fun j => j = i'))
|
||||
· apply congrArg
|
||||
rw [← filter_finRange l2 i']
|
||||
apply List.filter_congr (fun j _ => ?_)
|
||||
simpa using eq_comm
|
||||
trans List.filter (fun j => j = i')
|
||||
((List.finRange l2.length).filter (fun j => (l.AreDualInOther l2 i j)))
|
||||
· simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and]
|
||||
apply List.filter_congr (fun j _ => ?_)
|
||||
exact Bool.and_comm (decide (l.AreDualInOther l2 i j)) (decide (j = i'))
|
||||
· simp only [List.filter_filter, decide_eq_true_eq, Bool.decide_and]
|
||||
refine List.filter_congr (fun j _ => ?_)
|
||||
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true,
|
||||
Option.not_isSome, Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome,
|
||||
Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
simp only [Bool.and_iff_right_iff_imp, decide_eq_true_eq]
|
||||
intro hj
|
||||
have hj' := h.2.2 j hj
|
||||
apply Option.some_injective
|
||||
rw [hj']
|
||||
simp [i']
|
||||
apply And.intro
|
||||
· simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
rw [mem_withDual_iff_countId_gt_one] at h
|
||||
have h2 := countId_index_neq_zero l i
|
||||
omega
|
||||
· rw [countId_get_other, List.countP_eq_length_filter, ← h1]
|
||||
rfl
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma mem_withUniqueDualInOther_of_countId_eq_one (i : Fin l.length)
|
||||
(h : l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1) :
|
||||
i ∈ l.withUniqueDualInOther l2 := by
|
||||
have hw : i ∈ l.withDualInOther l2 := by
|
||||
rw [mem_withDualInOther_iff_countId_neq_zero, h.2]
|
||||
exact Nat.one_ne_zero
|
||||
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
apply And.intro
|
||||
· rw [mem_withDual_iff_countId_gt_one]
|
||||
omega
|
||||
· apply And.intro
|
||||
· rw [mem_withDualInOther_iff_countId_neq_zero]
|
||||
exact countId_neq_zero_of_mem_withDualInOther l l2 i hw
|
||||
· intro j hj
|
||||
have h2 := h.2
|
||||
rw [countId_get_other l l2, List.countP_eq_length_filter, List.length_eq_one] at h2
|
||||
obtain ⟨a, ha⟩ := h2
|
||||
have hj : j ∈ List.filter (fun j => decide (l.AreDualInOther l2 i j))
|
||||
(List.finRange l2.length) := by
|
||||
simpa using hj
|
||||
rw [ha] at hj
|
||||
simp only [List.mem_singleton] at hj
|
||||
subst hj
|
||||
have ht : (l.getDualInOther? l2 i).get ((mem_withInDualOther_iff_isSome l l2 i).mp hw) ∈
|
||||
(List.finRange l2.length).filter (fun j => decide (l.AreDualInOther l2 i j)) := by
|
||||
simp
|
||||
rw [ha] at ht
|
||||
simp only [List.mem_singleton] at ht
|
||||
subst ht
|
||||
exact Option.some_get ((mem_withInDualOther_iff_isSome l l2 i).mp hw)
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma mem_withUniqueDualInOther_iff_countId_eq_one (i : Fin l.length) :
|
||||
i ∈ l.withUniqueDualInOther l2 ↔ l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 :=
|
||||
Iff.intro (fun h => l.countId_eq_one_of_mem_withUniqueDualInOther l2 i h)
|
||||
(fun h => l.mem_withUniqueDualInOther_of_countId_eq_one l2 i h)
|
||||
|
||||
/-!
|
||||
|
||||
## getDual? and countId
|
||||
|
||||
-/
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_countId (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
l.countId (l.val[Fin.val ((l.getDual? i).get h)]) = l.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l.idMap ((l.getDual? i).get h) = _
|
||||
simp only [getDual?_get_id, List.get_eq_getElem]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_countId_right (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l2.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l2.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l2.idMap ((l.getDualInOther? l2 i).get h) = _
|
||||
simp only [getDualInOther?_id, List.get_eq_getElem]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_countId_left (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l.countId (l2.val[Fin.val ((l.getDualInOther? l2 i).get h)]) = l.countId (l.val.get i) := by
|
||||
apply countId_congr
|
||||
change l2.idMap ((l.getDualInOther? l2 i).get h) = _
|
||||
simp only [getDualInOther?_id, List.get_eq_getElem]
|
||||
rfl
|
||||
|
||||
lemma getDual?_isSome_of_countId_eq_two {i : Fin l.length}
|
||||
(h : l.countId (l.val.get i) = 2) : (l.getDual? i).isSome := by
|
||||
rw [← l.mem_withUniqueDual_iff_countId_eq_two] at h
|
||||
exact mem_withUniqueDual_isSome l i h
|
||||
|
||||
/-!
|
||||
|
||||
## Filters
|
||||
|
||||
-/
|
||||
|
||||
lemma filter_id_of_countId_eq_zero {i : Fin l.length} (h1 : l.countId (l.val.get i) = 0) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) = [] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h1
|
||||
exact h1
|
||||
|
||||
lemma filter_id_of_countId_eq_zero' {I : Index X} (h1 : l.countId I = 0) :
|
||||
l.val.filter (fun J => I.id = J.id) = [] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_zero] at h1
|
||||
exact h1
|
||||
|
||||
lemma filter_id_of_countId_eq_one {i : Fin l.length} (h1 : l.countId (l.val.get i) = 1) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) = [l.val.get i] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_one] at h1
|
||||
obtain ⟨J, hJ⟩ := h1
|
||||
have hme : (l.val.get i) ∈ List.filter (fun J => decide ((l.val.get i).id = J.id)) l.val := by
|
||||
simp only [List.get_eq_getElem, List.mem_filter, decide_True, and_true]
|
||||
exact List.getElem_mem l.val (↑i) i.isLt
|
||||
rw [hJ] at hme
|
||||
simp only [List.get_eq_getElem, List.mem_singleton] at hme
|
||||
erw [hJ]
|
||||
exact congrFun (congrArg List.cons (id (Eq.symm hme))) []
|
||||
|
||||
lemma filter_id_of_countId_eq_one_mem {I : Index X} (hI : I ∈ l.val) (h : l.countId I = 1) :
|
||||
l.val.filter (fun J => I.id = J.id) = [I] := by
|
||||
rw [countId_eq_length_filter, List.length_eq_one] at h
|
||||
obtain ⟨J, hJ⟩ := h
|
||||
have hme : I ∈ List.filter (fun J => decide (I.id = J.id)) l.val := by
|
||||
simp only [List.mem_filter, decide_True, and_true]
|
||||
exact hI
|
||||
rw [hJ] at hme
|
||||
simp only [List.mem_singleton] at hme
|
||||
erw [hJ]
|
||||
exact congrFun (congrArg List.cons (id (Eq.symm hme))) []
|
||||
|
||||
lemma filter_id_of_countId_eq_two {i : Fin l.length}
|
||||
(h : l.countId (l.val.get i) = 2) :
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) =
|
||||
[l.val.get i, l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))] ∨
|
||||
l.val.filter (fun J => (l.val.get i).id = J.id) =
|
||||
[l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i] := by
|
||||
have hc := h
|
||||
rw [countId_eq_length_filter] at hc
|
||||
by_cases hi : i < ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
· refine Or.inl (List.Sublist.eq_of_length ?_ ?_).symm
|
||||
· have h1 : [l.val.get i, l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))].filter
|
||||
(fun J => (l.val.get i).id = J.id) = [l.val.get i, l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))] := by
|
||||
simp only [List.get_eq_getElem, decide_True, List.filter_cons_of_pos, List.cons.injEq,
|
||||
List.filter_eq_self, List.mem_singleton, decide_eq_true_eq, forall_eq, true_and]
|
||||
change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
exact Eq.symm (getDual?_get_id l i (getDual?_isSome_of_countId_eq_two l h))
|
||||
rw [← h1]
|
||||
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
|
||||
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
|
||||
refine ⟨⟨⟨![i, (l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)],
|
||||
List.nodup_ofFn.mp ?_⟩, ?_⟩, ?_⟩
|
||||
· simpa using Fin.ne_of_lt hi
|
||||
· intro a b
|
||||
fin_cases a, b
|
||||
· simp [hi]
|
||||
· simp only [List.get_eq_getElem, List.length_cons, List.length_singleton, Nat.reduceAdd,
|
||||
List.length_nil, Fin.zero_eta, Fin.isValue, Function.Embedding.coeFn_mk,
|
||||
Matrix.cons_val_zero, Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, Fin.zero_le,
|
||||
iff_true]
|
||||
exact Fin.le_of_lt hi
|
||||
· simp [hi]
|
||||
· simp [hi]
|
||||
· intro a
|
||||
fin_cases a <;> rfl
|
||||
· rw [hc]
|
||||
rfl
|
||||
· have hi' : ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)) < i := by
|
||||
have h1 := l.getDual?_get_areDualInSelf i (getDual?_isSome_of_countId_eq_two l h)
|
||||
simp only [AreDualInSelf] at h1
|
||||
have h2 := h1.1
|
||||
omega
|
||||
refine Or.inr (List.Sublist.eq_of_length ?_ ?_).symm
|
||||
· have h1 : [l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)),
|
||||
l.val.get i].filter (fun J => (l.val.get i).id = J.id) = [l.val.get
|
||||
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i,] := by
|
||||
simp only [List.get_eq_getElem, List.filter_eq_self, List.mem_cons, List.mem_singleton,
|
||||
decide_eq_true_eq, forall_eq_or_imp, forall_eq, and_true]
|
||||
apply And.intro
|
||||
· change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
|
||||
exact Eq.symm (getDual?_get_id l i (getDual?_isSome_of_countId_eq_two l h))
|
||||
· exact And.symm (if_false_right.mp trivial)
|
||||
rw [← h1]
|
||||
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
|
||||
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
|
||||
refine ⟨⟨⟨![(l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h), i], ?_⟩, ?_⟩, ?_⟩
|
||||
· refine List.nodup_ofFn.mp ?_
|
||||
simp only [List.get_eq_getElem, List.length_singleton, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
List.length_nil, List.ofFn_succ, Fin.isValue, Matrix.cons_val_zero, Matrix.cons_val_succ,
|
||||
Matrix.cons_val_fin_one, List.ofFn_zero, List.nodup_cons, List.mem_singleton,
|
||||
List.not_mem_nil, not_false_eq_true, List.nodup_nil, and_self, and_true]
|
||||
exact Fin.ne_of_lt hi'
|
||||
· intro a b
|
||||
fin_cases a, b
|
||||
· simp [hi']
|
||||
· simp only [List.get_eq_getElem, List.length_cons, List.length_singleton, Nat.reduceAdd,
|
||||
List.length_nil, Fin.zero_eta, Fin.isValue, Function.Embedding.coeFn_mk,
|
||||
Matrix.cons_val_zero, Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, Fin.zero_le,
|
||||
iff_true]
|
||||
exact Fin.le_of_lt hi'
|
||||
· simp [hi']
|
||||
· simp [hi']
|
||||
· intro a
|
||||
fin_cases a <;> rfl
|
||||
· rw [hc]
|
||||
rfl
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,440 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Basic
|
||||
/-!
|
||||
|
||||
# Indices which are dual in an index list
|
||||
|
||||
Given an list of indices we say two indices are dual if they have the same id.
|
||||
|
||||
For example the `0`, `2` and `3` index in `l₁ := ['ᵘ¹', 'ᵘ²', 'ᵤ₁', 'ᵘ¹']` are pairwise dual to
|
||||
one another. The `1` (`'ᵘ²'`) index is not dual to any other index in the list.
|
||||
|
||||
We also define the notion of dual indices in different lists. For example,
|
||||
the `1` index in `l₁` is dual to the `1` and the `4` indices in
|
||||
`l₂ := ['ᵘ³', 'ᵘ²', 'ᵘ⁴', 'ᵤ₂']`.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-!
|
||||
|
||||
## Definitions
|
||||
|
||||
-/
|
||||
|
||||
/-- Two indices are dual if they are not equivalent, but have the same id. -/
|
||||
def AreDualInSelf (i j : Fin l.length) : Prop :=
|
||||
i ≠ j ∧ l.idMap i = l.idMap j
|
||||
|
||||
instance (i j : Fin l.length) : Decidable (l.AreDualInSelf i j) :=
|
||||
instDecidableAnd
|
||||
|
||||
/-- Two indices in different `IndexLists` are dual to one another if they have the same `id`. -/
|
||||
def AreDualInOther (i : Fin l.length) (j : Fin l2.length) :
|
||||
Prop := l.idMap i = l2.idMap j
|
||||
|
||||
instance {l : IndexList X} {l2 : IndexList X} (i : Fin l.length) (j : Fin l2.length) :
|
||||
Decidable (l.AreDualInOther l2 i j) := (l.idMap i).decEq (l2.idMap j)
|
||||
|
||||
/-- Given an `i`, if a dual exists in the same list,
|
||||
outputs the first such dual, otherwise outputs `none`. -/
|
||||
def getDual? (i : Fin l.length) : Option (Fin l.length) :=
|
||||
Fin.find (fun j => l.AreDualInSelf i j)
|
||||
|
||||
/-- Given an `i`, if a dual exists in the other list,
|
||||
outputs the first such dual, otherwise outputs `none`. -/
|
||||
def getDualInOther? (i : Fin l.length) : Option (Fin l2.length) :=
|
||||
Fin.find (fun j => l.AreDualInOther l2 i j)
|
||||
|
||||
/-- The finite set of indices of an index list which have a dual in that index list. -/
|
||||
def withDual : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => (l.getDual? i).isSome) Finset.univ
|
||||
|
||||
/-- The finite set of indices of an index list which do not have a dual. -/
|
||||
def withoutDual : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => (l.getDual? i).isNone) Finset.univ
|
||||
|
||||
/-- The finite set of indices of an index list which have a dual in another provided index list. -/
|
||||
def withDualInOther : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => (l.getDualInOther? l2 i).isSome) Finset.univ
|
||||
|
||||
/-- The finite set of indices of an index list which have a unique dual in that index list. -/
|
||||
def withUniqueDual : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => i ∈ l.withDual ∧
|
||||
∀ j, l.AreDualInSelf i j → j = l.getDual? i) Finset.univ
|
||||
|
||||
instance (i j : Option (Fin l.length)) : Decidable (i < j) :=
|
||||
match i, j with
|
||||
| none, none => isFalse (fun a => a)
|
||||
| none, some _ => isTrue (by trivial)
|
||||
| some _, none => isFalse (fun a => a)
|
||||
| some i, some j => Fin.decLt i j
|
||||
|
||||
/-- The finite set of those indices of `l` which have a unique dual, and for which
|
||||
that dual is greater-then (determined by the ordering on `Fin`) then the index itself. -/
|
||||
def withUniqueDualLT : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => i < l.getDual? i) l.withUniqueDual
|
||||
|
||||
/-- The finite set of those indices of `l` which have a unique dual, and for which
|
||||
that dual is less-then (determined by the ordering on `Fin`) then the index itself. -/
|
||||
def withUniqueDualGT : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => ¬ i < l.getDual? i) l.withUniqueDual
|
||||
|
||||
/-- The finite set of indices of an index list which have a unique dual in another, provided, index
|
||||
list. -/
|
||||
def withUniqueDualInOther : Finset (Fin l.length) :=
|
||||
Finset.filter (fun i => i ∉ l.withDual ∧ i ∈ l.withDualInOther l2
|
||||
∧ (∀ j, l.AreDualInOther l2 i j → j = l.getDualInOther? l2 i)) Finset.univ
|
||||
|
||||
/-!
|
||||
|
||||
## Basic properties
|
||||
|
||||
-/
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
@[simp, nolint simpNF]
|
||||
lemma mem_withDual_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
|
||||
i ∈ l.withDual := by
|
||||
simp only [withUniqueDual, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
simpa using h.1
|
||||
|
||||
lemma mem_withDual_of_withUniqueDual (i : l.withUniqueDual) :
|
||||
i.1 ∈ l.withDual :=
|
||||
mem_withDual_of_mem_withUniqueDual l (↑i) i.2
|
||||
|
||||
lemma not_mem_withDual_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
|
||||
i.1 ∉ l.withDual := by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
|
||||
true_and] at hi
|
||||
exact hi.2.1
|
||||
|
||||
lemma mem_withDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther l2) :
|
||||
i.1 ∈ l.withDualInOther l2 := by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualInOther, Finset.univ_eq_attach, Finset.mem_filter, Finset.mem_attach,
|
||||
true_and] at hi
|
||||
exact hi.2.2.1
|
||||
|
||||
@[simp]
|
||||
lemma withDual_isSome (i : l.withDual) : (l.getDual? i).isSome := by
|
||||
have hx := i.2
|
||||
simp only [Finset.mem_filter, withDual] at hx
|
||||
exact hx.2
|
||||
|
||||
@[simp]
|
||||
lemma mem_withDual_iff_isSome (i : Fin l.length) : i ∈ l.withDual ↔ (l.getDual? i).isSome := by
|
||||
simp [withDual]
|
||||
|
||||
@[simp]
|
||||
lemma mem_withUniqueDual_isSome (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
|
||||
(l.getDual? i).isSome := by
|
||||
simpa [withDual] using mem_withDual_of_mem_withUniqueDual l i h
|
||||
|
||||
@[simp]
|
||||
lemma mem_withInDualOther_iff_isSome (i : Fin l.length) :
|
||||
i ∈ l.withDualInOther l2 ↔ (l.getDualInOther? l2 i).isSome := by
|
||||
simp only [withDualInOther, getDualInOther?, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
|
||||
@[simp]
|
||||
lemma mem_withUniqueDualInOther_isSome (i : Fin l.length) (hi : i ∈ l.withUniqueDualInOther l2) :
|
||||
(l.getDualInOther? l2 i).isSome := by
|
||||
simp only [withUniqueDualInOther, Finset.mem_filter, Finset.mem_univ, true_and] at hi
|
||||
exact (mem_withInDualOther_iff_isSome l l2 i).mp hi.2.1
|
||||
|
||||
lemma not_mem_withDual_iff_isNone (i : Fin l.length) :
|
||||
i ∉ l.withDual ↔ (l.getDual? i).isNone := by
|
||||
simp only [mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none]
|
||||
|
||||
lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j := by
|
||||
simp only [withDual, getDual?, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
exact Fin.isSome_find_iff
|
||||
|
||||
lemma mem_withInDualOther_iff_exists :
|
||||
i ∈ l.withDualInOther l2 ↔ ∃ (j : Fin l2.length), l.AreDualInOther l2 i j := by
|
||||
simp only [withDualInOther, getDualInOther?, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
exact Fin.isSome_find_iff
|
||||
|
||||
lemma withDual_disjoint_withoutDual : Disjoint l.withDual l.withoutDual := by
|
||||
rw [Finset.disjoint_iff_ne]
|
||||
intro a ha b hb
|
||||
by_contra hn
|
||||
subst hn
|
||||
simp_all only [withDual, Finset.mem_filter, Finset.mem_univ, true_and, withoutDual,
|
||||
Option.isNone_iff_eq_none, Option.isSome_none, Bool.false_eq_true]
|
||||
|
||||
lemma not_mem_withDual_of_mem_withoutDual (i : Fin l.length) (h : i ∈ l.withoutDual) :
|
||||
i ∉ l.withDual := by
|
||||
have h1 := l.withDual_disjoint_withoutDual
|
||||
exact Finset.disjoint_right.mp h1 h
|
||||
|
||||
lemma withDual_union_withoutDual : l.withDual ∪ l.withoutDual = Finset.univ := by
|
||||
rw [Finset.eq_univ_iff_forall]
|
||||
intro i
|
||||
by_cases h : (l.getDual? i).isSome
|
||||
· simp [withDual, Finset.mem_filter, Finset.mem_univ, h]
|
||||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h
|
||||
simp [withoutDual, Finset.mem_filter, Finset.mem_univ, h]
|
||||
|
||||
lemma mem_withUniqueDual_of_mem_withUniqueDualLt (i : Fin l.length) (h : i ∈ l.withUniqueDualLT) :
|
||||
i ∈ l.withUniqueDual := by
|
||||
simp only [withUniqueDualLT, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
exact h.1
|
||||
|
||||
lemma mem_withUniqueDual_of_mem_withUniqueDualGt (i : Fin l.length) (h : i ∈ l.withUniqueDualGT) :
|
||||
i ∈ l.withUniqueDual := by
|
||||
simp only [withUniqueDualGT, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
exact h.1
|
||||
|
||||
lemma withUniqueDualLT_disjoint_withUniqueDualGT :
|
||||
Disjoint l.withUniqueDualLT l.withUniqueDualGT := by
|
||||
rw [Finset.disjoint_iff_inter_eq_empty]
|
||||
exact @Finset.filter_inter_filter_neg_eq (Fin l.length) _ _ _ _ _
|
||||
|
||||
lemma withUniqueDualLT_union_withUniqueDualGT :
|
||||
l.withUniqueDualLT ∪ l.withUniqueDualGT = l.withUniqueDual :=
|
||||
Finset.filter_union_filter_neg_eq _ _
|
||||
/-!
|
||||
## Are dual properties
|
||||
-/
|
||||
|
||||
namespace AreDualInSelf
|
||||
|
||||
variable {l l2 : IndexList X} {i j : Fin l.length}
|
||||
|
||||
instance (i j : Fin l.length) : Decidable (l.AreDualInSelf i j) :=
|
||||
instDecidableAnd
|
||||
|
||||
@[symm]
|
||||
lemma symm (h : l.AreDualInSelf i j) : l.AreDualInSelf j i := by
|
||||
simp only [AreDualInSelf] at h ⊢
|
||||
exact ⟨h.1.symm, h.2.symm⟩
|
||||
|
||||
@[simp]
|
||||
lemma self_false (i : Fin l.length) : ¬ l.AreDualInSelf i i := by
|
||||
simp [AreDualInSelf]
|
||||
|
||||
@[simp]
|
||||
lemma append_inl_inl : (l ++ l2).AreDualInSelf (appendEquiv (Sum.inl i)) (appendEquiv (Sum.inl j))
|
||||
↔ l.AreDualInSelf i j := by
|
||||
simp [AreDualInSelf]
|
||||
|
||||
@[simp]
|
||||
lemma append_inr_inr (l l2 : IndexList X) (i j : Fin l2.length) :
|
||||
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inr i)) (appendEquiv (Sum.inr j))
|
||||
↔ l2.AreDualInSelf i j := by
|
||||
simp [AreDualInSelf]
|
||||
|
||||
@[simp]
|
||||
lemma append_inl_inr (l l2 : IndexList X) (i : Fin l.length) (j : Fin l2.length) :
|
||||
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inl i)) (appendEquiv (Sum.inr j)) =
|
||||
l.AreDualInOther l2 i j := by
|
||||
simp [AreDualInSelf, AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma append_inr_inl (l l2 : IndexList X) (i : Fin l2.length) (j : Fin l.length) :
|
||||
(l ++ l2).AreDualInSelf (appendEquiv (Sum.inr i)) (appendEquiv (Sum.inl j)) =
|
||||
l2.AreDualInOther l i j := by
|
||||
simp [AreDualInSelf, AreDualInOther]
|
||||
|
||||
end AreDualInSelf
|
||||
|
||||
namespace AreDualInOther
|
||||
|
||||
variable {l l2 l3 : IndexList X} {i : Fin l.length} {j : Fin l2.length}
|
||||
|
||||
instance {l : IndexList X} {l2 : IndexList X} (i : Fin l.length) (j : Fin l2.length) :
|
||||
Decidable (l.AreDualInOther l2 i j) := (l.idMap i).decEq (l2.idMap j)
|
||||
|
||||
@[symm]
|
||||
lemma symm (h : l.AreDualInOther l2 i j) : l2.AreDualInOther l j i := by
|
||||
rw [AreDualInOther] at h ⊢
|
||||
exact h.symm
|
||||
|
||||
@[simp]
|
||||
lemma append_of_inl (i : Fin l.length) (j : Fin l3.length) :
|
||||
(l ++ l2).AreDualInOther l3 (appendEquiv (Sum.inl i)) j ↔ l.AreDualInOther l3 i j := by
|
||||
simp [AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma append_of_inr (i : Fin l2.length) (j : Fin l3.length) :
|
||||
(l ++ l2).AreDualInOther l3 (appendEquiv (Sum.inr i)) j ↔ l2.AreDualInOther l3 i j := by
|
||||
simp [AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma of_append_inl (i : Fin l.length) (j : Fin l2.length) :
|
||||
l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inl j)) ↔ l.AreDualInOther l2 i j := by
|
||||
simp [AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma of_append_inr (i : Fin l.length) (j : Fin l3.length) :
|
||||
l.AreDualInOther (l2 ++ l3) i (appendEquiv (Sum.inr j)) ↔ l.AreDualInOther l3 i j := by
|
||||
simp [AreDualInOther]
|
||||
|
||||
end AreDualInOther
|
||||
|
||||
/-!
|
||||
|
||||
## Basic properties of getDual?
|
||||
|
||||
-/
|
||||
|
||||
lemma getDual?_isSome_iff_exists (i : Fin l.length) :
|
||||
(l.getDual? i).isSome ↔ ∃ j, l.AreDualInSelf i j := by
|
||||
rw [getDual?, Fin.isSome_find_iff]
|
||||
|
||||
lemma getDual?_of_areDualInSelf (h : l.AreDualInSelf i j) :
|
||||
l.getDual? j = i ∨ l.getDual? i = j ∨ l.getDual? j = l.getDual? i := by
|
||||
have h3 : (l.getDual? i).isSome := by
|
||||
simpa [getDual?, Fin.isSome_find_iff] using ⟨j, h⟩
|
||||
obtain ⟨k, hk⟩ := Option.isSome_iff_exists.mp h3
|
||||
rw [hk]
|
||||
rw [getDual?, Fin.find_eq_some_iff] at hk
|
||||
by_cases hik : i < k
|
||||
· apply Or.inl
|
||||
rw [getDual?, Fin.find_eq_some_iff]
|
||||
apply And.intro h.symm
|
||||
intro k' hk'
|
||||
by_cases hik' : i = k'
|
||||
· exact Fin.ge_of_eq (id (Eq.symm hik'))
|
||||
· have hik'' : l.AreDualInSelf i k' := by
|
||||
simp only [AreDualInSelf, ne_eq, hik', not_false_eq_true, true_and]
|
||||
simp_all [AreDualInSelf]
|
||||
have hk'' := hk.2 k' hik''
|
||||
exact (lt_of_lt_of_le hik hk'').le
|
||||
· by_cases hjk : j ≤ k
|
||||
· apply Or.inr
|
||||
apply Or.inl
|
||||
have hj := hk.2 j h
|
||||
simp only [Option.some.injEq]
|
||||
exact Fin.le_antisymm hj hjk
|
||||
· apply Or.inr
|
||||
apply Or.inr
|
||||
rw [getDual?, Fin.find_eq_some_iff]
|
||||
apply And.intro
|
||||
· simp_all [AreDualInSelf]
|
||||
exact Fin.ne_of_gt hjk
|
||||
intro k' hk'
|
||||
by_cases hik' : i = k'
|
||||
· subst hik'
|
||||
exact Lean.Omega.Fin.le_of_not_lt hik
|
||||
· have hik'' : l.AreDualInSelf i k' := by
|
||||
simp only [AreDualInSelf, ne_eq, hik', not_false_eq_true, true_and]
|
||||
simp_all only [AreDualInSelf, ne_eq, and_imp, not_lt, not_le]
|
||||
exact hk.2 k' hik''
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_neq_self (i : Fin l.length) : ¬ l.getDual? i = some i := by
|
||||
intro h
|
||||
simp only [getDual?] at h
|
||||
rw [Fin.find_eq_some_iff] at h
|
||||
simp [AreDualInSelf] at h
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_get_neq_self (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
¬ (l.getDual? i).get h = i := by
|
||||
have h1 := l.getDual?_neq_self i
|
||||
by_contra hn
|
||||
have h' : l.getDual? i = some i := by
|
||||
nth_rewrite 2 [← hn]
|
||||
exact Option.eq_some_of_isSome h
|
||||
exact h1 h'
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_get_id (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
l.idMap ((l.getDual? i).get h) = l.idMap i := by
|
||||
have h1 : l.getDual? i = some ((l.getDual? i).get h) := Option.eq_some_of_isSome h
|
||||
nth_rewrite 1 [getDual?] at h1
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
simp only [AreDualInSelf, ne_eq, and_imp] at h1
|
||||
exact h1.1.2.symm
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_get_areDualInSelf (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
l.AreDualInSelf ((l.getDual? i).get h) i := by
|
||||
simp [AreDualInSelf]
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_areDualInSelf_get (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
l.AreDualInSelf i ((l.getDual? i).get h) :=
|
||||
AreDualInSelf.symm (getDual?_get_areDualInSelf l i h)
|
||||
|
||||
@[simp]
|
||||
lemma getDual?_getDual?_get_isSome (i : Fin l.length) (h : (l.getDual? i).isSome) :
|
||||
(l.getDual? ((l.getDual? i).get h)).isSome := by
|
||||
rw [getDual?, Fin.isSome_find_iff]
|
||||
exact ⟨i, getDual?_get_areDualInSelf l i h⟩
|
||||
|
||||
/-!
|
||||
|
||||
## Basic properties of getDualInOther?
|
||||
|
||||
-/
|
||||
|
||||
lemma getDualInOther?_isSome_iff_exists (i : Fin l.length) :
|
||||
(l.getDualInOther? l2 i).isSome ↔ ∃ j, l.AreDualInOther l2 i j := by
|
||||
rw [getDualInOther?, Fin.isSome_find_iff]
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_id (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l2.idMap ((l.getDualInOther? l2 i).get h) = l.idMap i := by
|
||||
have h1 : l.getDualInOther? l2 i = some ((l.getDualInOther? l2 i).get h) :=
|
||||
Option.eq_some_of_isSome h
|
||||
nth_rewrite 1 [getDualInOther?] at h1
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
simp only [AreDualInOther] at h1
|
||||
exact h1.1.symm
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_get_areDualInOther (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l2.AreDualInOther l ((l.getDualInOther? l2 i).get h) i := by
|
||||
simp [AreDualInOther]
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_areDualInOther_get (i : Fin l.length) (h : (l.getDualInOther? l2 i).isSome) :
|
||||
l.AreDualInOther l2 i ((l.getDualInOther? l2 i).get h) :=
|
||||
AreDualInOther.symm (getDualInOther?_get_areDualInOther l l2 i h)
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_getDualInOther?_get_isSome (i : Fin l.length)
|
||||
(h : (l.getDualInOther? l2 i).isSome) :
|
||||
(l2.getDualInOther? l ((l.getDualInOther? l2 i).get h)).isSome := by
|
||||
rw [getDualInOther?, Fin.isSome_find_iff]
|
||||
exact ⟨i, getDualInOther?_get_areDualInOther l l2 i h⟩
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_self_isSome (i : Fin l.length) :
|
||||
(l.getDualInOther? l i).isSome := by
|
||||
simp only [getDualInOther?]
|
||||
rw [@Fin.isSome_find_iff]
|
||||
exact ⟨i, rfl⟩
|
||||
|
||||
@[simp]
|
||||
lemma getDualInOther?_getDualInOther?_get_self (i : Fin l.length) :
|
||||
l.getDualInOther? l ((l.getDualInOther? l i).get (getDualInOther?_self_isSome l i)) =
|
||||
some ((l.getDualInOther? l i).get (getDualInOther?_self_isSome l i)) := by
|
||||
nth_rewrite 1 [getDualInOther?]
|
||||
rw [Fin.find_eq_some_iff]
|
||||
simp only [AreDualInOther, getDualInOther?_id, true_and]
|
||||
intro j hj
|
||||
have h1 := Option.eq_some_of_isSome (getDualInOther?_self_isSome l i)
|
||||
nth_rewrite 1 [getDualInOther?] at h1
|
||||
rw [Fin.find_eq_some_iff] at h1
|
||||
apply h1.2 j
|
||||
simpa [AreDualInOther] using hj
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,276 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.CountId
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Equivalences between finsets.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-!
|
||||
|
||||
## withoutDualEquiv
|
||||
-/
|
||||
|
||||
/-- An equivalence from `Fin l.withoutDual.card` to `l.withoutDual` determined by
|
||||
the order on `l.withoutDual` inherted from `Fin`. -/
|
||||
def withoutDualEquiv : Fin l.withoutDual.card ≃ l.withoutDual :=
|
||||
(Finset.orderIsoOfFin l.withoutDual (by rfl)).toEquiv
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
lemma list_ofFn_withoutDualEquiv_eq_sort :
|
||||
List.ofFn (Subtype.val ∘ l.withoutDualEquiv) = l.withoutDual.sort (fun i j => i ≤ j) := by
|
||||
rw [@List.ext_get_iff]
|
||||
refine And.intro ?_ (fun n h1 h2 => ?_)
|
||||
· simp only [List.length_ofFn, Finset.length_sort]
|
||||
· simp only [List.get_eq_getElem, List.getElem_ofFn, Function.comp_apply]
|
||||
rfl
|
||||
|
||||
/-- An equivalence splitting the indices of an index list `l` into those indices
|
||||
which have a dual in `l` and those which do not have a dual in `l`. -/
|
||||
def dualEquiv : l.withDual ⊕ Fin l.withoutDual.card ≃ Fin l.length :=
|
||||
(Equiv.sumCongr (Equiv.refl _) l.withoutDualEquiv).trans <|
|
||||
(Equiv.Finset.union _ _ l.withDual_disjoint_withoutDual).trans <|
|
||||
Equiv.subtypeUnivEquiv (by simp [withDual_union_withoutDual])
|
||||
|
||||
/-!
|
||||
|
||||
## getDualEquiv
|
||||
|
||||
-/
|
||||
|
||||
/-- An equivalence from `l.withUniqueDual` to itself obtained by taking the dual index. -/
|
||||
def getDualEquiv : l.withUniqueDual ≃ l.withUniqueDual where
|
||||
toFun x := ⟨(l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2), by
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two]
|
||||
simpa using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2⟩
|
||||
invFun x := ⟨(l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2), by
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two]
|
||||
simpa using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2)
|
||||
have h1 : l.countId (l.val.get d) = 2 := by
|
||||
simpa [d] using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l.AreDualInSelf d j)).length = 1 := by
|
||||
rw [countId_get, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
right_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDual? x).get (l.mem_withUniqueDual_isSome x.1 x.2)
|
||||
have h1 : l.countId (l.val.get d) = 2 := by
|
||||
simpa [d] using l.countId_eq_two_of_mem_withUniqueDual x.1 x.2
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l.AreDualInSelf d j)).length = 1 := by
|
||||
rw [countId_get, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma getDual?_getDualEquiv (i : l.withUniqueDual) : l.getDual? (l.getDualEquiv i) = i := by
|
||||
have h1 := (Equiv.apply_symm_apply l.getDualEquiv i).symm
|
||||
nth_rewrite 2 [h1]
|
||||
nth_rewrite 2 [getDualEquiv]
|
||||
simp only [Equiv.coe_fn_mk, Option.some_get]
|
||||
rfl
|
||||
|
||||
/-- An equivalence from `l.withUniqueDualInOther l2 ` to
|
||||
`l2.withUniqueDualInOther l` obtained by taking the dual index. -/
|
||||
def getDualInOtherEquiv : l.withUniqueDualInOther l2 ≃ l2.withUniqueDualInOther l where
|
||||
toFun x := ⟨(l.getDualInOther? l2 x).get (l.mem_withUniqueDualInOther_isSome l2 x.1 x.2), by
|
||||
rw [mem_withUniqueDualInOther_iff_countId_eq_one]
|
||||
simpa using And.comm.mpr (l.countId_eq_one_of_mem_withUniqueDualInOther l2 x.1 x.2)⟩
|
||||
invFun x := ⟨(l2.getDualInOther? l x).get (l2.mem_withUniqueDualInOther_isSome l x.1 x.2), by
|
||||
rw [mem_withUniqueDualInOther_iff_countId_eq_one]
|
||||
simpa using And.comm.mpr (l2.countId_eq_one_of_mem_withUniqueDualInOther l x.1 x.2)⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
let d := (l.getDualInOther? l2 x).get (l.mem_withUniqueDualInOther_isSome l2 x.1 x.2)
|
||||
have h1 : l.countId (l2.val.get d) = 1 := by
|
||||
simpa [d] using (l.countId_eq_one_of_mem_withUniqueDualInOther l2 x.1 x.2).1
|
||||
have h2 : ((List.finRange l.length).filter (fun j => l2.AreDualInOther l d j)).length = 1 := by
|
||||
rw [countId_get_other, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
right_inv x := SetCoe.ext <| by
|
||||
let d := (l2.getDualInOther? l x).get (l2.mem_withUniqueDualInOther_isSome l x.1 x.2)
|
||||
have h1 : l2.countId (l.val.get d) = 1 := by
|
||||
simpa [d] using (l2.countId_eq_one_of_mem_withUniqueDualInOther l x.1 x.2).1
|
||||
have h2 : ((List.finRange l2.length).filter (fun j => l.AreDualInOther l2 d j)).length = 1 := by
|
||||
rw [countId_get_other, List.countP_eq_length_filter] at h1
|
||||
omega
|
||||
obtain ⟨a, ha⟩ := List.length_eq_one.mp h2
|
||||
trans a
|
||||
· rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
· symm
|
||||
rw [← List.mem_singleton, ← ha]
|
||||
simp [d]
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
@[simp]
|
||||
lemma getDualInOtherEquiv_self_refl : l.getDualInOtherEquiv l = Equiv.refl _ := by
|
||||
apply Equiv.ext
|
||||
intro x
|
||||
simp only [getDualInOtherEquiv, Equiv.coe_fn_mk, Equiv.refl_apply]
|
||||
apply Subtype.eq
|
||||
simp only
|
||||
have hx2 := x.2
|
||||
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
|
||||
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, getDualInOther?_self_isSome,
|
||||
true_and, Finset.mem_filter, Finset.mem_univ] at hx2
|
||||
apply Option.some_injective
|
||||
rw [hx2.2 x.1 (by rfl)]
|
||||
exact Option.some_get (getDualInOtherEquiv.proof_1 l l x)
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
@[simp]
|
||||
lemma getDualInOtherEquiv_symm : (l.getDualInOtherEquiv l2).symm = l2.getDualInOtherEquiv l := by
|
||||
rfl
|
||||
|
||||
/-- An equivalence casting `withUniqueDualInOther` based on an equality of the left index list. -/
|
||||
def withUniqueDualCastLeft (h : l = l3) :
|
||||
l.withUniqueDualInOther l2 ≃ l3.withUniqueDualInOther l2 where
|
||||
toFun x := ⟨Fin.cast (by rw [h]) x.1, by subst h; exact x.prop⟩
|
||||
invFun x := ⟨Fin.cast (by rw [h]) x.1, by subst h; exact x.prop⟩
|
||||
left_inv x := SetCoe.ext rfl
|
||||
right_inv x := SetCoe.ext rfl
|
||||
|
||||
/-- An equivalence casting `withUniqueDualInOther` based on an equality of the right index list. -/
|
||||
def withUniqueDualCastRight (h : l2 = l3) :
|
||||
l.withUniqueDualInOther l2 ≃ l.withUniqueDualInOther l3 where
|
||||
toFun x := ⟨x.1, by subst h; exact x.prop⟩
|
||||
invFun x := ⟨x.1, by subst h; exact x.prop⟩
|
||||
left_inv x := SetCoe.ext rfl
|
||||
right_inv x := SetCoe.ext rfl
|
||||
|
||||
/-- An equivalence casting `withUniqueDualInOther` based on an equality of both index lists. -/
|
||||
def withUniqueDualCast {l1 l2 l1' l2' : IndexList X} (h : l1 = l1') (h2 : l2 = l2') :
|
||||
l1.withUniqueDualInOther l2 ≃ l1'.withUniqueDualInOther l2' where
|
||||
toFun x := ⟨Fin.cast (by rw [h]) x.1, by subst h h2; exact x.prop⟩
|
||||
invFun x := ⟨Fin.cast (by rw [h]) x.1, by subst h h2; exact x.prop⟩
|
||||
left_inv x := SetCoe.ext rfl
|
||||
right_inv x := SetCoe.ext rfl
|
||||
omit [IndexNotation X] [Fintype X]
|
||||
lemma getDualInOtherEquiv_cast_left (h : l = l3) :
|
||||
l.getDualInOtherEquiv l2 = ((withUniqueDualCastLeft l l2 l3 h).trans
|
||||
(l3.getDualInOtherEquiv l2)).trans (withUniqueDualCastRight l2 l l3 h).symm := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
lemma getDualInOtherEquiv_cast_right (h : l2 = l3) :
|
||||
l.getDualInOtherEquiv l2 = ((withUniqueDualCastRight l l2 l3 h).trans
|
||||
(l.getDualInOtherEquiv l3)).trans (withUniqueDualCastLeft l2 l l3 h).symm := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
lemma getDualInOtherEquiv_cast {l1 l2 l1' l2' : IndexList X} (h : l1 = l1') (h2 : l2 = l2') :
|
||||
l1.getDualInOtherEquiv l2 = ((withUniqueDualCast h h2).trans
|
||||
(l1'.getDualInOtherEquiv l2')).trans (withUniqueDualCast h2 h).symm := by
|
||||
subst h h2
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalences involving `withUniqueDualLT` and `withUniqueDualGT`.
|
||||
|
||||
-/
|
||||
|
||||
omit [DecidableEq X]
|
||||
/-! TODO: Replace with a mathlib lemma. -/
|
||||
lemma option_not_lt (i j : Option (Fin l.length)) : i < j → i ≠ j → ¬ j < i := by
|
||||
match i, j with
|
||||
| none, none =>
|
||||
exact fun a _ _ => a
|
||||
| none, some k =>
|
||||
exact fun _ _ a => a
|
||||
| some i, none =>
|
||||
exact fun h _ _ => h
|
||||
| some i, some j =>
|
||||
intro h h'
|
||||
simp_all
|
||||
change i < j at h
|
||||
change ¬ j < i
|
||||
exact Fin.lt_asymm h
|
||||
|
||||
/-! TODO: Replace with a mathlib lemma. -/
|
||||
lemma lt_option_of_not (i j : Option (Fin l.length)) : ¬ j < i → i ≠ j → i < j := by
|
||||
match i, j with
|
||||
| none, none =>
|
||||
exact fun a b => a (a (a (b rfl)))
|
||||
| none, some k =>
|
||||
exact fun _ _ => trivial
|
||||
| some i, none =>
|
||||
exact fun a _ => a trivial
|
||||
| some i, some j =>
|
||||
intro h h'
|
||||
simp_all
|
||||
change ¬ j < i at h
|
||||
change i < j
|
||||
omega
|
||||
|
||||
/-- The equivalence between `l.withUniqueDualLT` and `l.withUniqueDualGT` obtained by
|
||||
taking an index to its dual. -/
|
||||
def withUniqueDualLTEquivGT : l.withUniqueDualLT ≃ l.withUniqueDualGT where
|
||||
toFun i := ⟨l.getDualEquiv ⟨i, l.mem_withUniqueDual_of_mem_withUniqueDualLt i i.2⟩, by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualGT, Finset.mem_filter, Finset.coe_mem, true_and]
|
||||
simp only [withUniqueDualLT, Finset.mem_filter] at hi
|
||||
apply option_not_lt
|
||||
· rw [getDual?_getDualEquiv]
|
||||
simpa only [getDualEquiv, Equiv.coe_fn_mk, Option.some_get] using hi.2
|
||||
· exact getDual?_neq_self l _⟩
|
||||
invFun i := ⟨l.getDualEquiv.symm ⟨i, l.mem_withUniqueDual_of_mem_withUniqueDualGt i i.2⟩, by
|
||||
have hi := i.2
|
||||
simp only [withUniqueDualLT, Finset.mem_filter, Finset.coe_mem, true_and, gt_iff_lt]
|
||||
simp only [withUniqueDualGT, Finset.mem_filter] at hi
|
||||
apply lt_option_of_not
|
||||
· erw [getDual?_getDualEquiv]
|
||||
simpa [getDualEquiv] using hi.2
|
||||
· symm
|
||||
exact getDual?_neq_self l _⟩
|
||||
left_inv x := SetCoe.ext <| by
|
||||
simp
|
||||
right_inv x := SetCoe.ext <| by
|
||||
simp
|
||||
|
||||
/-- An equivalence from `l.withUniqueDualLT ⊕ l.withUniqueDualLT` to `l.withUniqueDual`.
|
||||
The first `l.withUniqueDualLT` are taken to themselves, whilst the second copy are
|
||||
taken to their dual. -/
|
||||
def withUniqueLTGTEquiv : l.withUniqueDualLT ⊕ l.withUniqueDualLT ≃ l.withUniqueDual := by
|
||||
apply (Equiv.sumCongr (Equiv.refl _) l.withUniqueDualLTEquivGT).trans
|
||||
apply (Equiv.Finset.union _ _ l.withUniqueDualLT_disjoint_withUniqueDualGT).trans
|
||||
apply (Equiv.subtypeEquivRight (by simp only [l.withUniqueDualLT_union_withUniqueDualGT,
|
||||
implies_true]))
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,764 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.CountId
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
import Mathlib.Data.Fin.Tuple.Basic
|
||||
/-!
|
||||
|
||||
# Normalizing an index list.
|
||||
|
||||
The normalization of an index list, is the corresponding index list
|
||||
where all `id` are set to `0, 1, 2, ...` determined by
|
||||
the order of indices without duals, followed by the order of indices with duals.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-! TODO: Replace with Mathlib lemma. -/
|
||||
lemma dedup_map_of_injective' {α : Type} [DecidableEq α] (f : α → ℕ)
|
||||
(l : List α) (hf : ∀ I ∈ l, ∀ J ∈ l, f I = f J ↔ I = J) :
|
||||
(l.map f).dedup = l.dedup.map f := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
simp only [List.map_cons]
|
||||
by_cases h : I ∈ l
|
||||
· rw [List.dedup_cons_of_mem h, List.dedup_cons_of_mem (List.mem_map_of_mem f h), ih]
|
||||
intro I' hI' J hJ
|
||||
have hI'2 : I' ∈ I :: l := List.mem_cons_of_mem I hI'
|
||||
have hJ2 : J ∈ I :: l := List.mem_cons_of_mem I hJ
|
||||
exact hf I' hI'2 J hJ2
|
||||
· rw [List.dedup_cons_of_not_mem h, List.dedup_cons_of_not_mem, ih, List.map_cons]
|
||||
· intro I' hI' J hJ
|
||||
have hI'2 : I' ∈ I :: l := List.mem_cons_of_mem I hI'
|
||||
have hJ2 : J ∈ I :: l := List.mem_cons_of_mem I hJ
|
||||
exact hf I' hI'2 J hJ2
|
||||
· simp_all only [List.mem_cons, or_true, implies_true, true_implies, forall_eq_or_imp,
|
||||
true_and, List.mem_map, not_exists, not_and, List.forall_mem_ne', not_false_eq_true]
|
||||
|
||||
/-! TODO: Replace with Mathlib lemma. -/
|
||||
lemma indexOf_map {α : Type} [DecidableEq α]
|
||||
(f : α → ℕ) (l : List α) (n : α) (hf : ∀ I ∈ l, ∀ J, f I = f J ↔ I = J) :
|
||||
l.indexOf n = (l.map f).indexOf (f n) := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
have ih' : (∀ I ∈ l, ∀ J, f I = f J ↔ I = J) := by
|
||||
intro I' hI' J'
|
||||
exact hf I' (List.mem_cons_of_mem I hI') J'
|
||||
simp only [List.map_cons]
|
||||
rw [← lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat]
|
||||
by_cases h : I = n
|
||||
· rw [l.indexOf_cons_eq h, List.indexOf_cons_eq]
|
||||
exact congrArg f h
|
||||
· rw [l.indexOf_cons_ne h, List.indexOf_cons_ne]
|
||||
· rw [lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat, ih ih']
|
||||
· exact (hf I (List.mem_cons_self I l) n).ne.mpr h
|
||||
|
||||
lemma indexOf_map_fin {α : Type} {m : ℕ} [DecidableEq α]
|
||||
(f : α → (Fin m)) (l : List α) (n : α) (hf : ∀ I ∈ l, ∀ J, f I = f J ↔ I = J) :
|
||||
l.indexOf n = (l.map f).indexOf (f n) := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
have ih' : (∀ I ∈ l, ∀ J, f I = f J ↔ I = J) := by
|
||||
intro I' hI' J'
|
||||
exact hf I' (List.mem_cons_of_mem I hI') J'
|
||||
simp only [List.map_cons]
|
||||
by_cases h : I = n
|
||||
· rw [l.indexOf_cons_eq h, List.indexOf_cons_eq]
|
||||
exact congrArg f h
|
||||
· rw [l.indexOf_cons_ne h, List.indexOf_cons_ne]
|
||||
· exact congrArg Nat.succ (ih ih')
|
||||
· exact (hf I (List.mem_cons_self I l) n).ne.mpr h
|
||||
|
||||
lemma indexOf_map' {α : Type} [DecidableEq α]
|
||||
(f : α → ℕ) (g : α → α) (l : List α) (n : α)
|
||||
(hf : ∀ I ∈ l, ∀ J, f I = f J ↔ g I = g J)
|
||||
(hg : ∀ I ∈ l, g I = I)
|
||||
(hs : ∀ I, f I = f (g I)) :
|
||||
l.indexOf (g n) = (l.map f).indexOf (f n) := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
have ih' : (∀ I ∈ l, ∀ J, f I = f J ↔ g I = g J) := by
|
||||
intro I' hI' J'
|
||||
exact hf I' (List.mem_cons_of_mem I hI') J'
|
||||
simp only [List.map_cons]
|
||||
rw [← lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat]
|
||||
by_cases h : I = g n
|
||||
· rw [l.indexOf_cons_eq h, List.indexOf_cons_eq]
|
||||
rw [h, ← hs]
|
||||
· rw [l.indexOf_cons_ne h, List.indexOf_cons_ne]
|
||||
· rw [lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat, ih ih']
|
||||
intro I hI
|
||||
apply hg
|
||||
exact List.mem_cons_of_mem _ hI
|
||||
· have hi'n := hf I (List.mem_cons_self I l) n
|
||||
refine (Iff.ne (id (Iff.symm hi'n))).mp ?_
|
||||
rw [hg]
|
||||
· exact h
|
||||
· exact List.mem_cons_self I l
|
||||
|
||||
lemma filter_dedup {α : Type} [DecidableEq α] (l : List α) (p : α → Prop) [DecidablePred p] :
|
||||
(l.filter p).dedup = (l.dedup.filter p) := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
by_cases h : p I
|
||||
· by_cases hd : I ∈ l
|
||||
· rw [List.filter_cons_of_pos (by simpa using h), List.dedup_cons_of_mem hd,
|
||||
List.dedup_cons_of_mem, ih]
|
||||
simp [hd, h]
|
||||
· rw [List.filter_cons_of_pos (by simpa using h), List.dedup_cons_of_not_mem hd,
|
||||
List.dedup_cons_of_not_mem, ih, List.filter_cons_of_pos]
|
||||
· exact decide_eq_true h
|
||||
· simp [hd]
|
||||
· by_cases hd : I ∈ l
|
||||
· rw [List.filter_cons_of_neg (by simpa using h), List.dedup_cons_of_mem hd, ih]
|
||||
· rw [List.filter_cons_of_neg (by simpa using h), List.dedup_cons_of_not_mem hd, ih]
|
||||
rw [List.filter_cons_of_neg]
|
||||
simpa using h
|
||||
|
||||
lemma findIdx?_map {α β : Type} (p : α → Bool)
|
||||
(f : β → α) (l : List β) :
|
||||
List.findIdx? p (List.map f l) = List.findIdx? (p ∘ f) l := by
|
||||
induction l with
|
||||
| nil => rfl
|
||||
| cons I l ih =>
|
||||
simp only [List.map_cons, List.findIdx?_cons, zero_add, List.findIdx?_succ, Function.comp_apply]
|
||||
exact congrArg (ite (p (f I) = true) (some 0)) (congrArg (Option.map fun i => i + 1) ih)
|
||||
|
||||
lemma findIdx?_on_finRange {n : ℕ} (p : Fin n.succ → Prop) [DecidablePred p]
|
||||
(hp : ¬ p 0) :
|
||||
List.findIdx? p (List.finRange n.succ) =
|
||||
Option.map (fun i => i + 1) (List.findIdx? (p ∘ Fin.succ) (List.finRange n)) := by
|
||||
rw [List.finRange_succ_eq_map]
|
||||
simp only [Nat.succ_eq_add_one, hp, Nat.reduceAdd, List.finRange_zero, List.map_nil,
|
||||
List.concat_eq_append, List.nil_append, List.findIdx?_cons, decide_eq_true_eq, zero_add,
|
||||
List.findIdx?_succ, List.findIdx?_nil, Option.map_none', ite_eq_right_iff, imp_false]
|
||||
simp only [↓reduceIte]
|
||||
rw [findIdx?_map]
|
||||
|
||||
lemma findIdx?_on_finRange_eq_findIdx {n : ℕ} (p : Fin n → Prop) [DecidablePred p] :
|
||||
List.findIdx? p (List.finRange n) = Option.map Fin.val (Fin.find p) := by
|
||||
induction n with
|
||||
| zero => rfl
|
||||
| succ n ih =>
|
||||
by_cases h : p 0
|
||||
· trans some 0
|
||||
· rw [List.finRange_succ_eq_map]
|
||||
simp only [Nat.succ_eq_add_one, List.findIdx?_cons, h, decide_True, ↓reduceIte]
|
||||
· symm
|
||||
simp only [Option.map_eq_some']
|
||||
use 0
|
||||
simp only [Fin.val_zero, and_true]
|
||||
rw [@Fin.find_eq_some_iff]
|
||||
simp [h]
|
||||
· rw [findIdx?_on_finRange _ h]
|
||||
erw [ih (p ∘ Fin.succ)]
|
||||
simp only [Option.map_map]
|
||||
have h1 : Option.map Fin.succ (Fin.find (p ∘ Fin.succ)) =
|
||||
Fin.find p := by
|
||||
by_cases hs : (Fin.find p).isSome
|
||||
· rw [@Option.isSome_iff_exists] at hs
|
||||
obtain ⟨i, hi⟩ := hs
|
||||
rw [hi]
|
||||
have hn : i ≠ 0 := by
|
||||
rw [@Fin.find_eq_some_iff] at hi
|
||||
by_contra hn
|
||||
simp_all
|
||||
simp only [Nat.succ_eq_add_one, Option.map_eq_some']
|
||||
use i.pred hn
|
||||
simp only [Fin.succ_pred, and_true]
|
||||
rw [@Fin.find_eq_some_iff] at hi ⊢
|
||||
simp_all only [Function.comp_apply, Fin.succ_pred, true_and]
|
||||
exact fun j hj => (Fin.pred_le_iff hn).mpr (hi.2 j.succ hj)
|
||||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at hs
|
||||
rw [hs]
|
||||
simp only [Nat.succ_eq_add_one, Option.map_eq_none']
|
||||
rw [@Fin.find_eq_none_iff] at hs ⊢
|
||||
exact fun i => hs i.succ
|
||||
rw [← h1]
|
||||
exact Eq.symm (Option.map_map Fin.val Fin.succ (Fin.find (p ∘ Fin.succ)))
|
||||
|
||||
/-!
|
||||
|
||||
## idList
|
||||
|
||||
-/
|
||||
|
||||
/--
|
||||
The list of elements of `Fin l.length` `i` is in `idListFin` if
|
||||
`l.idMap i` appears for the first time in the `i`th spot.
|
||||
|
||||
For example, for the list `['ᵘ¹', 'ᵘ²', 'ᵤ₁']` the `idListFin` is `[0, 1]`.
|
||||
-/
|
||||
def idListFin : List (Fin l.length) := ((List.finRange l.length).filter
|
||||
(fun i => (List.finRange l.length).findIdx? (fun j => l.idMap i = l.idMap j) = i))
|
||||
|
||||
omit [IndexNotation X] [Fintype X]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma idListFin_getDualInOther? : l.idListFin =
|
||||
(List.finRange l.length).filter (fun i => l.getDualInOther? l i = i) := by
|
||||
rw [idListFin]
|
||||
apply List.filter_congr
|
||||
intro x _
|
||||
simp only [decide_eq_decide]
|
||||
rw [findIdx?_on_finRange_eq_findIdx]
|
||||
change Option.map Fin.val (l.getDualInOther? l x) = Option.map Fin.val (some x) ↔ _
|
||||
exact Function.Injective.eq_iff (Option.map_injective (@Fin.val_injective l.length))
|
||||
|
||||
/-- The list of ids keeping only the first appearance of each id. -/
|
||||
def idList : List ℕ := List.map l.idMap l.idListFin
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma idList_getDualInOther? : l.idList =
|
||||
List.map l.idMap ((List.finRange l.length).filter (fun i => l.getDualInOther? l i = i)) := by
|
||||
rw [idList, idListFin_getDualInOther?]
|
||||
|
||||
lemma mem_idList_of_mem {I : Index X} (hI : I ∈ l.val) : I.id ∈ l.idList := by
|
||||
simp only [idList_getDualInOther?, List.mem_map, List.mem_filter, List.mem_finRange,
|
||||
decide_eq_true_eq, true_and]
|
||||
have hI : l.val.indexOf I < l.val.length := List.indexOf_lt_length.mpr hI
|
||||
have hI' : I = l.val.get ⟨l.val.indexOf I, hI⟩ := Eq.symm (List.indexOf_get hI)
|
||||
rw [hI']
|
||||
use (l.getDualInOther? l ⟨l.val.indexOf I, hI⟩).get (
|
||||
getDualInOther?_self_isSome l ⟨List.indexOf I l.val, hI⟩)
|
||||
apply And.intro
|
||||
· exact getDualInOther?_getDualInOther?_get_self l ⟨List.indexOf I l.val, hI⟩
|
||||
· simp only [getDualInOther?_id, List.get_eq_getElem, List.getElem_indexOf]
|
||||
exact Eq.symm ((fun {a b} => Nat.succ_inj'.mp) (congrArg Nat.succ (congrArg Index.id hI')))
|
||||
|
||||
lemma idMap_mem_idList (i : Fin l.length) : l.idMap i ∈ l.idList := by
|
||||
have h : l.val.get i ∈ l.val := by
|
||||
simp only [List.get_eq_getElem]
|
||||
exact List.getElem_mem l.val (↑i) i.isLt
|
||||
exact l.mem_idList_of_mem h
|
||||
|
||||
@[simp]
|
||||
lemma idList_indexOf_mem {I J : Index X} (hI : I ∈ l.val) (hJ : J ∈ l.val) :
|
||||
l.idList.indexOf I.id = l.idList.indexOf J.id ↔ I.id = J.id := by
|
||||
rw [← lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat]
|
||||
exact List.indexOf_inj (l.mem_idList_of_mem hI) (l.mem_idList_of_mem hJ)
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma idList_indexOf_get (i : Fin l.length) :
|
||||
l.idList.indexOf (l.idMap i) = l.idListFin.indexOf ((l.getDualInOther? l i).get
|
||||
(getDualInOther?_self_isSome l i)) := by
|
||||
rw [idList]
|
||||
simp only [idListFin_getDualInOther?]
|
||||
rw [← indexOf_map' l.idMap (fun i => (l.getDualInOther? l i).get
|
||||
(getDualInOther?_self_isSome l i))]
|
||||
· intro i _ j
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· simp [getDualInOther?, AreDualInOther, h]
|
||||
· trans l.idMap ((l.getDualInOther? l i).get
|
||||
(getDualInOther?_self_isSome l i))
|
||||
· exact Eq.symm (getDualInOther?_id l l i (getDualInOther?_self_isSome l i))
|
||||
· trans l.idMap ((l.getDualInOther? l j).get
|
||||
(getDualInOther?_self_isSome l j))
|
||||
· exact congrArg l.idMap h
|
||||
· exact getDualInOther?_id l l j (getDualInOther?_self_isSome l j)
|
||||
· intro i hi
|
||||
simp only [List.mem_filter, List.mem_finRange, decide_eq_true_eq, true_and] at hi
|
||||
exact Option.get_of_mem (getDualInOther?_self_isSome l i) hi
|
||||
· exact fun i => Eq.symm (getDualInOther?_id l l i (getDualInOther?_self_isSome l i))
|
||||
|
||||
/-!
|
||||
|
||||
## GetDualCast
|
||||
|
||||
-/
|
||||
|
||||
/-- Two `IndexList`s are related by `GetDualCast` if they are the same length,
|
||||
and have the same dual structure.
|
||||
|
||||
For example, `['ᵘ¹', 'ᵘ²', 'ᵤ₁']` and `['ᵤ₄', 'ᵘ⁵', 'ᵤ₄']` are related by `GetDualCast`,
|
||||
but `['ᵘ¹', 'ᵘ²', 'ᵤ₁']` and `['ᵘ¹', 'ᵘ²', 'ᵤ₃']` are not. -/
|
||||
def GetDualCast : Prop := l.length = l2.length ∧ ∀ (h : l.length = l2.length),
|
||||
l.getDual? = Option.map (Fin.cast h.symm) ∘ l2.getDual? ∘ Fin.cast h
|
||||
|
||||
namespace GetDualCast
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma symm (h : GetDualCast l l2) : GetDualCast l2 l := by
|
||||
apply And.intro h.1.symm
|
||||
intro h'
|
||||
rw [h.2 h.1]
|
||||
funext i
|
||||
simp only [Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, Option.map_map]
|
||||
have h1 (h : l.length = l2.length) : Fin.cast h ∘ Fin.cast h.symm = id := rfl
|
||||
rw [h1]
|
||||
simp
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma getDual?_isSome_iff (h : GetDualCast l l2) (i : Fin l.length) :
|
||||
(l.getDual? i).isSome ↔ (l2.getDual? (Fin.cast h.1 i)).isSome := by
|
||||
simp [h.2 h.1]
|
||||
|
||||
omit [DecidableEq X]
|
||||
lemma getDual?_get (h : GetDualCast l l2) (i : Fin l.length) (h1 : (l.getDual? i).isSome) :
|
||||
(l.getDual? i).get h1 = Fin.cast h.1.symm ((l2.getDual? (Fin.cast h.1 i)).get
|
||||
((getDual?_isSome_iff h i).mp h1)) := by
|
||||
apply Option.some_inj.mp
|
||||
simp only [Option.some_get]
|
||||
rw [← Option.map_some']
|
||||
simp only [Option.some_get]
|
||||
simp only [h.2 h.1, Function.comp_apply, Option.map_eq_some']
|
||||
|
||||
lemma areDualInSelf_of (h : GetDualCast l l2) (i j : Fin l.length) (hA : l.AreDualInSelf i j) :
|
||||
l2.AreDualInSelf (Fin.cast h.1 i) (Fin.cast h.1 j) := by
|
||||
have hn : (l.getDual? j).isSome := by
|
||||
rw [getDual?_isSome_iff_exists]
|
||||
exact ⟨i, hA.symm⟩
|
||||
have hni : (l.getDual? i).isSome := by
|
||||
rw [getDual?_isSome_iff_exists]
|
||||
exact ⟨j, hA⟩
|
||||
rcases l.getDual?_of_areDualInSelf hA with h1 | h1 | h1
|
||||
· have h1' : (l.getDual? j).get hn = i := Option.get_of_mem hn h1
|
||||
rw [h.getDual?_get] at h1'
|
||||
rw [← h1']
|
||||
simp
|
||||
· have h1' : (l.getDual? i).get hni = j := Option.get_of_mem hni h1
|
||||
rw [h.getDual?_get] at h1'
|
||||
rw [← h1']
|
||||
simp
|
||||
· have h1' : (l.getDual? j).get hn = (l.getDual? i).get hni := by
|
||||
apply Option.some_inj.mp
|
||||
simp [h1]
|
||||
rw [h.getDual?_get, h.getDual?_get] at h1'
|
||||
have h1 : ((l2.getDual? (Fin.cast h.1 j)).get ((getDual?_isSome_iff h j).mp hn))
|
||||
= ((l2.getDual? (Fin.cast h.1 i)).get ((getDual?_isSome_iff h i).mp hni)) := by
|
||||
simpa [Fin.ext_iff] using h1'
|
||||
simp only [AreDualInSelf, ne_eq]
|
||||
apply And.intro
|
||||
· simp only [AreDualInSelf, ne_eq] at hA
|
||||
simpa [Fin.ext_iff] using hA.1
|
||||
trans l2.idMap ((l2.getDual? (Fin.cast h.1 j)).get ((getDual?_isSome_iff h j).mp hn))
|
||||
· trans l2.idMap ((l2.getDual? (Fin.cast h.1 i)).get ((getDual?_isSome_iff h i).mp hni))
|
||||
· exact Eq.symm (getDual?_get_id l2 (Fin.cast h.left i) ((getDual?_isSome_iff h i).mp hni))
|
||||
· exact congrArg l2.idMap (id (Eq.symm h1))
|
||||
· exact getDual?_get_id l2 (Fin.cast h.left j) ((getDual?_isSome_iff h j).mp hn)
|
||||
|
||||
lemma areDualInSelf_iff (h : GetDualCast l l2) (i j : Fin l.length) : l.AreDualInSelf i j ↔
|
||||
l2.AreDualInSelf (Fin.cast h.1 i) (Fin.cast h.1 j) := by
|
||||
apply Iff.intro
|
||||
· exact areDualInSelf_of h i j
|
||||
· intro h'
|
||||
rw [← show Fin.cast h.1.symm (Fin.cast h.1 i) = i by rfl]
|
||||
rw [← show Fin.cast h.1.symm (Fin.cast h.1 j) = j by rfl]
|
||||
exact areDualInSelf_of h.symm _ _ h'
|
||||
|
||||
lemma idMap_eq_of (h : GetDualCast l l2) (i j : Fin l.length) (hm : l.idMap i = l.idMap j) :
|
||||
l2.idMap (Fin.cast h.1 i) = l2.idMap (Fin.cast h.1 j) := by
|
||||
by_cases h1 : i = j
|
||||
· exact congrArg l2.idMap (congrArg (Fin.cast h.left) h1)
|
||||
have h1' : l.AreDualInSelf i j := by
|
||||
simp only [AreDualInSelf, ne_eq]
|
||||
exact ⟨h1, hm⟩
|
||||
rw [h.areDualInSelf_iff] at h1'
|
||||
simpa using h1'.2
|
||||
|
||||
lemma idMap_eq_iff (h : GetDualCast l l2) (i j : Fin l.length) :
|
||||
l.idMap i = l.idMap j ↔ l2.idMap (Fin.cast h.1 i) = l2.idMap (Fin.cast h.1 j) := by
|
||||
apply Iff.intro
|
||||
· exact idMap_eq_of h i j
|
||||
· intro h'
|
||||
rw [← show Fin.cast h.1.symm (Fin.cast h.1 i) = i by rfl]
|
||||
rw [← show Fin.cast h.1.symm (Fin.cast h.1 j) = j by rfl]
|
||||
exact idMap_eq_of h.symm _ _ h'
|
||||
|
||||
lemma iff_idMap_eq : GetDualCast l l2 ↔
|
||||
l.length = l2.length ∧ ∀ (h : l.length = l2.length) i j,
|
||||
l.idMap i = l.idMap j ↔ l2.idMap (Fin.cast h i) = l2.idMap (Fin.cast h j) := by
|
||||
refine Iff.intro (fun h => And.intro h.1 (fun _ i j => idMap_eq_iff h i j))
|
||||
(fun h => And.intro h.1 (fun h' => ?_))
|
||||
funext i
|
||||
simp only [Function.comp_apply]
|
||||
by_cases h1 : (l.getDual? i).isSome
|
||||
· have h1' : (l2.getDual? (Fin.cast h.1 i)).isSome := by
|
||||
simp only [getDual?, Fin.isSome_find_iff] at h1 ⊢
|
||||
obtain ⟨j, hj⟩ := h1
|
||||
use (Fin.cast h.1 j)
|
||||
simp only [AreDualInSelf, ne_eq] at hj ⊢
|
||||
rw [← h.2]
|
||||
simpa [Fin.ext_iff] using hj
|
||||
have h2 := Option.eq_some_of_isSome h1'
|
||||
rw [Option.eq_some_of_isSome h1', Option.map_some']
|
||||
nth_rewrite 1 [getDual?] at h2 ⊢
|
||||
rw [Fin.find_eq_some_iff] at h2 ⊢
|
||||
apply And.intro
|
||||
· apply And.intro
|
||||
· simp only [AreDualInSelf, ne_eq, getDual?_get_id, and_true, and_imp] at h2
|
||||
simpa [Fin.ext_iff] using h2.1
|
||||
· rw [h.2 h.1]
|
||||
simp
|
||||
· intro j hj
|
||||
apply h2.2 (Fin.cast h.1 j)
|
||||
simp only [AreDualInSelf, ne_eq] at hj ⊢
|
||||
apply And.intro
|
||||
· simpa [Fin.ext_iff] using hj.1
|
||||
· rw [← h.2]
|
||||
simpa using hj.2
|
||||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h1
|
||||
rw [h1]
|
||||
symm
|
||||
refine Option.map_eq_none'.mpr ?_
|
||||
rw [getDual?, Fin.find_eq_none_iff] at h1 ⊢
|
||||
simp only [AreDualInSelf, ne_eq, not_and]
|
||||
intro j hij
|
||||
have h1j := h1 (Fin.cast h.1.symm j)
|
||||
simp only [AreDualInSelf, ne_eq, not_and] at h1j
|
||||
rw [h.2 h.1] at h1j
|
||||
apply h1j
|
||||
exact ne_of_apply_ne (Fin.cast h') hij
|
||||
|
||||
lemma refl (l : IndexList X) : GetDualCast l l := by
|
||||
rw [iff_idMap_eq]
|
||||
simp
|
||||
|
||||
@[trans]
|
||||
lemma trans (h : GetDualCast l l2) (h' : GetDualCast l2 l3) : GetDualCast l l3 := by
|
||||
rw [iff_idMap_eq] at h h' ⊢
|
||||
refine And.intro (h.1.trans h'.1) (fun hn i j => ?_)
|
||||
rw [h.2 h.1, h'.2 h'.1]
|
||||
rfl
|
||||
|
||||
lemma getDualInOther?_get (h : GetDualCast l l2) (i : Fin l.length) :
|
||||
(l.getDualInOther? l i).get (getDualInOther?_self_isSome l i) = (Fin.cast h.1.symm)
|
||||
((l2.getDualInOther? l2 (Fin.cast h.1 i)).get
|
||||
(getDualInOther?_self_isSome l2 (Fin.cast h.left i))) := by
|
||||
apply Option.some_inj.mp
|
||||
simp only [Option.some_get]
|
||||
erw [Fin.find_eq_some_iff]
|
||||
have h1 := Option.eq_some_of_isSome (getDualInOther?_self_isSome l2 (Fin.cast h.left i))
|
||||
erw [Fin.find_eq_some_iff] at h1
|
||||
apply And.intro
|
||||
· simp only [AreDualInOther]
|
||||
rw [idMap_eq_iff h]
|
||||
simp
|
||||
· intro j hj
|
||||
apply h1.2 (Fin.cast h.1 j)
|
||||
simp only [AreDualInOther]
|
||||
exact (idMap_eq_iff h i j).mp hj
|
||||
|
||||
lemma countId_cast (h : GetDualCast l l2) (i : Fin l.length) :
|
||||
l.countId (l.val.get i) = l2.countId (l2.val.get (Fin.cast h.1 i)) := by
|
||||
rw [countId_get, countId_get]
|
||||
simp only [add_left_inj]
|
||||
have h1 : List.finRange l2.length = List.map (Fin.cast h.1) (List.finRange l.length) := by
|
||||
rw [List.ext_get_iff]
|
||||
simp [h.1]
|
||||
rw [h1]
|
||||
rw [List.countP_map]
|
||||
apply List.countP_congr
|
||||
intro j _
|
||||
simp only [decide_eq_true_eq, Function.comp_apply]
|
||||
exact areDualInSelf_iff h i j
|
||||
|
||||
lemma idListFin_cast (h : GetDualCast l l2) :
|
||||
List.map (Fin.cast h.1) l.idListFin = l2.idListFin := by
|
||||
rw [idListFin_getDualInOther?]
|
||||
have h1 : List.finRange l.length = List.map (Fin.cast h.1.symm) (List.finRange l2.length) := by
|
||||
rw [List.ext_get_iff]
|
||||
simp [h.1]
|
||||
rw [h1]
|
||||
rw [List.filter_map]
|
||||
have h1 : Fin.cast h.1 ∘ Fin.cast h.1.symm = id := rfl
|
||||
simp only [List.map_map, h1, List.map_id]
|
||||
rw [idListFin_getDualInOther?]
|
||||
apply List.filter_congr
|
||||
intro i _
|
||||
simp only [getDualInOther?, AreDualInOther, Fin.find_eq_some_iff, true_and, Function.comp_apply,
|
||||
decide_eq_decide]
|
||||
refine Iff.intro (fun h' j hj => ?_) (fun h' j hj => ?_)
|
||||
· simpa using h' (Fin.cast h.1.symm j) (by rw [h.idMap_eq_iff]; exact hj)
|
||||
· simpa using h' (Fin.cast h.1 j) (by rw [h.symm.idMap_eq_iff]; exact hj)
|
||||
|
||||
lemma idList_indexOf (h : GetDualCast l l2) (i : Fin l.length) :
|
||||
l2.idList.indexOf (l2.idMap (Fin.cast h.1 i)) = l.idList.indexOf (l.idMap i) := by
|
||||
rw [idList_indexOf_get, idList_indexOf_get, ← idListFin_cast h, getDualInOther?_get h.symm]
|
||||
simp only [Fin.cast_trans, Fin.cast_eq_self]
|
||||
rw [indexOf_map_fin (Fin.cast h.1)]
|
||||
exact fun _ _ _ => eq_iff_eq_of_cmp_eq_cmp rfl
|
||||
|
||||
end GetDualCast
|
||||
|
||||
/-!
|
||||
|
||||
## Normalized index list
|
||||
|
||||
-/
|
||||
|
||||
/--
|
||||
Given an index list `l`, the corresponding index list where id's are set to
|
||||
sequentially lowest values.
|
||||
|
||||
E.g. on `['ᵘ¹', 'ᵘ³', 'ᵤ₁']` this gives `['ᵘ⁰', 'ᵘ¹', 'ᵤ₀']`.
|
||||
|
||||
-/
|
||||
def normalize : IndexList X where
|
||||
val := List.ofFn (fun i => ⟨l.colorMap i, l.idList.indexOf (l.idMap i)⟩)
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma normalize_eq_map :
|
||||
l.normalize.val = List.map (fun I => ⟨I.toColor, l.idList.indexOf I.id⟩) l.val := by
|
||||
have hl : l.val = List.map l.val.get (List.finRange l.length) := by
|
||||
simp [length]
|
||||
rw [hl]
|
||||
simp only [normalize, List.map_map]
|
||||
exact List.ofFn_eq_map
|
||||
|
||||
omit [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma normalize_length : l.normalize.length = l.length := by
|
||||
simp [normalize, idList, length]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma normalize_val_length : l.normalize.val.length = l.val.length := by
|
||||
simp [normalize, idList, length]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma normalize_colorMap : l.normalize.colorMap = l.colorMap ∘ Fin.cast l.normalize_length := by
|
||||
funext x
|
||||
simp [colorMap, normalize, Index.toColor]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma colorMap_normalize :
|
||||
l.colorMap = l.normalize.colorMap ∘ Fin.cast l.normalize_length.symm := by
|
||||
funext x
|
||||
simp [colorMap, normalize, Index.toColor]
|
||||
|
||||
@[simp]
|
||||
lemma normalize_countId (i : Fin l.normalize.length) :
|
||||
l.normalize.countId l.normalize.val[Fin.val i] =
|
||||
l.countId (l.val.get ⟨i, by simpa using i.2⟩) := by
|
||||
rw [countId, countId]
|
||||
simp only [Index.id, l.normalize_eq_map, List.getElem_map, List.countP_map, List.get_eq_getElem]
|
||||
apply List.countP_congr
|
||||
intro J hJ
|
||||
simp only [Function.comp_apply, decide_eq_true_eq]
|
||||
trans List.indexOf (l.val.get ⟨i, by simpa using i.2⟩).id l.idList = List.indexOf J.id l.idList
|
||||
· rfl
|
||||
· simp only [List.get_eq_getElem]
|
||||
rw [idList_indexOf_mem]
|
||||
· rfl
|
||||
· exact List.getElem_mem l.val _ _
|
||||
· exact hJ
|
||||
|
||||
@[simp]
|
||||
lemma normalize_countId' (i : Fin l.length) (hi : i.1 < l.normalize.length) :
|
||||
l.normalize.countId (l.normalize.val[Fin.val i]) = l.countId (l.val.get i) := by
|
||||
rw [countId, countId]
|
||||
simp only [Index.id, l.normalize_eq_map, List.getElem_map, List.countP_map, List.get_eq_getElem]
|
||||
apply List.countP_congr
|
||||
intro J hJ
|
||||
simp only [Function.comp_apply, decide_eq_true_eq]
|
||||
trans List.indexOf (l.val.get i).id l.idList = List.indexOf J.id l.idList
|
||||
· rfl
|
||||
· simp only [List.get_eq_getElem]
|
||||
rw [idList_indexOf_mem]
|
||||
· rfl
|
||||
· exact List.getElem_mem l.val _ _
|
||||
· exact hJ
|
||||
|
||||
@[simp]
|
||||
lemma normalize_countId_mem (I : Index X) (h : I ∈ l.val) :
|
||||
l.normalize.countId (I.toColor, List.indexOf I.id l.idList) = l.countId I := by
|
||||
have h1 : l.val.indexOf I < l.val.length := List.indexOf_lt_length.mpr h
|
||||
have h2 : I = l.val.get ⟨l.val.indexOf I, h1⟩ := Eq.symm (List.indexOf_get h1)
|
||||
have h3 : (I.toColor, List.indexOf I.id l.idList) = l.normalize.val.get
|
||||
⟨l.val.indexOf I, by simpa using h1⟩ := by
|
||||
simp only [List.get_eq_getElem, l.normalize_eq_map, List.getElem_map, List.getElem_indexOf]
|
||||
rw [h3]
|
||||
simp only [List.get_eq_getElem]
|
||||
trans l.countId (l.val.get ⟨l.val.indexOf I, h1⟩)
|
||||
· rw [← normalize_countId']
|
||||
· exact congrArg l.countId (id (Eq.symm h2))
|
||||
|
||||
lemma normalize_filter_countId_eq_one :
|
||||
List.map Index.id (l.normalize.val.filter (fun I => l.normalize.countId I = 1))
|
||||
= List.map l.idList.indexOf (List.map Index.id (l.val.filter (fun I => l.countId I = 1))) := by
|
||||
nth_rewrite 1 [l.normalize_eq_map]
|
||||
rw [List.filter_map]
|
||||
simp only [List.map_map]
|
||||
congr 1
|
||||
apply List.filter_congr
|
||||
intro I hI
|
||||
simp only [Function.comp_apply, decide_eq_decide]
|
||||
rw [normalize_countId_mem]
|
||||
exact hI
|
||||
|
||||
omit [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma normalize_idMap_apply (i : Fin l.normalize.length) :
|
||||
l.normalize.idMap i = l.idList.indexOf (l.val.get ⟨i, by simpa using i.2⟩).id := by
|
||||
simp [idMap, normalize, Index.id]
|
||||
|
||||
@[simp]
|
||||
lemma normalize_getDualCast_self : GetDualCast l.normalize l := by
|
||||
rw [GetDualCast.iff_idMap_eq]
|
||||
apply And.intro l.normalize_length
|
||||
intro h i j
|
||||
simp only [normalize_idMap_apply, List.get_eq_getElem]
|
||||
refine Iff.intro (fun h' => ?_) (fun h' => ?_)
|
||||
· refine (List.indexOf_inj (idMap_mem_idList l (Fin.cast h i))
|
||||
(idMap_mem_idList l (Fin.cast h j))).mp ?_
|
||||
rw [lawful_beq_subsingleton instBEqOfDecidableEq instBEqNat]
|
||||
exact h'
|
||||
· exact congrFun (congrArg List.indexOf h') l.idList
|
||||
|
||||
@[simp]
|
||||
lemma self_getDualCast_normalize : GetDualCast l l.normalize := by
|
||||
exact GetDualCast.symm l.normalize_getDualCast_self
|
||||
|
||||
lemma normalize_filter_countId_not_eq_one :
|
||||
List.map Index.id (l.normalize.val.filter (fun I => ¬ l.normalize.countId I = 1))
|
||||
= List.map l.idList.indexOf
|
||||
(List.map Index.id (l.val.filter (fun I => ¬ l.countId I = 1))) := by
|
||||
nth_rewrite 1 [l.normalize_eq_map]
|
||||
rw [List.filter_map]
|
||||
simp only [decide_not, List.map_map]
|
||||
congr 1
|
||||
apply List.filter_congr
|
||||
intro I hI
|
||||
simp only [Function.comp_apply, decide_eq_decide]
|
||||
rw [normalize_countId_mem]
|
||||
exact hI
|
||||
|
||||
namespace GetDualCast
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
lemma to_normalize (h : GetDualCast l l2) : GetDualCast l.normalize l2.normalize := by
|
||||
apply l.normalize_getDualCast_self.trans
|
||||
apply h.trans
|
||||
exact l2.self_getDualCast_normalize
|
||||
|
||||
lemma normalize_idMap_eq (h : GetDualCast l l2) :
|
||||
l.normalize.idMap = l2.normalize.idMap ∘ Fin.cast h.to_normalize.1 := by
|
||||
funext i
|
||||
simp only [normalize_idMap_apply, List.get_eq_getElem, Function.comp_apply, Fin.coe_cast]
|
||||
change List.indexOf (l.idMap (Fin.cast l.normalize_length i)) l.idList = _
|
||||
rw [← idList_indexOf h]
|
||||
rfl
|
||||
|
||||
lemma iff_normalize : GetDualCast l l2 ↔
|
||||
l.normalize.length = l2.normalize.length ∧ ∀ (h : l.normalize.length = l2.normalize.length),
|
||||
l.normalize.idMap = l2.normalize.idMap ∘ Fin.cast h := by
|
||||
refine Iff.intro (fun h => And.intro h.to_normalize.1 (fun _ => normalize_idMap_eq h))
|
||||
(fun h => l.self_getDualCast_normalize.trans
|
||||
(trans (iff_idMap_eq.mpr (And.intro h.1 (fun h' i j => ?_))) l2.normalize_getDualCast_self))
|
||||
rw [h.2 h.1]
|
||||
rfl
|
||||
|
||||
end GetDualCast
|
||||
|
||||
@[simp]
|
||||
lemma normalize_normalize : l.normalize.normalize = l.normalize := by
|
||||
refine ext_colorMap_idMap (normalize_length l.normalize) ?hi (normalize_colorMap l.normalize)
|
||||
exact (l.normalize_getDualCast_self).normalize_idMap_eq
|
||||
|
||||
/-!
|
||||
|
||||
## Equality of normalized forms
|
||||
|
||||
-/
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma normalize_length_eq_of_eq_length (h : l.length = l2.length) :
|
||||
l.normalize.length = l2.normalize.length := by
|
||||
rw [normalize_length, normalize_length, h]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma normalize_colorMap_eq_of_eq_colorMap (h : l.length = l2.length)
|
||||
(hc : l.colorMap = l2.colorMap ∘ Fin.cast h) :
|
||||
l.normalize.colorMap = l2.normalize.colorMap ∘
|
||||
Fin.cast (normalize_length_eq_of_eq_length l l2 h) := by
|
||||
simp only [normalize_colorMap, hc]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Reindexing
|
||||
|
||||
-/
|
||||
|
||||
/--
|
||||
Two `ColorIndexList` are said to be reindexes of one another if they:
|
||||
1. have the same length.
|
||||
2. every corresponding index has the same color, and duals which correspond.
|
||||
|
||||
Note: This does not allow for reordering of indices.
|
||||
-/
|
||||
def Reindexing : Prop := l.length = l2.length ∧
|
||||
∀ (h : l.length = l2.length), l.colorMap = l2.colorMap ∘ Fin.cast h ∧
|
||||
l.getDual? = Option.map (Fin.cast h.symm) ∘ l2.getDual? ∘ Fin.cast h
|
||||
|
||||
namespace Reindexing
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma iff_getDualCast : Reindexing l l2 ↔ GetDualCast l l2
|
||||
∧ ∀ (h : l.length = l2.length), l.colorMap = l2.colorMap ∘ Fin.cast h := by
|
||||
refine Iff.intro (fun h => ⟨⟨h.1, fun h' => (h.2 h').2⟩, fun h' => (h.2 h').1⟩) (fun h => ?_)
|
||||
refine And.intro h.1.1 (fun h' => And.intro (h.2 h') (h.1.2 h'))
|
||||
|
||||
lemma iff_normalize : Reindexing l l2 ↔ l.normalize = l2.normalize := by
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
· rw [iff_getDualCast, GetDualCast.iff_normalize] at h
|
||||
refine ext_colorMap_idMap h.1.1 (h.1.2 (h.1.1)) ?refine_1.hc
|
||||
· refine normalize_colorMap_eq_of_eq_colorMap _ _ ?_ ?_
|
||||
· simpa using h.1.1
|
||||
· apply h.2
|
||||
· rw [iff_getDualCast, GetDualCast.iff_normalize, h]
|
||||
simp only [normalize_length, Fin.cast_refl, Function.comp_id, imp_self, and_self, true_and]
|
||||
rw [colorMap_normalize, colorMap_cast h]
|
||||
intro h'
|
||||
funext i
|
||||
simp
|
||||
|
||||
lemma refl (l : IndexList X) : Reindexing l l :=
|
||||
iff_normalize.mpr rfl
|
||||
|
||||
@[symm]
|
||||
lemma symm (h : Reindexing l l2) : Reindexing l2 l := by
|
||||
rw [iff_normalize] at h ⊢
|
||||
exact h.symm
|
||||
|
||||
@[trans]
|
||||
lemma trans (h : Reindexing l l2) (h' : Reindexing l2 l3) : Reindexing l l3 := by
|
||||
rw [iff_normalize] at h h' ⊢
|
||||
exact h.trans h'
|
||||
|
||||
end Reindexing
|
||||
|
||||
@[simp]
|
||||
lemma normalize_reindexing_self : Reindexing l.normalize l := by
|
||||
rw [Reindexing.iff_normalize]
|
||||
exact l.normalize_normalize
|
||||
|
||||
@[simp]
|
||||
lemma self_reindexing_normalize : Reindexing l l.normalize := by
|
||||
refine Reindexing.symm (normalize_reindexing_self l)
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,176 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.CountId
|
||||
import HepLean.Tensors.IndexNotation.IndexList.Contraction
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
/-!
|
||||
|
||||
# withDuals equal to withUniqueDuals
|
||||
|
||||
In a permissible list of indices every index which has a dual has a unique dual.
|
||||
This corresponds to the condition that `l.withDual = l.withUniqueDual`.
|
||||
|
||||
We prove lemmas relating to this condition.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-- The conditon on an index list which is true if and only if for every index with a dual,
|
||||
that dual is unique. -/
|
||||
def OnlyUniqueDuals : Prop := l.withUniqueDual = l.withDual
|
||||
|
||||
namespace OnlyUniqueDuals
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
omit [IndexNotation X] [Fintype X]
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma iff_unique_forall :
|
||||
l.OnlyUniqueDuals ↔
|
||||
∀ (i : l.withDual) j, l.AreDualInSelf i j → j = l.getDual? i := by
|
||||
refine Iff.intro (fun h i j hj => ?_) (fun h => ?_)
|
||||
· rw [OnlyUniqueDuals, @Finset.ext_iff] at h
|
||||
simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and, and_iff_left_iff_imp] at h
|
||||
refine h i ?_ j hj
|
||||
exact withDual_isSome l i
|
||||
· refine Finset.ext (fun i => ?_)
|
||||
refine Iff.intro (fun hi => mem_withDual_of_mem_withUniqueDual l i hi) (fun hi => ?_)
|
||||
· simp only [withUniqueDual, mem_withDual_iff_isSome, Finset.mem_filter, Finset.mem_univ,
|
||||
true_and]
|
||||
exact And.intro ((mem_withDual_iff_isSome l i).mp hi) (fun j hj => h ⟨i, hi⟩ j hj)
|
||||
|
||||
omit [DecidableEq X] in
|
||||
lemma iff_countId_leq_two :
|
||||
l.OnlyUniqueDuals ↔ ∀ i, l.countId (l.val.get i) ≤ 2 := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· by_cases hi : i ∈ l.withDual
|
||||
· rw [← h, mem_withUniqueDual_iff_countId_eq_two] at hi
|
||||
rw [hi]
|
||||
· rw [mem_withDual_iff_countId_gt_one] at hi
|
||||
exact Nat.le_of_not_ge hi
|
||||
· refine Finset.ext (fun i => ?_)
|
||||
rw [mem_withUniqueDual_iff_countId_eq_two, mem_withDual_iff_countId_gt_one]
|
||||
have hi := h i
|
||||
omega
|
||||
|
||||
lemma iff_countId_leq_two' : l.OnlyUniqueDuals ↔ ∀ I, l.countId I ≤ 2 := by
|
||||
rw [iff_countId_leq_two]
|
||||
refine Iff.intro (fun h I => ?_) (fun h i => h (l.val.get i))
|
||||
by_cases hI : l.countId I = 0
|
||||
· rw [hI]
|
||||
exact Nat.zero_le 2
|
||||
· obtain ⟨I', hI1', hI2'⟩ := countId_neq_zero_mem l I hI
|
||||
rw [countId_congr _ hI2']
|
||||
have hi : List.indexOf I' l.val < l.length := List.indexOf_lt_length.mpr hI1'
|
||||
have hI' : I' = l.val.get ⟨List.indexOf I' l.val, hi⟩ := (List.indexOf_get hi).symm
|
||||
rw [hI']
|
||||
exact h ⟨List.indexOf I' l.val, hi⟩
|
||||
|
||||
/-!
|
||||
|
||||
## On appended lists
|
||||
|
||||
-/
|
||||
|
||||
lemma inl (h : (l ++ l2).OnlyUniqueDuals) : l.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
rw [countId_append] at hI
|
||||
exact Nat.le_of_add_right_le hI
|
||||
|
||||
lemma inr (h : (l ++ l2).OnlyUniqueDuals) : l2.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
rw [countId_append] at hI
|
||||
exact le_of_add_le_right hI
|
||||
|
||||
lemma symm' (h : (l ++ l2).OnlyUniqueDuals) : (l2 ++ l).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
rw [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
omega
|
||||
|
||||
lemma symm : (l ++ l2).OnlyUniqueDuals ↔ (l2 ++ l).OnlyUniqueDuals := by
|
||||
refine Iff.intro symm' symm'
|
||||
|
||||
lemma swap (h : (l ++ l2 ++ l3).OnlyUniqueDuals) : (l2 ++ l ++ l3).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
intro I
|
||||
have hI := h I
|
||||
simp only [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
omega
|
||||
|
||||
/-!
|
||||
|
||||
## Contractions
|
||||
|
||||
-/
|
||||
|
||||
lemma contrIndexList (h : l.OnlyUniqueDuals) : l.contrIndexList.OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h ⊢
|
||||
exact fun I => Nat.le_trans (countId_contrIndexList_le_countId l I) (h I)
|
||||
|
||||
lemma contrIndexList_left (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l.contrIndexList ++ l2).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h1 ⊢
|
||||
intro I
|
||||
have hI := h1 I
|
||||
simp only [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
exact add_le_of_add_le_right hI h1
|
||||
|
||||
lemma contrIndexList_right (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l ++ l2.contrIndexList).OnlyUniqueDuals := by
|
||||
rw [symm] at h1 ⊢
|
||||
exact contrIndexList_left h1
|
||||
|
||||
lemma contrIndexList_append (h1 : (l ++ l2).OnlyUniqueDuals) :
|
||||
(l.contrIndexList ++ l2.contrIndexList).OnlyUniqueDuals := by
|
||||
rw [iff_countId_leq_two'] at h1 ⊢
|
||||
intro I
|
||||
have hI := h1 I
|
||||
simp only [countId_append] at hI
|
||||
simp only [countId_append, ge_iff_le]
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
have h2 := countId_contrIndexList_le_countId l2 I
|
||||
omega
|
||||
|
||||
end OnlyUniqueDuals
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma countId_of_OnlyUniqueDuals (h : l.OnlyUniqueDuals) (I : Index X) :
|
||||
l.countId I ≤ 2 := by
|
||||
rw [OnlyUniqueDuals.iff_countId_leq_two'] at h
|
||||
exact h I
|
||||
|
||||
omit [IndexNotation X] [Fintype X] in
|
||||
lemma countId_eq_two_ofcontrIndexList_left_of_OnlyUniqueDuals
|
||||
(h : (l ++ l2).OnlyUniqueDuals) (I : Index X) (h' : (l.contrIndexList ++ l2).countId I = 2) :
|
||||
(l ++ l2).countId I = 2 := by
|
||||
simp only [countId_append]
|
||||
have h1 := countId_contrIndexList_le_countId l I
|
||||
have h3 := countId_of_OnlyUniqueDuals _ h I
|
||||
simp only [countId_append] at h3 h'
|
||||
omega
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,99 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.CountId
|
||||
import HepLean.Tensors.IndexNotation.IndexList.Contraction
|
||||
import Mathlib.Algebra.Order.Ring.Nat
|
||||
import Mathlib.Data.Finset.Sort
|
||||
import Mathlib.Tactic.FinCases
|
||||
/-!
|
||||
|
||||
# Subperm for index lists.
|
||||
|
||||
-/
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
namespace IndexList
|
||||
|
||||
variable {X : Type} [DecidableEq X]
|
||||
variable (l l2 l3 : IndexList X)
|
||||
|
||||
/-- An IndexList `l` is a subpermutation of `l2` if there are no duals in `l`, and
|
||||
every element of `l` has a unique dual in `l2`. -/
|
||||
def Subperm : Prop := l.withUniqueDualInOther l2 = Finset.univ
|
||||
|
||||
namespace Subperm
|
||||
|
||||
variable {l l2 l3 : IndexList X}
|
||||
|
||||
lemma iff_countId_fin : l.Subperm l2 ↔
|
||||
∀ i, l.countId (l.val.get i) = 1 ∧ l2.countId (l.val.get i) = 1 := by
|
||||
refine Iff.intro (fun h i => ?_) (fun h => ?_)
|
||||
· have hi : i ∈ l.withUniqueDualInOther l2 := by
|
||||
rw [h]
|
||||
exact Finset.mem_univ i
|
||||
exact countId_eq_one_of_mem_withUniqueDualInOther l l2 i hi
|
||||
· rw [Subperm, Finset.eq_univ_iff_forall]
|
||||
exact fun x => mem_withUniqueDualInOther_of_countId_eq_one l l2 x (h x)
|
||||
|
||||
lemma iff_countId : l.Subperm l2 ↔
|
||||
∀ I, l.countId I ≤ 1 ∧ (l.countId I = 1 → l2.countId I = 1) := by
|
||||
rw [iff_countId_fin]
|
||||
refine Iff.intro (fun h I => ?_) (fun h i => ?_)
|
||||
· by_cases h' : l.countId I = 0
|
||||
· simp_all
|
||||
· obtain ⟨I', hI'1, hI'2⟩ := l.countId_neq_zero_mem I h'
|
||||
rw [countId_congr _ hI'2, countId_congr _ hI'2]
|
||||
have hi' : l.val.indexOf I' < l.val.length := List.indexOf_lt_length.mpr hI'1
|
||||
have hIi' : I' = l.val.get ⟨l.val.indexOf I', hi'⟩ := (List.indexOf_get (hi')).symm
|
||||
have hi' := h ⟨l.val.indexOf I', hi'⟩
|
||||
rw [← hIi'] at hi'
|
||||
rw [hi'.1, hi'.2]
|
||||
exact if_false_right.mp (congrFun rfl)
|
||||
· have hi := h (l.val.get i)
|
||||
have h1 : l.countId (l.val.get i) ≠ 0 := countId_index_neq_zero l i
|
||||
omega
|
||||
|
||||
lemma trans (h1 : l.Subperm l2) (h2 : l2.Subperm l3) : l.Subperm l3 := by
|
||||
rw [iff_countId] at *
|
||||
intro I
|
||||
have h1 := h1 I
|
||||
have h2 := h2 I
|
||||
omega
|
||||
|
||||
lemma fst_eq_contrIndexList (h : l.Subperm l2) : l.contrIndexList = l := by
|
||||
rw [iff_countId] at *
|
||||
apply ext
|
||||
simp only [contrIndexList, List.filter_eq_self, decide_eq_true_eq]
|
||||
intro I hI
|
||||
have h' := countId_mem l I hI
|
||||
have hI' := h I
|
||||
omega
|
||||
|
||||
lemma symm (h : l.Subperm l2) (hl : l.length = l2.length) : l2.Subperm l := by
|
||||
refine (Finset.card_eq_iff_eq_univ (l2.withUniqueDualInOther l)).mp ?_
|
||||
trans (Finset.map ⟨Subtype.val, Subtype.val_injective⟩ (Equiv.finsetCongr
|
||||
(l.getDualInOtherEquiv l2) Finset.univ)).card
|
||||
· apply congrArg
|
||||
simp [Finset.ext_iff]
|
||||
· trans l2.length
|
||||
· simp only [Finset.univ_eq_attach, Equiv.finsetCongr_apply, Finset.card_map,
|
||||
Finset.card_attach]
|
||||
rw [h, ← hl]
|
||||
exact Finset.card_fin l.length
|
||||
· exact Eq.symm (Finset.card_fin l2.length)
|
||||
|
||||
lemma contr_refl : l.contrIndexList.Subperm l.contrIndexList := by
|
||||
rw [iff_countId]
|
||||
intro I
|
||||
simp only [imp_self, and_true]
|
||||
exact countId_contrIndexList_le_one l I
|
||||
|
||||
end Subperm
|
||||
|
||||
end IndexList
|
||||
|
||||
end IndexNotation
|
|
@ -1,108 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.TensorIndex
|
||||
/-!
|
||||
|
||||
# Strings of indices
|
||||
|
||||
A string of indices e.g. `ᵘ¹²ᵤ₄₃` is the structure we usually see
|
||||
following a tensor symbol in index notation.
|
||||
|
||||
This file defines such an index string, and from it constructs a list of indices.
|
||||
|
||||
-/
|
||||
|
||||
open Lean
|
||||
open String
|
||||
|
||||
namespace IndexNotation
|
||||
|
||||
variable (X : Type) [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of an index string
|
||||
|
||||
-/
|
||||
|
||||
/-- Takes in a list of characters and splits it into a list of characters at those characters
|
||||
which are notation characters e.g. `'ᵘ'`. -/
|
||||
def stringToPreIndexList (l : List Char) : List (List Char) :=
|
||||
let indexHeads := l.filter (fun c => isNotationChar X c)
|
||||
let indexTails := l.splitOnP (fun c => isNotationChar X c)
|
||||
let l' := List.zip indexHeads indexTails.tail
|
||||
l'.map fun x => x.1 :: x.2
|
||||
|
||||
/-- A bool which is true given a list of characters if and only if everl element
|
||||
of the corresponding `stringToPreIndexList` correspondings to an index. -/
|
||||
def listCharIsIndexString (l : List Char) : Bool :=
|
||||
(stringToPreIndexList X l).all fun l => (listCharIndex X l && l ≠ [])
|
||||
|
||||
/-- A string of indices to be associated with a tensor. For example, `ᵘ⁰ᵤ₂₆₀ᵘ³`. -/
|
||||
def IndexString : Type := {s : String // listCharIsIndexString X s.toList = true}
|
||||
|
||||
namespace IndexString
|
||||
|
||||
/-!
|
||||
|
||||
## Constructing a list of indices from an index string
|
||||
|
||||
-/
|
||||
|
||||
variable {X : Type} [IndexNotation X] [Fintype X] [DecidableEq X]
|
||||
|
||||
/-- The character list associated with a index string. -/
|
||||
def toCharList (s : IndexString X) : List Char := s.val.toList
|
||||
|
||||
/-! TODO: Move. -/
|
||||
/-- Takes a list of characters to the correpsonding index if it exists else to `none`. -/
|
||||
def charListToOptionIndex (l : List Char) : Option (Index X) :=
|
||||
if h : listCharIndex X l && l ≠ [] then
|
||||
some (IndexRep.ofCharList l (by simpa using h)).toIndex
|
||||
else
|
||||
none
|
||||
|
||||
/-- The indices associated to an index string. -/
|
||||
def toIndexList (s : IndexString X) : IndexList X :=
|
||||
{val :=
|
||||
(stringToPreIndexList X s.toCharList).filterMap fun l => charListToOptionIndex l}
|
||||
|
||||
/-- The formation of an index list from a string `s` statisfying `listCharIndexStringBool`. -/
|
||||
def toIndexList' (s : String) (hs : listCharIsIndexString X s.toList = true) : IndexList X :=
|
||||
toIndexList ⟨s, hs⟩
|
||||
|
||||
end IndexString
|
||||
|
||||
end IndexNotation
|
||||
namespace TensorStructure
|
||||
|
||||
/-!
|
||||
|
||||
## Making a tensor index from an index string
|
||||
|
||||
-/
|
||||
|
||||
namespace TensorIndex
|
||||
variable {R : Type} [CommSemiring R] (𝓣 : TensorStructure R)
|
||||
variable {𝓣 : TensorStructure R} [IndexNotation 𝓣.Color] [Fintype 𝓣.Color] [DecidableEq 𝓣.Color]
|
||||
variable {n m : ℕ} {cn : Fin n → 𝓣.Color} {cm : Fin m → 𝓣.Color}
|
||||
|
||||
open IndexNotation ColorIndexList IndexString
|
||||
|
||||
/-- The construction of a tensor index from a tensor and a string satisfing conditions which are
|
||||
easy to check automatically. -/
|
||||
noncomputable def fromIndexString (T : 𝓣.Tensor cn) (s : String)
|
||||
(hs : listCharIsIndexString 𝓣.toTensorColor.Color s.toList = true)
|
||||
(hn : n = (toIndexList' s hs).length)
|
||||
(hD : IndexList.OnlyUniqueDuals (toIndexList' s hs))
|
||||
(hC : IndexList.ColorCond (toIndexList' s hs))
|
||||
(hd : TensorColor.ColorMap.DualMap (toIndexList' s hs).colorMap
|
||||
(cn ∘ Fin.cast hn.symm)) : 𝓣.TensorIndex :=
|
||||
TensorStructure.TensorIndex.mkDualMap T ⟨(toIndexList' s hs), hD, hC⟩ hn hd
|
||||
|
||||
end TensorIndex
|
||||
|
||||
end TensorStructure
|
|
@ -1,584 +0,0 @@
|
|||
/-
|
||||
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.Tensors.IndexNotation.IndexList.Color
|
||||
import HepLean.Tensors.IndexNotation.ColorIndexList.ContrPerm
|
||||
import HepLean.Tensors.IndexNotation.ColorIndexList.Append
|
||||
import HepLean.Tensors.Basic
|
||||
import HepLean.Tensors.RisingLowering
|
||||
import HepLean.Tensors.Contraction
|
||||
/-!
|
||||
|
||||
# The structure of a tensor with a string of indices
|
||||
|
||||
-/
|
||||
|
||||
/-! TODO: Introduce a way to change an index from e.g. `ᵘ¹` to `ᵘ²`.
|
||||
Would be nice to have a tactic that did this automatically. -/
|
||||
|
||||
namespace TensorStructure
|
||||
noncomputable section
|
||||
|
||||
open TensorColor
|
||||
open IndexNotation
|
||||
|
||||
variable {R : Type} [CommSemiring R] (𝓣 : TensorStructure R)
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color}
|
||||
{cW : W → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν η : 𝓣.Color}
|
||||
|
||||
variable [IndexNotation 𝓣.Color]
|
||||
|
||||
/-- The structure an tensor with a index specification e.g. `ᵘ¹ᵤ₂`. -/
|
||||
structure TensorIndex extends ColorIndexList 𝓣.toTensorColor where
|
||||
/-- The underlying tensor. -/
|
||||
tensor : 𝓣.Tensor toColorIndexList.colorMap'
|
||||
|
||||
namespace TensorIndex
|
||||
|
||||
open TensorColor ColorIndexList
|
||||
|
||||
variable {𝓣 : TensorStructure R} [IndexNotation 𝓣.Color] [DecidableEq 𝓣.Color]
|
||||
variable {n m : ℕ} {cn : Fin n → 𝓣.Color} {cm : Fin m → 𝓣.Color}
|
||||
|
||||
instance : Coe 𝓣.TensorIndex (ColorIndexList 𝓣.toTensorColor) where
|
||||
coe T := T.toColorIndexList
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma colormap_mapIso {T₁ T₂ : 𝓣.TensorIndex} (hi : T₁.toColorIndexList = T₂.toColorIndexList) :
|
||||
ColorMap.MapIso (Fin.castOrderIso (congrArg IndexList.length (congrArg toIndexList hi))).toEquiv
|
||||
T₁.colorMap' T₂.colorMap' := by
|
||||
cases T₁; cases T₂
|
||||
simp only [ColorMap.MapIso, RelIso.coe_fn_toEquiv]
|
||||
simp only at hi
|
||||
rename_i a b c d
|
||||
cases a
|
||||
cases c
|
||||
rename_i a1 a2 a3 a4 a5 a6
|
||||
cases a1
|
||||
cases a4
|
||||
simp only [ColorIndexList.mk.injEq, IndexList.mk.injEq] at hi
|
||||
subst hi
|
||||
rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma ext {T₁ T₂ : 𝓣.TensorIndex} (hi : T₁.toColorIndexList = T₂.toColorIndexList)
|
||||
(h : T₁.tensor = 𝓣.mapIso (Fin.castOrderIso (by simp [IndexList.length, hi])).toEquiv
|
||||
(colormap_mapIso hi.symm) T₂.tensor) : T₁ = T₂ := by
|
||||
cases T₁; cases T₂
|
||||
simp only at h
|
||||
simp_all
|
||||
simp only at hi
|
||||
subst hi
|
||||
simp_all
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma index_eq_of_eq {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ = T₂) :
|
||||
T₁.toColorIndexList = T₂.toColorIndexList := by
|
||||
cases h
|
||||
rfl
|
||||
|
||||
/-- Given an equality of `TensorIndex`, the isomorphism taking one underlying
|
||||
tensor to the other. -/
|
||||
@[simp]
|
||||
def tensorIso {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ = T₂) :
|
||||
𝓣.Tensor T₂.colorMap' ≃ₗ[R] 𝓣.Tensor T₁.colorMap' :=
|
||||
𝓣.mapIso (Fin.castOrderIso (by rw [index_eq_of_eq h])).toEquiv
|
||||
(colormap_mapIso (index_eq_of_eq h).symm)
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma tensor_eq_of_eq {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ = T₂) :
|
||||
T₁.tensor = tensorIso h T₂.tensor := by
|
||||
have hi := index_eq_of_eq h
|
||||
cases T₁
|
||||
cases T₂
|
||||
simp only at hi
|
||||
subst hi
|
||||
simpa using h
|
||||
|
||||
/-- The construction of a `TensorIndex` from a tensor, a IndexListColor, and a condition
|
||||
on the dual maps. -/
|
||||
def mkDualMap (T : 𝓣.Tensor cn) (l : ColorIndexList 𝓣.toTensorColor) (hn : n = l.1.length)
|
||||
(hd : ColorMap.DualMap l.colorMap' (cn ∘ Fin.cast hn.symm)) :
|
||||
𝓣.TensorIndex where
|
||||
toColorIndexList := l
|
||||
tensor :=
|
||||
𝓣.mapIso (Equiv.refl _) (ColorMap.DualMap.split_dual' (by simpa using hd)) <|
|
||||
𝓣.dualize (ColorMap.DualMap.split l.colorMap' (cn ∘ Fin.cast hn.symm)) <|
|
||||
(𝓣.mapIso (Fin.castOrderIso hn).toEquiv rfl T : 𝓣.Tensor (cn ∘ Fin.cast hn.symm))
|
||||
|
||||
/-!
|
||||
|
||||
## The contraction of indices
|
||||
|
||||
-/
|
||||
|
||||
/-- The contraction of indices in a `TensorIndex`. -/
|
||||
def contr (T : 𝓣.TensorIndex) : 𝓣.TensorIndex where
|
||||
toColorIndexList := T.toColorIndexList.contr
|
||||
tensor := 𝓣.mapIso (Equiv.refl _) T.contrEquiv_colorMapIso <|
|
||||
𝓣.contr T.toColorIndexList.contrEquiv T.contrEquiv_contrCond T.tensor
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma contr_tensor (T : 𝓣.TensorIndex) :
|
||||
T.contr.tensor = ((𝓣.mapIso (Equiv.refl _) T.contrEquiv_colorMapIso <|
|
||||
𝓣.contr T.toColorIndexList.contrEquiv T.contrEquiv_contrCond T.tensor)) := by
|
||||
rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
/-- Applying contr to a tensor whose indices has no contracts does not do anything. -/
|
||||
@[simp]
|
||||
lemma contr_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = ∅) :
|
||||
T.contr = T := by
|
||||
refine ext ?_ ?_
|
||||
· simp only [contr, ColorIndexList.contr]
|
||||
have hx := T.contrIndexList_of_withDual_empty h
|
||||
apply ColorIndexList.ext
|
||||
simp only [hx]
|
||||
· simp only [contr]
|
||||
cases T
|
||||
rename_i i T
|
||||
simp only
|
||||
refine PiTensorProduct.induction_on' T ?_ (by
|
||||
intro a b hx hy
|
||||
simp [map_add, add_mul, hx, hy])
|
||||
intro r f
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, mapIso_tprod,
|
||||
id_eq, eq_mpr_eq_cast, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv]
|
||||
apply congrArg
|
||||
have hEm : IsEmpty { x // x ∈ i.withUniqueDualLT } := by
|
||||
rw [Finset.isEmpty_coe_sort]
|
||||
rw [Finset.eq_empty_iff_forall_not_mem]
|
||||
intro x hx
|
||||
have hx' : x ∈ i.withUniqueDual := by
|
||||
exact Finset.mem_of_mem_filter x hx
|
||||
rw [← i.unique_duals] at h
|
||||
rw [h] at hx'
|
||||
simp_all only [Finset.not_mem_empty]
|
||||
erw [TensorStructure.contr_tprod_isEmpty]
|
||||
erw [mapIso_tprod]
|
||||
simp only [Equiv.refl_symm, Equiv.refl_apply, colorMap', mapIso_tprod, id_eq,
|
||||
OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv]
|
||||
apply congrArg
|
||||
funext l
|
||||
rw [← LinearEquiv.symm_apply_eq]
|
||||
simp only [colorModuleCast, Equiv.cast_symm, OrderIso.toEquiv_symm, RelIso.coe_fn_toEquiv,
|
||||
Function.comp_apply, LinearEquiv.coe_mk, Equiv.cast_apply, LinearEquiv.coe_symm_mk, cast_cast]
|
||||
apply cast_eq_iff_heq.mpr
|
||||
let hl := i.contrEquiv_on_withDual_empty l h
|
||||
exact congr_arg_heq f hl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma contr_tensor_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = ∅) :
|
||||
T.contr.tensor = tensorIso (contr_of_withDual_empty T h) T.tensor := by
|
||||
exact tensor_eq_of_eq (contr_of_withDual_empty T h)
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma contr_contr (T : 𝓣.TensorIndex) : T.contr.contr = T.contr :=
|
||||
T.contr.contr_of_withDual_empty (by simp [contr, ColorIndexList.contr])
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma contr_toColorIndexList (T : 𝓣.TensorIndex) :
|
||||
T.contr.toColorIndexList = T.toColorIndexList.contr := rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma contr_toIndexList (T : 𝓣.TensorIndex) :
|
||||
T.contr.toIndexList = T.toIndexList.contrIndexList := rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Equivalence relation on `TensorIndex`
|
||||
|
||||
-/
|
||||
|
||||
/-- An (equivalence) relation on two `TensorIndex`.
|
||||
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
|
||||
in the indices or moved to the tensor itself. These two descriptions are equivalent. -/
|
||||
def Rel (T₁ T₂ : 𝓣.TensorIndex) : Prop :=
|
||||
T₁.ContrPerm T₂ ∧ ∀ (h : T₁.ContrPerm T₂),
|
||||
T₁.contr.tensor = 𝓣.mapIso (contrPermEquiv h).symm (contrPermEquiv_colorMap_iso h) T₂.contr.tensor
|
||||
|
||||
namespace Rel
|
||||
|
||||
/-- Rel is reflexive. -/
|
||||
lemma refl (T : 𝓣.TensorIndex) : Rel T T := by
|
||||
apply And.intro
|
||||
· simp only [ContrPerm.refl]
|
||||
· simp only [ContrPerm.refl, contr_toColorIndexList, contr_tensor, contrPermEquiv_refl,
|
||||
Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply, imp_self]
|
||||
|
||||
/-- Rel is symmetric. -/
|
||||
lemma symm {T₁ T₂ : 𝓣.TensorIndex} (h : Rel T₁ T₂) : Rel T₂ T₁ := by
|
||||
apply And.intro h.1.symm
|
||||
intro h'
|
||||
rw [← mapIso_symm]
|
||||
· symm
|
||||
erw [LinearEquiv.symm_apply_eq]
|
||||
rw [h.2]
|
||||
· rfl
|
||||
exact h'.symm
|
||||
|
||||
/-- Rel is transitive. -/
|
||||
lemma trans {T₁ T₂ T₃ : 𝓣.TensorIndex} (h1 : Rel T₁ T₂) (h2 : Rel T₂ T₃) : Rel T₁ T₃ := by
|
||||
apply And.intro ((h1.1.trans h2.1))
|
||||
intro h
|
||||
change _ = (𝓣.mapIso (contrPermEquiv (h1.1.trans h2.1)).symm _) T₃.contr.tensor
|
||||
trans (𝓣.mapIso ((contrPermEquiv h1.1).trans (contrPermEquiv h2.1)).symm (by
|
||||
simp only [contrPermEquiv_trans, contrPermEquiv_symm, contr_toColorIndexList]
|
||||
have h1 := contrPermEquiv_colorMap_iso (ContrPerm.symm (ContrPerm.trans h1.left h2.left))
|
||||
rw [← ColorMap.MapIso.symm'] at h1
|
||||
exact h1)) T₃.contr.tensor
|
||||
· erw [← mapIso_trans]
|
||||
· simp only [LinearEquiv.trans_apply]
|
||||
apply (h1.2 h1.1).trans
|
||||
· apply congrArg
|
||||
exact h2.2 h2.1
|
||||
· congr 1
|
||||
simp only [contrPermEquiv_trans, contrPermEquiv_symm]
|
||||
|
||||
/-- Rel forms an equivalence relation. -/
|
||||
lemma isEquivalence : Equivalence (@Rel _ _ 𝓣 _ _) where
|
||||
refl := Rel.refl
|
||||
symm := Rel.symm
|
||||
trans := Rel.trans
|
||||
|
||||
/-- The equality of tensors corresponding to related tensor indices. -/
|
||||
lemma to_eq {T₁ T₂ : 𝓣.TensorIndex} (h : Rel T₁ T₂) :
|
||||
T₁.contr.tensor = 𝓣.mapIso (contrPermEquiv h.1).symm
|
||||
(contrPermEquiv_colorMap_iso h.1) T₂.contr.tensor := h.2 h.1
|
||||
|
||||
lemma of_withDual_empty {T₁ T₂ : 𝓣.TensorIndex} (h : T₁.ContrPerm T₂)
|
||||
(h1 : T₁.withDual = ∅) (h2 : T₂.withDual = ∅)
|
||||
(hT : T₁.tensor =
|
||||
𝓣.mapIso (permEquiv h h1 h2).symm (permEquiv_colorMap_iso h h1 h2) T₂.tensor) : Rel T₁ T₂ := by
|
||||
apply And.intro h
|
||||
intro h'
|
||||
rw [contr_tensor_of_withDual_empty T₁ h1, contr_tensor_of_withDual_empty T₂ h2, hT]
|
||||
simp only [contr_toColorIndexList, tensorIso, mapIso_mapIso, contrPermEquiv_symm]
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
apply mapIso_ext
|
||||
ext i
|
||||
simp only [permEquiv, Equiv.trans_apply, Equiv.symm_trans_apply, Equiv.symm_symm,
|
||||
IndexList.getDualInOtherEquiv_symm, Equiv.subtypeUnivEquiv_symm_apply,
|
||||
Equiv.subtypeUnivEquiv_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.coe_cast,
|
||||
contrPermEquiv]
|
||||
have hn := congrArg (fun x => x.toIndexList) (contr_of_withDual_empty T₁ h1)
|
||||
have hn2 := congrArg (fun x => x.toIndexList) (contr_of_withDual_empty T₂ h2)
|
||||
simp only [contr_toColorIndexList] at hn hn2
|
||||
rw [IndexList.getDualInOtherEquiv_cast hn2 hn]
|
||||
rfl
|
||||
|
||||
end Rel
|
||||
|
||||
/-- The structure of a Setoid on `𝓣.TensorIndex` induced by `Rel`. -/
|
||||
instance asSetoid : Setoid 𝓣.TensorIndex := ⟨Rel, Rel.isEquivalence⟩
|
||||
|
||||
/-- A tensor index is equivalent to its contraction. -/
|
||||
lemma rel_contr (T : 𝓣.TensorIndex) : T ≈ T.contr := by
|
||||
apply And.intro
|
||||
· simp only [contr_toColorIndexList, ContrPerm.contr_self]
|
||||
· intro h
|
||||
rw [tensor_eq_of_eq T.contr_contr]
|
||||
simp only [contr_toColorIndexList, colorMap', contrPermEquiv_self_contr, OrderIso.toEquiv_symm,
|
||||
Fin.symm_castOrderIso, mapIso_mapIso, tensorIso]
|
||||
trans 𝓣.mapIso (Equiv.refl _) (by rfl) T.contr.tensor
|
||||
· simp only [contr_toColorIndexList, mapIso_refl, LinearEquiv.refl_apply]
|
||||
· rfl
|
||||
|
||||
/-!
|
||||
|
||||
## Scalar multiplication of
|
||||
|
||||
-/
|
||||
|
||||
/-- The scalar multiplication of a `TensorIndex` by an element of `R`. -/
|
||||
instance : SMul R 𝓣.TensorIndex where
|
||||
smul := fun r T => {
|
||||
toColorIndexList := T.toColorIndexList
|
||||
tensor := r • T.tensor}
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma smul_index (r : R) (T : 𝓣.TensorIndex) : (r • T).toColorIndexList = T.toColorIndexList := rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma smul_tensor (r : R) (T : 𝓣.TensorIndex) : (r • T).tensor = r • T.tensor := rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma smul_contr (r : R) (T : 𝓣.TensorIndex) : (r • T).contr = r • T.contr := by
|
||||
refine ext rfl ?_
|
||||
simp only [contr, smul_index, smul_tensor, LinearMapClass.map_smul, Fin.castOrderIso_refl,
|
||||
OrderIso.refl_toEquiv, mapIso_refl, LinearEquiv.refl_apply]
|
||||
|
||||
lemma smul_rel {T₁ T₂ : 𝓣.TensorIndex} (h : T₁ ≈ T₂) (r : R) : r • T₁ ≈ r • T₂ := by
|
||||
apply And.intro h.1
|
||||
intro h1
|
||||
rw [tensor_eq_of_eq (smul_contr r T₁), tensor_eq_of_eq (smul_contr r T₂)]
|
||||
simp only [contr_toColorIndexList, smul_index, Fin.castOrderIso_refl, OrderIso.refl_toEquiv,
|
||||
mapIso_refl, smul_tensor, map_smul, LinearEquiv.refl_apply, contrPermEquiv_symm, tensorIso]
|
||||
apply congrArg
|
||||
exact h.2 h1
|
||||
|
||||
/-!
|
||||
|
||||
## Addition of allowed `TensorIndex`
|
||||
|
||||
-/
|
||||
|
||||
/-- The condition on tensors with indices for their addition to exists.
|
||||
This condition states that the the indices of one tensor are exact permutations of indices
|
||||
of another after contraction. This includes the id of the index and the color.
|
||||
|
||||
This condition is general enough to allow addition of e.g. `ψᵤ₁ᵤ₂ + φᵤ₂ᵤ₁`, but
|
||||
will NOT allow e.g. `ψᵤ₁ᵤ₂ + φᵘ²ᵤ₁`. -/
|
||||
def AddCond (T₁ T₂ : 𝓣.TensorIndex) : Prop := T₁.ContrPerm T₂
|
||||
|
||||
namespace AddCond
|
||||
|
||||
lemma to_PermContr {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
||||
T₁.toColorIndexList.ContrPerm T₂.toColorIndexList := h
|
||||
|
||||
@[symm]
|
||||
lemma symm {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) : AddCond T₂ T₁ := by
|
||||
rw [AddCond] at h
|
||||
exact h.symm
|
||||
|
||||
lemma refl (T : 𝓣.TensorIndex) : AddCond T T := ContrPerm.refl T.toColorIndexList
|
||||
|
||||
lemma trans {T₁ T₂ T₃ : 𝓣.TensorIndex} (h1 : AddCond T₁ T₂) (h2 : AddCond T₂ T₃) :
|
||||
AddCond T₁ T₃ := by
|
||||
rw [AddCond] at h1 h2
|
||||
exact h1.trans h2
|
||||
|
||||
lemma rel_left {T₁ T₁' T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₁ ≈ T₁') :
|
||||
AddCond T₁' T₂ := h'.1.symm.trans h
|
||||
|
||||
lemma rel_right {T₁ T₂ T₂' : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₂ ≈ T₂') :
|
||||
AddCond T₁ T₂' := h.trans h'.1
|
||||
|
||||
end AddCond
|
||||
|
||||
/-- The equivalence between indices after contraction given a `AddCond`. -/
|
||||
@[simp]
|
||||
def addCondEquiv {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
||||
Fin T₁.contr.length ≃ Fin T₂.contr.length := contrPermEquiv h
|
||||
|
||||
lemma addCondEquiv_colorMap {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
||||
ColorMap.MapIso (addCondEquiv h) T₁.contr.colorMap' T₂.contr.colorMap' :=
|
||||
contrPermEquiv_colorMap_iso' h
|
||||
|
||||
/-- The addition of two `TensorIndex` given the condition that, after contraction,
|
||||
their index lists are the same. -/
|
||||
def add (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
𝓣.TensorIndex where
|
||||
toColorIndexList := T₂.toColorIndexList.contr
|
||||
tensor := (𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) T₁.contr.tensor) + T₂.contr.tensor
|
||||
|
||||
/-- Notation for addition of tensor indices. -/
|
||||
notation:71 T₁ "+["h"]" T₂:72 => add T₁ T₂ h
|
||||
|
||||
namespace AddCond
|
||||
|
||||
lemma add_right {T₁ T₂ T₃ : 𝓣.TensorIndex} (h : AddCond T₁ T₃) (h' : AddCond T₂ T₃) :
|
||||
AddCond T₁ (T₂ +[h'] T₃) := by
|
||||
simpa only [AddCond, add] using h.rel_right T₃.rel_contr
|
||||
|
||||
lemma add_left {T₁ T₂ T₃ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : AddCond T₂ T₃) :
|
||||
AddCond (T₁ +[h] T₂) T₃ :=
|
||||
(add_right h'.symm h).symm
|
||||
|
||||
lemma of_add_right' {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₂ T₃} (h : AddCond T₁ (T₂ +[h'] T₃)) :
|
||||
AddCond T₁ T₃ := by
|
||||
change T₁.AddCond T₃.contr at h
|
||||
exact h.rel_right T₃.rel_contr.symm
|
||||
|
||||
lemma of_add_right {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₂ T₃} (h : AddCond T₁ (T₂ +[h'] T₃)) :
|
||||
AddCond T₁ T₂ := h.of_add_right'.trans h'.symm
|
||||
|
||||
lemma of_add_left {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂}
|
||||
(h : AddCond (T₁ +[h'] T₂) T₃) : AddCond T₂ T₃ :=
|
||||
(of_add_right' h.symm).symm
|
||||
|
||||
lemma of_add_left' {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂}
|
||||
(h : AddCond (T₁ +[h'] T₂) T₃) : AddCond T₁ T₃ :=
|
||||
(of_add_right h.symm).symm
|
||||
|
||||
lemma add_left_of_add_right {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₂ T₃}
|
||||
(h : AddCond T₁ (T₂ +[h'] T₃)) : AddCond (T₁ +[of_add_right h] T₂) T₃ := by
|
||||
have h0 := ((of_add_right' h).trans h'.symm)
|
||||
exact (h'.symm.add_right h0).symm
|
||||
|
||||
lemma add_right_of_add_left {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂}
|
||||
(h : AddCond (T₁ +[h'] T₂) T₃) : AddCond T₁ (T₂ +[of_add_left h] T₃) :=
|
||||
(add_left (of_add_left h) (of_add_left' h).symm).symm
|
||||
|
||||
lemma add_comm {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) :
|
||||
AddCond (T₁ +[h] T₂) (T₂ +[h.symm] T₁) := by
|
||||
apply add_right
|
||||
apply add_left
|
||||
exact h.symm
|
||||
|
||||
end AddCond
|
||||
|
||||
@[simp]
|
||||
lemma add_toColorIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
(add T₁ T₂ h).toColorIndexList = T₂.toColorIndexList.contr := rfl
|
||||
|
||||
@[simp]
|
||||
lemma add_tensor (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
(add T₁ T₂ h).tensor =
|
||||
(𝓣.mapIso (addCondEquiv h) (addCondEquiv_colorMap h) T₁.contr.tensor) + T₂.contr.tensor := rfl
|
||||
|
||||
/-- Scalar multiplication commutes with addition. -/
|
||||
lemma smul_add (r : R) (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
r • (T₁ +[h] T₂) = r • T₁ +[h] r • T₂ := by
|
||||
refine ext rfl ?_
|
||||
simp only [add, contr_toColorIndexList, addCondEquiv, smul_index, smul_tensor, _root_.smul_add,
|
||||
Fin.castOrderIso_refl, OrderIso.refl_toEquiv, mapIso_refl, map_add, LinearEquiv.refl_apply,
|
||||
tensorIso]
|
||||
rw [tensor_eq_of_eq (smul_contr r T₁), tensor_eq_of_eq (smul_contr r T₂)]
|
||||
simp only [smul_index, contr_toColorIndexList, Fin.castOrderIso_refl, OrderIso.refl_toEquiv,
|
||||
mapIso_refl, smul_tensor, map_smul, LinearEquiv.refl_apply, tensorIso]
|
||||
|
||||
lemma add_withDual_empty (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
(T₁ +[h] T₂).withDual = ∅ := by
|
||||
simp only [add_toColorIndexList]
|
||||
change T₂.toColorIndexList.contr.withDual = ∅
|
||||
simp only [ColorIndexList.contr, IndexList.contrIndexList_withDual]
|
||||
|
||||
@[simp]
|
||||
lemma contr_add (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
(T₁ +[h] T₂).contr = T₁ +[h] T₂ :=
|
||||
contr_of_withDual_empty (T₁ +[h] T₂) (add_withDual_empty T₁ T₂ h)
|
||||
|
||||
lemma contr_add_tensor (T₁ T₂ : 𝓣.TensorIndex) (h : AddCond T₁ T₂) :
|
||||
(T₁ +[h] T₂).contr.tensor =
|
||||
𝓣.mapIso (Fin.castOrderIso (by rw [index_eq_of_eq (contr_add T₁ T₂ h)])).toEquiv
|
||||
(colormap_mapIso (by simp)) (T₁ +[h] T₂).tensor :=
|
||||
tensor_eq_of_eq (contr_add T₁ T₂ h)
|
||||
|
||||
lemma add_comm {T₁ T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) : T₁ +[h] T₂ ≈ T₂ +[h.symm] T₁ := by
|
||||
apply And.intro h.add_comm
|
||||
intro h
|
||||
simp only [contr_toColorIndexList, add_toColorIndexList, contr_add_tensor, add_tensor,
|
||||
addCondEquiv, map_add, mapIso_mapIso, colorMap', contrPermEquiv_symm]
|
||||
rw [_root_.add_comm]
|
||||
apply Mathlib.Tactic.LinearCombination.add_pf
|
||||
· apply congrFun
|
||||
apply congrArg
|
||||
apply mapIso_ext
|
||||
rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr, contrPermEquiv_trans,
|
||||
contrPermEquiv_trans]
|
||||
· apply congrFun
|
||||
apply congrArg
|
||||
apply mapIso_ext
|
||||
rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr, contrPermEquiv_trans,
|
||||
contrPermEquiv_trans]
|
||||
|
||||
open AddCond in
|
||||
lemma add_rel_left {T₁ T₁' T₂ : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₁ ≈ T₁') :
|
||||
T₁ +[h] T₂ ≈ T₁' +[h.rel_left h'] T₂ := by
|
||||
apply And.intro (ContrPerm.refl _)
|
||||
intro h
|
||||
simp only [contr_add_tensor, add_tensor, map_add]
|
||||
apply Mathlib.Tactic.LinearCombination.add_pf
|
||||
· rw [h'.to_eq]
|
||||
simp only [contr_toColorIndexList, add_toColorIndexList, colorMap', addCondEquiv,
|
||||
contrPermEquiv_symm, mapIso_mapIso, contrPermEquiv_trans, contrPermEquiv_refl,
|
||||
Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply]
|
||||
· simp only [contr_toColorIndexList, add_toColorIndexList, colorMap', contrPermEquiv_refl,
|
||||
Equiv.refl_symm, mapIso_refl, LinearEquiv.refl_apply]
|
||||
|
||||
open AddCond in
|
||||
lemma add_rel_right {T₁ T₂ T₂' : 𝓣.TensorIndex} (h : AddCond T₁ T₂) (h' : T₂ ≈ T₂') :
|
||||
T₁ +[h] T₂ ≈ T₁ +[h.rel_right h'] T₂' :=
|
||||
(add_comm _).trans ((add_rel_left _ h').trans (add_comm _))
|
||||
|
||||
open AddCond in
|
||||
lemma add_assoc' {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₂ T₃} (h : AddCond T₁ (T₂ +[h'] T₃)) :
|
||||
T₁ +[h] (T₂ +[h'] T₃) = T₁ +[h'.of_add_right h] T₂ +[h'.add_left_of_add_right h] T₃ := by
|
||||
refine ext ?_ ?_
|
||||
· simp only [add_toColorIndexList, ColorIndexList.contr_contr]
|
||||
· simp only [add_toColorIndexList, add_tensor, contr_toColorIndexList, addCondEquiv,
|
||||
contr_add_tensor, map_add, mapIso_mapIso]
|
||||
rw [_root_.add_assoc]
|
||||
apply Mathlib.Tactic.LinearCombination.add_pf
|
||||
· apply congrFun
|
||||
apply congrArg
|
||||
apply mapIso_ext
|
||||
rw [← contrPermEquiv_self_contr, ← contrPermEquiv_self_contr]
|
||||
rw [contrPermEquiv_trans, contrPermEquiv_trans, contrPermEquiv_trans]
|
||||
· apply Mathlib.Tactic.LinearCombination.add_pf _ rfl
|
||||
apply congrFun
|
||||
apply congrArg
|
||||
apply mapIso_ext
|
||||
rw [← contrPermEquiv_self_contr, contrPermEquiv_trans, ← contrPermEquiv_self_contr,
|
||||
contrPermEquiv_trans, contrPermEquiv_trans]
|
||||
|
||||
open AddCond in
|
||||
lemma add_assoc {T₁ T₂ T₃ : 𝓣.TensorIndex} {h' : AddCond T₁ T₂} (h : AddCond (T₁ +[h'] T₂) T₃) :
|
||||
T₁ +[h'] T₂ +[h] T₃ = T₁ +[h'.add_right_of_add_left h] (T₂ +[h'.of_add_left h] T₃) := by
|
||||
rw [add_assoc']
|
||||
|
||||
/-!
|
||||
|
||||
## Product of `TensorIndex` when allowed
|
||||
|
||||
-/
|
||||
|
||||
/-! TODO: Show that the product is well defined with respect to Rel. -/
|
||||
|
||||
/-- 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
|
||||
`ColorIndexList`. -/
|
||||
def ProdCond (T₁ T₂ : 𝓣.TensorIndex) : Prop :=
|
||||
AppendCond T₁.toColorIndexList T₂.toColorIndexList
|
||||
|
||||
namespace ProdCond
|
||||
|
||||
variable {T₁ T₁' T₂ T₂' : 𝓣.TensorIndex}
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
lemma to_AppendCond (h : ProdCond T₁ T₂) :
|
||||
T₁.AppendCond T₂ := h
|
||||
|
||||
@[symm]
|
||||
lemma symm (h : ProdCond T₁ T₂) : ProdCond T₂ T₁ := h.to_AppendCond.symm
|
||||
|
||||
/-! TODO: Prove properties regarding the interaction of `ProdCond` and `Rel`. -/
|
||||
|
||||
end ProdCond
|
||||
|
||||
/-- The tensor product of two `TensorIndex`.
|
||||
|
||||
Note: By defualt contraction is NOT done before taking the products. -/
|
||||
def prod (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T₂) : 𝓣.TensorIndex where
|
||||
toColorIndexList := T₁ ++[h] T₂
|
||||
tensor := 𝓣.mapIso IndexList.appendEquiv (T₁.colorMap_sumELim T₂) <|
|
||||
𝓣.tensoratorEquiv _ _ (T₁.tensor ⊗ₜ[R] T₂.tensor)
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma prod_toColorIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T₂) :
|
||||
(prod T₁ T₂ h).toColorIndexList = T₁.toColorIndexList ++[h] T₂.toColorIndexList := rfl
|
||||
|
||||
omit [DecidableEq 𝓣.Color] in
|
||||
@[simp]
|
||||
lemma prod_toIndexList (T₁ T₂ : 𝓣.TensorIndex) (h : ProdCond T₁ T₂) :
|
||||
(prod T₁ T₂ h).toIndexList = T₁.toIndexList ++ T₂.toIndexList := rfl
|
||||
|
||||
end TensorIndex
|
||||
end
|
||||
end TensorStructure
|
|
@ -1,255 +0,0 @@
|
|||
/-
|
||||
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.Tensors.Basic
|
||||
import Mathlib.RepresentationTheory.Basic
|
||||
/-!
|
||||
|
||||
# Group actions on tensor structures
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open TensorProduct
|
||||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
/-! TODO: Add preservation of the unit, and the metric. -/
|
||||
|
||||
/-- A multiplicative action on a tensor structure (e.g. the action of the Lorentz
|
||||
group on real Lorentz tensors). -/
|
||||
class MulActionTensor (G : Type) [Monoid G] (𝓣 : TensorStructure R) where
|
||||
/-- For each color `μ` a representation of `G` on `ColorModule μ`. -/
|
||||
repColorModule : (μ : 𝓣.Color) → Representation R G (𝓣.ColorModule μ)
|
||||
/-- The contraction of a vector with its dual is invariant under the group action. -/
|
||||
contrDual_inv : ∀ μ g, 𝓣.contrDual μ ∘ₗ
|
||||
TensorProduct.map (repColorModule μ g) (repColorModule (𝓣.τ μ) g) = 𝓣.contrDual μ
|
||||
/-- The invariance of the metric under the group action. -/
|
||||
metric_inv : ∀ μ g, (TensorProduct.map (repColorModule μ g) (repColorModule μ g)) (𝓣.metric μ) =
|
||||
𝓣.metric μ
|
||||
|
||||
namespace MulActionTensor
|
||||
|
||||
variable {G H : Type} [Group G] [Group H]
|
||||
|
||||
variable (𝓣 : TensorStructure R) [MulActionTensor G 𝓣]
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z]
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
|
||||
/-!
|
||||
|
||||
# Instances of `MulActionTensor`
|
||||
|
||||
-/
|
||||
|
||||
/-- The `MulActionTensor` defined by restriction along a group homomorphism. -/
|
||||
def compHom (f : H →* G) : MulActionTensor H 𝓣 where
|
||||
repColorModule μ := MonoidHom.comp (repColorModule μ) f
|
||||
contrDual_inv μ h := by
|
||||
simp only [MonoidHom.coe_comp, Function.comp_apply]
|
||||
rw [contrDual_inv]
|
||||
metric_inv μ h := by
|
||||
simp only [MonoidHom.coe_comp, Function.comp_apply]
|
||||
rw [metric_inv]
|
||||
|
||||
/-- The trivial `MulActionTensor` defined via trivial representations. -/
|
||||
def trivial : MulActionTensor G 𝓣 where
|
||||
repColorModule μ := Representation.trivial R
|
||||
contrDual_inv μ g := ext rfl
|
||||
metric_inv μ g := by
|
||||
simp only [Representation.trivial, MonoidHom.one_apply, TensorProduct.map_one]
|
||||
rfl
|
||||
|
||||
end MulActionTensor
|
||||
|
||||
namespace TensorStructure
|
||||
open TensorStructure
|
||||
open MulActionTensor
|
||||
|
||||
variable {G : Type} [Group G]
|
||||
|
||||
variable (𝓣 : TensorStructure R) [MulActionTensor G 𝓣]
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z]
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
|
||||
/-!
|
||||
|
||||
# Equivariance properties involving modules
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma contrDual_equivariant_tmul (g : G) (x : 𝓣.ColorModule μ) (y : 𝓣.ColorModule (𝓣.τ μ)) :
|
||||
(𝓣.contrDual μ ((repColorModule μ g x) ⊗ₜ[R] (repColorModule (𝓣.τ μ) g y))) =
|
||||
𝓣.contrDual μ (x ⊗ₜ[R] y) := by
|
||||
trans (𝓣.contrDual μ ∘ₗ
|
||||
TensorProduct.map (repColorModule μ g) (repColorModule (𝓣.τ μ) g)) (x ⊗ₜ[R] y)
|
||||
· rfl
|
||||
· rw [contrDual_inv]
|
||||
|
||||
@[simp]
|
||||
lemma colorModuleCast_equivariant_apply (h : μ = ν) (g : G) (x : 𝓣.ColorModule μ) :
|
||||
(𝓣.colorModuleCast h) (repColorModule μ g x) =
|
||||
(repColorModule ν g) (𝓣.colorModuleCast h x) := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrRightAux_contrDual_equivariant_tmul (g : G) (m : 𝓣.ColorModule ν ⊗[R] 𝓣.ColorModule μ)
|
||||
(x : 𝓣.ColorModule (𝓣.τ μ)) : (contrRightAux (𝓣.contrDual μ))
|
||||
((TensorProduct.map (repColorModule ν g) (repColorModule μ g) m) ⊗ₜ[R]
|
||||
(repColorModule (𝓣.τ μ) g x)) =
|
||||
repColorModule ν g ((contrRightAux (𝓣.contrDual μ)) (m ⊗ₜ[R] x)) := by
|
||||
refine TensorProduct.induction_on m (by simp) ?_ ?_
|
||||
· intro y z
|
||||
simp [contrRightAux]
|
||||
· intro x z h1 h2
|
||||
simp [add_tmul, LinearMap.map_add, h1, h2]
|
||||
|
||||
/-!
|
||||
|
||||
## Representation of tensor products
|
||||
|
||||
-/
|
||||
|
||||
/-- The representation of the group `G` on the vector space of tensors. -/
|
||||
def rep : Representation R G (𝓣.Tensor cX) where
|
||||
toFun g := PiTensorProduct.map (fun x => repColorModule (cX x) g)
|
||||
map_one' := by
|
||||
simp_all only [_root_.map_one, PiTensorProduct.map_one]
|
||||
map_mul' g g' := by
|
||||
simp_all only [_root_.map_mul]
|
||||
exact PiTensorProduct.map_mul _ _
|
||||
|
||||
local infixl:101 " • " => 𝓣.rep
|
||||
|
||||
@[simp]
|
||||
lemma repColorModule_colorModuleCast (h : μ = ν) (g : G) :
|
||||
(repColorModule ν g) ∘ₗ (𝓣.colorModuleCast h).toLinearMap =
|
||||
(𝓣.colorModuleCast h).toLinearMap ∘ₗ (repColorModule μ g) := by
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp [colorModuleCast_equivariant_apply]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma rep_mapIso (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) :
|
||||
(𝓣.rep g) ∘ₗ (𝓣.mapIso e h).toLinearMap = (𝓣.mapIso e h).toLinearMap ∘ₗ 𝓣.rep g := by
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp only [LinearMap.compMultilinearMap_apply, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply]
|
||||
erw [mapIso_tprod]
|
||||
simp only [rep, MonoidHom.coe_mk, OneHom.coe_mk]
|
||||
change (PiTensorProduct.map fun x => (repColorModule (cY x)) g)
|
||||
((PiTensorProduct.tprod R) fun i => (𝓣.colorModuleCast _) (x (e.symm i))) =
|
||||
(𝓣.mapIso e h) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) x))
|
||||
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod, mapIso_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
subst h
|
||||
simp [colorModuleCast_equivariant_apply]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma rep_mapIso_apply (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) (x : 𝓣.Tensor cX) :
|
||||
(𝓣.mapIso e h) (g • x) = g • (𝓣.mapIso e h x) := by
|
||||
trans ((𝓣.rep g) ∘ₗ (𝓣.mapIso e h).toLinearMap) x
|
||||
· simp
|
||||
· rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y] in
|
||||
@[simp]
|
||||
lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (cX i)) :
|
||||
g • (PiTensorProduct.tprod R f) = PiTensorProduct.tprod R (fun x =>
|
||||
repColorModule (cX x) g (f x)) := by
|
||||
simp only [rep, MonoidHom.coe_mk, OneHom.coe_mk]
|
||||
change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) f) = _
|
||||
rw [PiTensorProduct.map_tprod]
|
||||
|
||||
/-!
|
||||
|
||||
## Group acting on tensor products
|
||||
|
||||
-/
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
lemma tensoratorEquiv_equivariant (g : G) :
|
||||
(𝓣.tensoratorEquiv cX cY) ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g)) = 𝓣.rep g ∘ₗ
|
||||
(𝓣.tensoratorEquiv cX cY).toLinearMap := by
|
||||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, map_tmul, rep_tprod,
|
||||
tensoratorEquiv_tmul_tprod]
|
||||
apply congrArg
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_equivariant_apply (g : G) (x : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cY) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((TensorProduct.map (𝓣.rep g) (𝓣.rep g)) x)
|
||||
= (𝓣.rep g) ((𝓣.tensoratorEquiv cX cY) x) := by
|
||||
trans ((𝓣.tensoratorEquiv cX cY) ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g))) x
|
||||
· rfl
|
||||
· rw [tensoratorEquiv_equivariant]
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
lemma rep_tensoratorEquiv_tmul (g : G) (x : 𝓣.Tensor cX) (y : 𝓣.Tensor cY) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((g • x) ⊗ₜ[R] (g • y)) =
|
||||
g • ((𝓣.tensoratorEquiv cX cY) (x ⊗ₜ[R] y)) := by
|
||||
nth_rewrite 1 [← tensoratorEquiv_equivariant_apply]
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
lemma rep_tensoratorEquiv_symm (g : G) :
|
||||
(𝓣.tensoratorEquiv cX cY).symm ∘ₗ 𝓣.rep g = (TensorProduct.map (𝓣.rep g) (𝓣.rep g)) ∘ₗ
|
||||
(𝓣.tensoratorEquiv cX cY).symm.toLinearMap := by
|
||||
rw [LinearEquiv.eq_comp_toLinearMap_symm, LinearMap.comp_assoc,
|
||||
LinearEquiv.toLinearMap_symm_comp_eq]
|
||||
exact Eq.symm (tensoratorEquiv_equivariant 𝓣 g)
|
||||
|
||||
omit [Fintype X] [Fintype Y] in
|
||||
@[simp]
|
||||
lemma rep_tensoratorEquiv_symm_apply (g : G) (x : 𝓣.Tensor (Sum.elim cX cY)) :
|
||||
(𝓣.tensoratorEquiv cX cY).symm ((𝓣.rep g) x) =
|
||||
(TensorProduct.map (𝓣.rep g) (𝓣.rep g)) ((𝓣.tensoratorEquiv cX cY).symm x) := by
|
||||
trans ((𝓣.tensoratorEquiv cX cY).symm ∘ₗ 𝓣.rep g) x
|
||||
· rfl
|
||||
· rw [rep_tensoratorEquiv_symm]
|
||||
rfl
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma rep_lid (g : G) : TensorProduct.lid R (𝓣.Tensor cX) ∘ₗ
|
||||
(TensorProduct.map (LinearMap.id) (𝓣.rep g)) = (𝓣.rep g) ∘ₗ
|
||||
(TensorProduct.lid R (𝓣.Tensor cX)).toLinearMap := by
|
||||
apply TensorProduct.ext'
|
||||
intro r y
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, map_tmul,
|
||||
LinearMap.id_coe, id_eq, lid_tmul, map_smul]
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma rep_lid_apply (g : G) (x : R ⊗[R] 𝓣.Tensor cX) :
|
||||
(TensorProduct.lid R (𝓣.Tensor cX)) ((TensorProduct.map (LinearMap.id) (𝓣.rep g)) x) =
|
||||
(𝓣.rep g) ((TensorProduct.lid R (𝓣.Tensor cX)).toLinearMap x) := by
|
||||
trans ((TensorProduct.lid R (𝓣.Tensor cX)) ∘ₗ (TensorProduct.map (LinearMap.id) (𝓣.rep g))) x
|
||||
· rfl
|
||||
· rw [rep_lid]
|
||||
rfl
|
||||
|
||||
end TensorStructure
|
||||
|
||||
end
|
|
@ -1,351 +0,0 @@
|
|||
/-
|
||||
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.Tensors.MulActionTensor
|
||||
/-!
|
||||
|
||||
# Rising and Lowering of indices
|
||||
|
||||
We use the term `dualize` to describe the more general version of rising and lowering of indices.
|
||||
|
||||
In particular, rising and lowering indices corresponds taking the color of that index
|
||||
to its dual.
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
open TensorProduct
|
||||
|
||||
namespace TensorColor
|
||||
|
||||
variable {𝓒 : TensorColor} [DecidableEq 𝓒.Color] [Fintype 𝓒.Color]
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
[Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P]
|
||||
|
||||
/-!
|
||||
|
||||
## Dual maps
|
||||
|
||||
-/
|
||||
|
||||
namespace ColorMap
|
||||
|
||||
variable (cX : 𝓒.ColorMap X)
|
||||
|
||||
/-- Given an equivalence `C ⊕ P ≃ X` the color map obtained by `cX` by dualising
|
||||
all indices in `C`. -/
|
||||
def partDual (e : C ⊕ P ≃ X) : 𝓒.ColorMap X :=
|
||||
(Sum.elim (𝓒.τ ∘ cX ∘ e ∘ Sum.inl) (cX ∘ e ∘ Sum.inr) ∘ e.symm)
|
||||
|
||||
/-- Two color maps are said to be dual if their quotents are dual. -/
|
||||
def DualMap (c₁ c₂ : 𝓒.ColorMap X) : Prop :=
|
||||
𝓒.colorQuot ∘ c₁ = 𝓒.colorQuot ∘ c₂
|
||||
|
||||
namespace DualMap
|
||||
|
||||
variable {c₁ c₂ c₃ : 𝓒.ColorMap X}
|
||||
variable {n : ℕ}
|
||||
|
||||
/-- The bool which if `𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i)` is true for all `i`. -/
|
||||
def boolFin (c₁ c₂ : 𝓒.ColorMap (Fin n)) : Bool :=
|
||||
(Fin.list n).all fun i => if 𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i) then true else false
|
||||
|
||||
omit [Fintype 𝓒.Color] in
|
||||
lemma boolFin_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin c₁ c₂ = true) :
|
||||
DualMap c₁ c₂ := by
|
||||
simp only [boolFin, Bool.if_false_right, Bool.and_true, List.all_eq_true, decide_eq_true_eq] at h
|
||||
simp only [DualMap]
|
||||
funext x
|
||||
have h2 {n : ℕ} (m : Fin n) : m ∈ Fin.list n := by
|
||||
have h1' : (Fin.list n)[m] = m := by
|
||||
erw [Fin.getElem_list]
|
||||
rfl
|
||||
rw [← h1']
|
||||
apply List.getElem_mem
|
||||
exact h x (h2 _)
|
||||
|
||||
/-- The bool which is ture if `𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i)` is true for all `i`. -/
|
||||
def boolFin' (c₁ c₂ : 𝓒.ColorMap (Fin n)) : Bool :=
|
||||
∀ (i : Fin n), 𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i)
|
||||
|
||||
omit [Fintype 𝓒.Color]
|
||||
lemma boolFin'_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin' c₁ c₂ = true) :
|
||||
DualMap c₁ c₂ := by
|
||||
simp only [boolFin', decide_eq_true_eq] at h
|
||||
simp only [DualMap]
|
||||
funext x
|
||||
exact h x
|
||||
|
||||
omit [DecidableEq 𝓒.Color] [Fintype X] [DecidableEq X] in
|
||||
lemma refl : DualMap c₁ c₁ := rfl
|
||||
|
||||
omit [DecidableEq 𝓒.Color] [Fintype X] [DecidableEq X] in
|
||||
lemma symm (h : DualMap c₁ c₂) : DualMap c₂ c₁ := by
|
||||
rw [DualMap] at h ⊢
|
||||
exact h.symm
|
||||
|
||||
omit [DecidableEq 𝓒.Color] [Fintype X] [DecidableEq X] in
|
||||
lemma trans (h : DualMap c₁ c₂) (h' : DualMap c₂ c₃) : DualMap c₁ c₃ := by
|
||||
rw [DualMap] at h h' ⊢
|
||||
exact h.trans h'
|
||||
|
||||
/-- The splitting of `X` given two color maps based on the equality of the color. -/
|
||||
def split (c₁ c₂ : 𝓒.ColorMap X) : { x // c₁ x ≠ c₂ x} ⊕ { x // c₁ x = c₂ x} ≃ X :=
|
||||
((Equiv.Set.sumCompl {x | c₁ x = c₂ x}).symm.trans (Equiv.sumComm _ _)).symm
|
||||
|
||||
omit [DecidableEq 𝓒.Color] [Fintype X] [DecidableEq X] in
|
||||
lemma dual_eq_of_neq (h : DualMap c₁ c₂) {x : X} (h' : c₁ x ≠ c₂ x) :
|
||||
𝓒.τ (c₁ x) = c₂ x := by
|
||||
rw [DualMap] at h
|
||||
have h1 := congrFun h x
|
||||
simp only [Function.comp_apply, colorQuot, Quotient.eq, HasEquiv.Equiv, Setoid.r, colorRel] at h1
|
||||
simp_all only [ne_eq, false_or]
|
||||
exact 𝓒.τ_involutive (c₂ x)
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma split_dual (h : DualMap c₁ c₂) : c₁.partDual (split c₁ c₂) = c₂ := by
|
||||
rw [partDual, Equiv.comp_symm_eq]
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x =>
|
||||
exact h.dual_eq_of_neq x.2
|
||||
| Sum.inr x =>
|
||||
exact x.2
|
||||
|
||||
omit [Fintype X] [DecidableEq X] in
|
||||
@[simp]
|
||||
lemma split_dual' (h : DualMap c₁ c₂) : c₂.partDual (split c₁ c₂) = c₁ := by
|
||||
rw [partDual, Equiv.comp_symm_eq]
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x =>
|
||||
change 𝓒.τ (c₂ x) = c₁ x
|
||||
rw [← h.dual_eq_of_neq x.2]
|
||||
exact (𝓒.τ_involutive (c₁ x))
|
||||
| Sum.inr x =>
|
||||
exact x.2.symm
|
||||
|
||||
end DualMap
|
||||
|
||||
end ColorMap
|
||||
end TensorColor
|
||||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
namespace TensorStructure
|
||||
|
||||
variable (𝓣 : TensorStructure R)
|
||||
|
||||
variable {d : ℕ} {X Y Y' Z W C P : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
|
||||
[Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P]
|
||||
{cX cX2 : 𝓣.ColorMap X} {cY : 𝓣.ColorMap Y} {cZ : 𝓣.ColorMap Z}
|
||||
{cW : 𝓣.ColorMap W} {cY' : 𝓣.ColorMap Y'} {μ ν: 𝓣.Color}
|
||||
|
||||
variable {G H : Type} [Group G] [Group H] [MulActionTensor G 𝓣]
|
||||
local infixl:101 " • " => 𝓣.rep
|
||||
|
||||
/-!
|
||||
|
||||
## Properties of the unit
|
||||
|
||||
-/
|
||||
|
||||
/-! TODO: Move -/
|
||||
lemma unit_lhs_eq (x : 𝓣.ColorModule μ) (y : 𝓣.ColorModule (𝓣.τ μ) ⊗[R] 𝓣.ColorModule μ) :
|
||||
contrLeftAux (𝓣.contrDual μ) (x ⊗ₜ[R] y) =
|
||||
(contrRightAux (𝓣.contrDual (𝓣.τ μ))) ((TensorProduct.comm R _ _) y
|
||||
⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm) x) := by
|
||||
refine TensorProduct.induction_on y (by rfl) ?_ (fun z1 z2 h1 h2 => ?_)
|
||||
· intro x1 x2
|
||||
simp only [contrRightAux, LinearEquiv.refl_toLinearMap, comm_tmul, colorModuleCast,
|
||||
Equiv.cast_symm, LinearEquiv.coe_mk, Equiv.cast_apply, LinearMap.coe_comp,
|
||||
LinearEquiv.coe_coe, Function.comp_apply, assoc_tmul, map_tmul, LinearMap.id_coe, id_eq,
|
||||
contrDual_symm', cast_cast, cast_eq, rid_tmul]
|
||||
rfl
|
||||
· simp only [map_add, add_tmul]
|
||||
rw [← h1, ← h2, tmul_add, LinearMap.map_add]
|
||||
|
||||
@[simp]
|
||||
lemma unit_lid : (contrRightAux (𝓣.contrDual (𝓣.τ μ))) ((TensorProduct.comm R _ _) (𝓣.unit μ)
|
||||
⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm) x) = x := by
|
||||
have h1 := 𝓣.unit_rid μ x
|
||||
rw [← unit_lhs_eq]
|
||||
exact h1
|
||||
|
||||
/-!
|
||||
|
||||
## Properties of the metric
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma metric_cast (h : μ = ν) :
|
||||
(TensorProduct.congr (𝓣.colorModuleCast h) (𝓣.colorModuleCast h)) (𝓣.metric μ) =
|
||||
𝓣.metric ν := by
|
||||
subst h
|
||||
erw [congr_refl_refl]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma metric_contrRight_unit (μ : 𝓣.Color) (x : 𝓣.ColorModule μ) :
|
||||
(contrRightAux (𝓣.contrDual μ)) (𝓣.metric μ ⊗ₜ[R]
|
||||
((contrRightAux (𝓣.contrDual (𝓣.τ μ)))
|
||||
(𝓣.metric (𝓣.τ μ) ⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm x)))) = x := by
|
||||
change (contrRightAux (𝓣.contrDual μ) ∘ₗ TensorProduct.map (LinearMap.id)
|
||||
(contrRightAux (𝓣.contrDual (𝓣.τ μ)))) (𝓣.metric μ
|
||||
⊗ₜ[R] 𝓣.metric (𝓣.τ μ) ⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm x)) = _
|
||||
rw [contrRightAux_comp]
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul,
|
||||
map_tmul, LinearMap.id_coe, id_eq]
|
||||
rw [𝓣.metric_dual]
|
||||
exact unit_lid 𝓣
|
||||
|
||||
/-!
|
||||
|
||||
## Dualizing
|
||||
|
||||
-/
|
||||
|
||||
/-- Takes a vector with index with dual color to a vector with index the underlying color.
|
||||
Obtained by contraction with the metric. -/
|
||||
def dualizeSymm (μ : 𝓣.Color) : 𝓣.ColorModule (𝓣.τ μ) →ₗ[R] 𝓣.ColorModule μ :=
|
||||
contrRightAux (𝓣.contrDual μ) ∘ₗ
|
||||
TensorProduct.lTensorHomToHomLTensor R _ _ _ (𝓣.metric μ ⊗ₜ[R] LinearMap.id)
|
||||
|
||||
/-- Takes a vector to a vector with the dual color index.
|
||||
Obtained by contraction with the metric. -/
|
||||
def dualizeFun (μ : 𝓣.Color) : 𝓣.ColorModule μ →ₗ[R] 𝓣.ColorModule (𝓣.τ μ) :=
|
||||
𝓣.dualizeSymm (𝓣.τ μ) ∘ₗ (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm).toLinearMap
|
||||
|
||||
/-- Equivalence between the module with a color `μ` and the module with color
|
||||
`𝓣.τ μ` obtained by contraction with the metric. -/
|
||||
def dualizeModule (μ : 𝓣.Color) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModule (𝓣.τ μ) := by
|
||||
refine LinearEquiv.ofLinear (𝓣.dualizeFun μ) (𝓣.dualizeSymm μ) ?_ ?_
|
||||
· apply LinearMap.ext
|
||||
intro x
|
||||
simp [dualizeFun, dualizeSymm, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq,
|
||||
contrDual_symm_contrRightAux_apply_tmul, metric_cast]
|
||||
· apply LinearMap.ext
|
||||
intro x
|
||||
simp only [dualizeSymm, dualizeFun, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply, lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq,
|
||||
metric_contrRight_unit]
|
||||
|
||||
@[simp]
|
||||
lemma dualizeModule_equivariant (g : G) :
|
||||
(𝓣.dualizeModule μ) ∘ₗ ((MulActionTensor.repColorModule μ) g) =
|
||||
(MulActionTensor.repColorModule (𝓣.τ μ) g) ∘ₗ (𝓣.dualizeModule μ) := by
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp only [dualizeModule, dualizeFun, dualizeSymm, LinearEquiv.ofLinear_toLinearMap,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, colorModuleCast_equivariant_apply,
|
||||
lTensorHomToHomLTensor_apply, LinearMap.id_coe, id_eq]
|
||||
nth_rewrite 1 [← MulActionTensor.metric_inv (𝓣.τ μ) g]
|
||||
exact contrRightAux_contrDual_equivariant_tmul 𝓣 g (𝓣.metric (𝓣.τ μ))
|
||||
((𝓣.colorModuleCast (dualizeFun.proof_3 𝓣 μ)) x)
|
||||
|
||||
@[simp]
|
||||
lemma dualizeModule_equivariant_apply (g : G) (x : 𝓣.ColorModule μ) :
|
||||
(𝓣.dualizeModule μ) ((MulActionTensor.repColorModule μ) g x) =
|
||||
(MulActionTensor.repColorModule (𝓣.τ μ) g) (𝓣.dualizeModule μ x) := by
|
||||
trans ((𝓣.dualizeModule μ) ∘ₗ ((MulActionTensor.repColorModule μ) g)) x
|
||||
· rfl
|
||||
· rw [dualizeModule_equivariant]
|
||||
rfl
|
||||
|
||||
/-- Dualizes the color of all indicies of a tensor by contraction with the metric. -/
|
||||
def dualizeAll : 𝓣.Tensor cX ≃ₗ[R] 𝓣.Tensor (𝓣.τ ∘ cX) := by
|
||||
refine LinearEquiv.ofLinear
|
||||
(PiTensorProduct.map (fun x => (𝓣.dualizeModule (cX x)).toLinearMap))
|
||||
(PiTensorProduct.map (fun x => (𝓣.dualizeModule (cX x)).symm.toLinearMap)) ?_ ?_
|
||||
all_goals
|
||||
apply LinearMap.ext
|
||||
refine fun x ↦ PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx a
|
||||
simp only [Function.comp_apply, map_add, hx, LinearMap.id_coe, id_eq, LinearMap.coe_comp]
|
||||
simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq])
|
||||
intro rx fx
|
||||
simp only [Function.comp_apply, PiTensorProduct.tprodCoeff_eq_smul_tprod,
|
||||
LinearMapClass.map_smul, LinearMap.coe_comp, LinearMap.id_coe, id_eq]
|
||||
apply congrArg
|
||||
change (PiTensorProduct.map _)
|
||||
((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx)) = _
|
||||
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod]
|
||||
apply congrArg
|
||||
simp
|
||||
omit [Fintype X] [DecidableEq X]
|
||||
@[simp]
|
||||
lemma dualizeAll_equivariant (g : G) : (𝓣.dualizeAll.toLinearMap) ∘ₗ (@rep R _ G _ 𝓣 _ X cX g)
|
||||
= 𝓣.rep g ∘ₗ (𝓣.dualizeAll.toLinearMap) := by
|
||||
apply LinearMap.ext
|
||||
intro x
|
||||
simp only [dualizeAll, Function.comp_apply, LinearEquiv.ofLinear_toLinearMap, LinearMap.coe_comp]
|
||||
refine PiTensorProduct.induction_on' x ?_ (by
|
||||
intro a b hx a
|
||||
simp only [map_add, hx]
|
||||
simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq])
|
||||
intro rx fx
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, LinearMapClass.map_smul, rep_tprod]
|
||||
apply congrArg
|
||||
change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) _) =
|
||||
(𝓣.rep g) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx))
|
||||
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod]
|
||||
simp
|
||||
|
||||
omit [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] in
|
||||
lemma dualize_cond (e : C ⊕ P ≃ X) :
|
||||
cX = Sum.elim (cX ∘ e ∘ Sum.inl) (cX ∘ e ∘ Sum.inr) ∘ e.symm := by
|
||||
rw [Equiv.eq_comp_symm]
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => rfl
|
||||
| Sum.inr x => rfl
|
||||
|
||||
omit [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] in
|
||||
lemma dualize_cond' (e : C ⊕ P ≃ X) :
|
||||
Sum.elim (𝓣.τ ∘ cX ∘ ⇑e ∘ Sum.inl) (cX ∘ ⇑e ∘ Sum.inr) =
|
||||
(Sum.elim (𝓣.τ ∘ cX ∘ ⇑e ∘ Sum.inl) (cX ∘ ⇑e ∘ Sum.inr) ∘ ⇑e.symm) ∘ ⇑e := by
|
||||
funext x
|
||||
match x with
|
||||
| Sum.inl x => simp
|
||||
| Sum.inr x => simp
|
||||
|
||||
/-- Given an equivalence `C ⊕ P ≃ X` dualizes those indices of a tensor which correspond to
|
||||
`C` whilst leaving the indices `P` invariant. -/
|
||||
def dualize (e : C ⊕ P ≃ X) : 𝓣.Tensor cX ≃ₗ[R] 𝓣.Tensor (cX.partDual e) :=
|
||||
𝓣.mapIso e.symm (𝓣.dualize_cond e) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv _ _).symm ≪≫ₗ
|
||||
TensorProduct.congr 𝓣.dualizeAll (LinearEquiv.refl _ _) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv _ _) ≪≫ₗ
|
||||
𝓣.mapIso e (𝓣.dualize_cond' e)
|
||||
|
||||
omit [Fintype C] [Fintype P] in
|
||||
/-- Dualizing indices is equivariant with respect to the group action. This is the
|
||||
applied version of this statement. -/
|
||||
@[simp]
|
||||
lemma dualize_equivariant_apply (e : C ⊕ P ≃ X) (g : G) (x : 𝓣.Tensor cX) :
|
||||
𝓣.dualize e (g • x) = g • (𝓣.dualize e x) := by
|
||||
simp only [dualize, TensorProduct.congr, LinearEquiv.refl_toLinearMap, LinearEquiv.refl_symm,
|
||||
LinearEquiv.trans_apply, rep_mapIso_apply, rep_tensoratorEquiv_symm_apply,
|
||||
LinearEquiv.ofLinear_apply]
|
||||
rw [← LinearMap.comp_apply (TensorProduct.map _ _), ← TensorProduct.map_comp]
|
||||
simp only [dualizeAll_equivariant, LinearMap.id_comp]
|
||||
have h1 {M N A B C : Type} [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid A]
|
||||
[AddCommMonoid B] [AddCommMonoid C] [Module R M] [Module R N] [Module R A] [Module R B]
|
||||
[Module R C] (f : M →ₗ[R] N) (g : A →ₗ[R] B) (h : B →ₗ[R] C) : TensorProduct.map (h ∘ₗ g) f
|
||||
= TensorProduct.map h f ∘ₗ TensorProduct.map g (LinearMap.id) :=
|
||||
ext rfl
|
||||
rw [h1, LinearMap.coe_comp, Function.comp_apply]
|
||||
simp only [tensoratorEquiv_equivariant_apply, rep_mapIso_apply]
|
||||
|
||||
end TensorStructure
|
||||
|
||||
end
|
Loading…
Add table
Add a link
Reference in a new issue