refactor: Linting
This commit is contained in:
parent
1233987bd7
commit
52e591fa7a
27 changed files with 53 additions and 57 deletions
|
@ -183,12 +183,11 @@ lemma QuadSols.ext {χ : ACCSystemQuad} {S T : χ.QuadSols} (h : S.val = T.val)
|
|||
|
||||
/-- An instance giving the properties and structures to define an action of `ℚ` on `QuadSols`. -/
|
||||
instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction ℚ χ.QuadSols where
|
||||
smul a S := ⟨a • S.toLinSols , by
|
||||
smul a S := ⟨a • S.toLinSols, by
|
||||
intro i
|
||||
erw [(χ.quadraticACCs i).map_smul]
|
||||
rw [S.quadSol i]
|
||||
simp only [mul_zero]
|
||||
⟩
|
||||
simp only [mul_zero]⟩
|
||||
mul_smul a b S := by
|
||||
apply QuadSols.ext
|
||||
exact mul_smul _ _ _
|
||||
|
|
|
@ -44,8 +44,7 @@ def linSolMap {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) :
|
|||
χ.LinSols →ₗ[ℚ] χ.LinSols where
|
||||
toFun S := ⟨G.rep g S.val, by
|
||||
intro i
|
||||
rw [G.linearInvariant, S.linearSol]
|
||||
⟩
|
||||
rw [G.linearInvariant, S.linearSol]⟩
|
||||
map_add' S T := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
exact (G.rep g).map_add' _ _
|
||||
|
@ -85,8 +84,7 @@ instance quadSolAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) :
|
|||
smul f S := ⟨G.linSolRep f S.1, by
|
||||
intro i
|
||||
simp only [linSolRep_apply_apply_val]
|
||||
rw [G.quadInvariant, S.quadSol]
|
||||
⟩
|
||||
rw [G.quadInvariant, S.quadSol]⟩
|
||||
mul_smul f1 f2 S := by
|
||||
apply ACCSystemQuad.QuadSols.ext
|
||||
change (G.rep.toFun (f1 * f2)) S.val = _
|
||||
|
|
|
@ -42,7 +42,7 @@ def toSMPlusH : MSSMCharges.Charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) :=
|
|||
/-- An equivalence between `Fin 18 ⊕ Fin 2 → ℚ` and `(Fin 18 → ℚ) × (Fin 2 → ℚ)`. -/
|
||||
@[simps!]
|
||||
def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) where
|
||||
toFun f := (f ∘ Sum.inl , f ∘ Sum.inr)
|
||||
toFun f := (f ∘ Sum.inl, f ∘ Sum.inr)
|
||||
invFun f := Sum.elim f.1 f.2
|
||||
left_inv f := Sum.elim_comp_inl_inr f
|
||||
right_inv _ := rfl
|
||||
|
@ -473,8 +473,7 @@ def AnomalyFreeMk (S : MSSMACC.Charges) (hg : accGrav S = 0)
|
|||
intro i
|
||||
simp at i
|
||||
match i with
|
||||
| 0 => exact hquad
|
||||
⟩ , by exact hcube ⟩
|
||||
| 0 => exact hquad⟩, hcube⟩
|
||||
|
||||
lemma AnomalyFreeMk_val (S : MSSMACC.Charges) (hg : accGrav S = 0)
|
||||
(hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0)
|
||||
|
@ -490,8 +489,7 @@ def AnomalyFreeQuadMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0) :
|
|||
intro i
|
||||
simp at i
|
||||
match i with
|
||||
| 0 => exact hquad
|
||||
⟩
|
||||
| 0 => exact hquad⟩
|
||||
|
||||
/-- A `Sol` from a `LinSol` satisfying the quadratic and cubic ACCs. -/
|
||||
@[simp]
|
||||
|
@ -501,13 +499,12 @@ def AnomalyFreeMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0)
|
|||
intro i
|
||||
simp at i
|
||||
match i with
|
||||
| 0 => exact hquad
|
||||
⟩ , by exact hcube ⟩
|
||||
| 0 => exact hquad⟩, hcube⟩
|
||||
|
||||
/-- A `Sol` from a `QuadSol` satisfying the cubic ACCs. -/
|
||||
@[simp]
|
||||
def AnomalyFreeMk'' (S : MSSMACC.QuadSols) (hcube : accCube S.val = 0) : MSSMACC.Sols :=
|
||||
⟨S , by exact hcube ⟩
|
||||
⟨S, hcube⟩
|
||||
|
||||
lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
|
||||
(hcube : accCube S.val = 0) :
|
||||
|
|
|
@ -52,7 +52,7 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (hR' : R
|
|||
rw [planeY₃B₃_val, planeY₃B₃_val] at h
|
||||
have h1 := congrArg (fun S => dot Y₃.val S) h
|
||||
have h2 := congrArg (fun S => dot B₃.val S) h
|
||||
simp only [ Fin.isValue, ACCSystemCharges.chargesAddCommMonoid_add, Fin.reduceFinMk] at h1 h2
|
||||
simp only [Fin.isValue, ACCSystemCharges.chargesAddCommMonoid_add, Fin.reduceFinMk] at h1 h2
|
||||
erw [dot.map_add₂, dot.map_add₂] at h1 h2
|
||||
erw [dot.map_add₂ Y₃.val (a' • Y₃.val + b' • B₃.val) (c' • R.val)] at h1
|
||||
erw [dot.map_add₂ B₃.val (a' • Y₃.val + b' • B₃.val) (c' • R.val)] at h2
|
||||
|
|
|
@ -105,7 +105,7 @@ lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔
|
|||
intro h
|
||||
rw [quadCoeff] at h
|
||||
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
simp only [ Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
||||
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
|
||||
|
@ -124,7 +124,7 @@ lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
|
|||
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
|
||||
intro h
|
||||
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
simp only [Fin.isValue, Fin.reduceFinMk , mul_eq_zero,
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero,
|
||||
OfNat.ofNat_ne_zero, or_self, false_or] at h
|
||||
rw [h.2.1, h.2.2]
|
||||
simp
|
||||
|
@ -159,7 +159,7 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
|||
intro h
|
||||
rw [cubicCoeff] at h
|
||||
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
simp only [ Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
||||
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
||||
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
|
||||
|
@ -232,7 +232,7 @@ lemma toSolNSQuad_eq_planeY₃B₃_on_α (R : MSSMACC.AnomalyFreePerp) :
|
|||
|
||||
/-- Given an `R` perpendicular to `Y₃` and `B₃`, an element of `Sols`. This map is
|
||||
not surjective. -/
|
||||
def toSolNS : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, _ , _) =>
|
||||
def toSolNS : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, _, _) =>
|
||||
a • AnomalyFreeMk'' (toSolNSQuad R) (toSolNSQuad_cube R)
|
||||
|
||||
/-- A map from `Sols` to `MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` which on elements of
|
||||
|
@ -268,7 +268,7 @@ def inLineEqToSol : InLineEq × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, c
|
|||
|
||||
/-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/
|
||||
def inLineEqProj (T : InLineEqSol) : InLineEq × ℚ × ℚ × ℚ :=
|
||||
(⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1⟩,
|
||||
(quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val,
|
||||
(quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val),
|
||||
(quadCoeff T.val)⁻¹ * (
|
||||
|
@ -319,7 +319,7 @@ lemma inQuadToSol_smul (R : InQuad) (c₁ c₂ c₃ d : ℚ) :
|
|||
|
||||
/-- On elements of `inQuadSol` a right-inverse to `inQuadToSol`. -/
|
||||
def inQuadProj (T : InQuadSol) : InQuad × ℚ × ℚ × ℚ :=
|
||||
(⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1⟩,
|
||||
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
|
||||
(cubicCoeff T.val)⁻¹ * (cubeTriLin T.val.val T.val.val B₃.val),
|
||||
(cubicCoeff T.val)⁻¹ * (- cubeTriLin T.val.val T.val.val Y₃.val),
|
||||
|
@ -368,7 +368,7 @@ lemma inQuadCubeToSol_smul (R : InQuadCube) (c₁ c₂ c₃ d : ℚ) :
|
|||
|
||||
/-- On elements of `inQuadCubeSol` a right-inverse to `inQuadCubeToSol`. -/
|
||||
def inQuadCubeProj (T : InQuadCubeSol) : InQuadCube × ℚ × ℚ × ℚ :=
|
||||
(⟨⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(⟨⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1⟩,
|
||||
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
|
||||
(inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2⟩,
|
||||
(dot Y₃.val B₃.val)⁻¹ * (dot Y₃.val T.val.val - dot B₃.val T.val.val),
|
||||
|
|
|
@ -67,7 +67,7 @@ def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
|||
rw [charges_eq_toSpecies_eq]
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
simp only [ Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
|
||||
simp only [Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
|
||||
rw [chargeMap_toSpecies, chargeMap_toSpecies]
|
||||
simp only [Pi.mul_apply, Pi.inv_apply]
|
||||
rw [chargeMap_toSpecies]
|
||||
|
|
|
@ -60,7 +60,7 @@ def asLinSols (j : Fin n) : (PureU1 n.succ).LinSols :=
|
|||
simp at i
|
||||
match i with
|
||||
| 0 =>
|
||||
simp only [ Fin.isValue, PureU1_linearACCs, accGrav,
|
||||
simp only [Fin.isValue, PureU1_linearACCs, accGrav,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, Fin.coe_eq_castSucc]
|
||||
rw [Fin.sum_univ_castSucc]
|
||||
rw [Finset.sum_eq_single j]
|
||||
|
|
|
@ -148,7 +148,7 @@ lemma not_hasBoundary_zero_le (hnot : ¬ (HasBoundary S)) (h0 : S (0 : Fin n.suc
|
|||
induction i
|
||||
rfl
|
||||
rename_i i hii
|
||||
have hnott := hnot ⟨i, by {simp at hi; linarith} ⟩
|
||||
have hnott := hnot ⟨i, by {simp at hi; linarith}⟩
|
||||
have hii := hii (by omega)
|
||||
erw [← hii] at hnott
|
||||
exact (val_le_zero hS (hnott h0)).symm
|
||||
|
|
|
@ -702,7 +702,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
|
|||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') (g : Fin n.succ → ℚ) (f : Fin n → ℚ)
|
||||
(h : S.val = P g + P! f):
|
||||
∃
|
||||
(g' : Fin n.succ → ℚ) (f' : Fin n → ℚ) ,
|
||||
(g' : Fin n.succ → ℚ) (f' : Fin n → ℚ),
|
||||
S'.val = P g' + P! f' ∧ P! f' = P! f +
|
||||
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by
|
||||
let X := P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j
|
||||
|
|
|
@ -37,11 +37,11 @@ open VectorLikeEvenPlane
|
|||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
in the basis through that point is in the cubic. -/
|
||||
def LineInCubic (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ),
|
||||
accCube (2 * n.succ) (a • P g + b • P! f) = 0
|
||||
|
||||
lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) :
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ),
|
||||
3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f)
|
||||
+ b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by
|
||||
intro g f hS a b
|
||||
|
@ -90,7 +90,7 @@ lemma lineInCubicPerm_permute {S : (PureU1 (2 * n.succ)).LinSols}
|
|||
|
||||
lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols}
|
||||
(LIC : LineInCubicPerm S) :
|
||||
∀ (j : Fin n) (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) ,
|
||||
∀ (j : Fin n) (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f),
|
||||
(S.val (δ!₂ j) - S.val (δ!₁ j))
|
||||
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
|
||||
intro j g f h
|
||||
|
|
|
@ -80,7 +80,7 @@ lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols}
|
|||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case our parameterization above will be able to recover this point. -/
|
||||
def GenericCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) ,
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
|
||||
lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
||||
|
@ -95,7 +95,7 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
|||
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) = 0`. -/
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) ,
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0
|
||||
|
||||
lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
||||
|
|
|
@ -658,7 +658,7 @@ noncomputable def basisaAsBasis :
|
|||
basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card
|
||||
|
||||
lemma span_basis (S : (PureU1 (2 * n.succ + 1)).LinSols) :
|
||||
∃ (g f : Fin n.succ → ℚ) , S.val = P g + P! f := by
|
||||
∃ (g f : Fin n.succ → ℚ), S.val = P g + P! f := by
|
||||
have h := (mem_span_range_iff_exists_fun ℚ).mp (Basis.mem_span basisaAsBasis S)
|
||||
obtain ⟨f, hf⟩ := h
|
||||
simp [basisaAsBasis] at hf
|
||||
|
|
|
@ -37,11 +37,11 @@ open VectorLikeOddPlane
|
|||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
in the basis through that point is in the cubic. -/
|
||||
def LineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||
∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ),
|
||||
accCube (2 * n + 1) (a • P g + b • P! f) = 0
|
||||
|
||||
lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) :
|
||||
∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) (a b : ℚ) ,
|
||||
∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) (a b : ℚ),
|
||||
3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f)
|
||||
+ b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by
|
||||
intro g f hS a b
|
||||
|
@ -85,7 +85,7 @@ lemma lineInCubicPerm_permute {S : (PureU1 (2 * n + 1)).LinSols}
|
|||
|
||||
lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols}
|
||||
(LIC : LineInCubicPerm S) :
|
||||
∀ (j : Fin n.succ) (g f : Fin n.succ → ℚ) (_ : S.val = Pa g f) ,
|
||||
∀ (j : Fin n.succ) (g f : Fin n.succ → ℚ) (_ : S.val = Pa g f),
|
||||
(S.val (δ!₂ j) - S.val (δ!₁ j))
|
||||
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
|
||||
intro j g f h
|
||||
|
|
|
@ -78,7 +78,7 @@ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
|
|||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case our parameterization above will be able to recover this point. -/
|
||||
def GenericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) ,
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
|
||||
lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
||||
|
@ -94,7 +94,7 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
|||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case we will show that S is zero if it is true for all permutations. -/
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) ,
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0
|
||||
|
||||
lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
||||
|
|
|
@ -116,7 +116,7 @@ lemma cubic_zero_E'_zero (S : linearParameters) (hc : accCube (S.asCharges) = 0)
|
|||
def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where
|
||||
toFun S := S.asLinear
|
||||
invFun S := ⟨SMCharges.Q S.val (0 : Fin 1), (SMCharges.U S.val (0 : Fin 1) -
|
||||
SMCharges.D S.val (0 : Fin 1))/2 ,
|
||||
SMCharges.D S.val (0 : Fin 1))/2,
|
||||
SMCharges.E S.val (0 : Fin 1)⟩
|
||||
left_inv S := by
|
||||
apply linearParameters.ext
|
||||
|
|
|
@ -164,7 +164,7 @@ def accSU3 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [ Pi.add_apply, mul_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue