refactor: Lint
This commit is contained in:
parent
62fdab3ace
commit
72cc8c5cdc
1 changed files with 195 additions and 154 deletions
|
@ -3,17 +3,6 @@ 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.Logic.Function.CompTypeclasses
|
||||
import Mathlib.Data.Real.Basic
|
||||
import Mathlib.Data.Fintype.BigOperators
|
||||
import Mathlib.Logic.Equiv.Fin
|
||||
import Mathlib.Tactic.FinCases
|
||||
import Mathlib.Logic.Equiv.Fintype
|
||||
import Mathlib.Algebra.Module.Pi
|
||||
import Mathlib.Algebra.Module.Equiv
|
||||
import Mathlib.Algebra.Module.LinearMap.Basic
|
||||
import Mathlib.LinearAlgebra.TensorProduct.Basic
|
||||
import Mathlib.LinearAlgebra.TensorProduct.Basis
|
||||
import Mathlib.LinearAlgebra.PiTensorProduct
|
||||
import Mathlib.RepresentationTheory.Basic
|
||||
/-!
|
||||
|
@ -42,12 +31,15 @@ structure PreTensorStructure (R : Type) [CommSemiring R] where
|
|||
Color : Type
|
||||
/-- To each color we associate a module. -/
|
||||
ColorModule : Color → Type
|
||||
/-- A map taking every color to it dual color. -/
|
||||
/-- A map taking every color to its dual color. -/
|
||||
τ : Color → Color
|
||||
/-- The map `τ` is an involution. -/
|
||||
τ_involutive : Function.Involutive τ
|
||||
/-- Each `ColorModule` has the structure of an additive commutative monoid. -/
|
||||
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 wiwth dual color. -/
|
||||
/-- The contraction of a vector with a vector with dual color. -/
|
||||
contrDual : ∀ μ, ColorModule μ ⊗[R] ColorModule (τ μ) →ₗ[R] R
|
||||
|
||||
namespace PreTensorStructure
|
||||
|
@ -56,20 +48,23 @@ variable (𝓣 : PreTensorStructure 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]
|
||||
{c c₂ : X → 𝓣.Color} {d : Y → 𝓣.Color} {b : Z → 𝓣.Color}
|
||||
{bw : W → 𝓣.Color} {d' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color}
|
||||
{cW : W → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.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 : X → 𝓣.Color) : Type := ⨂[R] x, 𝓣.ColorModule (c x)
|
||||
|
||||
instance : AddCommMonoid (𝓣.Tensor c) :=
|
||||
PiTensorProduct.instAddCommMonoid fun i => 𝓣.ColorModule (c i)
|
||||
instance : AddCommMonoid (𝓣.Tensor cX) :=
|
||||
PiTensorProduct.instAddCommMonoid fun i => 𝓣.ColorModule (cX i)
|
||||
|
||||
instance : Module R (𝓣.Tensor c) := PiTensorProduct.instModule
|
||||
instance : Module R (𝓣.Tensor cX) := PiTensorProduct.instModule
|
||||
|
||||
/-- Equivalence of `ColorModule` given an equality of colors. -/
|
||||
def colorModuleCast (h : μ = ν) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModule ν where
|
||||
toFun x := Equiv.cast (congrArg 𝓣.ColorModule h) x
|
||||
invFun x := (Equiv.cast (congrArg 𝓣.ColorModule h)).symm x
|
||||
|
@ -79,14 +74,11 @@ def colorModuleCast (h : μ = ν) : 𝓣.ColorModule μ ≃ₗ[R] 𝓣.ColorModu
|
|||
map_smul' x y := by
|
||||
subst h
|
||||
rfl
|
||||
left_inv x := by
|
||||
subst h
|
||||
rfl
|
||||
right_inv x := 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
|
||||
|
||||
lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M] {f g : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d →ₗ[R] M}
|
||||
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'
|
||||
|
@ -101,11 +93,10 @@ lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M] {f g
|
|||
simp at hx hy
|
||||
simp [map_add, tmul_add, hx, hy])
|
||||
intro ry fy
|
||||
simp
|
||||
simp only [PiTensorProduct.tprodCoeff_eq_smul_tprod, tmul_smul, LinearMapClass.map_smul]
|
||||
apply congrArg
|
||||
simp [TensorProduct.smul_tmul]
|
||||
apply congrArg
|
||||
exact h fx fy
|
||||
simp only [smul_tmul, tmul_smul, LinearMapClass.map_smul]
|
||||
exact congrArg (HSMul.hSMul rx) (h fx fy)
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -113,19 +104,20 @@ lemma tensorProd_piTensorProd_ext {M : Type} [AddCommMonoid M] [Module R M] {f g
|
|||
|
||||
-/
|
||||
|
||||
/-- An linear equivalence of tensor spaces given a color-preserving equivalence of indexing sets. -/
|
||||
def mapIso {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = d ∘ e) :
|
||||
𝓣.Tensor c ≃ₗ[R] 𝓣.Tensor d :=
|
||||
(PiTensorProduct.reindex R _ e) ≪≫ₗ
|
||||
(PiTensorProduct.congr (fun y => 𝓣.colorModuleCast (by rw [h]; simp)))
|
||||
|
||||
lemma mapIso_trans_cond {e : X ≃ Y} {e' : Y ≃ Z} (h : c = d ∘ e) (h' : d = b ∘ e') :
|
||||
c = b ∘ (e.trans e') := by
|
||||
lemma mapIso_trans_cond {e : X ≃ Y} {e' : Y ≃ Z} (h : cX = cY ∘ e) (h' : cY = cZ ∘ e') :
|
||||
cX = cZ ∘ (e.trans e') := by
|
||||
funext a
|
||||
subst h h'
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma mapIso_trans (e : X ≃ Y) (e' : Y ≃ Z) (h : c = d ∘ e) (h' : d = b ∘ e') :
|
||||
lemma mapIso_trans (e : X ≃ Y) (e' : Y ≃ Z) (h : cX = cY ∘ e) (h' : cY = cZ ∘ e') :
|
||||
(𝓣.mapIso e h ≪≫ₗ 𝓣.mapIso e' h') = 𝓣.mapIso (e.trans e') (𝓣.mapIso_trans_cond h h') := by
|
||||
refine LinearEquiv.toLinearMap_inj.mp ?_
|
||||
apply PiTensorProduct.ext
|
||||
|
@ -134,33 +126,34 @@ lemma mapIso_trans (e : X ≃ Y) (e' : Y ≃ Z) (h : c = d ∘ e) (h' : d = b
|
|||
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 (d x)) e')
|
||||
((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 (c x)) (e.trans e')) _)
|
||||
((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]
|
||||
|
||||
@[simp]
|
||||
lemma mapIso_mapIso (e : X ≃ Y) (e' : Y ≃ Z) (h : c = d ∘ e) (h' : d = b ∘ e')
|
||||
(T : 𝓣.Tensor c) :
|
||||
lemma mapIso_mapIso (e : X ≃ Y) (e' : Y ≃ Z) (h : cX = cY ∘ e) (h' : cY = cZ ∘ e')
|
||||
(T : 𝓣.Tensor cX) :
|
||||
(𝓣.mapIso e' h') (𝓣.mapIso e h T) = 𝓣.mapIso (e.trans e') (𝓣.mapIso_trans_cond h h') T := by
|
||||
rw [← LinearEquiv.trans_apply, mapIso_trans]
|
||||
|
||||
@[simp]
|
||||
lemma mapIso_symm (e : X ≃ Y) (h : c = d ∘ e) :
|
||||
(𝓣.mapIso e h).symm = 𝓣.mapIso e.symm ((Equiv.eq_comp_symm e d c).mpr h.symm) := by
|
||||
lemma mapIso_symm (e : X ≃ Y) (h : cX = cY ∘ e) :
|
||||
(𝓣.mapIso e h).symm = 𝓣.mapIso e.symm ((Equiv.eq_comp_symm e cY cX).mpr h.symm) := by
|
||||
refine LinearEquiv.toLinearMap_inj.mp ?_
|
||||
apply PiTensorProduct.ext
|
||||
apply MultilinearMap.ext
|
||||
intro x
|
||||
simp [mapIso, LinearMap.compMultilinearMap_apply, LinearEquiv.coe_coe,
|
||||
LinearEquiv.symm_apply_apply, PiTensorProduct.reindex_tprod]
|
||||
change (PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (c x)) e).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 (d x)) e.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
|
||||
|
@ -173,7 +166,7 @@ lemma mapIso_symm (e : X ≃ Y) (h : c = d ∘ e) :
|
|||
rw [Equiv.apply_symm_apply]
|
||||
|
||||
@[simp]
|
||||
lemma mapIso_refl : 𝓣.mapIso (Equiv.refl X) (rfl : c = c) = LinearEquiv.refl R _ := by
|
||||
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
|
||||
|
@ -187,14 +180,14 @@ lemma mapIso_refl : 𝓣.mapIso (Equiv.refl X) (rfl : c = c) = LinearEquiv.refl
|
|||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma mapIso_tprod {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = d ∘ e) (f : (i : X) → 𝓣.ColorModule (c i)) :
|
||||
(𝓣.mapIso e h) (PiTensorProduct.tprod R f) =
|
||||
lemma mapIso_tprod {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (h : c = d ∘ e)
|
||||
(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 [mapIso]
|
||||
change (PiTensorProduct.congr fun y => 𝓣.colorModuleCast _)
|
||||
((PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (c x)) e) ((PiTensorProduct.tprod R) f)) = _
|
||||
((PiTensorProduct.reindex R _ e) ((PiTensorProduct.tprod R) f)) = _
|
||||
rw [PiTensorProduct.reindex_tprod]
|
||||
simp only [PiTensorProduct.congr_tprod]
|
||||
exact PiTensorProduct.congr_tprod (fun y => 𝓣.colorModuleCast _) fun i => f (e.symm i)
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -205,17 +198,20 @@ 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)
|
||||
|
||||
def elimPureTensor (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d) : 𝓣.PureTensor (Sum.elim c d) :=
|
||||
/-- 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
|
||||
|
||||
@[simp]
|
||||
lemma elimPureTensor_update_right (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d)
|
||||
(y : Y) (r : 𝓣.ColorModule (d y)) : 𝓣.elimPureTensor p (Function.update q y r) =
|
||||
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
|
||||
|
@ -230,8 +226,8 @@ lemma elimPureTensor_update_right (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d
|
|||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma elimPureTensor_update_left (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d)
|
||||
(x : X) (r : 𝓣.ColorModule (c x)) : 𝓣.elimPureTensor (Function.update p x r) q =
|
||||
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
|
||||
|
@ -245,50 +241,58 @@ lemma elimPureTensor_update_left (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d)
|
|||
rfl
|
||||
| Sum.inr y => rfl
|
||||
|
||||
def inlPureTensor (p : 𝓣.PureTensor (Sum.elim c d)) : 𝓣.PureTensor c := fun xy => p (Sum.inl xy)
|
||||
/-- 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)
|
||||
|
||||
def inrPureTensor (p : 𝓣.PureTensor (Sum.elim c d)) : 𝓣.PureTensor d := fun xy => p (Sum.inr xy)
|
||||
/-- 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)
|
||||
|
||||
@[simp]
|
||||
lemma inlPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim c d)) (x : X)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim c d (Sum.inl x))) :
|
||||
𝓣.inlPureTensor (Function.update f (Sum.inl x) v1) =Function.update (𝓣.inlPureTensor f) x v1 := by
|
||||
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 [inlPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
next h =>
|
||||
subst h
|
||||
simp_all only
|
||||
next h => simp_all only
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma inrPureTensor_update_left [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim c d)) (x : X)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim c d (Sum.inl x))) :
|
||||
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]
|
||||
|
||||
@[simp]
|
||||
lemma inrPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim c d)) (y : Y)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim c d (Sum.inr y))) :
|
||||
𝓣.inrPureTensor (Function.update f (Sum.inr y) v1) =Function.update (𝓣.inrPureTensor f) y v1 := by
|
||||
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 [inrPureTensor, Function.update, Sum.inl.injEq, Sum.elim_inl]
|
||||
split
|
||||
next h =>
|
||||
subst h
|
||||
simp_all only
|
||||
next h => simp_all only
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma inlPureTensor_update_right [DecidableEq (X ⊕ Y)] (f : 𝓣.PureTensor (Sum.elim c d)) (y : Y)
|
||||
(v1 : 𝓣.ColorModule (Sum.elim c d (Sum.inr y))) :
|
||||
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]
|
||||
|
||||
def elimPureTensorMulLin : MultilinearMap R (fun i => 𝓣.ColorModule (c i))
|
||||
(MultilinearMap R (fun x => 𝓣.ColorModule (d x)) (𝓣.Tensor (Sum.elim c d))) where
|
||||
/-- 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
|
||||
|
@ -311,7 +315,10 @@ def elimPureTensorMulLin : MultilinearMap R (fun i => 𝓣.ColorModule (c i))
|
|||
-/
|
||||
|
||||
/-! TODO: Replace with dependent type version of `MultilinearMap.domCoprod` when in Mathlib. -/
|
||||
def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim c d x)) (𝓣.Tensor c ⊗[R] 𝓣.Tensor d) where
|
||||
/-- 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
|
||||
|
@ -323,7 +330,9 @@ def domCoprod : MultilinearMap R (fun x => 𝓣.ColorModule (Sum.elim c d x)) (
|
|||
| Sum.inl x => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
|
||||
| Sum.inr y => simp [TensorProduct.tmul_smul, TensorProduct.smul_tmul]
|
||||
|
||||
def tensoratorSymm : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d →ₗ[R] 𝓣.Tensor (Sum.elim c d) := by
|
||||
/-- 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 <|
|
||||
|
@ -332,10 +341,13 @@ def tensoratorSymm : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d →ₗ[R] 𝓣.Tensor (S
|
|||
map_smul' := fun r a ↦ by simp}
|
||||
|
||||
/-! TODO: Replace with dependent type version of `PiTensorProduct.tmulEquiv` when in Mathlib. -/
|
||||
def tensorator : 𝓣.Tensor (Sum.elim c d) →ₗ[R] 𝓣.Tensor c ⊗[R] 𝓣.Tensor d :=
|
||||
/-- 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
|
||||
|
||||
def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d ≃ₗ[R] 𝓣.Tensor (Sum.elim c d) :=
|
||||
/-- 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
|
||||
|
@ -356,23 +368,23 @@ def tensoratorEquiv (c : X → 𝓣.Color) (d : Y → 𝓣.Color) : 𝓣.Tensor
|
|||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp [tensorator, tensoratorSymm]
|
||||
change (PiTensorProduct.lift 𝓣.domCoprod) ((PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q)) =_
|
||||
change (PiTensorProduct.lift 𝓣.domCoprod)
|
||||
((PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q)) =_
|
||||
rw [PiTensorProduct.lift.tprod]
|
||||
simp [elimPureTensorMulLin]
|
||||
rfl)
|
||||
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_tmul_tprod (p : 𝓣.PureTensor c) (q : 𝓣.PureTensor d) :
|
||||
(𝓣.tensoratorEquiv c d) ((PiTensorProduct.tprod R) p ⊗ₜ[R] (PiTensorProduct.tprod R) q) =
|
||||
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 [tensoratorEquiv, tensorator, tensoratorSymm, domCoprod]
|
||||
change (PiTensorProduct.lift (𝓣.elimPureTensorMulLin p)) ((PiTensorProduct.tprod R) q) = _
|
||||
rw [PiTensorProduct.lift.tprod]
|
||||
simp [elimPureTensorMulLin]
|
||||
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
|
||||
|
||||
lemma tensoratorEquiv_mapIso_cond {e : X ≃ Y} {e' : Z ≃ Y} {e'' : W ≃ X}
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') (h'' : bW = c ∘ e'') :
|
||||
Sum.elim bW b = Sum.elim c d ∘ ⇑(e''.sumCongr e') := by
|
||||
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : bW = cX ∘ e'') :
|
||||
Sum.elim bW cZ = Sum.elim cX cY ∘ ⇑(e''.sumCongr e') := by
|
||||
subst h h' h''
|
||||
funext x
|
||||
match x with
|
||||
|
@ -381,14 +393,15 @@ lemma tensoratorEquiv_mapIso_cond {e : X ≃ Y} {e' : Z ≃ Y} {e'' : W ≃ X}
|
|||
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_mapIso (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') (h'' : bW = c ∘ e'') :
|
||||
(TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ (𝓣.tensoratorEquiv c d)
|
||||
= (𝓣.tensoratorEquiv bW b)
|
||||
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : bW = cX ∘ e'') :
|
||||
(TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ (𝓣.tensoratorEquiv cX cY)
|
||||
= (𝓣.tensoratorEquiv bW cZ)
|
||||
≪≫ₗ (𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h'')) := by
|
||||
apply LinearEquiv.toLinearMap_inj.mp
|
||||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp
|
||||
simp only [LinearEquiv.coe_coe, LinearEquiv.trans_apply, congr_tmul, mapIso_tprod,
|
||||
tensoratorEquiv_tmul_tprod, Equiv.sumCongr_symm, Equiv.sumCongr_apply]
|
||||
apply congrArg
|
||||
funext x
|
||||
match x with
|
||||
|
@ -397,11 +410,13 @@ lemma tensoratorEquiv_mapIso (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
|||
|
||||
@[simp]
|
||||
lemma tensoratorEquiv_mapIso_apply (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') (h'' : bW = c ∘ e'')
|
||||
(x : 𝓣.Tensor bW ⊗[R] 𝓣.Tensor b) :
|
||||
(𝓣.tensoratorEquiv c d) ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) x) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h'')) ((𝓣.tensoratorEquiv bW b) x) := by
|
||||
trans ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ (𝓣.tensoratorEquiv c d)) x
|
||||
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : cW = cX ∘ e'')
|
||||
(x : 𝓣.Tensor cW ⊗[R] 𝓣.Tensor cZ) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) x) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h''))
|
||||
((𝓣.tensoratorEquiv cW cZ) x) := by
|
||||
trans ((TensorProduct.congr (𝓣.mapIso e'' h'') (𝓣.mapIso e' h')) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv cX cY)) x
|
||||
rfl
|
||||
rw [tensoratorEquiv_mapIso]
|
||||
rfl
|
||||
|
@ -409,10 +424,11 @@ lemma tensoratorEquiv_mapIso_apply (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
|||
exact h
|
||||
|
||||
lemma tensoratorEquiv_mapIso_tmul (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') (h'' : bW = c ∘ e'')
|
||||
(x : 𝓣.Tensor bW) (y : 𝓣.Tensor b) :
|
||||
(𝓣.tensoratorEquiv c d) ((𝓣.mapIso e'' h'' x) ⊗ₜ[R] (𝓣.mapIso e' h' y)) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h'')) ((𝓣.tensoratorEquiv bW b) (x ⊗ₜ y)) := by
|
||||
(h : cX = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (h'' : cW = cX ∘ e'')
|
||||
(x : 𝓣.Tensor cW) (y : 𝓣.Tensor cZ) :
|
||||
(𝓣.tensoratorEquiv cX cY) ((𝓣.mapIso e'' h'' x) ⊗ₜ[R] (𝓣.mapIso e' h' y)) =
|
||||
(𝓣.mapIso (Equiv.sumCongr e'' e') (𝓣.tensoratorEquiv_mapIso_cond h h' h''))
|
||||
((𝓣.tensoratorEquiv cW cZ) (x ⊗ₜ y)) := by
|
||||
rw [← tensoratorEquiv_mapIso_apply]
|
||||
rfl
|
||||
exact e
|
||||
|
@ -424,6 +440,7 @@ lemma tensoratorEquiv_mapIso_tmul (e : X ≃ Y) (e' : Z ≃ Y) (e'' : W ≃ X)
|
|||
|
||||
-/
|
||||
|
||||
/-- 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 <|
|
||||
|
@ -431,9 +448,14 @@ def decompEmbedSet (f : Y ↪ X) :
|
|||
Equiv.sumCongr ((Equiv.subtypeEquivRight (by simp))) <|
|
||||
(Function.Embedding.toEquivRange f).symm
|
||||
|
||||
def decompEmbedColorLeft (c : X → 𝓣.Color) (f : Y ↪ X) : {x // x ∈ (Finset.image f Finset.univ)ᶜ} → 𝓣.Color :=
|
||||
/-- 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
|
||||
|
||||
|
@ -441,11 +463,13 @@ 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 based on an embedding. -/
|
||||
/-- 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 c ≃ₗ[R] 𝓣.Tensor (𝓣.decompEmbedColorLeft c f) ⊗[R] 𝓣.Tensor (c ∘ f) :=
|
||||
(𝓣.mapIso (decompEmbedSet f) (𝓣.decompEmbed_cond c f)) ≪≫ₗ
|
||||
(𝓣.tensoratorEquiv (𝓣.decompEmbedColorLeft c f) (𝓣.decompEmbedColorRight c f)).symm
|
||||
𝓣.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
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -453,11 +477,12 @@ def decompEmbed (f : Y ↪ X) :
|
|||
|
||||
-/
|
||||
|
||||
def pairProd : 𝓣.Tensor c ⊗[R] 𝓣.Tensor c₂ →ₗ[R]
|
||||
⨂[R] x, 𝓣.ColorModule (c x) ⊗[R] 𝓣.ColorModule (c₂ x) :=
|
||||
/-- 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 (c x)) (𝓣.ColorModule (c₂ x))))
|
||||
TensorProduct.mk R (𝓣.ColorModule (cX x)) (𝓣.ColorModule (cX2 x))))
|
||||
|
||||
lemma mkPiAlgebra_equiv (e : X ≃ Y) :
|
||||
(PiTensorProduct.lift (MultilinearMap.mkPiAlgebra R X R)) =
|
||||
|
@ -470,23 +495,24 @@ lemma mkPiAlgebra_equiv (e : X ≃ Y) :
|
|||
MultilinearMap.mkPiAlgebra_apply, LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply,
|
||||
PiTensorProduct.reindex_tprod, Equiv.prod_comp]
|
||||
|
||||
def contrAll' : 𝓣.Tensor c ⊗[R] 𝓣.Tensor (𝓣.τ ∘ c) →ₗ[R] R :=
|
||||
/-- 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 (c x))) ∘ₗ
|
||||
(PiTensorProduct.map (fun x => 𝓣.contrDual (cX x))) ∘ₗ
|
||||
(𝓣.pairProd)
|
||||
|
||||
lemma contrAll'_mapIso_cond {e : X ≃ Y} (h : c = d ∘ e) :
|
||||
𝓣.τ ∘ d = (𝓣.τ ∘ c) ∘ ⇑e.symm := by
|
||||
lemma contrAll'_mapIso_cond {e : X ≃ Y} (h : cX = cY ∘ e) :
|
||||
𝓣.τ ∘ cY = (𝓣.τ ∘ cX) ∘ ⇑e.symm := by
|
||||
subst h
|
||||
ext1 x
|
||||
simp
|
||||
exact (Equiv.eq_comp_symm e (𝓣.τ ∘ cY) (𝓣.τ ∘ cY ∘ ⇑e)).mpr rfl
|
||||
|
||||
@[simp]
|
||||
lemma contrAll'_mapIso (e : X ≃ Y) (h : c = d ∘ e) :
|
||||
𝓣.contrAll' ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e h) (LinearEquiv.refl R _)).toLinearMap =
|
||||
𝓣.contrAll' ∘ₗ (TensorProduct.congr (LinearEquiv.refl R _)
|
||||
(𝓣.mapIso e.symm (𝓣.contrAll'_mapIso_cond h))).toLinearMap := by
|
||||
lemma contrAll'_mapIso (e : X ≃ Y) (h : c = cY ∘ e) :
|
||||
𝓣.contrAll' ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e h) (LinearEquiv.refl R _)).toLinearMap =
|
||||
𝓣.contrAll' ∘ₗ (TensorProduct.congr (LinearEquiv.refl R _)
|
||||
(𝓣.mapIso e.symm (𝓣.contrAll'_mapIso_cond h))).toLinearMap := by
|
||||
apply TensorProduct.ext'
|
||||
refine fun x ↦
|
||||
PiTensorProduct.induction_on' x ?_ (by
|
||||
|
@ -502,67 +528,70 @@ lemma contrAll'_mapIso (e : X ≃ Y) (h : c = d ∘ e) :
|
|||
simp [contrAll']
|
||||
rw [mkPiAlgebra_equiv e]
|
||||
apply congrArg
|
||||
simp
|
||||
simp only [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply]
|
||||
apply congrArg
|
||||
rw [← LinearEquiv.symm_apply_eq]
|
||||
rw [PiTensorProduct.reindex_symm]
|
||||
rw [← PiTensorProduct.map_reindex]
|
||||
subst h
|
||||
simp
|
||||
simp only [Equiv.symm_symm_apply, Function.comp_apply]
|
||||
apply congrArg
|
||||
rw [pairProd, pairProd]
|
||||
simp
|
||||
simp only [Function.comp_apply, lift.tmul, LinearMapClass.map_smul, LinearMap.smul_apply]
|
||||
apply congrArg
|
||||
change _ = ((PiTensorProduct.map₂ fun x => TensorProduct.mk R (𝓣.ColorModule (d (e x))) (𝓣.ColorModule (𝓣.τ (d (e x)))))
|
||||
change _ = ((PiTensorProduct.map₂ fun x => TensorProduct.mk R (𝓣.ColorModule (cY (e x)))
|
||||
(𝓣.ColorModule (𝓣.τ (cY (e x)))))
|
||||
((PiTensorProduct.tprod R) fx))
|
||||
((𝓣.mapIso e.symm _) ((PiTensorProduct.tprod R) fy))
|
||||
rw [mapIso_tprod]
|
||||
simp
|
||||
simp only [Equiv.symm_symm_apply, Function.comp_apply]
|
||||
rw [PiTensorProduct.map₂_tprod_tprod]
|
||||
change (PiTensorProduct.reindex R (fun x => 𝓣.ColorModule (d x) ⊗[R] 𝓣.ColorModule (𝓣.τ (d x))) e.symm)
|
||||
(((PiTensorProduct.map₂ fun x => TensorProduct.mk R (𝓣.ColorModule (d x)) (𝓣.ColorModule (𝓣.τ (d x))))
|
||||
change (PiTensorProduct.reindex R (_) e.symm)
|
||||
(((PiTensorProduct.map₂ _)
|
||||
((PiTensorProduct.tprod R) fun i => (𝓣.colorModuleCast _) (fx (e.symm i))))
|
||||
((PiTensorProduct.tprod R) fy)) = _
|
||||
rw [PiTensorProduct.map₂_tprod_tprod]
|
||||
simp
|
||||
simp only [Equiv.symm_symm_apply, Function.comp_apply, mk_apply]
|
||||
erw [PiTensorProduct.reindex_tprod]
|
||||
apply congrArg
|
||||
funext i
|
||||
simp
|
||||
simp only [Equiv.symm_symm_apply]
|
||||
congr
|
||||
simp [colorModuleCast]
|
||||
apply cast_eq_iff_heq.mpr
|
||||
rw [Equiv.symm_apply_apply]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll'_mapIso_tmul (e : X ≃ Y) (h : c = d ∘ e) (x : 𝓣.Tensor c) (y : 𝓣.Tensor (𝓣.τ ∘ d)) :
|
||||
𝓣.contrAll' ((𝓣.mapIso e h) x ⊗ₜ[R] y) = 𝓣.contrAll' (x ⊗ₜ[R] (𝓣.mapIso e.symm (𝓣.contrAll'_mapIso_cond h) y)) := by
|
||||
lemma contrAll'_mapIso_tmul (e : X ≃ Y) (h : c = cY ∘ e) (x : 𝓣.Tensor c)
|
||||
(y : 𝓣.Tensor (𝓣.τ ∘ cY)) : 𝓣.contrAll' ((𝓣.mapIso e h) x ⊗ₜ[R] y) =
|
||||
𝓣.contrAll' (x ⊗ₜ[R] (𝓣.mapIso e.symm (𝓣.contrAll'_mapIso_cond h) 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 {c : X → 𝓣.Color} {d : Y → 𝓣.Color}
|
||||
(e : X ≃ Y) (h : c = 𝓣.τ ∘ d ∘ e) : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d →ₗ[R] R :=
|
||||
𝓣.contrAll' ∘ₗ (TensorProduct.congr (LinearEquiv.refl _ _)
|
||||
(𝓣.mapIso e.symm (by subst h; funext a; simp; rw [𝓣.τ_involutive]))).toLinearMap
|
||||
|
||||
lemma contrAll_symm_cond {e : X ≃ Y} (h : c = 𝓣.τ ∘ d ∘ e) :
|
||||
d = 𝓣.τ ∘ c ∘ ⇑e.symm := by
|
||||
lemma contrAll_symm_cond {e : X ≃ Y} (h : c = 𝓣.τ ∘ cY ∘ e) :
|
||||
cY = 𝓣.τ ∘ c ∘ ⇑e.symm := by
|
||||
subst h
|
||||
ext1 x
|
||||
simp
|
||||
simp only [Function.comp_apply, Equiv.apply_symm_apply]
|
||||
rw [𝓣.τ_involutive]
|
||||
|
||||
lemma contrAll_mapIso_right_cond {e : X ≃ Y} {e' : Z ≃ Y}
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') : c = 𝓣.τ ∘ b ∘ ⇑(e.trans e'.symm) := by
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') : c = 𝓣.τ ∘ cZ ∘ ⇑(e.trans e'.symm) := by
|
||||
subst h h'
|
||||
ext1 x
|
||||
simp only [Function.comp_apply, Equiv.coe_trans, Equiv.apply_symm_apply]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_right_tmul (e : X ≃ Y) (e' : Z ≃ Y)
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') (x : 𝓣.Tensor c) (z : 𝓣.Tensor b) :
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') (x : 𝓣.Tensor c) (z : 𝓣.Tensor cZ) :
|
||||
𝓣.contrAll e h (x ⊗ₜ[R] (𝓣.mapIso e' h' z)) =
|
||||
𝓣.contrAll (e.trans e'.symm) (𝓣.contrAll_mapIso_right_cond h h') (x ⊗ₜ[R] z) := by
|
||||
rw [contrAll, contrAll]
|
||||
|
@ -572,7 +601,7 @@ lemma contrAll_mapIso_right_tmul (e : X ≃ Y) (e' : Z ≃ Y)
|
|||
|
||||
@[simp]
|
||||
lemma contrAll_comp_mapIso_right (e : X ≃ Y) (e' : Z ≃ Y)
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = d ∘ e') : 𝓣.contrAll e h ∘ₗ
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = cY ∘ e') : 𝓣.contrAll e h ∘ₗ
|
||||
(TensorProduct.congr (LinearEquiv.refl R (𝓣.Tensor c)) (𝓣.mapIso e' h')).toLinearMap
|
||||
= 𝓣.contrAll (e.trans e'.symm) (𝓣.contrAll_mapIso_right_cond h h') := by
|
||||
apply TensorProduct.ext'
|
||||
|
@ -580,14 +609,14 @@ lemma contrAll_comp_mapIso_right (e : X ≃ Y) (e' : Z ≃ Y)
|
|||
exact 𝓣.contrAll_mapIso_right_tmul e e' h h' x y
|
||||
|
||||
lemma contrAll_mapIso_left_cond {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = c ∘ e') : b = 𝓣.τ ∘ d ∘ ⇑(e'.trans e) := by
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = c ∘ e') : cZ = 𝓣.τ ∘ cY ∘ ⇑(e'.trans e) := by
|
||||
subst h h'
|
||||
ext1 x
|
||||
simp only [Function.comp_apply, Equiv.coe_trans, Equiv.apply_symm_apply]
|
||||
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_left_tmul {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = c ∘ e') (x : 𝓣.Tensor b) (y : 𝓣.Tensor d) :
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = c ∘ e') (x : 𝓣.Tensor cZ) (y : 𝓣.Tensor cY) :
|
||||
𝓣.contrAll e h ((𝓣.mapIso e' h' x) ⊗ₜ[R] y) =
|
||||
𝓣.contrAll (e'.trans e) (𝓣.contrAll_mapIso_left_cond h h') (x ⊗ₜ[R] y) := by
|
||||
rw [contrAll, contrAll]
|
||||
|
@ -597,9 +626,9 @@ lemma contrAll_mapIso_left_tmul {e : X ≃ Y} {e' : Z ≃ X}
|
|||
|
||||
@[simp]
|
||||
lemma contrAll_mapIso_left {e : X ≃ Y} {e' : Z ≃ X}
|
||||
(h : c = 𝓣.τ ∘ d ∘ e) (h' : b = c ∘ e') :
|
||||
(h : c = 𝓣.τ ∘ cY ∘ e) (h' : cZ = c ∘ e') :
|
||||
𝓣.contrAll e h ∘ₗ
|
||||
(TensorProduct.congr (𝓣.mapIso e' h') (LinearEquiv.refl R (𝓣.Tensor d))).toLinearMap
|
||||
(TensorProduct.congr (𝓣.mapIso e' h') (LinearEquiv.refl R (𝓣.Tensor cY))).toLinearMap
|
||||
= 𝓣.contrAll (e'.trans e) (𝓣.contrAll_mapIso_left_cond h h') := by
|
||||
apply TensorProduct.ext'
|
||||
intro x y
|
||||
|
@ -607,7 +636,10 @@ lemma contrAll_mapIso_left {e : X ≃ Y} {e' : Z ≃ X}
|
|||
|
||||
end PreTensorStructure
|
||||
|
||||
/-! TODO: Add unit here. -/
|
||||
/-- A `PreTensorStructure` with the additional constraint that `contrDua` is symmetric. -/
|
||||
structure TensorStructure (R : Type) [CommSemiring R] extends PreTensorStructure R where
|
||||
/-- The symmetry condition on `contrDua`. -/
|
||||
contrDual_symm : ∀ μ,
|
||||
(contrDual μ) ∘ₗ (TensorProduct.comm R (ColorModule (τ μ)) (ColorModule μ)).toLinearMap =
|
||||
(contrDual (τ μ)) ∘ₗ (TensorProduct.congr (LinearEquiv.refl _ _)
|
||||
|
@ -625,9 +657,12 @@ variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [De
|
|||
|
||||
end TensorStructure
|
||||
|
||||
/-- A `TensorStructure` with a group action. -/
|
||||
structure GroupTensorStructure (R : Type) [CommSemiring R]
|
||||
(G : Type) [Group G] extends 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 μ
|
||||
|
||||
|
@ -641,10 +676,11 @@ variable (𝓣 : GroupTensorStructure R G)
|
|||
|
||||
variable {d : ℕ} {X Y Y' Z : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
|
||||
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z]
|
||||
{c c₂ : X → 𝓣.Color} {d : Y → 𝓣.Color} {b : Z → 𝓣.Color} {d' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
{cX cX2 : X → 𝓣.Color} {cY : Y → 𝓣.Color} {cZ : Z → 𝓣.Color} {cY' : Y' → 𝓣.Color} {μ ν: 𝓣.Color}
|
||||
|
||||
def rep : Representation R G (𝓣.Tensor c) where
|
||||
toFun g := PiTensorProduct.map (fun x => 𝓣.repColorModule (c x) g)
|
||||
/-- 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
|
||||
|
@ -654,7 +690,8 @@ def rep : Representation R G (𝓣.Tensor c) where
|
|||
local infixl:78 " • " => 𝓣.rep
|
||||
|
||||
lemma repColorModule_colorModuleCast_apply (h : μ = ν) (g : G) (x : 𝓣.ColorModule μ) :
|
||||
(𝓣.repColorModule ν g) ((𝓣.colorModuleCast h) x) = (𝓣.colorModuleCast h) ((𝓣.repColorModule μ g) x) := by
|
||||
(𝓣.repColorModule ν g) ((𝓣.colorModuleCast h) x) =
|
||||
(𝓣.colorModuleCast h) ((𝓣.repColorModule μ g) x) := by
|
||||
subst h
|
||||
simp [colorModuleCast]
|
||||
|
||||
|
@ -667,17 +704,18 @@ lemma repColorModule_colorModuleCast (h : μ = ν) (g : G) :
|
|||
simp [repColorModule_colorModuleCast_apply]
|
||||
|
||||
@[simp]
|
||||
lemma rep_mapIso (e : X ≃ Y) (h : c = d ∘ e) (g : G) :
|
||||
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
|
||||
simp only [LinearMap.compMultilinearMap_apply, LinearMap.coe_comp, LinearEquiv.coe_coe,
|
||||
Function.comp_apply]
|
||||
erw [mapIso_tprod]
|
||||
simp [rep, repColorModule_colorModuleCast_apply]
|
||||
change (PiTensorProduct.map fun x => (𝓣.repColorModule (d x)) g)
|
||||
change (PiTensorProduct.map fun x => (𝓣.repColorModule (cY x)) g)
|
||||
((PiTensorProduct.tprod R) fun i => (𝓣.colorModuleCast _) (x (e.symm i))) =
|
||||
(𝓣.mapIso e h) ((PiTensorProduct.map fun x => (𝓣.repColorModule (c x)) g) ((PiTensorProduct.tprod R) x))
|
||||
(𝓣.mapIso e h) ((PiTensorProduct.map _) ((PiTensorProduct.tprod R) x))
|
||||
rw [PiTensorProduct.map_tprod, PiTensorProduct.map_tprod]
|
||||
rw [mapIso_tprod]
|
||||
apply congrArg
|
||||
|
@ -686,18 +724,18 @@ lemma rep_mapIso (e : X ≃ Y) (h : c = d ∘ e) (g : G) :
|
|||
simp [repColorModule_colorModuleCast_apply]
|
||||
|
||||
@[simp]
|
||||
lemma rep_mapIso_apply (e : X ≃ Y) (h : c = d ∘ e) (g : G) (x : 𝓣.Tensor c) :
|
||||
lemma rep_mapIso_apply (e : X ≃ Y) (h : cX = cY ∘ e) (g : G) (x : 𝓣.Tensor cX) :
|
||||
g • (𝓣.mapIso e h x) = (𝓣.mapIso e h) (g • x) := by
|
||||
trans ((𝓣.rep g) ∘ₗ (𝓣.mapIso e h).toLinearMap) x
|
||||
rfl
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (c i)) :
|
||||
lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (cX i)) :
|
||||
g • (PiTensorProduct.tprod R f) = PiTensorProduct.tprod R (fun x =>
|
||||
𝓣.repColorModule (c x) g (f x)) := by
|
||||
𝓣.repColorModule (cX x) g (f x)) := by
|
||||
simp [rep]
|
||||
change (PiTensorProduct.map fun x => (𝓣.repColorModule (c x)) g) ((PiTensorProduct.tprod R) f) = _
|
||||
change (PiTensorProduct.map _) ((PiTensorProduct.tprod R) f) = _
|
||||
rw [PiTensorProduct.map_tprod]
|
||||
|
||||
/-!
|
||||
|
@ -707,26 +745,29 @@ lemma rep_tprod (g : G) (f : (i : X) → 𝓣.ColorModule (c i)) :
|
|||
-/
|
||||
|
||||
lemma rep_tensoratorEquiv (g : G) :
|
||||
(𝓣.tensoratorEquiv c d) ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g)) = 𝓣.rep g ∘ₗ
|
||||
(𝓣.tensoratorEquiv c d).toLinearMap := by
|
||||
(𝓣.tensoratorEquiv cX cY) ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g)) = 𝓣.rep g ∘ₗ
|
||||
(𝓣.tensoratorEquiv cX cY).toLinearMap := by
|
||||
apply tensorProd_piTensorProd_ext
|
||||
intro p q
|
||||
simp
|
||||
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
|
||||
|
||||
lemma rep_tensoratorEquiv_apply (g : G) (x : 𝓣.Tensor c ⊗[R] 𝓣.Tensor d) :
|
||||
(𝓣.tensoratorEquiv c d) ((TensorProduct.map (𝓣.rep g) (𝓣.rep g)) x) = (𝓣.rep g) ((𝓣.tensoratorEquiv c d) x) := by
|
||||
trans ((𝓣.tensoratorEquiv c d) ∘ₗ (TensorProduct.map (𝓣.rep g) (𝓣.rep g))) x
|
||||
lemma rep_tensoratorEquiv_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 [rep_tensoratorEquiv]
|
||||
rfl
|
||||
|
||||
lemma rep_tensoratorEquiv_tmul (g : G) (x : 𝓣.Tensor c) (y : 𝓣.Tensor d) :
|
||||
(𝓣.tensoratorEquiv c d) ((g • x) ⊗ₜ[R] (g • y)) = g • ((𝓣.tensoratorEquiv c d) (x ⊗ₜ[R] y)) := by
|
||||
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 [← rep_tensoratorEquiv_apply]
|
||||
rfl
|
||||
|
||||
|
@ -759,7 +800,7 @@ lemma contrAll_rep {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃ Y) (
|
|||
(PiTensorProduct.map _) ((PiTensorProduct.map₂ _ _) _)
|
||||
rw [PiTensorProduct.map₂_tprod_tprod, PiTensorProduct.map₂_tprod_tprod, PiTensorProduct.map_tprod,
|
||||
PiTensorProduct.map_tprod]
|
||||
simp
|
||||
simp only [mk_apply]
|
||||
apply congrArg
|
||||
funext x
|
||||
rw [← repColorModule_colorModuleCast_apply]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue