refactor: Golfing
This commit is contained in:
parent
d9f19aa028
commit
0346bf192b
9 changed files with 42 additions and 116 deletions
|
@ -45,10 +45,8 @@ def toSMPlusH : MSSMCharges.charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) :=
|
|||
def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) where
|
||||
toFun f := (f ∘ Sum.inl , f ∘ Sum.inr)
|
||||
invFun f := Sum.elim f.1 f.2
|
||||
left_inv f := by
|
||||
aesop
|
||||
right_inv f := by
|
||||
aesop
|
||||
left_inv f := Sum.elim_comp_inl_inr f
|
||||
right_inv _ := rfl
|
||||
|
||||
/-- An equivalence between `MSSMCharges.charges` and `(Fin 18 → ℚ) × (Fin 2 → ℚ)`. This
|
||||
splits the charges up into the SM and the additional ones for the MSSM. -/
|
||||
|
@ -74,8 +72,8 @@ corresponding SM species of charges. -/
|
|||
@[simps!]
|
||||
def toSMSpecies (i : Fin 6) : MSSMCharges.charges →ₗ[ℚ] MSSMSpecies.charges where
|
||||
toFun S := (Prod.fst ∘ toSpecies) S i
|
||||
map_add' _ _ := by aesop
|
||||
map_smul' _ _ := by aesop
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
||||
lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin 3 → ℚ) × (Fin 2 → ℚ)) :
|
||||
(toSMSpecies i) (toSpecies.symm f) = f.1 i := by
|
||||
|
@ -98,16 +96,16 @@ abbrev N := toSMSpecies 5
|
|||
/-- The charge `Hd`. -/
|
||||
@[simps!]
|
||||
def Hd : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
toFun S := S ⟨18, by simp⟩
|
||||
map_add' _ _ := by aesop
|
||||
map_smul' _ _ := by aesop
|
||||
toFun S := S ⟨18, Nat.lt_of_sub_eq_succ rfl⟩
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
||||
/-- The charge `Hu`. -/
|
||||
@[simps!]
|
||||
def Hu : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
toFun S := S ⟨19, by simp⟩
|
||||
map_add' _ _ := by aesop
|
||||
map_smul' _ _ := by aesop
|
||||
toFun S := S ⟨19, Nat.lt_of_sub_eq_succ rfl⟩
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
||||
lemma charges_eq_toSpecies_eq (S T : MSSMCharges.charges) :
|
||||
S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by
|
||||
|
@ -333,10 +331,7 @@ lemma accQuad_ext {S T : (MSSMCharges).charges}
|
|||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
ring_nf
|
||||
have h1 : ∀ j, ∑ i, (toSMSpecies j S i)^2 = ∑ i, (toSMSpecies j T i)^2 := by
|
||||
intro j
|
||||
erw [h]
|
||||
rfl
|
||||
have h1 : ∀ j, ∑ i, (toSMSpecies j S i)^2 = ∑ i, (toSMSpecies j T i)^2 := fun j => h j
|
||||
repeat rw [h1]
|
||||
rw [hd, hu]
|
||||
|
||||
|
|
|
@ -80,13 +80,7 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (hR' : R
|
|||
rw [← Module.add_smul]
|
||||
simp
|
||||
rw [← Module.add_smul] at h1i
|
||||
have hR : ∃ i, R.val i ≠ 0 := by
|
||||
by_contra h
|
||||
simp at h
|
||||
have h0 : R.val = 0 := by
|
||||
funext i
|
||||
apply h i
|
||||
exact hR' h0
|
||||
have hR : ∃ i, R.val i ≠ 0 := Function.ne_iff.mp hR'
|
||||
obtain ⟨i, hi⟩ := hR
|
||||
have h2 := congrArg (fun S => S i) h1i
|
||||
change _ = 0 at h2
|
||||
|
@ -242,12 +236,12 @@ lemma α₁_proj (T : MSSMACC.Sols) : α₁ (proj T.1.1) =
|
|||
lemma α₁_proj_zero (T : MSSMACC.Sols) (h1 : α₃ (proj T.1.1) = 0) :
|
||||
α₁ (proj T.1.1) = 0 := by
|
||||
rw [α₁_proj, h1]
|
||||
simp
|
||||
exact mul_eq_zero_of_left rfl ((dot B₃.val) T.val - (dot Y₃.val) T.val)
|
||||
|
||||
lemma α₂_proj_zero (T : MSSMACC.Sols) (h1 : α₃ (proj T.1.1) = 0) :
|
||||
α₂ (proj T.1.1) = 0 := by
|
||||
rw [α₂_proj, h1]
|
||||
simp
|
||||
exact mul_eq_zero_of_left rfl ((dot Y₃.val) T.val - 2 * (dot B₃.val) T.val)
|
||||
|
||||
end proj
|
||||
|
||||
|
|
|
@ -39,28 +39,18 @@ def chargeMap (f : permGroup) : MSSMCharges.charges →ₗ[ℚ] MSSMCharges.char
|
|||
map_add' S T := by
|
||||
simp only
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
apply And.intro
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
rw [(toSMSpecies i).map_add]
|
||||
rw [toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
||||
simp only
|
||||
erw [(toSMSpecies i).map_add]
|
||||
rfl
|
||||
apply And.intro
|
||||
rw [Hd.map_add, Hd_toSpecies_inv, Hd_toSpecies_inv, Hd_toSpecies_inv]
|
||||
rfl
|
||||
rw [Hu.map_add, Hu_toSpecies_inv, Hu_toSpecies_inv, Hu_toSpecies_inv]
|
||||
rfl
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
apply And.intro
|
||||
apply And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
rw [(toSMSpecies i).map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
apply And.intro
|
||||
rfl
|
||||
rfl
|
||||
|
||||
lemma chargeMap_toSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
toSMSpecies j (chargeMap f S) = toSMSpecies j S ∘ f j := by
|
||||
|
@ -75,27 +65,21 @@ def repCharges : Representation ℚ (permGroup) (MSSMCharges).charges where
|
|||
apply LinearMap.ext
|
||||
intro S
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
apply And.intro
|
||||
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]
|
||||
rw [chargeMap_toSpecies, chargeMap_toSpecies]
|
||||
simp only [Pi.mul_apply, Pi.inv_apply]
|
||||
rw [chargeMap_toSpecies]
|
||||
rfl
|
||||
apply And.intro
|
||||
rfl
|
||||
rfl
|
||||
map_one' := by
|
||||
apply LinearMap.ext
|
||||
intro S
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
apply And.intro
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
erw [toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
apply And.intro
|
||||
rfl
|
||||
rfl
|
||||
|
||||
lemma repCharges_toSMSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
toSMSpecies j (repCharges f S) = toSMSpecies j S ∘ f⁻¹ j := by
|
||||
|
@ -104,10 +88,8 @@ lemma repCharges_toSMSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin
|
|||
lemma toSpecies_sum_invariant (m : ℕ) (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j (repCharges f S)) i =
|
||||
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j S) i := by
|
||||
erw [repCharges_toSMSpecies]
|
||||
change ∑ i : Fin 3, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin 3, ((fun a => a ^ m) ∘ _) i
|
||||
refine Equiv.Perm.sum_comp _ _ _ ?_
|
||||
simp only [permGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ]
|
||||
rw [repCharges_toSMSpecies]
|
||||
exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S)
|
||||
|
||||
lemma Hd_invariant (f : permGroup) (S : MSSMCharges.charges) :
|
||||
Hd (repCharges f S) = Hd S := rfl
|
||||
|
@ -151,7 +133,7 @@ lemma accQuad_invariant (f : permGroup) (S : MSSMCharges.charges) :
|
|||
lemma accCube_invariant (f : permGroup) (S : MSSMCharges.charges) :
|
||||
accCube (repCharges f S) = accCube S :=
|
||||
accCube_ext
|
||||
(by simpa using toSpecies_sum_invariant 3 f S)
|
||||
(fun j => toSpecies_sum_invariant 3 f S j)
|
||||
(Hd_invariant f S)
|
||||
(Hu_invariant f S)
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue