chore: Bump to 4.11.0

This commit is contained in:
jstoobysmith 2024-09-04 06:28:46 -04:00
parent b6162217b7
commit 17f09022db
48 changed files with 404 additions and 137 deletions

View file

@ -62,22 +62,32 @@ 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'
@ -111,6 +121,9 @@ variable {e : (C ⊕ C) ⊕ P ≃ X} {e' : Y ≃ Z} {cX : ColorMap 𝓒 X} {cY :
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
@ -159,6 +172,7 @@ def pairProd : 𝓣.Tensor cX ⊗[R] 𝓣.Tensor cX2 →ₗ[R]
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) =
@ -167,6 +181,7 @@ lemma pairProd_tmul_tprod_tprod (fx : (i : X) → 𝓣.ColorModule (cX i))
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)) ∘ₗ
@ -229,13 +244,16 @@ lemma contrAll'_mapIso (e : X ≃ Y) (h : cX.MapIso e cY) :
refine fun x ↦
PiTensorProduct.induction_on' x ?_ (by
intro a b hx hy y
simp [map_add, add_tmul, hx, hy])
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 at hx hy
simp [map_add, tmul_add, 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,
@ -278,11 +296,14 @@ def contrAll (e : X ≃ Y) (h : cX.ContrAll e cY) : 𝓣.Tensor cX ⊗[R] 𝓣.T
𝓣.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) :
@ -291,6 +312,7 @@ lemma contrAll_mapIso_right_tmul (e : X ≃ Y) (e' : Z ≃ Y)
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 ∘ₗ
@ -300,6 +322,7 @@ lemma contrAll_comp_mapIso_right (e : X ≃ Y) (e' : Z ≃ Y)
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) :
@ -308,6 +331,7 @@ lemma contrAll_mapIso_left_tmul {e : X ≃ Y} {e' : Z ≃ X}
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) :
@ -399,6 +423,7 @@ lemma contrAll_rep_tmul {c : X → 𝓣.Color} {d : Y → 𝓣.Color} (e : X ≃
-/
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]
@ -419,6 +444,8 @@ def contr (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e) :
(𝓣.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]
@ -431,6 +458,7 @@ lemma contr_tprod (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrCond e) (f : (i : X)
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)) :
@ -441,6 +469,7 @@ lemma contr_tprod_isEmpty [IsEmpty C] (e : (C ⊕ C) ⊕ P ≃ X) (h : cX.ContrC
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)