refactor: Change case of type and props

This commit is contained in:
jstoobysmith 2024-06-26 11:54:02 -04:00
parent 18b83f582e
commit f7a638d32e
58 changed files with 695 additions and 696 deletions

View file

@ -60,7 +60,7 @@ lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/
def chargeToLinear (S : (SM n).charges) (hGrav : accGrav S = 0)
def chargeToLinear (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols :=
⟨S, by
intro i
@ -84,14 +84,14 @@ def quadToAF (S : (SM n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/
def chargeToQuad (S : (SM n).charges) (hGrav : accGrav S = 0)
def chargeToQuad (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SM n).QuadSols :=
linearToQuad $ chargeToLinear S hGrav hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/
def chargeToAF (S : (SM n).charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
def chargeToAF (S : (SM n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
(hSU3 : accSU3 S = 0) (hc : accCube S = 0) : (SM n).Sols :=
quadToAF (chargeToQuad S hGrav hSU2 hSU3) hc
@ -103,7 +103,7 @@ def linearToAF (S : (SM n).LinSols)
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
def perm (n : ) : ACCSystemGroupAction (SM n) where
group := permGroup n
group := PermGroup n
groupInst := inferInstance
rep := repCharges
linearInvariant := by