PhysLean/HepLean/Tensors/RisingLowering.lean

352 lines
14 KiB
Text
Raw Normal View History

2024-07-30 16:07:16 -04:00
/-
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
2024-07-30 16:07:16 -04:00
/-!
# 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)
2024-08-06 15:56:29 -04:00
/-- Given an equivalence `C ⊕ P ≃ X` the color map obtained by `cX` by dualising
2024-08-07 08:56:45 -04:00
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
2024-09-04 06:28:46 -04:00
omit [Fintype 𝓒.Color] in
lemma boolFin_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin c₁ c₂ = true) :
DualMap c₁ c₂ := by
2024-09-04 07:31:34 -04:00
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 _)
2024-08-16 16:15:39 -04:00
/-- The bool which is ture if `𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i)` is true for all `i`. -/
2024-08-16 15:56:18 -04:00
def boolFin' (c₁ c₂ : 𝓒.ColorMap (Fin n)) : Bool :=
∀ (i : Fin n), 𝓒.colorQuot (c₁ i) = 𝓒.colorQuot (c₂ i)
2024-09-04 06:28:46 -04:00
omit [Fintype 𝓒.Color]
2024-08-16 15:56:18 -04:00
lemma boolFin'_DualMap {c₁ c₂ : 𝓒.ColorMap (Fin n)} (h : boolFin' c₁ c₂ = true) :
DualMap c₁ c₂ := by
2024-09-04 07:31:34 -04:00
simp only [boolFin', decide_eq_true_eq] at h
simp only [DualMap]
2024-08-16 15:56:18 -04:00
funext x
exact h x
2024-09-04 06:28:46 -04:00
omit [DecidableEq 𝓒.Color] [Fintype X] [DecidableEq X] in
2024-08-30 13:40:32 -04:00
lemma refl : DualMap c₁ c₁ := rfl
2024-09-04 06:28:46 -04:00
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
2024-09-04 06:28:46 -04:00
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
2024-09-04 06:28:46 -04:00
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
2024-09-04 07:31:34 -04:00
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)
2024-09-04 06:36:42 -04:00
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
2024-09-04 06:36:42 -04:00
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
2024-07-30 16:07:16 -04:00
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}
2024-07-30 16:07:16 -04:00
variable {G H : Type} [Group G] [Group H] [MulActionTensor G 𝓣]
local infixl:101 " • " => 𝓣.rep
2024-07-30 16:07:16 -04:00
/-!
## Properties of the unit
-/
/-! TODO: Move -/
lemma unit_lhs_eq (x : 𝓣.ColorModule μ) (y : 𝓣.ColorModule (𝓣.τ μ) ⊗[R] 𝓣.ColorModule μ) :
2024-07-30 16:31:38 -04:00
contrLeftAux (𝓣.contrDual μ) (x ⊗ₜ[R] y) =
2024-07-30 16:07:16 -04:00
(contrRightAux (𝓣.contrDual (𝓣.τ μ))) ((TensorProduct.comm R _ _) y
⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm) x) := by
2024-08-30 13:40:32 -04:00
refine TensorProduct.induction_on y (by rfl) ?_ (fun z1 z2 h1 h2 => ?_)
2024-08-20 14:38:29 -04:00
· intro x1 x2
simp only [contrRightAux, LinearEquiv.refl_toLinearMap, comm_tmul, colorModuleCast,
2024-08-21 06:45:09 -04:00
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]
2024-08-20 14:38:29 -04:00
rfl
· simp only [map_add, add_tmul]
rw [← h1, ← h2, tmul_add, LinearMap.map_add]
2024-07-30 16:07:16 -04:00
@[simp]
lemma unit_lid : (contrRightAux (𝓣.contrDual (𝓣.τ μ))) ((TensorProduct.comm R _ _) (𝓣.unit μ)
2024-07-30 16:31:38 -04:00
⊗ₜ[R] (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm) x) = x := by
2024-07-30 16:07:16 -04:00
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]
2024-08-30 13:40:32 -04:00
rfl
2024-07-30 16:07:16 -04:00
@[simp]
2024-07-30 16:31:38 -04:00
lemma metric_contrRight_unit (μ : 𝓣.Color) (x : 𝓣.ColorModule μ) :
(contrRightAux (𝓣.contrDual μ)) (𝓣.metric μ ⊗ₜ[R]
2024-07-30 16:07:16 -04:00
((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]
2024-07-30 16:31:38 -04:00
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, assoc_symm_tmul,
map_tmul, LinearMap.id_coe, id_eq]
2024-07-30 16:07:16 -04:00
rw [𝓣.metric_dual]
2024-08-30 13:40:32 -04:00
exact unit_lid 𝓣
2024-07-30 16:07:16 -04:00
/-!
## Dualizing
-/
2024-07-30 16:31:38 -04:00
/-- 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 μ) ∘ₗ
2024-07-30 16:07:16 -04:00
TensorProduct.lTensorHomToHomLTensor R _ _ _ (𝓣.metric μ ⊗ₜ[R] LinearMap.id)
2024-07-30 16:31:38 -04:00
/-- Takes a vector to a vector with the dual color index.
Obtained by contraction with the metric. -/
def dualizeFun (μ : 𝓣.Color) : 𝓣.ColorModule μ →ₗ[R] 𝓣.ColorModule (𝓣.τ μ) :=
2024-07-30 16:07:16 -04:00
𝓣.dualizeSymm (𝓣.τ μ) ∘ₗ (𝓣.colorModuleCast (𝓣.τ_involutive μ).symm).toLinearMap
2024-07-30 16:31:38 -04:00
/-- Equivalence between the module with a color `μ` and the module with color
`𝓣.τ μ` obtained by contraction with the metric. -/
2024-07-30 16:07:16 -04:00
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]
2024-08-30 13:40:32 -04:00
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
2024-08-20 14:38:29 -04:00
· rfl
· rw [dualizeModule_equivariant]
rfl
2024-07-30 16:31:38 -04:00
/-- Dualizes the color of all indicies of a tensor by contraction with the metric. -/
2024-07-30 16:07:16 -04:00
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
2024-09-04 07:31:34 -04:00
simp only [Function.comp_apply, map_add, hx, LinearMap.id_coe, id_eq, LinearMap.coe_comp]
2024-07-30 16:07:16 -04:00
simp_all only [Function.comp_apply, LinearMap.coe_comp, LinearMap.id_coe, id_eq])
intro rx fx
2024-07-30 16:31:38 -04:00
simp only [Function.comp_apply, PiTensorProduct.tprodCoeff_eq_smul_tprod,
LinearMapClass.map_smul, LinearMap.coe_comp, LinearMap.id_coe, id_eq]
2024-07-30 16:07:16 -04:00
apply congrArg
change (PiTensorProduct.map _)
((PiTensorProduct.map _) ((PiTensorProduct.tprod R) fx)) = _
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod]
apply congrArg
simp
2024-09-04 06:36:42 -04:00
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
2024-09-04 07:31:34 -04:00
simp only [dualizeAll, Function.comp_apply, LinearEquiv.ofLinear_toLinearMap, LinearMap.coe_comp]
refine PiTensorProduct.induction_on' x ?_ (by
intro a b hx a
2024-09-04 07:31:34 -04:00
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
2024-09-04 06:28:46 -04:00
omit [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] in
2024-07-30 16:07:16 -04:00
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
2024-09-04 06:28:46 -04:00
omit [Fintype C] [DecidableEq C] [Fintype P] [DecidableEq P] in
2024-07-30 16:07:16 -04:00
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
2024-07-30 16:31:38 -04:00
/-- 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) :=
2024-07-30 16:07:16 -04:00
𝓣.mapIso e.symm (𝓣.dualize_cond e) ≪≫ₗ
(𝓣.tensoratorEquiv _ _).symm ≪≫ₗ
TensorProduct.congr 𝓣.dualizeAll (LinearEquiv.refl _ _) ≪≫ₗ
(𝓣.tensoratorEquiv _ _) ≪≫ₗ
𝓣.mapIso e (𝓣.dualize_cond' e)
2024-09-04 06:28:46 -04:00
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]
2024-07-30 16:07:16 -04:00
end TensorStructure
end