fix: Anomaly Cancellation Group Actions & others
This commit is contained in:
parent
a532c51d64
commit
e6045e5f58
4 changed files with 14 additions and 13 deletions
|
@ -89,16 +89,16 @@ instance quadSolAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) :
|
|||
rfl
|
||||
|
||||
lemma linSolRep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.QuadSols) : χ.quadSolsInclLinSols (G.quadSolAction.toFun S g) =
|
||||
(S : χ.QuadSols) : χ.quadSolsInclLinSols (G.quadSolAction.toFun _ _ S g) =
|
||||
G.linSolRep g (χ.quadSolsInclLinSols S) := rfl
|
||||
|
||||
lemma rep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun S g) =
|
||||
(S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun _ _ S g) =
|
||||
G.rep g (χ.quadSolsIncl S) := rfl
|
||||
|
||||
/-- The group action acting on solutions to the anomaly cancellation conditions. -/
|
||||
instance solAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : MulAction G.group χ.Sols where
|
||||
smul g S := ⟨G.quadSolAction.toFun S.1 g, by
|
||||
smul g S := ⟨G.quadSolAction.toFun _ _ S.1 g, by
|
||||
simp only [MulAction.toFun_apply]
|
||||
change χ.cubicACC (G.rep g S.val) = 0
|
||||
rw [G.cubicInvariant, S.cubicSol]⟩
|
||||
|
@ -114,14 +114,14 @@ instance solAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : MulAction G.
|
|||
rfl
|
||||
|
||||
lemma quadSolAction_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun S g) =
|
||||
G.quadSolAction.toFun (χ.solsInclQuadSols S) g := rfl
|
||||
(S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun _ _ S g) =
|
||||
G.quadSolAction.toFun _ _ (χ.solsInclQuadSols S) g := rfl
|
||||
|
||||
lemma linSolRep_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.Sols) : χ.solsInclLinSols (G.solAction.toFun S g) =
|
||||
(S : χ.Sols) : χ.solsInclLinSols (G.solAction.toFun _ _ S g) =
|
||||
G.linSolRep g (χ.solsInclLinSols S) := rfl
|
||||
|
||||
lemma rep_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.Sols) : χ.solsIncl (G.solAction.toFun S g) = G.rep g (χ.solsIncl S) := rfl
|
||||
(S : χ.Sols) : χ.solsIncl (G.solAction.toFun _ _ S g) = G.rep g (χ.solsIncl S) := rfl
|
||||
|
||||
end ACCSystemGroupAction
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue