refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-16 16:42:20 +00:00
parent ec69deaff2
commit c73ae1aae6
5 changed files with 21 additions and 21 deletions

View file

@ -29,7 +29,7 @@ noncomputable section
def pair : Discrete C ⥤ Rep k G := F ⊗ F
/-- The isomorphism between the image of `(pair F).obj` and
`(lift.obj F).obj (OverColor.mk ![c,c])`. -/
`(lift.obj F).obj (OverColor.mk ![c,c])`. -/
def pairIso (c : C) : (pair F).obj (Discrete.mk c) ≅ (lift.obj F).obj (OverColor.mk ![c,c]) := by
symm
apply ((lift.obj F).mapIso fin2Iso).trans
@ -50,11 +50,11 @@ def pairIso (c : C) : (pair F).obj (Discrete.mk c) ≅ (lift.obj F).obj (OverCol
/-- The functor taking `c` to `F c ⊗ F (τ c)`. -/
def pairτ (τ : C → C) : Discrete C ⥤ Rep k G :=
F ⊗ ((Discrete.functor (Discrete.mk ∘ τ) : Discrete C ⥤ Discrete C) ⋙ F)
F ⊗ ((Discrete.functor (Discrete.mk ∘ τ) : Discrete C ⥤ Discrete C) ⋙ F)
/-- The functor taking `c` to `F (τ c) ⊗ F c`. -/
def τPair (τ : C → C) : Discrete C ⥤ Rep k G :=
((Discrete.functor (Discrete.mk ∘ τ) : Discrete C ⥤ Discrete C) ⋙ F) ⊗ F
((Discrete.functor (Discrete.mk ∘ τ) : Discrete C ⥤ Discrete C) ⋙ F) ⊗ F
end
end Discrete

View file

@ -42,7 +42,7 @@ def mkIso {c1 c2 : X → C} (h : c1 = c2) : mk c1 ≅ mk c2 :=
/-- The isomorphism splitting a `mk c` for `Fin 2 → C` into the tensor product of
the `Fin 1 → C` maps defined by `c 0` and `c 1`. -/
def fin2Iso {c : Fin 2 → C} : mk c ≅ mk ![c 0] ⊗ mk ![c 1] :=by
def fin2Iso {c : Fin 2 → C} : mk c ≅ mk ![c 0] ⊗ mk ![c 1] := by
let e1 : Fin 2 ≃ Fin 1 ⊕ Fin 1 := (finSumFinEquiv (n := 1)).symm
apply (equivToIso e1).trans
apply (mkSum _).trans
@ -175,7 +175,6 @@ def contrPairFin1Fin1 (τ : C → C) (c : Fin 1 ⊕ Fin 1 → C)
variable {k : Type} [CommRing k] {G : Type} [Group G]
/-- The Isomorphism between a `Fin n.succ.succ → C` and the product containing an object in the
image of `contrPair` based on the given values. -/
def contrPairEquiv {n : } (τ : C → C) (c : Fin n.succ.succ → C) (i : Fin n.succ.succ)
@ -188,7 +187,6 @@ def contrPairEquiv {n : } (τ : C → C) (c : Fin n.succ.succ → C) (i : Fin
(contrPairFin1Fin1 τ ((c ∘ ⇑(finExtractTwo i j).symm) ∘ Sum.inl) (by simpa using h)) <|
mkIso (by ext x; simp)
/-- Given a function `c` from `Fin 1` to `C`, this function returns a morphism
from `mk c` to `mk ![c 0]`. --/
def permFinOne (c : Fin 1 → C) : mk c ⟶ mk ![c 0] :=
@ -197,7 +195,7 @@ def permFinOne (c : Fin 1 → C) : mk c ⟶ mk ![c 0] :=
fin_cases x
rfl)).hom
/-- This a function that takes a function `c` from `Fin 2` to `C` and
/-- This a function that takes a function `c` from `Fin 2` to `C` and
returns a morphism from `mk c` to `mk ![c 0, c 1]`. --/
def permFinTwo (c : Fin 2 → C) : mk c ⟶ mk ![c 0, c 1] :=
(mkIso (by
@ -213,6 +211,5 @@ def permFinThree (c : Fin 3 → C) : mk c ⟶ mk ![c 0, c 1, c 2] :=
fin_cases x <;>
rfl)).hom
end OverColor
end IndexNotation

View file

@ -666,14 +666,14 @@ between the vector space obtained by applying the lift of `F` and that obtained
`F`.
--/
def forgetLiftAppV (c : C) : ((lift.obj F).obj (OverColor.mk (fun (_ : Fin 1) => c))).V ≃ₗ[k]
(F.obj (Discrete.mk c)).V :=
(F.obj (Discrete.mk c)).V :=
(PiTensorProduct.subsingletonEquiv 0 :
(⨂[k] (_ : Fin 1), (F.obj (Discrete.mk c))) ≃ₗ[k] F.obj (Discrete.mk c) )
(⨂[k] (_ : Fin 1), (F.obj (Discrete.mk c))) ≃ₗ[k] F.obj (Discrete.mk c))
/-- The `forgetLiftAppV` function takes an object `c` of type `C` and returns a isomorphism
between the objects obtained by applying the lift of `F` and that obtained by applying
`F`. -/
def forgetLiftApp (c : C) : (lift.obj F).obj (OverColor.mk (fun (_ : Fin 1 ) => c))
`F`. -/
def forgetLiftApp (c : C) : (lift.obj F).obj (OverColor.mk (fun (_ : Fin 1) => c))
≅ F.obj (Discrete.mk c) :=
Action.mkIso (forgetLiftAppV F c).toModuleIso
(fun g => by