refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-03 07:32:46 +00:00
parent 24a20aef81
commit ce32218654
5 changed files with 70 additions and 55 deletions

View file

@ -75,8 +75,8 @@ lemma prodMatrix_hermitian (Φ1 Φ2 : HiggsField) (x : SpaceTime) :
/-- The map `prodMatrix` is a smooth function on spacetime. -/
lemma prodMatrix_smooth (Φ1 Φ2 : HiggsField) :
Smooth 𝓘(, SpaceTime) 𝓘(, Matrix (Fin 2) (Fin 2) ) (prodMatrix Φ1 Φ2) := by
rw [show 𝓘(, Matrix (Fin 2) (Fin 2) ) = modelWithCornersSelf (Fin 2 → Fin 2 → ) from rfl,
Smooth 𝓘(, SpaceTime) 𝓘(, Matrix (Fin 2) (Fin 2) ) (prodMatrix Φ1 Φ2) := by
rw [show 𝓘(, Matrix (Fin 2) (Fin 2) ) = modelWithCornersSelf (Fin 2 → Fin 2 → ) from rfl,
smooth_pi_space]
intro i
rw [smooth_pi_space]
@ -87,7 +87,7 @@ lemma prodMatrix_smooth (Φ1 Φ2 : HiggsField) :
informal_lemma prodMatrix_invariant where
math :≈ "The map ``prodMatrix is invariant under the simultanous action of ``gaugeAction
on the two Higgs fields."
on the two Higgs fields."
deps :≈ [``prodMatrix, ``gaugeAction]
informal_lemma prodMatrix_to_higgsField where

View file

@ -26,7 +26,6 @@ open SpaceTime
noncomputable section
/-!
## Some basic properties about SL(2, )
@ -36,7 +35,8 @@ Possibly to be moved to mathlib at some point.
lemma inverse_coe (M : SL(2, )) : M.1⁻¹ = (M⁻¹).1 := by
apply Matrix.inv_inj
simp
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero, not_false_eq_true,
nonsing_inv_nonsing_inv, SpecialLinearGroup.coe_inv]
have h1 : IsUnit M.1.det := by
simp
rw [Matrix.inv_adjugate M.1 h1]

View file

@ -18,7 +18,6 @@ https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrul
-/
namespace Fermion
noncomputable section
@ -28,7 +27,7 @@ open Complex
open TensorProduct
/-- The vector space ^2 carrying the fundamental representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_a.-/
In index notation corresponds to a Weyl fermion with indices ψ_a. -/
def leftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : LeftHandedModule) =>
@ -49,7 +48,7 @@ def leftHanded : Rep SL(2,) := Rep.of {
mulVec_mulVec]}
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ^a. -/
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ^a. -/
def altLeftHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltLeftHandedModule) =>
@ -72,7 +71,7 @@ def altLeftHanded : Rep SL(2,) := Rep.of {
exact transpose_mul _ _}
/-- The vector space ^2 carrying the conjugate representation of SL(2,C).
In index notation corresponds to a Weyl fermion with indices ψ_{dot a}.-/
In index notation corresponds to a Weyl fermion with indices ψ_{dot a}. -/
def rightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : RightHandedModule) =>
@ -93,7 +92,7 @@ def rightHanded : Rep SL(2,) := Rep.of {
/-- The vector space ^2 carrying the representation of SL(2,C) given by
M → (M⁻¹)^†.
In index notation this corresponds to a Weyl fermion with index `ψ^{dot a}`.-/
In index notation this corresponds to a Weyl fermion with index `ψ^{dot a}`. -/
def altRightHanded : Rep SL(2,) := Rep.of {
toFun := fun M => {
toFun := fun (ψ : AltRightHandedModule) =>
@ -153,7 +152,7 @@ lemma leftHandedToAlt_hom_apply (ψ : leftHanded) :
/-- The morphism from `altLeftHanded` to
`leftHanded` defined by multiplying an element of
altLeftHandedWeyl by the matrix `εₐ₁ₐ₂ = !![0, -1; 1, 0]`. -/
def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
def leftHandedAltTo : altLeftHanded ⟶ leftHanded where
hom := {
toFun := fun ψ =>
LeftHandedModule.toFin2Equiv.symm (!![0, -1; 1, 0] *ᵥ ψ.toFin2),
@ -293,7 +292,6 @@ def leftAltContraction : leftHanded ⊗ altLeftHanded ⟶ 𝟙_ (Rep SL(2,
rw [dotProduct_mulVec, vecMul_transpose, mulVec_mulVec]
simp
@[simp]
lemma leftAltContraction_hom_tmul (ψ : leftHanded) (φ : altLeftHanded) :
leftAltContraction.hom (ψ ⊗ₜ φ) = ψ.toFin2 ⬝ᵥ φ.toFin2 := by
rw [leftAltContraction]
@ -304,7 +302,7 @@ lemma leftAltContraction_hom_tmul (ψ : leftHanded) (φ : altLeftHanded) :
summing over components of altLeftHandedWeyl and leftHandedWeyl in the
standard basis (i.e. the dot product).
Physically, the contraction of a alt-left-handed Weyl fermion with a left-handed Weyl fermion.
In index notation this is φ^a ψ_a. -/
In index notation this is φ^a ψ_a. -/
def altLeftContraction : altLeftHanded ⊗ leftHanded ⟶ 𝟙_ (Rep SL(2,)) where
hom := TensorProduct.lift altLeftBi
comm M := by
@ -314,8 +312,7 @@ def altLeftContraction : altLeftHanded ⊗ leftHanded ⟶ 𝟙_ (Rep SL(2,
rw [dotProduct_mulVec, mulVec_transpose, vecMul_vecMul]
simp
@[simp]
lemma altLeftContraction_hom_tmul (φ : altLeftHanded) (ψ : leftHanded) :
lemma altLeftContraction_hom_tmul (φ : altLeftHanded) (ψ : leftHanded) :
altLeftContraction.hom (φ ⊗ₜ ψ) = φ.toFin2 ⬝ᵥ ψ.toFin2 := by
rw [altLeftContraction]
erw [TensorProduct.lift.tmul]

View file

@ -63,7 +63,7 @@ def toFin2Equiv : LeftHandedModule ≃ₗ[] (Fin 2 → ) where
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `LeftHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : LeftHandedModule) := toFin2Equiv ψ
end LeftHandedModule
@ -109,7 +109,7 @@ def toFin2Equiv : AltLeftHandedModule ≃ₗ[] (Fin 2 → ) where
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `AltLeftHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : AltLeftHandedModule) := toFin2Equiv ψ
end AltLeftHandedModule
@ -155,7 +155,7 @@ def toFin2Equiv : RightHandedModule ≃ₗ[] (Fin 2 → ) where
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `RightHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : RightHandedModule) := toFin2Equiv ψ
end RightHandedModule
@ -201,7 +201,7 @@ def toFin2Equiv : AltRightHandedModule ≃ₗ[] (Fin 2 → ) where
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `AltRightHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : AltRightHandedModule) := toFin2Equiv ψ
end AltRightHandedModule

View file

@ -18,8 +18,10 @@ import HepLean.SpaceTime.WeylFermion.Basic
namespace IndexNotation
open CategoryTheory
/-- The core of the category of Types over C. -/
def OverColor (C : Type) := CategoryTheory.Core (CategoryTheory.Over C)
/-- The instance of `OverColor C` as a groupoid. -/
instance (C : Type) : Groupoid (OverColor C) := coreCategory
namespace OverColor
@ -28,6 +30,7 @@ namespace Hom
variable {C : Type} {f g h : OverColor C}
/-- Given a hom in `OverColor C` the underlying equivalence between types. -/
def toEquiv (m : f ⟶ g) : f.left ≃ g.left where
toFun := m.hom.left
invFun := m.inv.left
@ -44,7 +47,8 @@ lemma toEquiv_comp (m : f ⟶ g) (n : g ⟶ h) : toEquiv (m ≫ n) = (toEquiv m)
lemma toEquiv_symm_apply (m : f ⟶ g) (i : g.left) :
f.hom ((toEquiv m).symm i) = g.hom i := by
sorry
simpa [toEquiv, types_comp] using congrFun m.inv.w i
end Hom
end OverColor
@ -60,13 +64,16 @@ open MatrixGroups
open Complex
open TensorProduct
open IndexNotation
open CategoryTheory
/-- The colors associated with complex representations of SL(2, ) of intrest to physics. -/
inductive Color
| upL : Color
| downL : Color
| upR : Color
| downR : Color
/-- The corresponding representations associated with a color. -/
def colorToRep (c : Color) : Rep SL(2, ) :=
match c with
| Color.upL => altLeftHanded
@ -74,6 +81,7 @@ def colorToRep (c : Color) : Rep SL(2, ) :=
| Color.upR => altRightHanded
| Color.downR => rightHanded
/-- The linear equivalence between `colorToRep c1` and `colorToRep c2` when `c1 = c2`. -/
def colorToRepCongr {c1 c2 : Color} (h : c1 = c2) : colorToRep c1 ≃ₗ[] colorToRep c2 where
toFun := Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)
invFun := (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)).symm
@ -86,50 +94,58 @@ def colorToRepCongr {c1 c2 : Color} (h : c1 = c2) : colorToRep c1 ≃ₗ[] co
left_inv x := Equiv.symm_apply_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
right_inv x := Equiv.apply_symm_apply (Equiv.cast (congrArg (CoeSort.coe ∘ colorToRep) h)) x
lemma colorToRepCongr_comm_ρ {c1 c2 : Color} (h : c1 = c2) (M : SL(2, )) (x : (colorToRep c1) ) :
lemma colorToRepCongr_comm_ρ {c1 c2 : Color} (h : c1 = c2) (M : SL(2, )) (x : (colorToRep c1)) :
(colorToRepCongr h) ((colorToRep c1).ρ M x) = (colorToRep c2).ρ M ((colorToRepCongr h) x) := by
subst h
rfl
def obj (f : OverColor Color) : Rep SL(2, ) := Rep.of {
namespace colorFun
/-- Given a object in `OverColor Color` the correpsonding tensor product of representations. -/
def obj' (f : OverColor Color) : Rep SL(2, ) := Rep.of {
toFun := fun M => PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M),
map_one' := by
simp
map_mul' := fun M N => by
simp
simp only [CategoryTheory.Functor.id_obj, _root_.map_mul]
ext x : 2
simp only [LinearMap.compMultilinearMap_apply, PiTensorProduct.map_tprod, LinearMap.mul_apply]}
lemma obj_ρ (f : OverColor Color) (M : SL(2, )) : (obj f).ρ M =
PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M) := rfl
lemma obj'_ρ (f : OverColor Color) (M : SL(2, )) : (obj' f).ρ M =
PiTensorProduct.map (fun x => (colorToRep (f.hom x)).ρ M) := rfl
lemma obj_ρ_tprod (f : OverColor Color) (M : SL(2, )) (x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
(obj f).ρ M ((PiTensorProduct.tprod ) x) =
lemma obj'_ρ_tprod (f : OverColor Color) (M : SL(2, ))
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
(obj' f).ρ M ((PiTensorProduct.tprod ) x) =
PiTensorProduct.tprod (fun i => (colorToRep (f.hom i)).ρ M (x i)) := by
rw [obj_ρ]
rw [obj'_ρ]
change (PiTensorProduct.map fun x => (colorToRep (f.hom x)).ρ M) ((PiTensorProduct.tprod ) x) =
(PiTensorProduct.tprod ) fun i => ((colorToRep (f.hom i)).ρ M) (x i)
rw [PiTensorProduct.map_tprod]
def morToLinearEquiv {f g : OverColor Color} (m : f ⟶ g) : (obj f).V ≃ₗ[] (obj g).V :=
/-- Given a morphism in `OverColor Color` the corresopnding linear equivalence between `obj' _`
induced by reindexing. -/
def mapToLinearEquiv' {f g : OverColor Color} (m : f ⟶ g) : (obj' f).V ≃ₗ[] (obj' g).V :=
(PiTensorProduct.reindex (fun x => colorToRep (f.hom x)) (OverColor.Hom.toEquiv m)).trans
(PiTensorProduct.congr (fun i => colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i)))
lemma morToLinearEquiv_tprod {f g : OverColor Color} (m : f ⟶ g)
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
morToLinearEquiv m (PiTensorProduct.tprod x) =
PiTensorProduct.tprod (fun i => (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i ))
lemma mapToLinearEquiv'_tprod {f g : OverColor Color} (m : f ⟶ g)
(x : (i : f.left) → CoeSort.coe (colorToRep (f.hom i))) :
mapToLinearEquiv' m (PiTensorProduct.tprod x) =
PiTensorProduct.tprod (fun i => (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i))
(x ((OverColor.Hom.toEquiv m).symm i))) := by
rw [morToLinearEquiv]
simp
rw [mapToLinearEquiv']
simp only [CategoryTheory.Functor.id_obj, LinearEquiv.trans_apply]
change (PiTensorProduct.congr fun i => colorToRepCongr _)
((PiTensorProduct.reindex (fun x => CoeSort.coe (colorToRep (f.hom x)))
(OverColor.Hom.toEquiv m)) ((PiTensorProduct.tprod ) x)) = _
rw [PiTensorProduct.reindex_tprod, PiTensorProduct.congr_tprod]
rfl
def mor {f g : OverColor Color} (m : f ⟶ g) : obj f ⟶ obj g where
hom := (morToLinearEquiv m).toLinearMap
/-- Given a morphism in `OverColor Color` the corresopnding map of representations induced by
reindexing. -/
def map' {f g : OverColor Color} (m : f ⟶ g) : obj' f ⟶ obj' g where
hom := (mapToLinearEquiv' m).toLinearMap
comm M := by
ext x : 2
refine PiTensorProduct.induction_on' x ?_ (by
@ -140,19 +156,21 @@ def mor {f g : OverColor Color} (m : f ⟶ g) : obj f ⟶ obj g where
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
change (morToLinearEquiv m) (((obj f).ρ M) ((PiTensorProduct.tprod ) x)) =
((obj g).ρ M) ((morToLinearEquiv m) ((PiTensorProduct.tprod ) x))
rw [morToLinearEquiv_tprod, obj_ρ_tprod]
erw [morToLinearEquiv_tprod, obj_ρ_tprod]
change (mapToLinearEquiv' m) (((obj' f).ρ M) ((PiTensorProduct.tprod ) x)) =
((obj' g).ρ M) ((mapToLinearEquiv' m) ((PiTensorProduct.tprod ) x))
rw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
erw [mapToLinearEquiv'_tprod, obj'_ρ_tprod]
apply congrArg
funext i
rw [colorToRepCongr_comm_ρ]
open CategoryTheory
end colorFun
/-- The functor between `OverColor Color` and `Rep SL(2, )` taking a map of colors
to the corresponding tensor product representation. -/
def colorFun : OverColor Color ⥤ Rep SL(2, ) where
obj := obj
map := mor
obj := colorFun.obj'
map := colorFun.map'
map_id f := by
simp only
ext x
@ -164,8 +182,8 @@ def colorFun : OverColor Color ⥤ Rep SL(2, ) where
simp only [CategoryTheory.Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod,
_root_.map_smul, Action.id_hom, ModuleCat.id_apply]
apply congrArg
rw [mor]
erw [morToLinearEquiv_tprod]
rw [colorFun.map']
erw [colorFun.mapToLinearEquiv'_tprod]
apply congrArg
funext i
rfl
@ -177,21 +195,21 @@ def colorFun : OverColor Color ⥤ Rep SL(2, ) where
simp only [CategoryTheory.Functor.id_obj, map_add, hx, ModuleCat.coe_comp,
Function.comp_apply, hy])
intro r x
simp
simp only [Functor.id_obj, PiTensorProduct.tprodCoeff_eq_smul_tprod, _root_.map_smul,
Action.comp_hom, ModuleCat.coe_comp, Function.comp_apply]
apply congrArg
rw [mor, mor, mor]
simp
change (morToLinearEquiv (CategoryTheory.CategoryStruct.comp f g)) ((PiTensorProduct.tprod ) x) =
(morToLinearEquiv g) ((morToLinearEquiv f) ((PiTensorProduct.tprod ) x))
rw [morToLinearEquiv_tprod, morToLinearEquiv_tprod]
erw [morToLinearEquiv_tprod]
rw [colorFun.map', colorFun.map', colorFun.map']
simp only
change (colorFun.mapToLinearEquiv' (CategoryTheory.CategoryStruct.comp f g))
((PiTensorProduct.tprod ) x) =
(colorFun.mapToLinearEquiv' g) ((colorFun.mapToLinearEquiv' f) ((PiTensorProduct.tprod ) x))
rw [colorFun.mapToLinearEquiv'_tprod, colorFun.mapToLinearEquiv'_tprod]
erw [colorFun.mapToLinearEquiv'_tprod]
apply congrArg
funext i
simp only [colorToRepCongr, Function.comp_apply, Equiv.cast_symm, LinearEquiv.coe_mk,
Equiv.cast_apply, cast_cast, cast_inj]
rfl
end
end Fermion