refactor: Lint

This commit is contained in:
jstoobysmith 2024-04-18 08:47:45 -04:00
parent 982281b0ff
commit 2ad7fc3cea
4 changed files with 48 additions and 33 deletions

View file

@ -71,7 +71,7 @@ def speciesEmbed (m n : ) :
0
map_add' S T := by
funext i
simp
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add]
by_cases hi : i.val < m
erw [dif_pos hi, dif_pos hi, dif_pos hi]
erw [dif_neg hi, dif_neg hi, dif_neg hi]
@ -116,13 +116,15 @@ lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).charg
lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).charges) (j : Fin 6) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (familyUniversal n S)) i =
n * (toSpecies j S ⟨0, by simp⟩) ^ m := by
simp
have h1 : (n : ) * (toSpecies j S ⟨0, by simp⟩) ^ m = ∑ _i : Fin n, (toSpecies j S ⟨0, by simp⟩) ^ m:= by
simp only [SMνSpecies_numberCharges, Function.comp_apply, toSpecies_apply, Fin.zero_eta,
Fin.isValue]
have h1 : (n : ) * (toSpecies j S ⟨0, by simp⟩) ^ m =
∑ _i : Fin n, (toSpecies j S ⟨0, by simp⟩) ^ m := by
rw [Fin.sum_const]
simp
erw [h1]
apply Finset.sum_congr
simp
simp only
intro i _
erw [toSpecies_familyUniversal]
@ -134,10 +136,10 @@ lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).charges)
(T : (SMνCharges n).charges) (j : Fin 6) :
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by
simp
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
rw [Finset.mul_sum]
apply Finset.sum_congr
simp
simp only
intro i _
erw [toSpecies_familyUniversal]
rfl
@ -146,41 +148,45 @@ lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).charges)
(T L : (SMνCharges n).charges) (j : Fin 6) :
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by
simp
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
rw [Finset.mul_sum]
apply Finset.sum_congr
simp
simp only
intro i _
erw [toSpecies_familyUniversal]
simp
simp only [SMνSpecies_numberCharges, Fin.zero_eta, Fin.isValue, toSpecies_apply]
ring
lemma familyUniversal_accGrav (S : (SMνCharges 1).charges) :
accGrav (familyUniversal n S) = n * (accGrav S) := by
rw [accGrav_decomp, accGrav_decomp]
repeat rw [sum_familyUniversal_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accSU2 (S : (SMνCharges 1).charges) :
accSU2 (familyUniversal n S) = n * (accSU2 S) := by
rw [accSU2_decomp, accSU2_decomp]
repeat rw [sum_familyUniversal_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accSU3 (S : (SMνCharges 1).charges) :
accSU3 (familyUniversal n S) = n * (accSU3 S) := by
rw [accSU3_decomp, accSU3_decomp]
repeat rw [sum_familyUniversal_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accYY (S : (SMνCharges 1).charges) :
accYY (familyUniversal n S) = n * (accYY S) := by
rw [accYY_decomp, accYY_decomp]
repeat rw [sum_familyUniversal_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges n).charges) :
@ -190,7 +196,8 @@ lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges
rw [quadBiLin_decomp]
repeat rw [sum_familyUniversal_two]
repeat rw [toSpecies_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj, sub_left_inj,
sub_right_inj]
ring
lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) :
@ -198,7 +205,8 @@ lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) :
rw [accQuad_decomp]
repeat erw [sum_familyUniversal]
rw [accQuad_decomp]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharges n).charges) :
@ -209,7 +217,7 @@ lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharg
rw [cubeTriLin_decomp]
repeat rw [sum_familyUniversal_three]
repeat rw [toSpecies_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj]
ring
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνCharges n).charges) :
@ -222,7 +230,7 @@ lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνChar
rw [familyUniversal_cubeTriLin]
repeat rw [sum_familyUniversal_two]
repeat rw [toSpecies_one]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply]
ring
lemma familyUniversal_accCube (S : (SMνCharges 1).charges) :
@ -230,7 +238,8 @@ lemma familyUniversal_accCube (S : (SMνCharges 1).charges) :
rw [accCube_decomp]
repeat erw [sum_familyUniversal]
rw [accCube_decomp]
simp
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
end SMRHN