refactor: Change structure of SM file

This commit is contained in:
jstoobysmith 2024-05-09 15:09:14 -04:00
parent 8fd0b63edb
commit 5af2eb4d8d
5 changed files with 300 additions and 264 deletions

View file

@ -0,0 +1,59 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.StandardModel.Basic
import Mathlib.Data.Complex.Exponential
import Mathlib.Geometry.Manifold.VectorBundle.Basic
import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection
import Mathlib.Geometry.Manifold.Instances.Real
import Mathlib.RepresentationTheory.Basic
import Mathlib.LinearAlgebra.Matrix.ToLin
import Mathlib.Analysis.InnerProductSpace.Adjoint
/-!
# Representations appearing in the Standard Model
This file defines the basic representations which appear in the Standard Model.
-/
universe v u
namespace StandardModel
open Manifold
open Matrix
open Complex
open ComplexConjugate
@[simps!]
noncomputable def repU1Map (g : unitary ) : unitaryGroup (Fin 2) :=
⟨g ^ 3 • 1, by
rw [mem_unitaryGroup_iff, smul_one_mul, show g = ⟨g.1, g.prop⟩ from rfl]
simp only [SubmonoidClass.mk_pow, Submonoid.mk_smul, star_smul, star_pow, RCLike.star_def,
star_one]
rw [smul_smul, ← mul_pow]
erw [(unitary.mem_iff.mp g.prop).2]
simp only [one_pow, one_smul]⟩
@[simps!]
noncomputable def repU1 : unitary →* unitaryGroup (Fin 2) where
toFun g := repU1Map g
map_mul' g h := by
simp only [repU1Map, Submonoid.mk_mul_mk, mul_smul_one, smul_smul, mul_comm, ← mul_pow]
map_one' := by
simp only [repU1Map, one_pow, one_smul, Submonoid.mk_eq_one]
@[simps!]
def fundamentalSU2 : specialUnitaryGroup (Fin 2) →* unitaryGroup (Fin 2) where
toFun g := ⟨g.1, g.prop.1⟩
map_mul' _ _ := Subtype.ext rfl
map_one' := Subtype.ext rfl
lemma repU1_fundamentalSU2_commute (u1 : unitary ) (g : specialUnitaryGroup (Fin 2) ) :
repU1 u1 * fundamentalSU2 g = fundamentalSU2 g * repU1 u1 := by
apply Subtype.ext
simp
end StandardModel