refactor: Lint

This commit is contained in:
jstoobysmith 2024-11-22 15:36:34 +00:00
parent 51158267d3
commit be64dbb8bb
4 changed files with 121 additions and 89 deletions

View file

@ -42,10 +42,10 @@ scoped[TensorSpecies] notation "⟪" ψ "," φ "⟫ₜₛ" => contractSelfField
/-- The map `contractSelfField` is equivariant with respect to the group action. -/
@[simp]
lemma contractSelfField_equivariant {S : TensorSpecies} {c : S.C} {g : S.G}
lemma contractSelfField_equivariant {S : TensorSpecies} {c : S.C} {g : S.G}
(ψ : S.FD.obj (Discrete.mk c)) (φ : S.FD.obj (Discrete.mk c)) :
⟪(S.FD.obj (Discrete.mk c)).ρ g ψ, (S.FD.obj (Discrete.mk c)).ρ g φ⟫ₜₛ = ⟪ψ, φ⟫ₜₛ := by
simpa using congrFun (congrArg (fun x => x.toFun) ((S.contractSelfHom c).comm g )) (ψ ⊗ₜ[S.k] φ)
simpa using congrFun (congrArg (fun x => x.toFun) ((S.contractSelfHom c).comm g)) (ψ ⊗ₜ[S.k] φ)
informal_lemma contractSelfField_non_degenerate where
math :≈ "The contraction of two vectors of the same color is non-degenerate.
@ -85,12 +85,12 @@ def IsNormZero {c : S.C} (ψ : S.FD.obj (Discrete.mk c)) : Prop := ⟪ψ, ψ⟫
/-- The zero vector has norm equal to zero. -/
@[simp]
lemma zero_isNormZero {c : S.C} : @IsNormZero S c 0 := by
lemma zero_isNormZero {c : S.C} : @IsNormZero S c 0 := by
simp only [IsNormZero, tmul_zero, map_zero]
/-- If a vector is norm-zero, then any scalar multiple of that vector is also norm-zero. -/
lemma smul_isNormZero_of_isNormZero {c : S.C} {ψ : S.FD.obj (Discrete.mk c)}
(h : S.IsNormZero ψ ) (a : S.k) : S.IsNormZero (a • ψ) := by
(h : S.IsNormZero ψ) (a : S.k) : S.IsNormZero (a • ψ) := by
simp only [IsNormZero, tmul_smul, map_smul, smul_tmul]
rw [h]
simp only [smul_eq_mul, mul_zero]

View file

@ -22,7 +22,6 @@ namespace TensorSpecies
variable (S : TensorSpecies)
/-- The isomorphism between the image of a map `Fin 1 ⊕ Fin 1 → S.C` contructed by `finExtractTwo`
under `S.F.obj`, and an object in the image of `OverColor.Discrete.pairτ S.FD`. -/
def contrFin1Fin1 {n : } (c : Fin n.succ.succ → S.C)