feat: More fixes

This commit is contained in:
jstoobysmith 2024-11-02 08:50:17 +00:00
parent 4df8663cbc
commit c9c9047a0c
31 changed files with 134 additions and 124 deletions

View file

@ -110,7 +110,7 @@ lemma δ!₃_δ₁0 : @δ!₃ n = δ₁ 0 := rfl
lemma δ!₄_δ₂Last: @δ!₄ n = δ₂ (Fin.last n) := by
rw [Fin.ext_iff]
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₂, Fin.natAdd_last, Fin.val_last]
omega
@ -169,7 +169,7 @@ lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) :
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd] at h2
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_addNat] at h2
omega
· rfl
@ -197,7 +197,7 @@ lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp only [Fin.coe_cast, Fin.coe_natAdd, Fin.coe_castAdd, add_right_inj] at h2
simp only [Fin.coe_cast, Fin.coe_natAdd, Fin.coe_castAdd, Fin.coe_addNat, add_right_inj] at h2
omega
· rfl
@ -254,11 +254,11 @@ lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp only [basis!AsCharges, succ_eq_add_one, PureU1_numberCharges]
split<;> rename_i h
· simp only [δ!₃, succ_eq_add_one, Fin.isValue, δ!₁, Fin.ext_iff, Fin.coe_cast, Fin.coe_castAdd,
Fin.coe_fin_one, Fin.coe_natAdd] at h
Fin.val_eq_zero, Fin.coe_natAdd] at h
omega
· split <;> rename_i h2
· simp only [δ!₃, succ_eq_add_one, Fin.isValue, δ!₂, Fin.ext_iff, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_fin_one, Fin.coe_natAdd] at h2
Fin.coe_castAdd, Fin.val_eq_zero, Fin.coe_natAdd] at h2
omega
· rfl
@ -266,12 +266,12 @@ lemma basis!_on_δ!₄ (j : Fin n) : basis!AsCharges j δ!₄ = 0 := by
simp only [basis!AsCharges, succ_eq_add_one, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ!₁, Fin.coe_castAdd, add_right_inj] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ!₂, Fin.coe_castAdd, add_right_inj] at h2
omega
· rfl
@ -650,7 +650,7 @@ lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n → ) :
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
FiniteDimensional.finrank (PureU1 (2 * n.succ)).LinSols := by
Module.finrank (PureU1 (2 * n.succ)).LinSols := by
erw [BasisLinear.finrank_AnomalyFreeLinear]
simp only [Fintype.card_sum, Fintype.card_fin, mul_eq]
exact split_odd n

View file

@ -149,7 +149,7 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
· simp [Fin.ext_iff, δ!₂, δ!₁]
· simp [Fin.ext_iff, δ!₂, δ!₄]
· simp only [Nat.succ_eq_add_one, δ!₁, δ!₄, Fin.isValue, ne_eq, Fin.ext_iff, Fin.coe_cast,
Fin.coe_natAdd, Fin.coe_castAdd, Fin.val_last, Fin.coe_fin_one, add_zero, add_right_inj]
Fin.coe_natAdd, Fin.coe_castAdd, Fin.val_last, Fin.val_eq_zero, add_zero, add_right_inj]
omega
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)

View file

@ -150,15 +150,15 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ)).group),
SpecialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun _ _ S M)) :
LineInCubicPerm S.1.1 :=
fun M => special_case_lineInCubic (h M)
theorem special_case {S : (PureU1 (2 * n.succ.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ)).group),
SpecialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun _ _ S M)) :
∃ (M : (FamilyPermutations (2 * n.succ.succ)).group),
((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M).1.1
((FamilyPermutations (2 * n.succ.succ)).solAction.toFun _ _ S M).1.1
∈ Submodule.span (Set.range basis) :=
lineInCubicPerm_in_plane S (special_case_lineInCubic_perm h)

View file

@ -166,7 +166,7 @@ lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
ConstAbsProp (S.val i, S.val j) := by
refine Prop_two ConstAbsProp Fin.zero_ne_one ?_
intro M
let S' := (FamilyPermutations 4).solAction.toFun S M
let S' := (FamilyPermutations 4).solAction.toFun _ _ S M
have hS' : LineInPlaneCond S'.1.1 :=
(lineInPlaneCond_perm hS M)
exact linesInPlane_four S' hS'

View file

@ -90,7 +90,7 @@ lemma δa₂_δ!₁ (j : Fin n) : δa₂ j = δ!₁ j.castSucc := by
lemma δa₃_δ₃ : @δa₃ n = δ₃ := by
rw [Fin.ext_iff]
simp only [succ_eq_add_one, δa₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd,
Fin.coe_fin_one, add_zero, δ₃]
Fin.val_eq_zero, add_zero, δ₃]
exact Nat.add_comm 1 n
lemma δa₃_δ!₁ : δa₃ = δ!₁ (Fin.last n) := by
@ -265,12 +265,12 @@ lemma basis_on_δ₃ (j : Fin n) : basisAsCharges j δ₃ = 0 := by
simp only [basisAsCharges, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₁] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₂] at h2
omega
· rfl
@ -279,12 +279,12 @@ lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp only [basis!AsCharges, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_fin_one, δ!₁,
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.val_eq_zero, δ!₁,
Fin.coe_natAdd] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_fin_one, δ!₂,
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.val_eq_zero, δ!₂,
Fin.coe_natAdd] at h2
omega
· rfl
@ -643,7 +643,7 @@ lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) :
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) =
FiniteDimensional.finrank (PureU1 (2 * n.succ + 1)).LinSols := by
Module.finrank (PureU1 (2 * n.succ + 1)).LinSols := by
erw [BasisLinear.finrank_AnomalyFreeLinear]
simp only [Fintype.card_sum, Fintype.card_fin]
exact Eq.symm (Nat.two_mul n.succ)

View file

@ -149,7 +149,7 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ + 1)).group),
SpecialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun _ _ S M)) :
LineInCubicPerm S.1.1 := by
intro M
have hM := special_case_lineInCubic (h M)
@ -157,9 +157,11 @@ lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols}
theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group),
SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun _ _ S M)) :
S.1.1 = 0 := by
have ht := special_case_lineInCubic_perm h
exact lineInCubicPerm_zero ht
end Odd
end PureU1

View file

@ -38,8 +38,6 @@ def chargeMap {n : } (f : PermGroup n) :
map_add' _ _ := rfl
map_smul' _ _:= rfl
open PureU1Charges in
/-- The representation of `permGroup` acting on the vector space of charges. -/
@[simp]
def permCharges {n : } : Representation (PermGroup n) (PureU1 n).Charges where