PhysLean/HepLean/Lorentz/Weyl/Modules.lean
2024-12-20 16:46:11 +00:00

208 lines
7.1 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
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 HepLean.Lorentz.SL2C.Basic
/-!
## Modules associated with Fermions
Weyl fermions live in the vector space `^2`, defined here as `Fin 2 → `.
However if we simply define the Module of Weyl fermions as `Fin 2 → ` we get casting problems,
where e.g. left-handed fermions can be cast to right-handed fermions etc.
To overcome this, for each type of Weyl fermion we define a structure that wraps `Fin 2 → `,
and these structures we define the instance of a module. This prevents casting between different
types of fermions.
-/
namespace Fermion
noncomputable section
section LeftHanded
/-- The module in which left handed fermions live. This is equivalent to `Fin 2 → `. -/
structure LeftHandedModule where
/-- The underlying value in `Fin 2 → `. -/
val : Fin 2 →
namespace LeftHandedModule
/-- The equivalence between `LeftHandedModule` and `Fin 2 → `. -/
def toFin2Fun : LeftHandedModule ≃ (Fin 2 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `LeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommMonoid LeftHandedModule := Equiv.addCommMonoid toFin2Fun
/-- The instance of `AddCommGroup` on `LeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommGroup LeftHandedModule := Equiv.addCommGroup toFin2Fun
/-- The instance of `Module` on `LeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : Module LeftHandedModule := Equiv.module toFin2Fun
/-- The linear equivalence between `LeftHandedModule` and `(Fin 2 → )`. -/
@[simps!]
def toFin2Equiv : LeftHandedModule ≃ₗ[] (Fin 2 → ) where
toFun := toFin2Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin2Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `LeftHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : LeftHandedModule) := toFin2Equiv ψ
end LeftHandedModule
end LeftHanded
section AltLeftHanded
/-- The module in which alt-left handed fermions live. This is equivalent to `Fin 2 → `. -/
structure AltLeftHandedModule where
/-- The underlying value in `Fin 2 → `. -/
val : Fin 2 →
namespace AltLeftHandedModule
/-- The equivalence between `AltLeftHandedModule` and `Fin 2 → `. -/
def toFin2Fun : AltLeftHandedModule ≃ (Fin 2 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `AltLeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommMonoid AltLeftHandedModule := Equiv.addCommMonoid toFin2Fun
/-- The instance of `AddCommGroup` on `AltLeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommGroup AltLeftHandedModule := Equiv.addCommGroup toFin2Fun
/-- The instance of `Module` on `AltLeftHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : Module AltLeftHandedModule := Equiv.module toFin2Fun
/-- The linear equivalence between `AltLeftHandedModule` and `(Fin 2 → )`. -/
@[simps!]
def toFin2Equiv : AltLeftHandedModule ≃ₗ[] (Fin 2 → ) where
toFun := toFin2Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin2Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `AltLeftHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : AltLeftHandedModule) := toFin2Equiv ψ
end AltLeftHandedModule
end AltLeftHanded
section RightHanded
/-- The module in which right handed fermions live. This is equivalent to `Fin 2 → `. -/
structure RightHandedModule where
/-- The underlying value in `Fin 2 → `. -/
val : Fin 2 →
namespace RightHandedModule
/-- The equivalence between `RightHandedModule` and `Fin 2 → `. -/
def toFin2Fun : RightHandedModule ≃ (Fin 2 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `RightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommMonoid RightHandedModule := Equiv.addCommMonoid toFin2Fun
/-- The instance of `AddCommGroup` on `RightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommGroup RightHandedModule := Equiv.addCommGroup toFin2Fun
/-- The instance of `Module` on `RightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : Module RightHandedModule := Equiv.module toFin2Fun
/-- The linear equivalence between `RightHandedModule` and `(Fin 2 → )`. -/
@[simps!]
def toFin2Equiv : RightHandedModule ≃ₗ[] (Fin 2 → ) where
toFun := toFin2Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin2Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `RightHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : RightHandedModule) := toFin2Equiv ψ
end RightHandedModule
end RightHanded
section AltRightHanded
/-- The module in which alt-right handed fermions live. This is equivalent to `Fin 2 → `. -/
structure AltRightHandedModule where
/-- The underlying value in `Fin 2 → `. -/
val : Fin 2 →
namespace AltRightHandedModule
/-- The equivalence between `AltRightHandedModule` and `Fin 2 → `. -/
def toFin2Fun : AltRightHandedModule ≃ (Fin 2 → ) where
toFun v := v.val
invFun f := ⟨f⟩
left_inv _ := rfl
right_inv _ := rfl
/-- The instance of `AddCommMonoid` on `AltRightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommMonoid AltRightHandedModule := Equiv.addCommMonoid toFin2Fun
/-- The instance of `AddCommGroup` on `AltRightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : AddCommGroup AltRightHandedModule := Equiv.addCommGroup toFin2Fun
/-- The instance of `Module` on `AltRightHandedModule` defined via its equivalence
with `Fin 2 → `. -/
instance : Module AltRightHandedModule := Equiv.module toFin2Fun
/-- The linear equivalence between `AltRightHandedModule` and `(Fin 2 → )`. -/
@[simps!]
def toFin2Equiv : AltRightHandedModule ≃ₗ[] (Fin 2 → ) where
toFun := toFin2Fun
map_add' := fun _ _ => rfl
map_smul' := fun _ _ => rfl
invFun := toFin2Fun.symm
left_inv := fun _ => rfl
right_inv := fun _ => rfl
/-- The underlying element of `Fin 2 → ` of a element in `AltRightHandedModule` defined
through the linear equivalence `toFin2Equiv`. -/
abbrev toFin2 (ψ : AltRightHandedModule) := toFin2Equiv ψ
end AltRightHandedModule
end AltRightHanded
end
end Fermion