commit
3b00f963bf
36 changed files with 204 additions and 148 deletions
|
@ -66,7 +66,7 @@ lemma ext_even (S T : Fin (2 * n.succ) → ℚ) (h1 : ∀ i, S (evenFst i) = T (
|
|||
· let j : Fin n.succ := ⟨i - n.succ, by omega⟩
|
||||
have h2 := h2 j
|
||||
have h3 : evenSnd j = i := by
|
||||
simp only [succ_eq_add_one, evenSnd, Fin.ext_iff, Fin.coe_cast, Fin.coe_natAdd]
|
||||
simp only [succ_eq_add_one, evenSnd, Fin.ext_iff, Fin.coe_cast, Fin.coe_natAdd, j]
|
||||
omega
|
||||
rw [h3] at h2
|
||||
exact h2
|
||||
|
@ -699,8 +699,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
|
|||
use f'
|
||||
change P! f' = _ at hf'
|
||||
erw [hf']
|
||||
simp only [and_self, and_true]
|
||||
change S'.val = P g + (P! f + _)
|
||||
simp only [and_self, and_true, X]
|
||||
rw [← add_assoc, ← h]
|
||||
apply swap!_as_add at hS
|
||||
exact hS
|
||||
|
|
|
@ -707,8 +707,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ + 1)).LinSols} (j : Fin n.succ)
|
|||
use f'
|
||||
change P! f' = _ at hf'
|
||||
erw [hf']
|
||||
simp only [and_self, and_true]
|
||||
change S'.val = P g + (P! f + _)
|
||||
simp only [and_self, and_true, X]
|
||||
rw [← add_assoc, ← hS1]
|
||||
apply swap!_as_add at hS
|
||||
exact hS
|
||||
|
|
|
@ -5,6 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import Mathlib.LinearAlgebra.UnitaryGroup
|
||||
import Mathlib.Data.Complex.Exponential
|
||||
import Mathlib.Data.Complex.Trigonometric
|
||||
import Mathlib.Tactic.Polyrith
|
||||
/-!
|
||||
# The CKM Matrix
|
||||
|
@ -321,9 +322,9 @@ lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) :
|
|||
head_fin_const, mul_zero]
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, zero_mul, add_zero, mul_zero,
|
||||
_root_.map_mul, abs_exp, mul_re, I_re, ofReal_re, I_im, ofReal_im, sub_self, Real.exp_zero,
|
||||
one_mul, mul_one, Fin.mk_one, cons_val_one, head_cons, zero_add, head_fin_const,
|
||||
Fin.reduceFinMk, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
_root_.map_mul, mul_re, I_re, ofReal_re, I_im, ofReal_im, sub_self, Real.exp_zero,
|
||||
one_mul, mul_one, Complex.abs_exp, Fin.mk_one, cons_val_one, head_cons, zero_add,
|
||||
head_fin_const, Fin.reduceFinMk, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
tail_val', head_val']
|
||||
all_goals change Complex.abs (0 * _ + _) = _
|
||||
all_goals simp [Complex.abs_exp]
|
||||
|
|
|
@ -47,7 +47,7 @@ lemma complexContrBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 1 ⊕ Fin 3) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [complexContrBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change (((LorentzGroup.toComplex (SL2C.toLorentzGroup M))) *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
simp only [mulVec_single, MulOpposite.op_one, Pi.smul_apply, transpose_apply, one_smul]
|
||||
|
||||
lemma complexContrBasis_ρ_val (M : SL(2,ℂ)) (v : complexContr) :
|
||||
((complexContr.ρ M) v).val =
|
||||
|
@ -74,7 +74,7 @@ lemma complexCoBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 1 ⊕ Fin 3) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [complexCoBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change ((LorentzGroup.toComplex (SL2C.toLorentzGroup M))⁻¹ᵀ *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
simp only [mulVec_single, MulOpposite.op_one, transpose_transpose, Pi.smul_apply, one_smul]
|
||||
|
||||
/-- The standard basis of complex covariant Lorentz vectors indexed by `Fin 4`. -/
|
||||
def complexCoBasisFin4 : Basis (Fin 4) ℂ complexCo :=
|
||||
|
|
|
@ -67,7 +67,7 @@ def contrMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexContr wh
|
|||
simp only [LorentzGroup.toComplex_mul_minkowskiMatrix_mul_transpose]
|
||||
|
||||
lemma contrMetric_apply_one : contrMetric.hom (1 : ℂ) = contrMetricVal := by
|
||||
change contrMetric.hom.hom.toFun (1 : ℂ) = contrMetricVal
|
||||
change (1 : ℂ) • contrMetricVal = contrMetricVal
|
||||
simp only [contrMetric, one_smul]
|
||||
|
||||
/-- The metric `ηᵢᵢ` as an element of `(complexCo ⊗ complexCo).V`. -/
|
||||
|
@ -121,7 +121,7 @@ def coMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexCo where
|
|||
LorentzGroup.toComplex_transpose_mul_minkowskiMatrix_mul_self]
|
||||
|
||||
lemma coMetric_apply_one : coMetric.hom (1 : ℂ) = coMetricVal := by
|
||||
change coMetric.hom.hom.toFun (1 : ℂ) = coMetricVal
|
||||
change (1 : ℂ) • coMetricVal = coMetricVal
|
||||
simp only [coMetric, one_smul]
|
||||
|
||||
/-!
|
||||
|
|
|
@ -66,8 +66,8 @@ def contrCoUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ complexCo where
|
|||
simp
|
||||
|
||||
lemma contrCoUnit_apply_one : contrCoUnit.hom (1 : ℂ) = contrCoUnitVal := by
|
||||
change contrCoUnit.hom.hom.toFun (1 : ℂ) = contrCoUnitVal
|
||||
simp only [contrCoUnit, one_smul]
|
||||
change (1 : ℂ) • contrCoUnitVal = contrCoUnitVal
|
||||
rw [one_smul]
|
||||
|
||||
/-- The co-contra unit for complex lorentz vectors. Usually denoted `δᵢⁱ`. -/
|
||||
def coContrUnitVal : (complexCo ⊗ complexContr).V :=
|
||||
|
@ -118,8 +118,9 @@ def coContrUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexCo ⊗ complexContr where
|
|||
simp
|
||||
|
||||
lemma coContrUnit_apply_one : coContrUnit.hom (1 : ℂ) = coContrUnitVal := by
|
||||
change coContrUnit.hom.hom.toFun (1 : ℂ) = coContrUnitVal
|
||||
simp only [coContrUnit, one_smul]
|
||||
change (1 : ℂ) • coContrUnitVal = coContrUnitVal
|
||||
rw [one_smul]
|
||||
|
||||
/-!
|
||||
|
||||
## Contraction of the units
|
||||
|
@ -151,7 +152,7 @@ lemma contr_contrCoUnit (x : complexCo) :
|
|||
repeat rw (config := { transparency := .instances }) [h1'']
|
||||
repeat rw [coContrContraction_basis']
|
||||
simp only [Fin.isValue, leftUnitor, ModuleCat.MonoidalCategory.leftUnitor, ModuleCat.of_coe,
|
||||
CategoryTheory.Iso.trans_hom, LinearEquiv.toModuleIso_hom_hom, ModuleCat.ofSelfIso_hom,
|
||||
CategoryTheory.Iso.trans_hom, ModuleCat.ofSelfIso_hom,
|
||||
CategoryTheory.Category.comp_id, Action.instMonoidalCategory_tensorUnit_V, ↓reduceIte,
|
||||
reduceCtorEq, zero_tmul, map_zero, smul_zero, add_zero, Sum.inr.injEq, one_ne_zero,
|
||||
Fin.reduceEq, zero_add, zero_ne_one]
|
||||
|
|
|
@ -189,7 +189,7 @@ def asConsTensor : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗
|
|||
/-- The map `𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ complexContr ⊗ leftHanded ⊗ rightHanded` corresponding
|
||||
to Pauli matrices, when evaluated on `1` corresponds to the tensor `PauliMatrix.asTensor`. -/
|
||||
lemma asConsTensor_apply_one : asConsTensor.hom (1 : ℂ) = asTensor := by
|
||||
change asConsTensor.hom.hom.toFun (1 : ℂ) = asTensor
|
||||
change (1 : ℂ) • asTensor = asTensor
|
||||
simp only [asConsTensor, one_smul]
|
||||
|
||||
end
|
||||
|
|
|
@ -58,12 +58,14 @@ def _root_.LorentzGroup.toNormOne (Λ : LorentzGroup d) : NormOne d :=
|
|||
lemma _root_.LorentzGroup.toNormOne_inl (Λ : LorentzGroup d) :
|
||||
(LorentzGroup.toNormOne Λ).val.val (Sum.inl 0) = Λ.1 (Sum.inl 0) (Sum.inl 0) := by
|
||||
simp only [Fin.isValue, LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte,
|
||||
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
|
||||
Finsupp.coe_mk, Matrix.mulVec_single, MulOpposite.op_one, Pi.smul_apply, Matrix.transpose_apply,
|
||||
one_smul]
|
||||
|
||||
lemma _root_.LorentzGroup.toNormOne_inr (Λ : LorentzGroup d) (i : Fin d) :
|
||||
(LorentzGroup.toNormOne Λ).val.val (Sum.inr i) = Λ.1 (Sum.inr i) (Sum.inl 0) := by
|
||||
simp only [LorentzGroup.toNormOne_coe_val, Finsupp.single, one_ne_zero, ↓reduceIte, Fin.isValue,
|
||||
Finsupp.coe_mk, Matrix.mulVec_single, mul_one]
|
||||
Finsupp.coe_mk, Matrix.mulVec_single, MulOpposite.op_one, Pi.smul_apply, Matrix.transpose_apply,
|
||||
one_smul]
|
||||
|
||||
lemma _root_.LorentzGroup.inl_inl_mul (Λ Λ' : LorentzGroup d) : (Λ * Λ').1 (Sum.inl 0) (Sum.inl 0) =
|
||||
⟪(LorentzGroup.toNormOne (LorentzGroup.transpose Λ)).1,
|
||||
|
|
|
@ -129,7 +129,7 @@ lemma toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (detM
|
|||
_ = 1 := detM
|
||||
have detD_one : D.det = 1 :=
|
||||
let z := x * conj y
|
||||
have k₀ : (M * E₂ * Mᴴ) 0 1 = z := by rw [he', he]; simp [E₂]
|
||||
have k₀ : (M * E₂ * Mᴴ) 0 1 = z := by rw [he', he]; simp [E₂, z]
|
||||
have k₁ : (M * E₃ * Mᴴ) 0 1 = ⟨-z.im, z.re⟩ :=
|
||||
calc
|
||||
_ = x * I * conj y := by rw [he', he]; simp [E₃]
|
||||
|
@ -140,7 +140,7 @@ lemma toSelfAdjointMap_det_one' {M : ℂ²ˣ²} (hM : M.IsUpperTriangular) (detM
|
|||
| 0, 1 => congrArg Complex.re k₁ | 1, 1 => congrArg Complex.im k₁
|
||||
calc D.det
|
||||
_ = normSq z := by simp [hD, z.normSq_apply]
|
||||
_ = normSq x * normSq y := by simp [x.normSq_mul]
|
||||
_ = normSq x * normSq y := by simp [x.normSq_mul, z]
|
||||
_ = 1 := detA_one
|
||||
|
||||
letI : Invertible D.det := detD_one ▸ invertibleOne
|
||||
|
|
|
@ -52,7 +52,7 @@ lemma leftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [leftBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply]
|
||||
change (M.1 *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, mul_one]
|
||||
simp only [mulVec_single, MulOpposite.op_one, Pi.smul_apply, transpose_apply, one_smul]
|
||||
|
||||
@[simp]
|
||||
lemma leftBasis_toFin2ℂ (i : Fin 2) : (leftBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
|
@ -97,7 +97,7 @@ lemma altLeftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [altLeftBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change ((M.1⁻¹)ᵀ *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
simp only [mulVec_single, MulOpposite.op_one, transpose_transpose, Pi.smul_apply, one_smul]
|
||||
|
||||
/-- The vector space ℂ^2 carrying the conjugate representation of SL(2,C).
|
||||
In index notation corresponds to a Weyl fermion with indices ψ^{dot a}. -/
|
||||
|
@ -134,7 +134,7 @@ lemma rightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [rightBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change (M.1.map star *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
simp [mulVec_single, transpose_apply, mul_one]
|
||||
|
||||
/-- The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)^†.
|
||||
|
@ -176,7 +176,7 @@ lemma altRightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
|||
rw [LinearMap.toMatrix_apply]
|
||||
simp only [altRightBasis, Basis.coe_ofEquivFun, Basis.ofEquivFun_repr_apply, transpose_apply]
|
||||
change ((M.1⁻¹).conjTranspose *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, transpose_apply, mul_one]
|
||||
simp [mulVec_single, transpose_apply, mul_one]
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -108,7 +108,7 @@ def leftMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ leftHanded where
|
|||
not_false_eq_true, mul_nonsing_inv, transpose_one, mul_one]
|
||||
|
||||
lemma leftMetric_apply_one : leftMetric.hom (1 : ℂ) = leftMetricVal := by
|
||||
change leftMetric.hom.hom.toFun (1 : ℂ) = leftMetricVal
|
||||
change (1 : ℂ) • leftMetricVal = leftMetricVal
|
||||
simp only [leftMetric, one_smul]
|
||||
|
||||
/-- The metric `εₐₐ` as an element of `(altLeftHanded ⊗ altLeftHanded).V`. -/
|
||||
|
@ -155,7 +155,7 @@ def altLeftMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ altLeftHande
|
|||
not_false_eq_true, mul_nonsing_inv, mul_one]
|
||||
|
||||
lemma altLeftMetric_apply_one : altLeftMetric.hom (1 : ℂ) = altLeftMetricVal := by
|
||||
change altLeftMetric.hom.hom.toFun (1 : ℂ) = altLeftMetricVal
|
||||
change (1 : ℂ) • altLeftMetricVal = altLeftMetricVal
|
||||
simp only [altLeftMetric, one_smul]
|
||||
|
||||
/-- The metric `ε^{dot a}^{dot a}` as an element of `(rightHanded ⊗ rightHanded).V`. -/
|
||||
|
@ -209,7 +209,7 @@ def rightMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ rightHanded ⊗ rightHanded wher
|
|||
rfl
|
||||
|
||||
lemma rightMetric_apply_one : rightMetric.hom (1 : ℂ) = rightMetricVal := by
|
||||
change rightMetric.hom.hom.toFun (1 : ℂ) = rightMetricVal
|
||||
change (1 : ℂ) • rightMetricVal = rightMetricVal
|
||||
simp only [rightMetric, one_smul]
|
||||
|
||||
/-- The metric `ε_{dot a}_{dot a}` as an element of `(altRightHanded ⊗ altRightHanded).V`. -/
|
||||
|
@ -267,7 +267,7 @@ def altRightMetric : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altRightHanded ⊗ altRightHa
|
|||
rfl
|
||||
|
||||
lemma altRightMetric_apply_one : altRightMetric.hom (1 : ℂ) = altRightMetricVal := by
|
||||
change altRightMetric.hom.hom.toFun (1 : ℂ) = altRightMetricVal
|
||||
change (1 : ℂ) • altRightMetricVal = altRightMetricVal
|
||||
simp only [altRightMetric, one_smul]
|
||||
|
||||
/-!
|
||||
|
|
|
@ -61,7 +61,7 @@ def leftAltLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ leftHanded ⊗ altLeftHanded
|
|||
simp
|
||||
|
||||
lemma leftAltLeftUnit_apply_one : leftAltLeftUnit.hom (1 : ℂ) = leftAltLeftUnitVal := by
|
||||
change leftAltLeftUnit.hom.hom.toFun (1 : ℂ) = leftAltLeftUnitVal
|
||||
change (1 : ℂ) • leftAltLeftUnitVal = leftAltLeftUnitVal
|
||||
simp only [leftAltLeftUnit, one_smul]
|
||||
|
||||
/-- The alt-left-left unit `δₐᵃ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
|
||||
|
@ -106,7 +106,7 @@ def altLeftLeftUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altLeftHanded ⊗ leftHanded
|
|||
|
||||
/-- Applying the morphism `altLeftLeftUnit` to `1` returns `altLeftLeftUnitVal`. -/
|
||||
lemma altLeftLeftUnit_apply_one : altLeftLeftUnit.hom (1 : ℂ) = altLeftLeftUnitVal := by
|
||||
change altLeftLeftUnit.hom.hom.toFun (1 : ℂ) = altLeftLeftUnitVal
|
||||
change (1 : ℂ) • altLeftLeftUnitVal = altLeftLeftUnitVal
|
||||
simp only [altLeftLeftUnit, one_smul]
|
||||
|
||||
/-- The right-alt-right unit `δ^{dot a}_{dot a}` as an element of
|
||||
|
@ -157,7 +157,7 @@ def rightAltRightUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ rightHanded ⊗ altRightHa
|
|||
simp
|
||||
|
||||
lemma rightAltRightUnit_apply_one : rightAltRightUnit.hom (1 : ℂ) = rightAltRightUnitVal := by
|
||||
change rightAltRightUnit.hom.hom.toFun (1 : ℂ) = rightAltRightUnitVal
|
||||
change (1 : ℂ) • rightAltRightUnitVal = rightAltRightUnitVal
|
||||
simp only [rightAltRightUnit, one_smul]
|
||||
|
||||
/-- The alt-right-right unit `δ_{dot a}^{dot a}` as an element of
|
||||
|
@ -206,7 +206,7 @@ def altRightRightUnit : 𝟙_ (Rep ℂ SL(2,ℂ)) ⟶ altRightHanded ⊗ rightHa
|
|||
simp
|
||||
|
||||
lemma altRightRightUnit_apply_one : altRightRightUnit.hom (1 : ℂ) = altRightRightUnitVal := by
|
||||
change altRightRightUnit.hom.hom.toFun (1 : ℂ) = altRightRightUnitVal
|
||||
change (1 : ℂ) • altRightRightUnitVal = altRightRightUnitVal
|
||||
simp only [altRightRightUnit, one_smul]
|
||||
|
||||
/-!
|
||||
|
@ -236,7 +236,7 @@ lemma contr_altLeftLeftUnit (x : leftHanded) :
|
|||
erw [h1, h1, h1, h1]
|
||||
repeat rw [leftAltContraction_basis]
|
||||
simp only [Fin.isValue, leftUnitor, ModuleCat.MonoidalCategory.leftUnitor, ModuleCat.of_coe,
|
||||
CategoryTheory.Iso.trans_hom, LinearEquiv.toModuleIso_hom_hom, ModuleCat.ofSelfIso_hom,
|
||||
CategoryTheory.Iso.trans_hom, ModuleCat.ofSelfIso_hom,
|
||||
CategoryTheory.Category.comp_id, Action.instMonoidalCategory_tensorUnit_V, Fin.val_zero,
|
||||
↓reduceIte, Fin.val_one, one_ne_zero, zero_tmul, map_zero, smul_zero, add_zero, zero_ne_one,
|
||||
zero_add]
|
||||
|
|
|
@ -38,7 +38,8 @@ lemma takeWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
|
|||
simpa using h ⟨0, by simp⟩ ⟨1, by simp⟩ (by simp [Fin.lt_def]) (by simpa using hPb)
|
||||
simp [hPb, hPa]
|
||||
· simp only [hPb, decide_false, List.nil_eq]
|
||||
simp_all only [List.length_cons, List.get_eq_getElem]
|
||||
simp_all only [List.length_cons, List.get_eq_getElem, List.tail_cons, decide_false,
|
||||
Bool.false_eq_true, not_false_eq_true, List.takeWhile_cons_of_neg, List.nil_eq]
|
||||
split
|
||||
· rfl
|
||||
· rfl
|
||||
|
@ -66,22 +67,24 @@ lemma dropWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
|
|||
List.length_singleton, zero_lt_one, Fin.zero_eta, Fin.isValue, List.get_eq_getElem,
|
||||
Fin.val_eq_zero, List.getElem_cons_zero, decide_eq_true_eq, forall_const, zero_le,
|
||||
Nat.sub_eq_zero_of_le, List.eraseIdx_zero, ite_not, List.nil_eq]
|
||||
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero,
|
||||
List.getElem_cons_zero, implies_true, decide_true, decide_false, List.tail_cons, ite_self]
|
||||
simp_all only [List.length_cons, List.length_nil, List.get_eq_getElem, Fin.val_eq_zero,
|
||||
List.getElem_cons_zero, implies_true, List.tail_cons, List.dropWhile_nil, decide_true,
|
||||
decide_false, ite_self]
|
||||
| a :: [], Nat.succ n, h => by
|
||||
simp only [List.dropWhile, List.eraseIdx_nil, List.takeWhile, Nat.succ_eq_add_one]
|
||||
rw [List.eraseIdx_of_length_le]
|
||||
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero,
|
||||
List.getElem_cons_zero, implies_true, ite_self]
|
||||
simp_all only [List.length_singleton, List.get_eq_getElem, Fin.val_eq_zero,
|
||||
List.getElem_cons_zero, implies_true]
|
||||
split
|
||||
next x heq =>
|
||||
simp_all only [decide_eq_true_eq, List.length_nil, List.length_singleton,
|
||||
add_tsub_cancel_right, zero_le]
|
||||
simp_all only [decide_eq_true_eq, decide_true, List.dropWhile_cons_of_pos, List.dropWhile_nil,
|
||||
List.length_singleton, le_add_iff_nonneg_left, zero_le, ↓reduceIte,
|
||||
add_tsub_cancel_right, List.eraseIdx_nil]
|
||||
next x heq =>
|
||||
simp_all only [decide_eq_false_iff_not, List.length_singleton, List.length_nil, tsub_zero,
|
||||
le_add_iff_nonneg_left, zero_le]
|
||||
simp_all only [decide_eq_false_iff_not, decide_false, Bool.false_eq_true, not_false_eq_true,
|
||||
List.dropWhile_cons_of_neg, List.length_nil, le_add_iff_nonneg_left,
|
||||
zero_le, ↓reduceIte, tsub_zero, List.eraseIdx_cons_succ, List.eraseIdx_nil]
|
||||
exact Nat.le_add_left [a].length n
|
||||
| a :: b :: l, 0, h => by
|
||||
simp only [List.dropWhile, List.takeWhile, nonpos_iff_eq_zero, List.length_eq_zero, zero_le,
|
||||
Nat.sub_eq_zero_of_le, List.eraseIdx_zero]
|
||||
|
@ -92,14 +95,14 @@ lemma dropWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] :
|
|||
· simp only [hPb, decide_false, nonpos_iff_eq_zero, List.length_eq_zero]
|
||||
simp_all only [List.length_cons, List.get_eq_getElem]
|
||||
simp_rw [decide_false]
|
||||
simp_all only [List.tail_cons, decide_false, Bool.false_eq_true, not_false_eq_true,
|
||||
List.dropWhile_cons_of_neg, nonpos_iff_eq_zero, List.length_eq_zero]
|
||||
split
|
||||
next h_1 =>
|
||||
simp [nonpos_iff_eq_zero]
|
||||
next h_1 =>
|
||||
simp_all only [nonpos_iff_eq_zero, List.length_eq_zero]
|
||||
split
|
||||
next x heq => rfl
|
||||
next x heq => simp_all only [not_true_eq_false]
|
||||
next x heq =>
|
||||
simp_all only [decide_eq_true_eq, List.length_singleton, nonpos_iff_eq_zero, one_ne_zero,
|
||||
↓reduceIte]
|
||||
next x heq => simp_all only [decide_eq_false_iff_not, List.length_nil,
|
||||
le_refl, ↓reduceIte, List.tail_cons]
|
||||
| a :: b :: l, Nat.succ n, h => by
|
||||
simp only [Nat.succ_eq_add_one, List.eraseIdx_cons_succ]
|
||||
by_cases hPb : P b
|
||||
|
@ -186,10 +189,11 @@ lemma orderedInsertPos_sigma {I : Type} {f : I → Type}
|
|||
split
|
||||
next x heq =>
|
||||
simp_all only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not,
|
||||
List.length_cons, decide_false, Bool.not_false]
|
||||
List.length_cons, List.map_cons, decide_false, Bool.not_false, List.takeWhile_cons_of_pos]
|
||||
next x heq =>
|
||||
simp_all only [Bool.not_eq_eq_eq_not, Bool.not_false, decide_eq_true_eq, List.length_nil,
|
||||
decide_true, Bool.not_true]
|
||||
List.map_cons, decide_true, Bool.not_true, Bool.false_eq_true, not_false_eq_true,
|
||||
List.takeWhile_cons_of_neg]
|
||||
|
||||
lemma orderedInsert_get_lt {I : Type} (le1 : I → I → Prop) [DecidableRel le1]
|
||||
(r : List I) (r0 : I) (i : ℕ)
|
||||
|
|
|
@ -133,7 +133,7 @@ lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [D
|
|||
· simp [h]
|
||||
conv_lhs => simp [h]
|
||||
rw [insertionSort_orderedInsert_append r a l1 l2]
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
simp only [List.cons_append, List.insertionSort]
|
||||
rw [orderedInsert_commute r a b h]
|
||||
|
||||
lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
|
@ -144,7 +144,7 @@ lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [D
|
|||
| a :: l1, l2 => by
|
||||
conv_lhs => simp
|
||||
rw [insertionSort_orderedInsert_append]
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
simp only [List.cons_append, List.insertionSort]
|
||||
rw [insertionSort_insertionSort_append r l1 l2]
|
||||
|
||||
lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r]
|
||||
|
@ -155,7 +155,7 @@ lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → P
|
|||
simp only [List.nil_append]
|
||||
exact insertionSort_insertionSort_append r l2 l3
|
||||
| a :: l1, l2, l3 => by
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
simp only [List.cons_append, List.insertionSort]
|
||||
rw [insertionSort_append_insertionSort_append r l1 l2 l3]
|
||||
|
||||
@[simp]
|
||||
|
@ -273,16 +273,17 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidable
|
|||
refine hbc ?_
|
||||
exact IsTrans.trans _ _ _ hrba hac
|
||||
have ha1 : b1.1 ≤ a2.1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
simp only [orderedInsertPos, decide_not, b1, b2]
|
||||
rw [ht]
|
||||
apply List.Sublist.length_le
|
||||
exact List.takeWhile_sublist fun c => !decide (r b c)
|
||||
simp only [decide_not, b1]
|
||||
exact List.takeWhile_sublist _
|
||||
have ha2 : a1.1 = a2.1 + 1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
simp only [orderedInsertPos, decide_not, a1, a2]
|
||||
rw [takeWhile_orderedInsert]
|
||||
exact hr
|
||||
have hb : b1.1 = b2.1 := by
|
||||
simp only [orderedInsertPos, decide_not]
|
||||
simp only [orderedInsertPos, decide_not, b1, b2]
|
||||
rw [takeWhile_orderedInsert']
|
||||
exact hr
|
||||
let n := ((insertionSortEquiv r l) ⟨n, by simpa using hn⟩)
|
||||
|
@ -338,7 +339,7 @@ lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Pro
|
|||
(insertionSortEquiv r (List.orderedInsert r a l1 ++ a2 :: l2) ⟨l1.length + 1, by
|
||||
simp⟩)
|
||||
= (finCongr (by
|
||||
simp only [List.insertionSort, List.append_eq, orderedInsert_length, List.length_cons,
|
||||
simp only [List.cons_append, List.insertionSort, orderedInsert_length, List.length_cons,
|
||||
List.length_insertionSort, List.length_append]
|
||||
omega))
|
||||
((insertionSortEquiv r (a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩)
|
||||
|
@ -408,8 +409,8 @@ lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [Decidabl
|
|||
(hl : l.Sorted r) →
|
||||
List.filter p (List.orderedInsert r a l) = List.orderedInsert r a (List.filter p l)
|
||||
| [], hl => by
|
||||
simp only [List.orderedInsert, List.filter_eq_self, List.mem_singleton, decide_eq_true_eq,
|
||||
forall_eq]
|
||||
simp only [List.orderedInsert, List.filter_nil, List.orderedInsert_nil, List.filter_eq_self,
|
||||
List.mem_singleton, decide_eq_true_eq, forall_eq]
|
||||
exact h
|
||||
| b :: l, hl => by
|
||||
simp only [List.orderedInsert]
|
||||
|
|
|
@ -34,8 +34,8 @@ namespace Equiv
|
|||
def sumEquivSigmalCond : Fin m ⊕ Fin n ≃ Σ b, cond b (Fin m) (Fin n) :=
|
||||
calc Fin m ⊕ Fin n
|
||||
_ ≃ Fin n ⊕ Fin m := sumComm ..
|
||||
_ ≃ Σ b, Bool.casesOn b (Fin n) (Fin m) := sumEquivSigmaBool ..
|
||||
_ ≃ Σ b, cond b (Fin m) (Fin n) := sigmaCongrRight fun | true | false => Equiv.refl _
|
||||
_ ≃ Σ b, bif b then (Fin m) else (Fin n) := sumEquivSigmaBool ..
|
||||
_ ≃ Σ b, cond b (Fin m) (Fin n) := sigmaCongrRight (fun | true | false => Equiv.refl _)
|
||||
|
||||
/-- The composition of `finSumFinEquiv` and `Equiv.sumEquivSigmalCond` used by
|
||||
`LinearMap.SchurTriangulationAux.of`. -/
|
||||
|
@ -183,10 +183,10 @@ protected noncomputable def SchurTriangulationAux.of
|
|||
have hB : ∀ s, bE s = B s.1 s.2
|
||||
| ⟨true, i⟩ => show bE ⟨true, i⟩ = bV i from
|
||||
show (int.collectedBasis fun b => (B b).toBasis).toOrthonormalBasis _ ⟨true, i⟩ = bV i
|
||||
by simp
|
||||
by simp [B]
|
||||
| ⟨false, j⟩ => show bE ⟨false, j⟩ = bW j from
|
||||
show (int.collectedBasis fun b => (B b).toBasis).toOrthonormalBasis _ ⟨false, j⟩ = bW j
|
||||
by simp
|
||||
by simp [B]
|
||||
have hf {bi i' bj j'} (hi : e i = ⟨bi, i'⟩) (hj : e j = ⟨bj, j'⟩) :=
|
||||
calc toMatrixOrthonormal basis f i j
|
||||
_ = toMatrixOrthonormal bE f (e i) (e j) := by
|
||||
|
|
|
@ -3,7 +3,7 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import Mathlib.Algebra.BigOperators.Group.Finset
|
||||
import Mathlib.Algebra.BigOperators.Group.Finset.Basic
|
||||
/-!
|
||||
|
||||
# Creation and annihilation parts of fields
|
||||
|
@ -66,7 +66,7 @@ lemma not_normalOrder_annihilate_iff_false (a : CreateAnnihilate) :
|
|||
|
||||
lemma sum_eq {M : Type} [AddCommMonoid M] (f : CreateAnnihilate → M) :
|
||||
∑ i, f i = f create + f annihilate := by
|
||||
change ∑ i in {create, annihilate}, f i = f create + f annihilate
|
||||
change ∑ i ∈ {create, annihilate}, f i = f create + f annihilate
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
|
|
|
@ -84,7 +84,7 @@ lemma koszulSignInsert_append_annihilate (φ' φ : 𝓕.CrAnFieldOp)
|
|||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' (φs ++ [φ]) =
|
||||
Wick.koszulSignInsert 𝓕.crAnStatistics normalOrderRel φ' φs
|
||||
| [] => by
|
||||
simp only [Wick.koszulSignInsert, normalOrderRel, hφ, ite_eq_left_iff,
|
||||
simp only [List.nil_append, Wick.koszulSignInsert, normalOrderRel, hφ, ite_eq_left_iff,
|
||||
CreateAnnihilate.not_normalOrder_annihilate_iff_false, ite_eq_right_iff, and_imp,
|
||||
IsEmpty.forall_iff]
|
||||
| φ'' :: φs => by
|
||||
|
@ -281,6 +281,7 @@ lemma normalOrderList_append_annihilate (φ : 𝓕.CrAnFieldOp)
|
|||
simp only [normalOrderList, List.insertionSort, List.append_eq]
|
||||
have hi := normalOrderList_append_annihilate φ hφ φs
|
||||
dsimp only [normalOrderList] at hi
|
||||
simp only [List.cons_append, List.insertionSort]
|
||||
rw [hi, orderedInsert_append_annihilate φ' φ hφ]
|
||||
|
||||
lemma normalOrder_swap_create_annihilate_fst (φc φa : 𝓕.CrAnFieldOp)
|
||||
|
|
|
@ -308,15 +308,15 @@ lemma orderedInsert_in_swap_eq_time {φ ψ φ': 𝓕.CrAnFieldOp} (h1 : crAnTime
|
|||
have h1 (b : 𝓕.CrAnFieldOp) : (crAnTimeOrderRel b φ) ↔ (crAnTimeOrderRel b ψ) :=
|
||||
Iff.intro (fun h => IsTrans.trans _ _ _ h h1) (fun h => IsTrans.trans _ _ _ h h2)
|
||||
by_cases h : crAnTimeOrderRel φ' φ
|
||||
· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 φ']
|
||||
· simp only [List.nil_append, List.orderedInsert, h, ↓reduceIte, ← h1 φ']
|
||||
use [φ'], φs'
|
||||
simp
|
||||
· simp only [List.orderedInsert, h, ↓reduceIte, ← h1 φ']
|
||||
· simp only [List.nil_append, List.orderedInsert, h, ↓reduceIte, ← h1 φ']
|
||||
use [], List.orderedInsert crAnTimeOrderRel φ' φs'
|
||||
simp
|
||||
| φ'' :: φs, φs' => by
|
||||
obtain ⟨l1, l2, hl⟩ := orderedInsert_in_swap_eq_time (φ' := φ') h1 h2 φs φs'
|
||||
simp only [List.orderedInsert, List.append_eq]
|
||||
simp only [List.cons_append, List.orderedInsert]
|
||||
rw [hl.1, hl.2]
|
||||
by_cases h : crAnTimeOrderRel φ' φ''
|
||||
· simp only [h, ↓reduceIte]
|
||||
|
@ -334,7 +334,7 @@ lemma crAnTimeOrderList_swap_eq_time {φ ψ : 𝓕.CrAnFieldOp}
|
|||
crAnTimeOrderList (φs ++ ψ :: φ :: φs') = l1 ++ ψ :: φ :: l2
|
||||
| [], φs' => by
|
||||
simp only [crAnTimeOrderList]
|
||||
simp only [List.insertionSort]
|
||||
simp only [List.nil_append, List.insertionSort]
|
||||
use List.takeWhile (fun b => ¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel φs'),
|
||||
List.dropWhile (fun b => ¬ crAnTimeOrderRel ψ b) (List.insertionSort crAnTimeOrderRel φs')
|
||||
apply And.intro
|
||||
|
@ -345,7 +345,7 @@ lemma crAnTimeOrderList_swap_eq_time {φ ψ : 𝓕.CrAnFieldOp}
|
|||
simpa using orderedInsert_swap_eq_time h2 h1 _
|
||||
| φ'' :: φs, φs' => by
|
||||
rw [crAnTimeOrderList, crAnTimeOrderList]
|
||||
simp only [List.insertionSort, List.append_eq]
|
||||
simp only [List.cons_append, List.insertionSort]
|
||||
obtain ⟨l1, l2, hl⟩ := crAnTimeOrderList_swap_eq_time h1 h2 φs φs'
|
||||
simp only [crAnTimeOrderList] at hl
|
||||
rw [hl.1, hl.2]
|
||||
|
|
|
@ -184,7 +184,7 @@ lemma ofList_append (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
|||
(if a = (if b = c then bosonic else fermionic) then bosonic else fermionic) =
|
||||
if (if a = b then bosonic else fermionic) = c then bosonic else fermionic := by
|
||||
fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl
|
||||
simp only [ofList, List.append_eq, Fin.isValue, ih, hab]
|
||||
simp only [List.cons_append, ofList, ih, hab]
|
||||
|
||||
lemma ofList_append_eq_mul (s : 𝓕 → FieldStatistic) (φs φs' : List 𝓕) :
|
||||
ofList s (φs ++ φs') = ofList s φs * ofList s φs' := by
|
||||
|
|
|
@ -42,6 +42,7 @@ lemma koszulSign_mul_self (l : List 𝓕) : koszulSign q le l * koszulSign q le
|
|||
|
||||
@[simp]
|
||||
lemma koszulSign_freeMonoid_of (φ : 𝓕) : koszulSign q le (FreeMonoid.of φ) = 1 := by
|
||||
change koszulSign q le [φ] = 1
|
||||
simp only [koszulSign, mul_one]
|
||||
rfl
|
||||
|
||||
|
@ -235,6 +236,9 @@ lemma koszulSign_eraseIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : List 𝓕)
|
|||
𝓢(q φs[↑n], ofList q (List.take (↑((insertionSortEquiv le φs) n)) (List.insertionSort le φs))))
|
||||
swap
|
||||
· simp only [Fin.getElem_fin]
|
||||
rw [Equiv.trans_apply, Equiv.trans_apply]
|
||||
simp only [instCommGroup.eq_1, mul_one, Fin.castOrderIso,
|
||||
Equiv.coe_fn_mk, Fin.cast_mk, Fin.eta, Fin.coe_cast]
|
||||
ring
|
||||
conv_rhs =>
|
||||
rhs
|
||||
|
@ -274,7 +278,7 @@ lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : (
|
|||
simp only [List.nil_append]
|
||||
exact koszulSign_swap_eq_rel_cons q le h1 h2 φs'
|
||||
| φ'' :: φs, φs' => by
|
||||
simp only [Wick.koszulSign, List.append_eq]
|
||||
simp only [List.cons_append, koszulSign]
|
||||
rw [koszulSign_swap_eq_rel h1 h2]
|
||||
congr 1
|
||||
apply Wick.koszulSignInsert_eq_perm
|
||||
|
@ -301,7 +305,7 @@ lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le]
|
|||
simp only [List.nil_append]
|
||||
exact koszulSign_eq_rel_eq_stat_append q le h1 h2 hq φs
|
||||
| φ'' :: φs', φs => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
simp only [List.cons_append, koszulSign]
|
||||
rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs]
|
||||
simp only [mul_eq_mul_right_iff]
|
||||
left
|
||||
|
@ -381,7 +385,7 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le]
|
|||
simp only [List.nil_append]
|
||||
exact koszulSign_of_append_eq_insertionSort_left q le φs φs'
|
||||
| φ'' :: φs'', φs, φs' => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
simp only [List.cons_append, koszulSign]
|
||||
rw [koszulSign_of_append_eq_insertionSort φs'' φs φs', ← mul_assoc]
|
||||
congr 2
|
||||
apply koszulSignInsert_eq_perm
|
||||
|
@ -430,7 +434,7 @@ lemma koszulSign_perm_eq [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : L
|
|||
simp only [List.nil_append]
|
||||
exact koszulSign_perm_eq_append q le φ φs φs' φs2 hp h
|
||||
| φ1 :: φs1, φs, φs', φs2, h, hp => by
|
||||
simp only [koszulSign, List.append_eq]
|
||||
simp only [List.cons_append, koszulSign]
|
||||
have ih := koszulSign_perm_eq φ φs1 φs φs' φs2 h hp
|
||||
rw [ih]
|
||||
congr 1
|
||||
|
|
|
@ -69,7 +69,7 @@ lemma koszulSignInsert_ge_forall_append (φs : List 𝓕) (φ' φ : 𝓕) (hi :
|
|||
induction φs with
|
||||
| nil => simp [koszulSignInsert, hi]
|
||||
| cons φ'' φs ih =>
|
||||
simp only [koszulSignInsert, Fin.isValue, List.append_eq]
|
||||
simp only [koszulSignInsert, List.cons_append]
|
||||
by_cases hr : le φ' φ''
|
||||
· rw [if_pos hr, if_pos hr, ih]
|
||||
· rw [if_neg hr, if_neg hr, ih]
|
||||
|
|
|
@ -181,6 +181,7 @@ lemma consAddContract_surjective_on_zero_contract (i : Fin n.succ)
|
|||
rcases h with h | h
|
||||
· obtain ⟨b, hb, rfl⟩ := h
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
simp only [succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and, c'] at hb
|
||||
exact hb
|
||||
· subst h
|
||||
rw [← h2]
|
||||
|
@ -203,9 +204,9 @@ lemma consAddContract_surjective_on_zero_contract (i : Fin n.succ)
|
|||
obtain ⟨y, rfl⟩ := (Fin.exists_succAbove_eq (x := y) (y := i)) (by omega)
|
||||
use {x, y}
|
||||
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Fin.val_succEmb,
|
||||
h, true_and]
|
||||
h, true_and, c']
|
||||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply]
|
||||
simp
|
||||
simpa using h
|
||||
|
||||
lemma consAddContract_bijection (i : Fin n.succ) :
|
||||
Function.Bijective (fun c => (⟨(consAddContract i c), by simp⟩ :
|
||||
|
|
|
@ -220,7 +220,7 @@ lemma insertAndContract_none_prod_contractions (φ : 𝓕.FieldOp) (φs : List
|
|||
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
|
||||
((φsΛ.insertAndContractNat i none).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
erw [← e1.prod_comp]
|
||||
rw [← e1.prod_comp]
|
||||
rfl
|
||||
|
||||
lemma insertAndContract_some_prod_contractions (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
|
@ -233,7 +233,7 @@ lemma insertAndContract_some_prod_contractions (φ : 𝓕.FieldOp) (φs : List
|
|||
((φsΛ.insertAndContractNat i (some j)).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
|
||||
erw [← e2.prod_comp]
|
||||
let e1 := Equiv.ofBijective (φsΛ.insertLiftSome i j) (insertLiftSome_bijective i j)
|
||||
erw [← e1.prod_comp]
|
||||
rw [← e1.prod_comp]
|
||||
rw [Fintype.prod_sum_type]
|
||||
simp only [Finset.univ_unique, PUnit.default_eq_unit, Nat.succ_eq_add_one, Finset.prod_singleton,
|
||||
Finset.univ_eq_attach]
|
||||
|
|
|
@ -276,7 +276,7 @@ def uncontractedIndexEquiv (c : WickContraction n) :
|
|||
Fin (c.uncontractedList).length ≃ c.uncontracted where
|
||||
toFun i := ⟨c.uncontractedList.get i, c.uncontractedList_get_mem_uncontracted i⟩
|
||||
invFun i := ⟨List.indexOf i.1 c.uncontractedList,
|
||||
List.indexOf_lt_length.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
|
||||
List.indexOf_lt_length_iff.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
|
||||
left_inv i := by
|
||||
ext
|
||||
exact List.get_indexOf (uncontractedList_nodup c) _
|
||||
|
|
|
@ -3,7 +3,7 @@ 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.Geometry.Manifold.SmoothManifoldWithCorners
|
||||
import Mathlib.Geometry.Manifold.IsManifold.Basic
|
||||
import Mathlib.Analysis.InnerProductSpace.PiL2
|
||||
import HepLean.Meta.TODO.Basic
|
||||
/-!
|
||||
|
|
|
@ -84,8 +84,8 @@ TODO "Make `HiggsBundle` an associated bundle."
|
|||
abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
|
||||
|
||||
/-- The instance of a smooth vector bundle with total space `HiggsBundle` and fiber `HiggsVec`. -/
|
||||
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
||||
Bundle.Trivial.smoothVectorBundle HiggsVec
|
||||
instance : ContMDiffVectorBundle ⊤ HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
||||
Bundle.Trivial.contMDiffVectorBundle HiggsVec
|
||||
|
||||
/-- A Higgs field is a smooth section of the Higgs bundle. -/
|
||||
abbrev HiggsField : Type := ContMDiffSection SpaceTime.asSmoothManifold HiggsVec ⊤ HiggsBundle
|
||||
|
|
|
@ -210,10 +210,12 @@ theorem rotate_fst_real_snd_zero (φ : HiggsVec) :
|
|||
coe_smul, mulVec_empty, add_zero, zero_add, one_smul]
|
||||
funext i
|
||||
fin_cases i
|
||||
· simp only [Fin.zero_eta, Fin.isValue, Pi.smul_apply, Function.comp_apply, cons_val_zero,
|
||||
tail_cons, head_cons, real_smul, mul_one]
|
||||
· simp only [Fin.mk_one, Fin.isValue, Pi.smul_apply, Function.comp_apply, cons_val_one, head_cons,
|
||||
tail_cons, smul_zero]
|
||||
· simp only [one_pow, mulVec_cons, Nat.succ_eq_add_one, Nat.reduceAdd, zero_smul, coe_smul,
|
||||
mulVec_empty, add_zero, zero_add, Fin.zero_eta, Fin.isValue, Pi.smul_apply, Function.comp_apply,
|
||||
cons_val_zero, tail_cons, head_cons, real_smul, mul_one, one_smul, P]
|
||||
· simp only [one_pow, mulVec_cons, Nat.succ_eq_add_one, Nat.reduceAdd, zero_smul, coe_smul,
|
||||
mulVec_empty, add_zero, zero_add, Fin.mk_one, Fin.isValue, Pi.smul_apply, Function.comp_apply,
|
||||
cons_val_one, head_cons, tail_cons, smul_zero, P]
|
||||
|
||||
/-- There exists a `g` in `GaugeGroupI` such that `rep g φ = φ'` iff `‖φ‖ = ‖φ'‖`. -/
|
||||
informal_lemma guage_orbit where
|
||||
|
|
|
@ -79,8 +79,7 @@ lemma pairIsoSep_tmul {c1 c2 : C} (x : F.obj (Discrete.mk c1)) (y : F.obj (Discr
|
|||
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj, mk_hom, mk_left, Functor.id_obj]
|
||||
simp only [LinearMap.coe_comp, Function.comp_apply, ModuleCat.MonoidalCategory.tensorHom_tmul,
|
||||
LinearEquiv.toModuleIso_inv_hom, LinearEquiv.coe_coe, forgetLiftAppV_symm_apply, mk_left,
|
||||
Functor.id_obj]
|
||||
LinearEquiv.coe_coe, forgetLiftAppV_symm_apply, mk_left, Functor.id_obj]
|
||||
change ((lift.obj F).map fin2Iso.inv).hom
|
||||
(((lift.obj F).map ((mkIso _).hom ⊗ (mkIso _).hom)).hom
|
||||
((Functor.LaxMonoidal.μ (lift.obj F).toFunctor (mk fun _ => c1) (mk fun _ => c2)).hom
|
||||
|
|
|
@ -177,8 +177,17 @@ def ε : 𝟙_ (Rep k G) ≅ objObj' F (𝟙_ (OverColor C)) :=
|
|||
simp only [objObj'_V_carrier, OverColor.instMonoidalCategoryStruct_tensorUnit_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_hom,
|
||||
Action.instMonoidalCategory_tensorUnit_V, Action.tensorUnit_ρ', Functor.id_obj,
|
||||
Category.id_comp, LinearEquiv.coe_coe]
|
||||
simp [objObj'_ρ_empty F g])
|
||||
Category.id_comp, LinearEquiv.coe_coe,
|
||||
ModuleCat.comp_apply]
|
||||
rw [ModuleCat.hom_comp, ModuleCat.hom_comp]
|
||||
simp only [objObj'_V_carrier, instMonoidalCategoryStruct_tensorUnit_left,
|
||||
instMonoidalCategoryStruct_tensorUnit_hom, LinearEquiv.toModuleIso_hom, ModuleCat.hom_ofHom,
|
||||
LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply]
|
||||
change _ = (objObj' F (𝟙_ (OverColor C))).ρ g ((PiTensorProduct.isEmptyEquiv Empty).symm x)
|
||||
simp only [objObj'_ρ_empty F g, instMonoidalCategoryStruct_tensorUnit_hom,
|
||||
instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj,
|
||||
PiTensorProduct.isEmptyEquiv_symm_apply, map_smul, LinearMap.id_coe, id_eq]
|
||||
rfl)
|
||||
|
||||
/-- An auxiliary equivalence, and trivial, of modules needed to define `μModEquiv`. -/
|
||||
def discreteSumEquiv {X Y : OverColor C} (i : X.left ⊕ Y.left) :
|
||||
|
@ -290,10 +299,13 @@ lemma μ_natural_left {X Y : OverColor C} (f : X ⟶ Y) (Z : OverColor C) :
|
|||
(PiTensorProduct.tprod k fun i => discreteSumEquiv' F i
|
||||
(HepLean.PiTensorProduct.elimPureTensor p q i))
|
||||
rw [objMap'_tprod]
|
||||
have h1 : (((objMap' F f).hom ▷ (objObj' F Z).V) (PiTensorProduct.tprod k p ⊗ₜ[k]
|
||||
PiTensorProduct.tprod k q)) = ((objMap' F f).hom
|
||||
((PiTensorProduct.tprod k) p) ⊗ₜ[k] ((PiTensorProduct.tprod k) q)) := by rfl
|
||||
erw [h1]
|
||||
change ((μ F Y Z).hom.hom.hom' ∘ₗ ((objMap' F f).hom ▷ (objObj' F Z).V).hom') _ = _
|
||||
simp only [objObj'_V_carrier, instMonoidalCategoryStruct_tensorObj_left,
|
||||
instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
|
||||
LinearMap.coe_comp, Function.comp_apply, Functor.id_obj]
|
||||
conv_lhs =>
|
||||
right
|
||||
change ((objMap' F f).hom ((PiTensorProduct.tprod k) p) ⊗ₜ[k] ((PiTensorProduct.tprod k) q))
|
||||
rw [objMap'_tprod]
|
||||
change (μ F Y Z).hom.hom (((PiTensorProduct.tprod k) fun i => (discreteFunctorMapEqIso F _)
|
||||
(p ((OverColor.Hom.toEquiv f).symm i))) ⊗ₜ[k] (PiTensorProduct.tprod k) q) = _
|
||||
|
@ -327,11 +339,12 @@ lemma μ_natural_right {X Y : OverColor C} (X' : OverColor C) (f : X ⟶ Y) :
|
|||
change _ = (objMap' F (X' ◁ f)).hom ((PiTensorProduct.tprod k) fun i =>
|
||||
(discreteSumEquiv' F i) (HepLean.PiTensorProduct.elimPureTensor p q i))
|
||||
rw [objMap'_tprod]
|
||||
have h1 : (((objObj' F X').V ◁ (objMap' F f).hom)
|
||||
((PiTensorProduct.tprod k) p ⊗ₜ[k] (PiTensorProduct.tprod k) q))
|
||||
= ((PiTensorProduct.tprod k) p ⊗ₜ[k] (objMap' F f).hom ((PiTensorProduct.tprod k) q)) := by
|
||||
rfl
|
||||
erw [h1]
|
||||
rw [ModuleCat.Hom.hom, ConcreteCategory.hom]
|
||||
simp only [ModuleCat.instConcreteCategoryLinearMapIdCarrier, LinearMap.coe_comp,
|
||||
Function.comp_apply]
|
||||
conv_lhs =>
|
||||
right
|
||||
change (PiTensorProduct.tprod k) p ⊗ₜ[k] (objMap' F f).hom ((PiTensorProduct.tprod k) q)
|
||||
rw [objMap'_tprod]
|
||||
change (μ F X' Y).hom.hom ((PiTensorProduct.tprod k) p ⊗ₜ[k] (PiTensorProduct.tprod k) fun i =>
|
||||
(discreteFunctorMapEqIso F _) (q ((OverColor.Hom.toEquiv f).symm i))) = _
|
||||
|
@ -530,7 +543,8 @@ lemma objMap'_comp {X Y Z : OverColor C} (f : X ⟶ Y) (g : Y ⟶ Z) :
|
|||
((PiTensorProduct.tprod k) x) =
|
||||
(mapToLinearEquiv' F g) ((mapToLinearEquiv' F f) ((PiTensorProduct.tprod k) x))
|
||||
rw [mapToLinearEquiv'_tprod, mapToLinearEquiv'_tprod]
|
||||
erw [mapToLinearEquiv'_tprod]
|
||||
conv_rhs =>
|
||||
erw [mapToLinearEquiv'_tprod F g]
|
||||
refine congrArg _ (funext (fun i => ?_))
|
||||
simp only [discreteFunctorMapEqIso, Functor.mapIso_hom, eqToIso.hom, Functor.mapIso_inv,
|
||||
eqToIso.inv, LinearEquiv.ofLinear_apply]
|
||||
|
@ -640,8 +654,10 @@ lemma mapApp'_unit : Functor.LaxMonoidal.ε (obj' F) ≫ mapApp' η (𝟙_ (Over
|
|||
simp only [obj', ε, instMonoidalCategoryStruct_tensorUnit_left, Functor.id_obj,
|
||||
instMonoidalCategoryStruct_tensorUnit_hom, objObj'_V_carrier,
|
||||
Action.instMonoidalCategory_tensorUnit_V, CategoryStruct.comp, Action.Hom.comp_hom,
|
||||
Action.mkIso_hom_hom, LinearEquiv.toModuleIso_hom_hom]
|
||||
rw [LinearMap.comp_apply]
|
||||
Action.mkIso_hom_hom]
|
||||
rw [ModuleCat.Hom.hom, ConcreteCategory.hom, ModuleCat.Hom.hom, ConcreteCategory.hom]
|
||||
simp only [ModuleCat.instConcreteCategoryLinearMapIdCarrier, LinearMap.coe_comp,
|
||||
Function.comp_apply]
|
||||
erw [PiTensorProduct.isEmptyEquiv_symm_apply, PiTensorProduct.isEmptyEquiv_symm_apply]
|
||||
simp only [map_smul]
|
||||
apply congrArg
|
||||
|
@ -660,6 +676,9 @@ lemma mapApp'_tensor (X Y : OverColor C) :
|
|||
instMonoidalCategoryStruct_tensorObj_hom, Functor.id_obj, CategoryStruct.comp,
|
||||
Action.Hom.comp_hom, Action.instMonoidalCategory_tensorObj_V, LinearMap.coe_comp,
|
||||
Function.comp_apply, Action.instMonoidalCategory_tensorHom_hom]
|
||||
rw [ModuleCat.Hom.hom, ConcreteCategory.hom, ModuleCat.Hom.hom, ConcreteCategory.hom]
|
||||
simp only [ModuleCat.instConcreteCategoryLinearMapIdCarrier, LinearMap.coe_comp,
|
||||
Function.comp_apply]
|
||||
erw [μ_tmul_tprod]
|
||||
erw [mapApp'_tprod]
|
||||
change _ = (μ F' X Y).hom.hom
|
||||
|
@ -825,7 +844,10 @@ def forgetLiftApp (c : C) : (lift.obj F).obj (OverColor.mk (fun (_ : Fin 1) => c
|
|||
(fun g => by
|
||||
refine ModuleCat.hom_ext ?_
|
||||
refine LinearMap.ext (fun x => ?_)
|
||||
simp only [forgetLiftAppV, Fin.isValue, LinearEquiv.toModuleIso_hom_hom]
|
||||
rw [ModuleCat.Hom.hom, ConcreteCategory.hom, ModuleCat.Hom.hom, ConcreteCategory.hom]
|
||||
simp only [ModuleCat.instConcreteCategoryLinearMapIdCarrier, LinearMap.coe_comp,
|
||||
Function.comp_apply]
|
||||
simp only [forgetLiftAppV, Fin.isValue]
|
||||
refine PiTensorProduct.induction_on' x (fun r x => ?_) <| fun x y hx hy => by
|
||||
simp_rw [map_add, hx, hy]
|
||||
simp only [CategoryStruct.comp, Fin.isValue, Functor.id_obj,
|
||||
|
|
|
@ -249,8 +249,12 @@ lemma contr_two_two_inner_tprod (c : S.C) (x : S.F.obj (OverColor.mk ![c, c]))
|
|||
Monoidal.tensorUnit_obj, Action.instMonoidalCategory_whiskerRight_hom,
|
||||
Action.instMonoidalCategory_associator_inv_hom, Action.instMonoidalCategory_associator_hom_hom,
|
||||
F_def]
|
||||
enter [2, 2]
|
||||
erw [OverColor.lift.μ_tmul_tprod S.FD]
|
||||
conv_lhs =>
|
||||
enter [2]
|
||||
rw (config := { transparency := .instances }) [OverColor.lift.map_tprod]
|
||||
conv_lhs =>
|
||||
rw (config := { transparency := .instances }) [contrMap_tprod]
|
||||
congr 1
|
||||
/- The contraction. -/
|
||||
|
|
|
@ -70,7 +70,7 @@ lemma contrFin1Fin1_inv_tmul {n : ℕ} (c : Fin n.succ.succ → S.C)
|
|||
((OverColor.forgetLiftApp S.FD (S.τ (c i))).inv.hom y))))) = _
|
||||
simp only [Nat.succ_eq_add_one, Action.instMonoidalCategory_tensorObj_V, Equivalence.symm_inverse,
|
||||
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
forgetLiftApp, Action.mkIso_inv_hom, LinearEquiv.toModuleIso_inv_hom, Fin.isValue]
|
||||
forgetLiftApp, Action.mkIso_inv_hom, Fin.isValue]
|
||||
erw [OverColor.forgetLiftAppV_symm_apply,
|
||||
OverColor.forgetLiftAppV_symm_apply S.FD (S.τ (c i))]
|
||||
change ((OverColor.lift.obj S.FD).map (OverColor.mkSum
|
||||
|
@ -200,8 +200,12 @@ lemma contrMap_tprod {n : ℕ} (c : Fin n.succ.succ → S.C)
|
|||
(((lift.obj S.FD).map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom).hom
|
||||
((PiTensorProduct.tprod S.k) x)))))) = _
|
||||
rw [lift.map_tprod]
|
||||
erw [lift.map_tprod]
|
||||
erw [lift.μIso_inv_tprod]
|
||||
conv_lhs =>
|
||||
enter [2, 2, 2, 2]
|
||||
rw (transparency := .instances) [lift.map_tprod]
|
||||
conv_lhs =>
|
||||
enter [2, 2, 2]
|
||||
rw (transparency := .instances) [lift.μIso_inv_tprod]
|
||||
change (λ_ ((lift.obj S.FD).obj (OverColor.mk (c ∘ i.succAbove ∘ j.succAbove)))).hom.hom
|
||||
(((S.contr.app { as := c i }).hom ▷ ((lift.obj S.FD).obj
|
||||
(OverColor.mk (c ∘ i.succAbove ∘ j.succAbove))).V)
|
||||
|
@ -247,12 +251,12 @@ lemma contrMap_tprod {n : ℕ} (c : Fin n.succ.succ → S.C)
|
|||
simp only [Nat.succ_eq_add_one, Fin.isValue, HepLean.Fin.finExtractTwo_symm_inl_inr_apply]
|
||||
simp [h]
|
||||
/- The tensor. -/
|
||||
· erw [lift.map_tprod]
|
||||
· conv_lhs => erw [lift.map_tprod]
|
||||
apply congrArg
|
||||
funext d
|
||||
simp only [mk_hom, Function.comp_apply, lift.discreteFunctorMapEqIso, Functor.mapIso_hom,
|
||||
eqToIso.hom, Functor.mapIso_inv, eqToIso.inv, eqToIso_refl, Functor.mapIso_refl, Iso.refl_hom,
|
||||
Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply]
|
||||
Action.id_hom, Iso.refl_inv, LinearEquiv.ofLinear_apply, ]
|
||||
change (S.FD.map (eqToHom _)).hom
|
||||
((x ((HepLean.Fin.finExtractTwo i j).symm (Sum.inr (d))))) = _
|
||||
simp only [Nat.succ_eq_add_one]
|
||||
|
|
|
@ -39,9 +39,9 @@ lemma contrFin1Fin1_naturality {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
(perm_contr_cond S h σ)).hom
|
||||
≫ ((Discrete.pairτ S.FD S.τ).map (Discrete.eqToHom (Hom.toEquiv_comp_inv_apply σ i) :
|
||||
(Discrete.mk (c ((Hom.toEquiv σ).symm i))) ⟶ (Discrete.mk (c1 i)))) := by
|
||||
erw [← CategoryTheory.Iso.eq_comp_inv]
|
||||
rw [← CategoryTheory.Iso.eq_comp_inv]
|
||||
rw [CategoryTheory.Category.assoc]
|
||||
erw [← CategoryTheory.Iso.inv_comp_eq]
|
||||
rw [← CategoryTheory.Iso.inv_comp_eq]
|
||||
ext1
|
||||
apply ModuleCat.hom_ext
|
||||
apply TensorProduct.ext'
|
||||
|
@ -57,7 +57,10 @@ lemma contrFin1Fin1_naturality {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
extractOne_homToEquiv, Fin.isValue, mk_hom, finExtractTwo_symm_inl_inr_apply,
|
||||
Discrete.mk.injEq]
|
||||
erw [perm_contr_cond S h σ]))).hom y))
|
||||
· apply congrArg
|
||||
· rw [ModuleCat.Hom.hom, ConcreteCategory.hom]
|
||||
simp only [ModuleCat.instConcreteCategoryLinearMapIdCarrier, LinearMap.coe_comp,
|
||||
Function.comp_apply]
|
||||
apply congrArg
|
||||
have h1' {α :Type} {a b c d : α} (hab : a= b) (hcd : c = d) (h : a = d) : b = c := by
|
||||
rw [← hab, hcd]
|
||||
exact h
|
||||
|
@ -108,7 +111,7 @@ lemma contrIso_comm_aux_1 {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
change ((S.F.map σ) ≫ (S.F.map (equivToIso (HepLean.Fin.finExtractTwo i j)).hom) ≫
|
||||
(S.F.map (mkSum (c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm)).hom)).hom X = _
|
||||
rw [← Functor.map_comp, ← Functor.map_comp]
|
||||
erw [extractTwo_finExtractTwo]
|
||||
rw [extractTwo_finExtractTwo]
|
||||
simp only [Nat.succ_eq_add_one, extractOne_homToEquiv, Functor.map_comp, Action.comp_hom,
|
||||
ModuleCat.hom_comp, Function.comp_apply]
|
||||
rfl
|
||||
|
@ -128,8 +131,9 @@ lemma contrIso_comm_aux_2 {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
(OverColor.mk ((c1 ∘ ⇑(HepLean.Fin.finExtractTwo i j).symm) ∘ Sum.inr))).inv =
|
||||
(Functor.Monoidal.μIso S.F _ _).inv ≫
|
||||
(S.F.map (extractTwoAux' i j σ) ⊗ S.F.map (extractTwoAux i j σ)) := by
|
||||
erw [CategoryTheory.IsIso.comp_inv_eq, CategoryTheory.Category.assoc]
|
||||
erw [CategoryTheory.IsIso.eq_inv_comp]
|
||||
apply (CategoryTheory.IsIso.comp_inv_eq _).mpr
|
||||
rw [CategoryTheory.Category.assoc]
|
||||
apply (CategoryTheory.IsIso.eq_inv_comp _).mpr
|
||||
exact (Functor.LaxMonoidal.μ_natural S.F (extractTwoAux' i j σ) (extractTwoAux i j σ)).symm
|
||||
exact congrArg (λ f => Action.Hom.hom f) h1
|
||||
|
||||
|
@ -170,7 +174,8 @@ lemma contrIso_comm_aux_5 {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
(S.F.map (mkIso (contrIso.proof_1 S c ((Hom.toEquiv σ).symm i)
|
||||
((HepLean.Fin.finExtractOnePerm ((Hom.toEquiv σ).symm i) (Hom.toEquiv σ)).symm j))).hom).hom)
|
||||
≫ (S.contrIsoComm σ).hom := by
|
||||
erw [← CategoryTheory.MonoidalCategory.tensor_comp (f₁ := (S.F.map (extractTwoAux' i j σ)).hom)]
|
||||
conv_lhs =>
|
||||
erw [← CategoryTheory.MonoidalCategory.tensor_comp (f₁ := (S.F.map (extractTwoAux' i j σ)).hom)]
|
||||
rw [contrIso_comm_aux_3 S σ]
|
||||
rw [contrFin1Fin1_naturality S h σ]
|
||||
simp [contrIsoComm]
|
||||
|
@ -233,8 +238,8 @@ lemma contrMap_naturality {n : ℕ} {c c1 : Fin n.succ.succ → S.C}
|
|||
have h1 : 𝟙_ (Rep S.k S.G) ◁ S.F.map (extractTwo i j σ) = 𝟙 _ ⊗ S.F.map (extractTwo i j σ) := by
|
||||
rfl
|
||||
rw [h1, ← tensor_comp]
|
||||
erw [CategoryTheory.Category.id_comp, CategoryTheory.Category.comp_id]
|
||||
erw [CategoryTheory.Category.comp_id]
|
||||
rw [CategoryTheory.Category.id_comp]
|
||||
erw [CategoryTheory.Category.comp_id, CategoryTheory.Category.comp_id]
|
||||
rw [S.contr.naturality]
|
||||
rfl
|
||||
|
||||
|
|
|
@ -425,9 +425,16 @@ lemma prod_contrMap_tprod (p : (i : (𝟭 Type).obj (OverColor.mk c1).left) →
|
|||
simp only [TensorSpecies.F_def]
|
||||
conv_rhs => rw [lift.obj_μ_tprod_tmul]
|
||||
simp only [TensorProduct.smul_tmul, TensorProduct.tmul_smul, map_smul]
|
||||
conv_lhs => rw [lift.obj_μ_tprod_tmul]
|
||||
conv_rhs => erw [lift.map_tprod]
|
||||
conv_rhs => erw [contrMap, TensorSpecies.contrMap_tprod]
|
||||
conv_lhs =>
|
||||
enter [2, 2]
|
||||
rw [lift.obj_μ_tprod_tmul]
|
||||
conv_rhs =>
|
||||
enter [2, 2]
|
||||
erw [lift.map_tprod]
|
||||
conv_rhs =>
|
||||
enter [2]
|
||||
rw [contrMap]
|
||||
erw [TensorSpecies.contrMap_tprod]
|
||||
simp only [TensorProduct.smul_tmul, TensorProduct.tmul_smul, map_smul]
|
||||
congr 1
|
||||
/- The contraction. -/
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue