Theory Two_Catoid
section ‹2-Catoids›
theory Two_Catoid
imports Catoids.Catoid
begin
text‹We define 2-catoids and in particular (strict) 2-categories as local functional 2-catoids. With Isabelle
we first need to make two copies of catoids for the 0-structure and 1-structure.›
subsection ‹0-Structures and 1-structures.›
class multimagma0 =
fixes mcomp0 :: "'a ⇒ 'a ⇒ 'a set" (infixl "⊙⇩0" 70)
begin
sublocale mm0: multimagma mcomp0.
abbreviation "Δ⇩0 ≡ mm0.Δ"
abbreviation conv0 :: "'a set ⇒ 'a set ⇒ 'a set" (infixl "*⇩0" 70) where
"X *⇩0 Y ≡ mm0.conv X Y"
lemma "X *⇩0 Y = (⋃x ∈ X. ⋃y ∈ Y. x ⊙⇩0 y)"
by (simp add: mm0.conv_def)
end
class multimagma1 =
fixes mcomp1 :: "'a ⇒ 'a ⇒ 'a set" (infixl "⊙⇩1" 70)
begin
sublocale mm1: multimagma mcomp1.
abbreviation "Δ⇩1 ≡ mm1.Δ"
abbreviation conv1 :: "'a set ⇒ 'a set ⇒ 'a set" (infixl "*⇩1" 70) where
"X *⇩1 Y ≡ mm1.conv X Y"
end
class multisemigroup0 = multimagma0 +
assumes assoc: "(⋃v ∈ y ⊙⇩0 z. x ⊙⇩0 v) = (⋃v ∈ x ⊙⇩0 y. v ⊙⇩0 z)"
sublocale multisemigroup0 ⊆ msg0: multisemigroup mcomp0
by (unfold_locales, simp add: local.assoc)
class multisemigroup1 = multimagma1 +
assumes assoc: "(⋃v ∈ y ⊙⇩1 z. x ⊙⇩1 v) = (⋃v ∈ x ⊙⇩1 y. v ⊙⇩1 z)"
sublocale multisemigroup1 ⊆ msg1: multisemigroup mcomp1
by (unfold_locales, simp add: local.assoc)
class st_multimagma0 = multimagma0 +
fixes σ⇩0 :: "'a ⇒ 'a"
and τ⇩0 :: "'a ⇒ 'a"
assumes Dst0: "x ⊙⇩0 y ≠ {} ⟹ τ⇩0 x = σ⇩0 y"
and src0_absorb [simp]: "σ⇩0 x ⊙⇩0 x = {x}"
and tgt0_absorb [simp]: "x ⊙⇩0 τ⇩0 x = {x}"
begin