refactor: Linting
This commit is contained in:
parent
1233987bd7
commit
52e591fa7a
27 changed files with 53 additions and 57 deletions
|
@ -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]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue