refactor: Lint

This commit is contained in:
jstoobysmith 2024-12-20 13:34:49 +00:00
parent a5e0f3ceac
commit e4dafbd291
5 changed files with 9 additions and 9 deletions

View file

@ -87,7 +87,7 @@ lemma toList_get (a : CreateAnnilateSect f l) :
simp [tail]
@[simp]
lemma toList_grade :
lemma toList_grade :
FieldStatistic.ofList (fun i => q i.fst) a.toList = fermionic ↔
FieldStatistic.ofList q l = fermionic := by
induction l with
@ -308,7 +308,7 @@ variable {𝓕 : Type} {f : 𝓕 → Type} (q : 𝓕 → FieldStatistic) (le :
{l : List 𝓕} (a : CreateAnnilateSect f l)
lemma toList_koszulSignInsert (x : (i : 𝓕) × f i) :
koszulSignInsert (fun i => q i.fst) (fun i j => le i.fst j.fst) x a.toList =
koszulSignInsert (fun i => q i.fst) (fun i j => le i.fst j.fst) x a.toList =
koszulSignInsert q le x.1 l := by
induction l with
| nil => simp [koszulSignInsert]

View file

@ -167,7 +167,7 @@ lemma ofListLift_expand (f : 𝓕 → Type) [∀ i, Fintype (f i)] (x : ) :
lemma koszulOrder_ofListLift {f : 𝓕 → Type} [∀ i, Fintype (f i)]
(l : List 𝓕) (x : ) :
koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1) (ofListLift f l x) =
koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1) (ofListLift f l x) =
sumFiber f (koszulOrder q le (ofList l x)) := by
rw [koszulOrder_ofList]
rw [map_smul]

View file

@ -24,7 +24,8 @@ open FieldStatistic
is zero.
This can be thought as as a condtion on the operator algebra `A` as much as it can
on `F`. -/
class OperatorMap {A : Type} [Semiring A] [Algebra A] (q : I → FieldStatistic) (le1 : I → I → Prop)
class OperatorMap {A : Type} [Semiring A] [Algebra A]
(q : I → FieldStatistic) (le1 : I → I → Prop)
[DecidableRel le1] (F : FreeAlgebra I →ₐ[] A) : Prop where
superCommute_mem_center : ∀ i j, F (superCommute q (FreeAlgebra.ι i) (FreeAlgebra.ι j)) ∈
Subalgebra.center A
@ -255,7 +256,8 @@ lemma le_all_mul_koszulOrder_ofList_expand {I : Type}
exact fun j => hi j
lemma le_all_mul_koszulOrder_ofListLift_expand {I : Type} {f : I → Type} [∀ i, Fintype (f i)]
(q : I → FieldStatistic) (r : List I) (x : ) (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
(q : I → FieldStatistic) (r : List I) (x : )
(le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1]
[IsTotal (Σ i, f i) le1] [IsTrans (Σ i, f i) le1]
(i : (Σ i, f i)) (hi : ∀ (j : (Σ i, f i)), le1 j i)
{A : Type} [Semiring A] [Algebra A]

View file

@ -103,7 +103,6 @@ lemma insertSign_succ_cons (n : ) (r0 r1 : 𝓕) (r : List 𝓕) : insertSign
simp only [insertSign, List.take_succ_cons]
rw [superCommuteCoef_cons]
lemma insertSign_insert_gt (n m : ) (r0 r1 : 𝓕) (r : List 𝓕) (hn : n < m) :
insertSign q n r0 (List.insertIdx m r1 r) = insertSign q n r0 r := by
rw [insertSign, insertSign]

View file

@ -276,9 +276,8 @@ lemma ofList_ofList_superCommute (la lb : List 𝓕) (xa xb : ) :
rw [superCommute_ofList_ofList_superCommuteCoef]
abel
lemma ofListLift_ofList_superCommute' (l : List 𝓕) (r : List 𝓕) (x y : ) :
ofList r y * ofList l x = superCommuteCoef q l r • (ofList l x * ofList r y)
ofList r y * ofList l x = superCommuteCoef q l r • (ofList l x * ofList r y)
- superCommuteCoef q l r • superCommute q (ofList l x) (ofList r y) := by
nth_rewrite 2 [ofList_ofList_superCommute q]
rw [superCommuteCoef]
@ -340,7 +339,7 @@ lemma ofList_ofListLift_superCommute (l : List (Σ i, f i)) (r : List 𝓕) (x y
abel
lemma ofListLift_ofList_superCommute (l : List (Σ i, f i)) (r : List 𝓕) (x y : ) :
ofListLift f r y * ofList l x = superCommuteLiftCoef q l r • (ofList l x * ofListLift f r y)
ofListLift f r y * ofList l x = superCommuteLiftCoef q l r • (ofList l x * ofListLift f r y)
- superCommuteLiftCoef q l r •
superCommute (fun i => q i.1) (ofList l x) (ofListLift f r y) := by
rw [ofList_ofListLift_superCommute, superCommuteLiftCoef]