feat: More fixes
This commit is contained in:
parent
4df8663cbc
commit
c9c9047a0c
31 changed files with 134 additions and 124 deletions
|
@ -37,16 +37,16 @@ def BL₁ : (PlusU1 1).Sols where
|
|||
intro i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
| 1 => rfl
|
||||
| 2 => rfl
|
||||
| 3 => rfl
|
||||
| 0 => with_unfolding_all rfl
|
||||
| 1 => with_unfolding_all rfl
|
||||
| 2 => with_unfolding_all rfl
|
||||
| 3 => with_unfolding_all rfl
|
||||
quadSol := by
|
||||
intro i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
cubicSol := by rfl
|
||||
| 0 => with_unfolding_all rfl
|
||||
cubicSol := by with_unfolding_all rfl
|
||||
|
||||
/-- $B - L$ in the $n$-family case. -/
|
||||
@[simps!]
|
||||
|
@ -67,7 +67,7 @@ lemma on_quadBiLin (S : (PlusU1 n).Charges) :
|
|||
|
||||
lemma on_quadBiLin_AFL (S : (PlusU1 n).LinSols) : quadBiLin (BL n).val S.val = 0 := by
|
||||
rw [on_quadBiLin, YYsol S, SU2Sol S, SU3Sol S]
|
||||
rfl
|
||||
with_unfolding_all rfl
|
||||
|
||||
lemma add_AFL_quad (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
||||
accQuad (a • S.val + b • (BL n).val) = a ^ 2 * accQuad S.val := by
|
||||
|
@ -100,7 +100,7 @@ lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
|
|||
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
|
||||
cubeTriLin (BL n).val (BL n).val S.val = 0 := by
|
||||
rw [on_cubeTriLin, gravSol S, SU3Sol S]
|
||||
rfl
|
||||
with_unfolding_all rfl
|
||||
|
||||
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
||||
accCube (a • S.val + b • (BL n).val) =
|
||||
|
|
|
@ -57,8 +57,8 @@ theorem plane_exists_dim_le_7 {n : ℕ} (hn : ExistsPlane n) : n ≤ 7 := by
|
|||
obtain ⟨B, hB⟩ := exists_plane_exists_basis hn
|
||||
have h1 := LinearIndependent.fintype_card_le_finrank hB
|
||||
simp only [Fintype.card_sum, Fintype.card_fin] at h1
|
||||
rw [show FiniteDimensional.finrank ℚ (PlusU1 3).Charges = 18 from
|
||||
FiniteDimensional.finrank_fin_fun ℚ] at h1
|
||||
rw [show Module.finrank ℚ (PlusU1 3).Charges = 18 from
|
||||
Module.finrank_fin_fun ℚ] at h1
|
||||
exact Nat.le_of_add_le_add_left h1
|
||||
|
||||
end PlusU1
|
||||
|
|
|
@ -35,16 +35,16 @@ def Y₁ : (PlusU1 1).Sols where
|
|||
intro i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
| 1 => rfl
|
||||
| 2 => rfl
|
||||
| 3 => rfl
|
||||
| 0 => with_unfolding_all rfl
|
||||
| 1 => with_unfolding_all rfl
|
||||
| 2 => with_unfolding_all rfl
|
||||
| 3 => with_unfolding_all rfl
|
||||
quadSol := by
|
||||
intro i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
cubicSol := by rfl
|
||||
| 0 => with_unfolding_all rfl
|
||||
cubicSol := by with_unfolding_all rfl
|
||||
|
||||
/-- The hypercharge for `n` family. -/
|
||||
@[simps!]
|
||||
|
@ -96,7 +96,7 @@ lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
|
|||
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
|
||||
cubeTriLin (Y n).val (Y n).val S.val = 0 := by
|
||||
rw [on_cubeTriLin, YYsol S]
|
||||
rfl
|
||||
with_unfolding_all rfl
|
||||
|
||||
lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
|
||||
cubeTriLin (Y n).val S S = 6 * accQuad S := by
|
||||
|
@ -109,7 +109,7 @@ lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
|
|||
lemma on_cubeTriLin'_ALQ (S : (PlusU1 n).QuadSols) :
|
||||
cubeTriLin (Y n).val S.val S.val = 0 := by
|
||||
rw [on_cubeTriLin', quadSol S]
|
||||
rfl
|
||||
with_unfolding_all rfl
|
||||
|
||||
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
||||
accCube (a • S.val + b • (Y n).val) =
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue