feat: More results regarding index notation.

This commit is contained in:
jstoobysmith 2024-08-06 08:10:47 -04:00
parent a36afa9212
commit cef7e574ca
12 changed files with 424 additions and 32 deletions

View file

@ -44,13 +44,16 @@ structure TensorColor where
namespace TensorColor
variable (𝓒 : TensorColor)
variable (𝓒 : TensorColor) [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
variable {d : } {X Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y] [DecidableEq Y]
[Fintype Y'] [DecidableEq Y'] [Fintype Z] [DecidableEq Z] [Fintype W] [DecidableEq W]
/-- A relation on colors which is true if the two colors are equal or are duals. -/
def colorRel (μ ν : 𝓒.Color) : Prop := μ = ν μ = 𝓒ν
instance : Decidable (colorRel 𝓒 μ ν) :=
Or.decidable
/-- An equivalence relation on colors which is true if the two colors are equal or are duals. -/
lemma colorRel_equivalence : Equivalence 𝓒.colorRel where
refl := by
@ -87,6 +90,11 @@ instance colorSetoid : Setoid 𝓒.Color := ⟨𝓒.colorRel, 𝓒.colorRel_equi
def colorQuot (μ : 𝓒.Color) : Quotient 𝓒.colorSetoid :=
Quotient.mk 𝓒.colorSetoid μ
instance (μ ν : 𝓒.Color) : Decidable (μ ≈ ν) :=
Or.decidable
instance : DecidableEq (Quotient 𝓒.colorSetoid) :=
instDecidableEqQuotientOfDecidableEquiv
end TensorColor
noncomputable section