feat: Add results about solution planes

This commit is contained in:
jstoobysmith 2024-04-19 09:57:30 -04:00
parent 137c73c9bb
commit e710c9278e
5 changed files with 743 additions and 2 deletions

View file

@ -4,6 +4,8 @@ Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.LinearMaps
import Mathlib.Algebra.Module.Basic
import Mathlib.LinearAlgebra.FiniteDimensional
/-!
# Basic set up for anomaly cancellation conditions
@ -57,6 +59,9 @@ instance chargesModule (χ : ACCSystemCharges) : Module χ.charges :=
instance ChargesAddCommGroup (χ : ACCSystemCharges) : AddCommGroup χ.charges :=
Module.addCommMonoidToAddCommGroup
instance (χ : ACCSystemCharges) : Module.Finite χ.charges :=
FiniteDimensional.finiteDimensional_pi
end ACCSystemCharges
/-- The type of charges plus the linear ACCs. -/
@ -230,6 +235,10 @@ lemma Sols.ext {χ : ACCSystem} {S T : χ.Sols} (h : S.val = T.val) :
cases' S
simp_all only
/-- We say a charge S is a solution if it extends to a solution. -/
def isSolution (χ : ACCSystem) (S : χ.charges) : Prop :=
∃ (sol : χ.Sols), sol.val = S
/-- An instance giving the properties and structures to define an action of `` on `Sols`. -/
instance solsMulAction (χ : ACCSystem) : MulAction χ.Sols where
smul a S := ⟨a • S.toQuadSols , by