Theory CZH_ECAT_Comma
section‹Comma categories›
theory CZH_ECAT_Comma
imports
CZH_ECAT_NTCF
CZH_ECAT_Simple
begin
subsection‹Background›
named_theorems cat_comma_cs_simps
named_theorems cat_comma_cs_intros
subsection‹Comma category›
subsubsection‹Definition and elementary properties›
text‹
See Exercise 1.3.vi in \<^cite>‹"riehl_category_2016"› or
Chapter II-6 in \<^cite>‹"mac_lane_categories_2010"›.
›
definition cat_comma_Obj :: "V ⇒ V ⇒ V"
where "cat_comma_Obj 𝔊 ℌ ≡ set
{
[a, b, f]⇩∘ | a b f.
a ∈⇩∘ 𝔊⦇HomDom⦈⦇Obj⦈ ∧
b ∈⇩∘ ℌ⦇HomDom⦈⦇Obj⦈ ∧
f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘𝔊⦇HomCod⦈⇙ ℌ⦇ObjMap⦈⦇b⦈
}"
lemma small_cat_comma_Obj[simp]:
"small
{
[a, b, f]⇩∘ | a b f.
a ∈⇩∘ 𝔄⦇Obj⦈ ∧ b ∈⇩∘ 𝔅⦇Obj⦈ ∧ f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈
}"
(is ‹small ?abfs›)
proof-
define Q where
"Q i = (if i = 0 then 𝔄⦇Obj⦈ else if i = 1⇩ℕ then 𝔅⦇Obj⦈ else ℭ⦇Arr⦈)"
for i
have "?abfs ⊆ elts (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
unfolding Q_def
proof
(
intro subsetI,
unfold mem_Collect_eq,
elim exE conjE,
intro vproductI;
simp only:
)
fix a b f show "𝒟⇩∘ [a, b, f]⇩∘ = set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: three nat_omega_simps)
qed (force simp: nat_omega_simps)+
then show "small ?abfs" by (rule down)
qed
definition cat_comma_Hom :: "V ⇒ V ⇒ V ⇒ V ⇒ V"
where "cat_comma_Hom 𝔊 ℌ A B ≡ set
{
[A, B, [g, h]⇩∘]⇩∘ | g h.
A ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
B ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
g : A⦇0⦈ ↦⇘𝔊⦇HomDom⦈⇙ B⦇0⦈ ∧
h : A⦇1⇩ℕ⦈ ↦⇘ℌ⦇HomDom⦈⇙ B⦇1⇩ℕ⦈ ∧
B⦇2⇩ℕ⦈ ∘⇩A⇘𝔊⦇HomCod⦈⇙ 𝔊⦇ArrMap⦈⦇g⦈ =
ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘𝔊⦇HomCod⦈⇙ A⦇2⇩ℕ⦈
}"
lemma small_cat_comma_Hom[simp]: "small
{
[A, B, [g, h]⇩∘]⇩∘ | g h.
A ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
B ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
g : A⦇0⦈ ↦⇘𝔄⇙ B⦇0⦈ ∧
h : A⦇1⇩ℕ⦈ ↦⇘𝔅⇙ B⦇1⇩ℕ⦈ ∧
B⦇2⇩ℕ⦈ ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ A⦇2⇩ℕ⦈
}"
(is ‹small ?abf_a'b'f'_gh›)
proof-
define Q where
"Q i =
(
if i = 0
then cat_comma_Obj 𝔊 ℌ
else if i = 1⇩ℕ then cat_comma_Obj 𝔊 ℌ else 𝔄⦇Arr⦈ ×⇩∙ 𝔅⦇Arr⦈
)"
for i
have "?abf_a'b'f'_gh ⊆ elts (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
unfolding Q_def
proof
(
intro subsetI,
unfold mem_Collect_eq,
elim exE conjE,
intro vproductI;
simp only:
)
fix a b f show "𝒟⇩∘ [a, b, f]⇩∘ = ZFC_in_HOL.set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: three nat_omega_simps)
qed (force simp : nat_omega_simps)+
then show "small ?abf_a'b'f'_gh" by (rule down)
qed
definition cat_comma_Arr :: "V ⇒ V ⇒ V"
where "cat_comma_Arr 𝔊 ℌ ≡
(
⋃⇩∘A∈⇩∘cat_comma_Obj 𝔊 ℌ. ⋃⇩∘B∈⇩∘cat_comma_Obj 𝔊 ℌ.
cat_comma_Hom 𝔊 ℌ A B
)"
definition cat_comma_composable :: "V ⇒ V ⇒ V"
where "cat_comma_composable 𝔊 ℌ ≡ set
{
[[B, C, G]⇩∘, [A, B, F]⇩∘]⇩∘ | A B C G F.
[B, C, G]⇩∘ ∈⇩∘ cat_comma_Arr 𝔊 ℌ ∧ [A, B, F]⇩∘ ∈⇩∘ cat_comma_Arr 𝔊 ℌ
}"
lemma small_cat_comma_composable[simp]:
shows "small
{
[[B, C, G]⇩∘, [A, B, F]⇩∘]⇩∘ | A B C G F.
[B, C, G]⇩∘ ∈⇩∘ cat_comma_Arr 𝔊 ℌ ∧ [A, B, F]⇩∘ ∈⇩∘ cat_comma_Arr 𝔊 ℌ
}"
(is ‹small ?S›)
proof(rule down)
show "?S ⊆ elts (cat_comma_Arr 𝔊 ℌ ×⇩∙ cat_comma_Arr 𝔊 ℌ)" by auto
qed
definition cat_comma :: "V ⇒ V ⇒ V" (‹(_ ⇩C⇩F↓⇩C⇩F _)› [1000, 1000] 999)
where "𝔊 ⇩C⇩F↓⇩C⇩F ℌ =
[
cat_comma_Obj 𝔊 ℌ,
cat_comma_Arr 𝔊 ℌ,
(λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇0⦈),
(λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇1⇩ℕ⦈),
(
λGF∈⇩∘cat_comma_composable 𝔊 ℌ.
[
GF⦇1⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇1⇩ℕ⦈,
[
GF⦇0⦈⦇2⇩ℕ⦈⦇0⦈ ∘⇩A⇘𝔊⦇HomDom⦈⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈ ∘⇩A⇘ℌ⦇HomDom⦈⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈
]⇩∘
]⇩∘
),
(
λA∈⇩∘cat_comma_Obj 𝔊 ℌ.
[A, A, [𝔊⦇HomDom⦈⦇CId⦈⦇A⦇0⦈⦈, ℌ⦇HomDom⦈⦇CId⦈⦇A⦇1⇩ℕ⦈⦈]⇩∘]⇩∘
)
]⇩∘"
text‹Components.›
lemma cat_comma_components:
shows "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈ = cat_comma_Obj 𝔊 ℌ"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈ = cat_comma_Arr 𝔊 ℌ"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈ = (λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇0⦈)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈ = (λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇1⇩ℕ⦈)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Comp⦈ =
(
λGF∈⇩∘cat_comma_composable 𝔊 ℌ.
[
GF⦇1⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇1⇩ℕ⦈,
[
GF⦇0⦈⦇2⇩ℕ⦈⦇0⦈ ∘⇩A⇘𝔊⦇HomDom⦈⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈ ∘⇩A⇘ℌ⦇HomDom⦈⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈
]⇩∘
]⇩∘
)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈ =
(
λA∈⇩∘cat_comma_Obj 𝔊 ℌ.
[A, A, [𝔊⦇HomDom⦈⦇CId⦈⦇A⦇0⦈⦈, ℌ⦇HomDom⦈⦇CId⦈⦇A⦇1⇩ℕ⦈⦈]⇩∘]⇩∘
)"
unfolding cat_comma_def dg_field_simps by (simp_all add: nat_omega_simps)
context
fixes α 𝔄 𝔅 ℭ 𝔊 ℌ
assumes 𝔊: "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and ℌ: "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
begin
interpretation 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule 𝔊)
interpretation ℌ: is_functor α 𝔅 ℭ ℌ by (rule ℌ)
lemma cat_comma_Obj_def':
"cat_comma_Obj 𝔊 ℌ ≡ set
{
[a, b, f]⇩∘ | a b f.
a ∈⇩∘ 𝔄⦇Obj⦈ ∧ b ∈⇩∘ 𝔅⦇Obj⦈ ∧ f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈
}"
unfolding cat_comma_Obj_def cat_cs_simps by simp
lemma cat_comma_Hom_def':
"cat_comma_Hom 𝔊 ℌ A B ≡ set
{
[A, B, [g, h]⇩∘]⇩∘ | g h.
A ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
B ∈⇩∘ cat_comma_Obj 𝔊 ℌ ∧
g : A⦇0⦈ ↦⇘𝔄⇙ B⦇0⦈ ∧
h : A⦇1⇩ℕ⦈ ↦⇘𝔅⇙ B⦇1⇩ℕ⦈ ∧
B⦇2⇩ℕ⦈ ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ A⦇2⇩ℕ⦈
}"
unfolding cat_comma_Hom_def cat_cs_simps by simp
lemma cat_comma_components':
shows "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈ = cat_comma_Obj 𝔊 ℌ"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈ = cat_comma_Arr 𝔊 ℌ"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈ = (λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇0⦈)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈ = (λF∈⇩∘cat_comma_Arr 𝔊 ℌ. F⦇1⇩ℕ⦈)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Comp⦈ =
(
λGF∈⇩∘cat_comma_composable 𝔊 ℌ.
[
GF⦇1⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇1⇩ℕ⦈,
[
GF⦇0⦈⦇2⇩ℕ⦈⦇0⦈ ∘⇩A⇘𝔄⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇0⦈,
GF⦇0⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈ ∘⇩A⇘𝔅⇙ GF⦇1⇩ℕ⦈⦇2⇩ℕ⦈⦇1⇩ℕ⦈
]⇩∘
]⇩∘
)"
and "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈ =
(λA∈⇩∘cat_comma_Obj 𝔊 ℌ. [A, A, [𝔄⦇CId⦈⦇A⦇0⦈⦈, 𝔅⦇CId⦈⦇A⦇1⇩ℕ⦈⦈]⇩∘]⇩∘)"
unfolding cat_comma_components cat_cs_simps by simp_all
end
subsubsection‹Objects›
lemma cat_comma_ObjI[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "A = [a, b, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
shows "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
using assms(4-6)
unfolding cat_comma_Obj_def'[OF assms(1,2)] assms(3) cat_comma_components
by simp
lemma cat_comma_ObjD[dest]:
assumes "[a, b, f]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "a ∈⇩∘ 𝔄⦇Obj⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
using assms
unfolding
cat_comma_components'[OF assms(2,3)] cat_comma_Obj_def'[OF assms(2,3)]
by auto
lemma cat_comma_ObjE[elim]:
assumes "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
obtains a b f where "A = [a, b, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
using assms
unfolding
cat_comma_components'[OF assms(2,3)] cat_comma_Obj_def'[OF assms(2,3)]
by auto
subsubsection‹Arrows›
lemma cat_comma_HomI[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "F = [A, B, [g, h]⇩∘]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
shows "F ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
using assms(1,2,6-10)
unfolding cat_comma_Hom_def'[OF assms(1,2)] assms(3-5)
by
(
intro in_set_CollectI exI conjI small_cat_comma_Hom,
unfold cat_comma_components'(1,2)[OF assms(1,2), symmetric],
(
cs_concl
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
)
(clarsimp simp: nat_omega_simps)+
lemma cat_comma_HomE[elim]:
assumes "F ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
obtains a b f a' b' f' g h
where "F = [A, B, [g, h]⇩∘]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
using assms(1)
by
(
unfold
cat_comma_components'[OF assms(2,3)] cat_comma_Hom_def'[OF assms(2,3)],
elim in_small_setE;
(unfold mem_Collect_eq, elim exE conjE cat_comma_ObjE[OF _ assms(2,3)])?,
insert that,
all‹
(unfold cat_comma_components'(1,2)[OF assms(2,3), symmetric],
elim cat_comma_ObjE[OF _ assms(2,3)]) | -
›
)
(auto simp: nat_omega_simps)
lemma cat_comma_HomD[dest]:
assumes "[[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘ ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
using assms(1) by (force elim!: cat_comma_HomE[OF _ assms(2,3)])+
lemma cat_comma_ArrI[cat_comma_cs_intros]:
assumes "F ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
shows "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
using assms
unfolding cat_comma_components cat_comma_Arr_def
by (intro vifunionI)
lemma cat_comma_ArrE[elim]:
assumes "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
obtains A B
where "F ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
using assms unfolding cat_comma_components cat_comma_Arr_def by auto
lemma cat_comma_ArrD[dest]:
assumes "[A, B, F]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "[A, B, F]⇩∘ ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
proof-
from assms obtain C D
where "[A, B, F]⇩∘ ∈⇩∘ cat_comma_Hom 𝔊 ℌ C D"
and "C ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "D ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by (elim cat_comma_ArrE)
moreover from cat_comma_HomE[OF this(1) assms(2,3)] have "A = C" and "B = D"
by auto
ultimately show "[A, B, F]⇩∘ ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
qed
subsubsection‹Domain›
lemma cat_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈)"
unfolding cat_comma_components by simp
lemma cat_comma_Dom_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
unfolding cat_comma_components by simp
lemma cat_comma_Dom_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
shows "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈⦇ABF⦈ = A"
using assms(2) unfolding assms(1) cat_comma_components by simp
lemma cat_comma_Dom_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈) ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
proof(rule vsv.vsv_vrange_vsubset)
fix ABF assume "ABF ∈⇩∘ 𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈)"
then have "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_prems cs_shallow cs_simp: cat_comma_cs_simps)
then obtain A B
where ABF: "ABF ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and A: "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and B: "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
from this(1) obtain a b f a' b' f' g h
where "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (elim cat_comma_HomE[OF _ assms(1,2)])
from ABF this A B show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈⦇ABF⦈ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
qed (auto intro: cat_comma_cs_intros)
subsubsection‹Codomain›
lemma cat_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈)"
unfolding cat_comma_components by simp
lemma cat_comma_Cod_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
unfolding cat_comma_components by simp
lemma cat_comma_Cod_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
shows "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈⦇ABF⦈ = B"
using assms(2)
unfolding assms(1) cat_comma_components
by (simp add: nat_omega_simps)
lemma cat_comma_Cod_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈) ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
proof(rule vsv.vsv_vrange_vsubset)
fix ABF assume "ABF ∈⇩∘ 𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈)"
then have "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_prems cs_shallow cs_simp: cat_comma_cs_simps)
then obtain A B
where F: "ABF ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
and A: "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and B: "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
from this(1) obtain a b f a' b' f' g h
where "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (elim cat_comma_HomE[OF _ assms(1,2)])
from F this A B show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈⦇ABF⦈ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
qed (auto intro: cat_comma_cs_intros)
subsubsection‹Arrow with a domain and a codomain›
lemma cat_comma_is_arrI[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "ABF = [A, B, F]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "F = [g, h]⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
shows "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
proof(intro is_arrI)
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
from assms(7-11) show "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
unfolding assms(3-6)
by
(
cs_concl
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
with assms(7-11) show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈⦇ABF⦈ = A" "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈⦇ABF⦈ = B"
unfolding assms(3-6) by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)+
qed
lemma cat_comma_is_arrD[dest]:
assumes "[[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘ :
[a, b, f]⇩∘ ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ [a', b', f']⇩∘"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
proof-
note F_is_arrD = is_arrD[OF assms(1)]
note F_cat_comma_ArrD = cat_comma_ArrD[OF F_is_arrD(1) assms(2,3)]
show "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (intro cat_comma_HomD[OF F_cat_comma_ArrD(1) assms(2,3)])+
qed
lemma cat_comma_is_arrE[elim]:
assumes "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
obtains a b f a' b' f' g h
where "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
proof-
note F_is_arrD = is_arrD[OF assms(1)]
from F_is_arrD(1) obtain C D
where "ABF ∈⇩∘ cat_comma_Hom 𝔊 ℌ C D"
and "C ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "D ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
from this(1) obtain a b f a' b' f' g h
where F_def: "ABF = [C, D, [g, h]⇩∘]⇩∘"
and "C = [a, b, f]⇩∘"
and "D = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (elim cat_comma_HomE[OF _ assms(2,3)])
with that show ?thesis
by (metis F_is_arrD(1,2,3) cat_comma_Cod_app cat_comma_Dom_app)
qed
subsubsection‹Composition›
lemma cat_comma_composableI:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "ABCGF = [BCG, ABF]⇩∘"
and "BCG : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
and "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
shows "ABCGF ∈⇩∘ cat_comma_composable 𝔊 ℌ"
proof-
from assms(1,2,5) obtain a b f a' b' f' gh
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, gh]⇩∘"
and "A = [a, b, f]⇩∘"
and "B = [a', b', f']⇩∘"
by auto
with assms(1,2,4) obtain a'' b'' f'' g'h'
where BCG_def: "BCG = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, g'h']⇩∘"
and "B = [a', b', f']⇩∘"
and "C = [a'', b'', f'']⇩∘"
by auto
from is_arrD(1)[OF assms(4)] have "BCG ∈⇩∘ cat_comma_Arr 𝔊 ℌ"
unfolding cat_comma_components'(2)[OF assms(1,2)].
moreover from is_arrD(1)[OF assms(5)] have "ABF ∈⇩∘ cat_comma_Arr 𝔊 ℌ"
unfolding cat_comma_components'(2)[OF assms(1,2)].
ultimately show ?thesis
unfolding assms(3) ABF_def BCG_def cat_comma_composable_def
by simp
qed
lemma cat_comma_composableE[elim]:
assumes "ABCGF ∈⇩∘ cat_comma_composable 𝔊 ℌ"
and "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
obtains BCG ABF A B C
where "ABCGF = [BCG, ABF]⇩∘"
and "BCG : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
and "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
proof-
from assms(1) obtain A B C G F
where ABCGF_def: "ABCGF = [[B, C, G]⇩∘, [A, B, F]⇩∘]⇩∘"
and BCG: "[B, C, G]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
and ABF: "[A, B, F]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
unfolding cat_comma_composable_def
by (auto simp: cat_comma_components'[OF assms(2,3)])
note BCG = cat_comma_ArrD[OF BCG assms(2,3)]
and ABF = cat_comma_ArrD[OF ABF assms(2,3)]
from ABF(1) assms(2,3) obtain a b f a' b' f' g h
where "[A, B, F]⇩∘ = [A, B, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and F_def: "F = [g, h]⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [cat_comma_cs_simps]:
"f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with BCG(1) assms(2,3) obtain a'' b'' f'' g' h'
where g'h'_def: "[B, C, G]⇩∘ = [B, C, [g', h']⇩∘]⇩∘"
and C_def: "C = [a'', b'', f'']⇩∘"
and G_def: "G = [g', h']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_comma_cs_simps]:
"f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
from F_def have "F = [g, h]⇩∘" by simp
from assms(2,3) g h f f' g' h' f'' have
"[B, C, G]⇩∘ : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
unfolding ABCGF_def F_def G_def A_def B_def C_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_is_arrI
)+
moreover from assms(2,3) g h f f' g' h' f'' have
"[A, B, F]⇩∘ : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
unfolding ABCGF_def F_def G_def A_def B_def C_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_is_arrI
)+
ultimately show ?thesis using that ABCGF_def by auto
qed
lemma cat_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Comp⦈)"
unfolding cat_comma_components by auto
lemma cat_comma_Comp_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Comp⦈) = cat_comma_composable 𝔊 ℌ"
unfolding cat_comma_components by auto
lemma cat_comma_Comp_app[cat_comma_cs_simps]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "G = [B, C, [g', h']⇩∘]⇩∘"
and "F = [A, B, [g, h]⇩∘]⇩∘"
and "G : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
and "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
shows "G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F = [A, C, [g' ∘⇩A⇘𝔄⇙ g, h' ∘⇩A⇘𝔅⇙ h]⇩∘]⇩∘"
using assms(1,2,5,6)
unfolding cat_comma_components'[OF assms(1,2)] assms(3,4)
by
(
cs_concl
cs_simp: omega_of_set V_cs_simps vfsequence_simps
cs_intro: nat_omega_intros V_cs_intros cat_comma_composableI TrueI
)
lemma cat_comma_Comp_is_arr[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "BCG : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
and "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
shows "BCG ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
from assms(1,2,4) obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [symmetric, cat_cs_simps]:
"f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with assms(1,2,3) obtain a'' b'' f'' g' h'
where BCG_def: "BCG = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, [g', h']⇩∘]⇩∘"
and C_def: "C = [a'', b'', f'']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_cs_simps]: "f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
from g' have 𝔊g': "𝔊⦇ArrMap⦈⦇g'⦈ : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ 𝔊⦇ObjMap⦈⦇a''⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
note [cat_cs_simps] =
category.cat_assoc_helper[
where ℭ=ℭ and h=f'' and g=‹𝔊⦇ArrMap⦈⦇g'⦈› and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
category.cat_assoc_helper[
where ℭ=ℭ and h=f and g=‹ℌ⦇ArrMap⦈⦇h⦈› and q=‹f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈›
]
from assms(1,2,3,4) g h f f' g' h' f'' show ?thesis
unfolding ABF_def BCG_def A_def B_def C_def
by (intro cat_comma_is_arrI[OF assms(1,2)])
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros
)+
qed
subsubsection‹Identity›
lemma cat_comma_CId_vsv[cat_comma_cs_intros]: "vsv (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈)"
unfolding cat_comma_components by simp
lemma cat_comma_CId_vdomain[cat_comma_cs_simps]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
unfolding cat_comma_components'[OF assms(1,2)] by simp
lemma cat_comma_CId_app[cat_comma_cs_simps]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "A = [a, b ,f]⇩∘"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
shows "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇A⦈ = [A, A, [𝔄⦇CId⦈⦇a⦈, 𝔅⦇CId⦈⦇b⦈]⇩∘]⇩∘"
proof-
from assms(4)[unfolded assms(3), unfolded cat_comma_components'[OF assms(1,2)]]
have "[a, b, f]⇩∘ ∈⇩∘ cat_comma_Obj 𝔊 ℌ".
then show ?thesis
unfolding cat_comma_components'(6)[OF assms(1,2)] assms(3)
by (simp add: nat_omega_simps)
qed
subsubsection‹‹Hom›-set›
lemma cat_comma_Hom:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
shows "Hom (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) A B = cat_comma_Hom 𝔊 ℌ A B"
proof(intro vsubset_antisym vsubsetI, unfold in_Hom_iff)
fix ABF assume "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
with assms(1,2) show "ABF ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
by (elim cat_comma_is_arrE[OF _ assms(1,2)], intro cat_comma_HomI) force+
next
fix ABF assume "ABF ∈⇩∘ cat_comma_Hom 𝔊 ℌ A B"
with assms(1,2) show "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
by (elim cat_comma_HomE[OF _ assms(1,2)], intro cat_comma_is_arrI) force+
qed
subsubsection‹Comma category is a category›
lemma category_cat_comma[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "category α (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret 𝔉: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
show ?thesis
proof(rule categoryI')
show "vfsequence (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)" unfolding cat_comma_def by auto
show "vcard (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) = 6⇩ℕ"
unfolding cat_comma_def by (simp add: nat_omega_simps)
show "ℛ⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Dom⦈) ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by (rule cat_comma_Dom_vrange[OF assms])
show "ℛ⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Cod⦈) ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by (rule cat_comma_Cod_vrange[OF assms])
show "(GF ∈⇩∘ 𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Comp⦈)) ⟷
(∃g f b c a. GF = [g, f]⇩∘ ∧ g : b ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ c ∧ f : a ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ b)"
for GF
proof(intro iffI; (elim exE conjE)?; (simp only: cat_comma_Comp_vdomain)?)
assume prems: "GF ∈⇩∘ cat_comma_composable 𝔊 ℌ"
with assms obtain G F abf a'b'f' a''b''f''
where "GF = [G, F]⇩∘"
and "G : a'b'f' ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ a''b''f''"
and "F : abf ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ a'b'f'"
by auto
with assms show "∃g f b c a.
GF = [g, f]⇩∘ ∧ g : b ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ c ∧ f : a ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ b"
by auto
qed (use assms in ‹cs_concl cs_shallow cs_intro: cat_comma_composableI›)
from assms show "𝒟⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
from assms show "G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
if "G : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C" and "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
for B C G A F
using that by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
from assms show
"H ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F =
H ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ (G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F)"
if "H : C ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ D"
and "G : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C"
and "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
for C D H B G A F
using assms that
proof-
from that(3) assms obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [cat_cs_simps]: "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with that(2) assms obtain a'' b'' f'' g' h'
where G_def: "G = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, [g', h']⇩∘]⇩∘"
and C_def: "C = [a'', b'', f'']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_cs_simps]:
"f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
with that(1) assms obtain a''' b''' f''' g'' h''
where H_def: "H = [[a'', b'', f'']⇩∘, [a''', b''', f''']⇩∘, [g'', h'']⇩∘]⇩∘"
and D_def: "D = [a''', b''', f''']⇩∘"
and g'': "g'' : a'' ↦⇘𝔄⇙ a'''"
and h'': "h'' : b'' ↦⇘𝔅⇙ b'''"
and f''': "f''' : 𝔊⦇ObjMap⦈⦇a'''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'''⦈"
and [cat_cs_simps]:
"f''' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g''⦈ = ℌ⦇ArrMap⦈⦇h''⦈ ∘⇩A⇘ℭ⇙ f''"
by auto
note [cat_cs_simps] =
category.cat_assoc_helper[
where ℭ=ℭ
and h=f''
and g=‹𝔊⦇ArrMap⦈⦇g'⦈›
and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
category.cat_assoc_helper[
where ℭ=ℭ
and h=f''
and g=‹𝔊⦇ArrMap⦈⦇g'⦈›
and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
category.cat_assoc_helper[
where ℭ=ℭ
and h=f'''
and g=‹𝔊⦇ArrMap⦈⦇g''⦈›
and q=‹ℌ⦇ArrMap⦈⦇h''⦈ ∘⇩A⇘ℭ⇙ f''›
]
from assms that g h f f' g' h' f'' g'' h'' f''' show ?thesis
unfolding F_def G_def H_def A_def B_def C_def D_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇A⦈ : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ A"
if "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" for A
using that
by (elim cat_comma_ObjE[OF _ assms(1)]; (simp only:)?)
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇B⦈ ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F = F"
if "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" for A B F
using that
by (elim cat_comma_is_arrE[OF _ assms]; (simp only:)?)
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
show "F ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇B⦈ = F"
if "F : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C" for B C F
using that
by (elim cat_comma_is_arrE[OF _ assms]; (simp only:)?)
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈ ⊆⇩∘ Vset α"
proof(intro vsubsetI, elim cat_comma_ObjE[OF _ assms])
fix F a b f assume prems:
"F = [a, b, f]⇩∘"
"a ∈⇩∘ 𝔄⦇Obj⦈"
"b ∈⇩∘ 𝔅⦇Obj⦈"
"f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
from prems(2-4) show "F ∈⇩∘ Vset α"
unfolding prems(1) by (cs_concl cs_intro: cat_cs_intros V_cs_intros)
qed
show "(⋃⇩∘a∈⇩∘A. ⋃⇩∘b∈⇩∘B. Hom (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) a b) ∈⇩∘ Vset α"
if "A ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "B ⊆⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and "A ∈⇩∘ Vset α"
and "B ∈⇩∘ Vset α"
for A B
proof-
define A0 where "A0 = ℛ⇩∘ (λF∈⇩∘A. F⦇0⦈)"
define A1 where "A1 = ℛ⇩∘ (λF∈⇩∘A. F⦇1⇩ℕ⦈)"
define B0 where "B0 = ℛ⇩∘ (λF∈⇩∘B. F⦇0⦈)"
define B1 where "B1 = ℛ⇩∘ (λF∈⇩∘B. F⦇1⇩ℕ⦈)"
define A0B0 where "A0B0 = (⋃⇩∘a∈⇩∘A0. ⋃⇩∘b∈⇩∘B0. Hom 𝔄 a b)"
define A1B1 where "A1B1 = (⋃⇩∘a∈⇩∘A1. ⋃⇩∘b∈⇩∘B1. Hom 𝔅 a b)"
have A0B0: "A0B0 ∈⇩∘ Vset α"
unfolding A0B0_def
proof(rule 𝔊.HomDom.cat_Hom_vifunion_in_Vset; (intro vsubsetI)?)
show "A0 ∈⇩∘ Vset α"
unfolding A0_def
proof(intro vrange_vprojection_in_VsetI that(3))
fix F assume "F ∈⇩∘ A"
with that(1) have "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" by auto
with assms obtain a b f where F_def: "F = [a, b, f]⇩∘" by auto
show "vsv F" unfolding F_def by auto
show "0 ∈⇩∘ 𝒟⇩∘ F" unfolding F_def by simp
qed auto
show "B0 ∈⇩∘ Vset α"
unfolding B0_def
proof(intro vrange_vprojection_in_VsetI that(4))
fix F assume "F ∈⇩∘ B"
with that(2) have "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" by auto
with assms obtain a b f where F_def: "F = [a, b, f]⇩∘" by auto
show "vsv F" unfolding F_def by auto
show "0 ∈⇩∘ 𝒟⇩∘ F" unfolding F_def by simp
qed auto
next
fix a assume "a ∈⇩∘ A0"
with that(1) obtain F
where a_def: "a = F⦇0⦈" and "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
unfolding A0_def by force
with assms obtain b f where "F = [a, b, f]⇩∘" and "a ∈⇩∘ 𝔄⦇Obj⦈" by auto
then show "a ∈⇩∘ 𝔄⦇Obj⦈" unfolding a_def by simp
next
fix a assume "a ∈⇩∘ B0"
with that(2) obtain F
where a_def: "a = F⦇0⦈" and "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
unfolding B0_def by force
with assms obtain b f where "F = [a, b, f]⇩∘" and "a ∈⇩∘ 𝔄⦇Obj⦈" by auto
then show "a ∈⇩∘ 𝔄⦇Obj⦈" unfolding a_def by simp
qed
have A1B1: "A1B1 ∈⇩∘ Vset α"
unfolding A1B1_def
proof(rule 𝔉.HomDom.cat_Hom_vifunion_in_Vset; (intro vsubsetI)?)
show "A1 ∈⇩∘ Vset α"
unfolding A1_def
proof(intro vrange_vprojection_in_VsetI that(3))
fix F assume "F ∈⇩∘ A"
with that(1) have "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" by auto
with assms obtain a b f where F_def: "F = [a, b, f]⇩∘" by auto
show "vsv F" unfolding F_def by auto
show "1⇩ℕ ∈⇩∘ 𝒟⇩∘ F" unfolding F_def by (simp add: nat_omega_simps)
qed auto
show "B1 ∈⇩∘ Vset α"
unfolding B1_def
proof(intro vrange_vprojection_in_VsetI that(4))
fix F assume "F ∈⇩∘ B"
with that(2) have "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" by auto
with assms obtain a b f where F_def: "F = [a, b, f]⇩∘" by auto
show "vsv F" unfolding F_def by auto
show "1⇩ℕ ∈⇩∘ 𝒟⇩∘ F" unfolding F_def by (simp add: nat_omega_simps)
qed auto
next
fix b assume "b ∈⇩∘ A1"
with that(1) obtain F
where b_def: "b = F⦇1⇩ℕ⦈" and "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
unfolding A1_def by force
with assms obtain a f where "F = [a, b, f]⇩∘" and "b ∈⇩∘ 𝔅⦇Obj⦈"
by (auto simp: nat_omega_simps)
then show "b ∈⇩∘ 𝔅⦇Obj⦈" unfolding b_def by simp
next
fix b assume "b ∈⇩∘ B1"
with that(2) obtain F
where b_def: "b = F⦇1⇩ℕ⦈" and "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
unfolding B1_def by force
with assms obtain a f where "F = [a, b, f]⇩∘" and "b ∈⇩∘ 𝔅⦇Obj⦈"
by (auto simp: nat_omega_simps)
then show "b ∈⇩∘ 𝔅⦇Obj⦈" unfolding b_def by simp
qed
define Q where
"Q i = (if i = 0 then A else if i = 1⇩ℕ then B else (A0B0 ×⇩∙ A1B1))"
for i
have
"(⋃⇩∘a∈⇩∘A. ⋃⇩∘b∈⇩∘B.
Hom (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) a b) ⊆⇩∘ (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
proof
(
intro vsubsetI,
elim vifunionE,
unfold in_Hom_iff,
intro vproductI ballI
)
fix abf a'b'f' F assume prems:
"abf ∈⇩∘ A" "a'b'f' ∈⇩∘ B" "F : abf ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ a'b'f'"
from prems(3) assms obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and abf_def: "abf = [a, b, f]⇩∘"
and a'b'f'_def: "a'b'f' = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
have gh: "[g, h]⇩∘ ∈⇩∘ A0B0 ×⇩∙ A1B1"
unfolding A0B0_def A1B1_def
proof
(
intro ftimesI2 vifunionI,
unfold in_Hom_iff A0_def B0_def A1_def B1_def
)
from prems(1) show "a ∈⇩∘ ℛ⇩∘ (λF∈⇩∘A. F⦇0⦈)"
by (intro vsv.vsv_vimageI2'[where a=abf]) (simp_all add: abf_def)
from prems(2) show "a' ∈⇩∘ ℛ⇩∘ (λF∈⇩∘B. F⦇0⦈)"
by (intro vsv.vsv_vimageI2'[where a=a'b'f'])
(simp_all add: a'b'f'_def)
from prems(1) show "b ∈⇩∘ ℛ⇩∘ (λF∈⇩∘A. F⦇1⇩ℕ⦈)"
by (intro vsv.vsv_vimageI2'[where a=abf])
(simp_all add: nat_omega_simps abf_def)
from prems(2) show "b' ∈⇩∘ ℛ⇩∘ (λF∈⇩∘B. F⦇1⇩ℕ⦈)"
by (intro vsv.vsv_vimageI2'[where a=a'b'f'])
(simp_all add: nat_omega_simps a'b'f'_def)
qed (intro g h)+
show "vsv F" unfolding F_def by auto
show "𝒟⇩∘ F = set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: F_def three nat_omega_simps)
fix i assume "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}"
then consider ‹i = 0› | ‹i = 1⇩ℕ› | ‹i = 2⇩ℕ› by auto
from this prems show "F⦇i⦈ ∈⇩∘ Q i"
by cases
(simp_all add: F_def Q_def gh abf_def a'b'f'_def nat_omega_simps)
qed
moreover have "(∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i) ∈⇩∘ Vset α"
proof(rule Limit_vproduct_in_VsetI)
show "set {0, 1⇩ℕ, 2⇩ℕ} ∈⇩∘ Vset α"
by (cs_concl cs_shallow cs_intro: V_cs_intros)
from A0B0 A1B1 assms(1,2) that(3,4) show
"Q i ∈⇩∘ Vset α" if "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}"
for i
by (simp_all add: Q_def Limit_ftimes_in_VsetI nat_omega_simps)
qed auto
ultimately show "(⋃⇩∘a∈⇩∘A. ⋃⇩∘b∈⇩∘B. Hom (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) a b) ∈⇩∘ Vset α" by auto
qed
qed (auto simp: cat_comma_cs_simps intro: cat_comma_cs_intros)
qed
subsubsection‹Tiny comma category›
lemma tiny_category_cat_comma[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
shows "tiny_category α (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)"
proof-
interpret 𝔊: is_tm_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_tm_functor α 𝔅 ℭ ℌ by (rule assms(2))
note 𝔊 = 𝔊.is_functor_axioms
and ℌ = ℌ.is_functor_axioms
interpret category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
show ?thesis
proof(intro tiny_categoryI' category_cat_comma)
have vrange_𝔊: "ℛ⇩∘ (𝔊⦇ObjMap⦈) ∈⇩∘ Vset α"
by (simp add: vrange_in_VsetI 𝔊.tm_cf_ObjMap_in_Vset)
moreover have vrange_ℌ: "ℛ⇩∘ (ℌ⦇ObjMap⦈) ∈⇩∘ Vset α"
by (simp add: vrange_in_VsetI ℌ.tm_cf_ObjMap_in_Vset)
ultimately have UU_Hom_in_Vset:
"(⋃⇩∘a∈⇩∘ℛ⇩∘ (𝔊⦇ObjMap⦈). ⋃⇩∘b∈⇩∘ℛ⇩∘ (ℌ⦇ObjMap⦈). Hom ℭ a b) ∈⇩∘ Vset α"
using 𝔊.cf_ObjMap_vrange ℌ.cf_ObjMap_vrange
by (auto intro: 𝔊.HomCod.cat_Hom_vifunion_in_Vset)
define Q where
"Q i =
(
if i = 0
then 𝔄⦇Obj⦈
else
if i = 1⇩ℕ
then 𝔅⦇Obj⦈
else (⋃⇩∘a∈⇩∘ℛ⇩∘ (𝔊⦇ObjMap⦈). ⋃⇩∘b∈⇩∘ℛ⇩∘ (ℌ⦇ObjMap⦈). Hom ℭ a b)
)"
for i
have "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈ ⊆⇩∘ (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
proof(intro vsubsetI)
fix A assume "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
then obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by (elim cat_comma_ObjE[OF _ 𝔊 ℌ])
from f have f:
"f ∈⇩∘ (⋃⇩∘a∈⇩∘ℛ⇩∘ (𝔊⦇ObjMap⦈). ⋃⇩∘b∈⇩∘ℛ⇩∘ (ℌ⦇ObjMap⦈). Hom ℭ a b)"
by (intro vifunionI, unfold in_Hom_iff)
(
simp_all add:
a b
ℌ.ObjMap.vsv_vimageI2
ℌ.cf_ObjMap_vdomain
𝔊.ObjMap.vsv_vimageI2
𝔊.cf_ObjMap_vdomain
)
show "A ∈⇩∘ (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
proof(intro vproductI, unfold Ball_def; (intro allI impI)?)
show "𝒟⇩∘ A = set {0, 1⇩ℕ, 2⇩ℕ}"
unfolding A_def by (simp add: three nat_omega_simps)
fix i assume "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}"
then consider ‹i = 0› | ‹i = 1⇩ℕ› | ‹i = 2⇩ℕ› by auto
from this a b f show "A⦇i⦈ ∈⇩∘ Q i"
unfolding A_def Q_def by cases (simp_all add: nat_omega_simps)
qed (auto simp: A_def)
qed
moreover have "(∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i) ∈⇩∘ Vset α"
proof(rule Limit_vproduct_in_VsetI)
show "set {0, 1⇩ℕ, 2⇩ℕ} ∈⇩∘ Vset α"
unfolding three[symmetric] by simp
from this show "Q i ∈⇩∘ Vset α" if "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}" for i
using that assms(1,2) UU_Hom_in_Vset
by
(
simp_all add:
Q_def
𝔊.HomDom.tiny_cat_Obj_in_Vset
ℌ.HomDom.tiny_cat_Obj_in_Vset
nat_omega_simps
)
qed auto
ultimately show [simp]: "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈ ∈⇩∘ Vset α" by auto
define Q where
"Q i =
(
if i = 0
then 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈
else
if i = 1⇩ℕ
then 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈
else 𝔄⦇Arr⦈ ×⇩∙ 𝔅⦇Arr⦈
)"
for i
have "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈ ⊆⇩∘ (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
proof(intro vsubsetI)
fix F assume "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain abf a'b'f' where F: "F : abf ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ a'b'f'"
by (auto intro: is_arrI)
with assms obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and abf_def: "abf = [a, b, f]⇩∘"
and a'b'f'_def: "a'b'f' = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
from g h have "[g, h]⇩∘ ∈⇩∘ 𝔄⦇Arr⦈ ×⇩∙ 𝔅⦇Arr⦈" by auto
show "F ∈⇩∘ (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
proof(intro vproductI, unfold Ball_def; (intro allI impI)?)
show "𝒟⇩∘ F = set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: F_def three nat_omega_simps)
fix i assume "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}"
then consider ‹i = 0› | ‹i = 1⇩ℕ› | ‹i = 2⇩ℕ› by auto
from this F g h show "F⦇i⦈ ∈⇩∘ Q i"
unfolding Q_def F_def abf_def[symmetric] a'b'f'_def[symmetric]
by cases (auto simp: nat_omega_simps)
qed (auto simp: F_def)
qed
moreover have "(∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i) ∈⇩∘ Vset α"
proof(rule Limit_vproduct_in_VsetI)
show "set {0, 1⇩ℕ, 2⇩ℕ} ∈⇩∘ Vset α"
by (simp add: three[symmetric] nat_omega_simps)
moreover have "𝔄⦇Arr⦈ ×⇩∙ 𝔅⦇Arr⦈ ∈⇩∘ Vset α"
by
(
auto intro!:
Limit_ftimes_in_VsetI
𝔊.𝒵_β 𝒵_def
𝔊.HomDom.tiny_cat_Arr_in_Vset
ℌ.HomDom.tiny_cat_Arr_in_Vset
)
ultimately show "Q i ∈⇩∘ Vset α" if "i ∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}" for i
using that assms(1,2) UU_Hom_in_Vset
by
(
simp_all add:
Q_def
𝔊.HomDom.tiny_cat_Obj_in_Vset
ℌ.HomDom.tiny_cat_Obj_in_Vset
nat_omega_simps
)
qed auto
ultimately show "𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈ ∈⇩∘ Vset α" by auto
qed (rule 𝔊, rule ℌ)
qed
subsection‹Opposite comma category functor›
subsubsection‹Background›
text‹
See \<^cite>‹"noauthor_wikipedia_2001"›\footnote{
\url{https://en.wikipedia.org/wiki/Opposite_category}
} for background information.
›
subsubsection‹Object flip›
definition op_cf_commma_obj_flip :: "V ⇒ V ⇒ V"
where "op_cf_commma_obj_flip 𝔊 ℌ =
(λA∈⇩∘(𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈. [A⦇1⇩ℕ⦈, A⦇0⦈, A⦇2⇩ℕ⦈]⇩∘)"
text‹Elementary properties.›
mk_VLambda op_cf_commma_obj_flip_def
|vsv op_cf_commma_obj_flip_vsv[cat_comma_cs_intros]|
|vdomain op_cf_commma_obj_flip_vdomain[cat_comma_cs_simps]|
|app op_cf_commma_obj_flip_app'|
lemma op_cf_commma_obj_flip_app[cat_comma_cs_simps]:
assumes "A = [a, b, f]⇩∘" and "A ∈⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈"
shows "op_cf_commma_obj_flip 𝔊 ℌ⦇A⦈ = [b, a, f]⇩∘"
using assms unfolding op_cf_commma_obj_flip_def by (simp add: nat_omega_simps)
lemma op_cf_commma_obj_flip_v11[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "v11 (op_cf_commma_obj_flip 𝔊 ℌ)"
proof(rule vsv.vsv_valeq_v11I, unfold op_cf_commma_obj_flip_vdomain)
fix A B assume prems:
"A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
"B ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
"op_cf_commma_obj_flip 𝔊 ℌ⦇A⦈ = op_cf_commma_obj_flip 𝔊 ℌ⦇B⦈"
from prems(1,2) assms obtain a b f a' b' f'
where A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
by (elim cat_comma_ObjE[OF _ assms])
from prems(3,1,2) show "A = B"
by (simp_all add: A_def B_def op_cf_commma_obj_flip_app nat_omega_simps)
qed (auto intro: op_cf_commma_obj_flip_vsv)
lemma op_cf_commma_obj_flip_vrange[cat_comma_cs_simps]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (op_cf_commma_obj_flip 𝔊 ℌ) = (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Obj⦈"
proof(intro vsubset_antisym)
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
show "ℛ⇩∘ (op_cf_commma_obj_flip 𝔊 ℌ) ⊆⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Obj⦈"
proof
(
intro vsv.vsv_vrange_vsubset op_cf_commma_obj_flip_vsv,
unfold cat_comma_cs_simps
)
fix A assume "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
then obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by (elim cat_comma_ObjE[OF _ assms])
from a b f show
"op_cf_commma_obj_flip 𝔊 ℌ⦇A⦈ ∈⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Obj⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
show "(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Obj⦈ ⊆⇩∘ ℛ⇩∘ (op_cf_commma_obj_flip 𝔊 ℌ)"
proof(intro vsubsetI)
fix B assume "B ∈⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Obj⦈"
then obtain a b f
where B_def: "B = [b, a, f]⇩∘"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by
(
elim cat_comma_ObjE[
OF _ ℌ.is_functor_op 𝔊.is_functor_op, unfolded cat_op_simps
]
)
from a b f have B_def: "B = op_cf_commma_obj_flip 𝔊 ℌ⦇a, b, f⦈⇩∙"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps B_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)
from a b f have "[a, b, f]⇩∘ ∈⇩∘ 𝒟⇩∘ (op_cf_commma_obj_flip 𝔊 ℌ)"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
with op_cf_commma_obj_flip_vsv show "B ∈⇩∘ ℛ⇩∘ (op_cf_commma_obj_flip 𝔊 ℌ)"
unfolding B_def by auto
qed
qed
subsubsection‹Definition and elementary properties›
definition op_cf_comma :: "V ⇒ V ⇒ V"
where "op_cf_comma 𝔊 ℌ =
[
op_cf_commma_obj_flip 𝔊 ℌ,
(
λABF∈⇩∘(𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Arr⦈.
[
op_cf_commma_obj_flip 𝔊 ℌ⦇ABF⦇1⇩ℕ⦈⦈,
op_cf_commma_obj_flip 𝔊 ℌ⦇ABF⦇0⦈⦈,
[ABF⦇2⇩ℕ⦈⦇1⇩ℕ⦈, ABF⦇2⇩ℕ⦈⦇0⇩ℕ⦈]⇩∘
]⇩∘
),
op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ),
(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)
]⇩∘"
text‹Components.›
lemma op_cf_comma_components:
shows [cat_comma_cs_simps]:
"op_cf_comma 𝔊 ℌ⦇ObjMap⦈ = op_cf_commma_obj_flip 𝔊 ℌ"
and "op_cf_comma 𝔊 ℌ⦇ArrMap⦈ =
(
λABF∈⇩∘(𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Arr⦈.
[
op_cf_commma_obj_flip 𝔊 ℌ⦇ABF⦇1⇩ℕ⦈⦈,
op_cf_commma_obj_flip 𝔊 ℌ⦇ABF⦇0⦈⦈,
[ABF⦇2⇩ℕ⦈⦇1⇩ℕ⦈, ABF⦇2⇩ℕ⦈⦇0⇩ℕ⦈]⇩∘
]⇩∘
)"
and [cat_comma_cs_simps]:
"op_cf_comma 𝔊 ℌ⦇HomDom⦈ = op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)"
and [cat_comma_cs_simps]:
"op_cf_comma 𝔊 ℌ⦇HomCod⦈ = (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)"
unfolding op_cf_comma_def dghm_field_simps by (simp_all add: nat_omega_simps)
subsubsection‹Arrow map›
mk_VLambda op_cf_comma_components(2)
|vsv op_cf_comma_ArrMap_vsv[cat_comma_cs_intros]|
|vdomain op_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]|
|app op_cf_comma_ArrMap_app'|
lemma op_cf_comma_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
shows "op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ =
[
op_cf_commma_obj_flip 𝔊 ℌ⦇a', b', f'⦈⇩∙,
op_cf_commma_obj_flip 𝔊 ℌ⦇a, b, f⦈⇩∙,
[h, g]⇩∘
]⇩∘"
using assms(2) by (simp add: assms(1) op_cf_comma_ArrMap_app' nat_omega_simps)
lemma op_cf_comma_ArrMap_v11[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "v11 (op_cf_comma 𝔊 ℌ⦇ArrMap⦈)"
proof
(
rule vsv.vsv_valeq_v11I,
unfold op_cf_comma_ArrMap_vdomain,
intro op_cf_comma_ArrMap_vsv
)
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
fix ABF ABF' assume prems:
"ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
"ABF' ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
"op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ = op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF'⦈"
from prems(1) obtain A B where ABF: "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
from prems(2) obtain A' B' where ABF': "ABF' : A' ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B'" by auto
from ABF obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "h : b ↦⇘𝔅⇙ b'"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (elim cat_comma_is_arrE[OF _ assms])
from ABF' obtain a'' b'' f'' a''' b''' f''' g' h'
where ABF'_def: "ABF' = [[a'', b'', f'']⇩∘, [a''', b''', f''']⇩∘, [g', h']⇩∘]⇩∘"
and A'_def: "A' = [a'', b'', f'']⇩∘"
and B'_def: "B' = [a''', b''', f''']⇩∘"
and "g' : a'' ↦⇘𝔄⇙ a'''"
and "h' : b'' ↦⇘𝔅⇙ b'''"
and "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and "f''' : 𝔊⦇ObjMap⦈⦇a'''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'''⦈"
and "f''' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f''"
by (elim cat_comma_is_arrE[OF _ assms])
from ABF ABF' have abf:
"[a, b, f]⇩∘ ∈⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈"
"[a', b', f']⇩∘ ∈⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈"
"[a'', b'', f'']⇩∘ ∈⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈"
"[a''', b''', f''']⇩∘ ∈⇩∘ (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈"
unfolding ABF_def ABF'_def A_def B_def A'_def B'_def by auto
note v11_injective = v11.v11_injective[
OF op_cf_commma_obj_flip_v11, OF assms, unfolded cat_comma_cs_simps
]
from
prems(3,1,2) assms
op_cf_commma_obj_flip_v11
v11_injective[OF abf(1,3)]
v11_injective[OF abf(2,4)]
show "ABF = ABF'"
by
(
simp_all add:
ABF_def ABF'_def op_cf_comma_ArrMap_app' nat_omega_simps
)
qed
lemma op_cf_comma_ArrMap_is_arr:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
shows "op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ :
op_cf_commma_obj_flip 𝔊 ℌ⦇B⦈ ↦⇘(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⇙
op_cf_commma_obj_flip 𝔊 ℌ⦇A⦈"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms(3) obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and f'g_hf: "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by (elim cat_comma_is_arrE[OF _ assms(1,2)])
from g h f f' f'g_hf show ?thesis
unfolding ABF_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
lemma op_cf_comma_ArrMap_is_arr':
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B"
and "A' = op_cf_commma_obj_flip 𝔊 ℌ⦇B⦈"
and "B' = op_cf_commma_obj_flip 𝔊 ℌ⦇A⦈"
shows "op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ : A' ↦⇘(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⇙ B'"
using assms(1-3) unfolding assms(4,5) by (intro op_cf_comma_ArrMap_is_arr)
lemma op_cf_comma_ArrMap_vrange[cat_comma_cs_simps]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (op_cf_comma 𝔊 ℌ⦇ArrMap⦈) = (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Arr⦈"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
interpret op_𝔊ℌ: category α ‹(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)›
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros cat_op_intros)
show ?thesis
proof(intro vsubset_antisym)
show "ℛ⇩∘ (op_cf_comma 𝔊 ℌ⦇ArrMap⦈) ⊆⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Arr⦈"
proof
(
intro vsv.vsv_vrange_vsubset op_cf_comma_ArrMap_vsv,
unfold cat_comma_cs_simps
)
fix ABF assume prems: "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain A B where ABF: "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
from op_cf_comma_ArrMap_is_arr[OF assms this] show
"op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ ∈⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Arr⦈"
by auto
qed
show "(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Arr⦈ ⊆⇩∘ ℛ⇩∘ (op_cf_comma 𝔊 ℌ⦇ArrMap⦈)"
proof(intro vsubsetI)
fix ABF assume prems: "ABF ∈⇩∘ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇Arr⦈"
then obtain A B where ABF: "ABF : A ↦⇘(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⇙ B"
by auto
then obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a' ↦⇘𝔅⇙ a"
and h: "h : b' ↦⇘𝔄⇙ b"
and f: "f : 𝔊⦇ObjMap⦈⦇b⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇a⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇b'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇a'⦈"
and f'g_hf: "f' ∘⇩A⇘op_cat ℭ⇙ ℌ⦇ArrMap⦈⦇g⦈ = 𝔊⦇ArrMap⦈⦇h⦈ ∘⇩A⇘op_cat ℭ⇙ f"
by
(
elim cat_comma_is_arrE[
OF _ ℌ.is_functor_op 𝔊.is_functor_op, unfolded cat_op_simps
]
)
from f'g_hf g h f f' have gf'_fh:
"ℌ⦇ArrMap⦈⦇g⦈ ∘⇩A⇘ℭ⇙ f' = f ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇h⦈"
by
(
cs_prems cs_shallow
cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros
)
with g h f f' have
"[[b', a', f']⇩∘, [b, a, f]⇩∘, [h, g]⇩∘]⇩∘ ∈⇩∘ 𝒟⇩∘ (op_cf_comma 𝔊 ℌ⦇ArrMap⦈)"
"ABF = op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇[b', a', f']⇩∘, [b, a, f]⇩∘, [h, g]⇩∘⦈⇩∙"
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps ABF_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
with op_cf_comma_ArrMap_vsv show "ABF ∈⇩∘ ℛ⇩∘ (op_cf_comma 𝔊 ℌ⦇ArrMap⦈)"
by auto
qed
qed
qed
subsubsection‹Opposite comma category functor is an isomorphism of categories›
lemma op_cf_comma_is_iso_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "op_cf_comma 𝔊 ℌ :
op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
show ?thesis
proof(intro is_iso_functorI' is_functorI')
show "vfsequence (op_cf_comma 𝔊 ℌ)"
unfolding op_cf_comma_def by simp
show "vcard (op_cf_comma 𝔊 ℌ) = 4⇩ℕ"
unfolding op_cf_comma_def by (simp add: nat_omega_simps)
from assms show "op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇ABF⦈ :
op_cf_comma 𝔊 ℌ⦇ObjMap⦈⦇A⦈ ↦⇘(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⇙
op_cf_comma 𝔊 ℌ⦇ObjMap⦈⦇B⦈"
if "ABF : A ↦⇘op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⇙ B" for A B ABF
using that
unfolding cat_op_simps
by
(
cs_concl cs_shallow
cs_intro: op_cf_comma_ArrMap_is_arr' cs_simp: cat_comma_cs_simps
)
show
"op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇G ∘⇩A⇘op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⇙ F⦈ =
op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇G⦈ ∘⇩A⇘(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⇙
op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇F⦈"
if "G : B ↦⇘op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⇙ C"
and "F : A ↦⇘op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⇙ B"
for B C G A F
proof-
note G = that(1)[unfolded cat_op_simps]
note F = that(2)[unfolded cat_op_simps]
from assms G obtain a b f a' b' f' g h
where G_def: "G = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and C_def: "C = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [symmetric, cat_comma_cs_simps]:
"f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with assms F obtain a'' b'' f'' g' h'
where F_def: "F = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, [g', h']⇩∘]⇩∘"
and A_def: "A = [a'', b'', f'']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_comma_cs_simps]:
"f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
note [cat_comma_cs_simps] =
category.cat_assoc_helper[
where ℭ=ℭ and h=f'' and g=‹𝔊⦇ArrMap⦈⦇g'⦈› and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
from assms that g h f f' g' h' f' f'' show ?thesis
unfolding cat_op_simps G_def C_def B_def F_def A_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
show
"op_cf_comma 𝔊 ℌ⦇ArrMap⦈⦇op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇CId⦈⦇C⦈⦈ =
(op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)⦇CId⦈⦇op_cf_comma 𝔊 ℌ⦇ObjMap⦈⦇C⦈⦈"
if "C ∈⇩∘ op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)⦇Obj⦈" for C
proof-
from that[unfolded cat_op_simps] assms obtain a b f
where C_def: "C = [a, b, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by auto
from a b f that show ?thesis
unfolding cat_op_simps C_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
qed
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_op_intros
)+
qed
lemma op_cf_comma_is_iso_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔄' = op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)"
and "𝔅' = (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)"
shows "op_cf_comma 𝔊 ℌ : 𝔄' ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ 𝔅'"
using assms(1,2) unfolding assms(3,4) by (rule op_cf_comma_is_iso_functor)
lemma op_cf_comma_is_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "op_cf_comma 𝔊 ℌ :
op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇘α⇙ (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)"
by (rule is_iso_functorD(1)[OF op_cf_comma_is_iso_functor[OF assms]])
lemma op_cf_comma_is_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔄' = op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ)"
and "𝔅' = (op_cf ℌ) ⇩C⇩F↓⇩C⇩F (op_cf 𝔊)"
shows "op_cf_comma 𝔊 ℌ : 𝔄' ↦↦⇩C⇘α⇙ 𝔅'"
using assms(1,2) unfolding assms(3,4) by (rule op_cf_comma_is_functor)
subsection‹Projections for a comma category›
subsubsection‹Definitions and elementary properties›
text‹See Chapter II-6 in \<^cite>‹"mac_lane_categories_2010"›.›
definition cf_comma_proj_left :: "V ⇒ V ⇒ V" (‹(_ ⇩C⇩F⨅ _)› [1000, 1000] 999)
where "𝔊 ⇩C⇩F⨅ ℌ =
[
(λa∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈. a⦇0⦈),
(λf∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈. f⦇2⇩ℕ⦈⦇0⦈),
𝔊 ⇩C⇩F↓⇩C⇩F ℌ,
𝔊⦇HomDom⦈
]⇩∘"
definition cf_comma_proj_right :: "V ⇒ V ⇒ V" (‹(_ ⨅⇩C⇩F _)› [1000, 1000] 999)
where "𝔊 ⨅⇩C⇩F ℌ =
[
(λa∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈. a⦇1⇩ℕ⦈),
(λf∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈. f⦇2⇩ℕ⦈⦇1⇩ℕ⦈),
𝔊 ⇩C⇩F↓⇩C⇩F ℌ,
ℌ⦇HomDom⦈
]⇩∘"
text‹Components.›
lemma cf_comma_proj_left_components:
shows "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ = (λa∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈. a⦇0⦈)"
and "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ = (λf∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈. f⦇2⇩ℕ⦈⦇0⦈)"
and "𝔊 ⇩C⇩F⨅ ℌ⦇HomDom⦈ = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
and "𝔊 ⇩C⇩F⨅ ℌ⦇HomCod⦈ = 𝔊⦇HomDom⦈"
unfolding cf_comma_proj_left_def dghm_field_simps
by (simp_all add: nat_omega_simps)
lemma cf_comma_proj_right_components:
shows "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ = (λa∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈. a⦇1⇩ℕ⦈)"
and "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ = (λf∈⇩∘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈. f⦇2⇩ℕ⦈⦇1⇩ℕ⦈)"
and "𝔊 ⨅⇩C⇩F ℌ⦇HomDom⦈ = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
and "𝔊 ⨅⇩C⇩F ℌ⦇HomCod⦈ = ℌ⦇HomDom⦈"
unfolding cf_comma_proj_right_def dghm_field_simps
by (simp_all add: nat_omega_simps)
context
fixes α 𝔄 𝔅 ℭ 𝔊 ℌ
assumes 𝔊: "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and ℌ: "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
begin
interpretation 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule 𝔊)
interpretation ℌ: is_functor α 𝔅 ℭ ℌ by (rule ℌ)
lemmas cf_comma_proj_left_components' =
cf_comma_proj_left_components[of 𝔊 ℌ, unfolded 𝔊.cf_HomDom]
lemmas cf_comma_proj_right_components' =
cf_comma_proj_right_components[of 𝔊 ℌ, unfolded ℌ.cf_HomDom]
lemmas [cat_comma_cs_simps] =
cf_comma_proj_left_components'(3,4)
cf_comma_proj_right_components'(3,4)
end
subsubsection‹Object map›
mk_VLambda cf_comma_proj_left_components(1)
|vsv cf_comma_proj_left_ObjMap_vsv[cat_comma_cs_intros]|
|vdomain cf_comma_proj_left_ObjMap_vdomain[cat_comma_cs_simps]|
mk_VLambda cf_comma_proj_right_components(1)
|vsv cf_comma_proj_right_ObjMap_vsv[cat_comma_cs_intros]|
|vdomain cf_comma_proj_right_ObjMap_vdomain[cat_comma_cs_simps]|
lemma cf_comma_proj_left_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a, b, f]⇩∘" and "[a, b, f]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
shows "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇A⦈ = a"
using assms(2) unfolding assms(1) cf_comma_proj_left_components by simp
lemma cf_comma_proj_right_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a, b, f]⇩∘" and "[a, b, f]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
shows "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇A⦈ = b"
using assms(2)
unfolding assms(1) cf_comma_proj_right_components
by (simp add: nat_omega_simps)
lemma cf_comma_proj_left_ObjMap_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈) ⊆⇩∘ 𝔄⦇Obj⦈"
proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps)
fix A assume prems: "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
with assms obtain a b f where A_def: "A = [a, b, f]⇩∘" and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
by auto
from assms prems a show "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇A⦈ ∈⇩∘ 𝔄⦇Obj⦈"
unfolding A_def by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
qed (auto intro: cat_comma_cs_intros)
lemma cf_comma_proj_right_ObjMap_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈) ⊆⇩∘ 𝔅⦇Obj⦈"
proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps)
fix A assume prems: "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
with assms obtain a b f where A_def: "A = [a, b, f]⇩∘" and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
by auto
from assms prems b show "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇A⦈ ∈⇩∘ 𝔅⦇Obj⦈"
unfolding A_def by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
qed (auto intro: cat_comma_cs_intros)
subsubsection‹Arrow map›
mk_VLambda cf_comma_proj_left_components(2)
|vsv cf_comma_proj_left_ArrMap_vsv[cat_comma_cs_intros]|
|vdomain cf_comma_proj_left_ArrMap_vdomain[cat_comma_cs_simps]|
mk_VLambda cf_comma_proj_right_components(2)
|vsv cf_comma_proj_right_ArrMap_vsv[cat_comma_cs_intros]|
|vdomain cf_comma_proj_right_ArrMap_vdomain[cat_comma_cs_simps]|
lemma cf_comma_proj_left_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, [g, h]⇩∘]⇩∘" and "[A, B, [g, h]⇩∘]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
shows "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇ABF⦈ = g"
using assms(2)
unfolding assms(1) cf_comma_proj_left_components
by (simp add: nat_omega_simps)
lemma cf_comma_proj_right_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "[A, B, [g, h]⇩∘]⇩∘ ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
shows "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇ABF⦈ = h"
using assms(2)
unfolding assms(1) cf_comma_proj_right_components
by (simp add: nat_omega_simps)
lemma cf_comma_proj_left_ArrMap_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈) ⊆⇩∘ 𝔄⦇Arr⦈"
proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps)
from assms interpret category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
fix F assume prems: "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain A B where "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
by auto
from assms prems g show "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇F⦈ ∈⇩∘ 𝔄⦇Arr⦈"
unfolding F_def
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
qed (auto intro: cat_comma_cs_intros)
lemma cf_comma_proj_right_ArrMap_vrange:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈) ⊆⇩∘ 𝔅⦇Arr⦈"
proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps)
from assms interpret category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
fix F assume prems: "F ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and h: "h : b ↦⇘𝔅⇙ b'"
by auto
from assms prems h show "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇F⦈ ∈⇩∘ 𝔅⦇Arr⦈"
unfolding F_def
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
qed (auto intro: cat_comma_cs_intros)
subsubsection‹Projections for a comma category are functors›
lemma cf_comma_proj_left_is_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "𝔊 ⇩C⇩F⨅ ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇘α⇙ 𝔄"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
from assms interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
show ?thesis
proof(rule is_functorI')
show "vfsequence (𝔊 ⇩C⇩F⨅ ℌ)"
unfolding cf_comma_proj_left_def by auto
show "vcard (𝔊 ⇩C⇩F⨅ ℌ) = 4⇩ℕ"
unfolding cf_comma_proj_left_def by (simp add: nat_omega_simps)
from assms show "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈) ⊆⇩∘ 𝔄⦇Obj⦈"
by (rule cf_comma_proj_left_ObjMap_vrange)
show "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇F⦈ : 𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇A⦈ ↦⇘𝔄⇙ 𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇B⦈"
if "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" for A B F
proof-
from assms that obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
by auto
from that g show
"𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇F⦈ : 𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇A⦈ ↦⇘𝔄⇙ 𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇B⦈"
unfolding F_def A_def B_def
by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
qed
show
"𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F⦈ =
𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇G⦈ ∘⇩A⇘𝔄⇙ 𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇F⦈"
if "G : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C" and "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" for B C G A F
proof-
from assms that(2) obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [cat_cs_simps]: "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with that(1) assms obtain a'' b'' f'' g' h'
where G_def: "G = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, [g', h']⇩∘]⇩∘"
and C_def: "C = [a'', b'', f'']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_cs_simps]: "f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
note [cat_cs_simps] =
category.cat_assoc_helper
[
where ℭ=ℭ
and h=f''
and g=‹𝔊⦇ArrMap⦈⦇g'⦈›
and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
category.cat_assoc_helper
[
where ℭ=ℭ
and h=f
and g=‹ℌ⦇ArrMap⦈⦇h⦈›
and q=‹f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈›
]
from assms that g g' h h' f f' f'' show ?thesis
unfolding F_def G_def A_def B_def C_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_comma_cs_intros cat_cs_intros
)
qed
show "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇A⦈⦈ = 𝔄⦇CId⦈⦇𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇A⦈⦈"
if "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" for A
proof-
from assms that obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by auto
from assms that this(2-4) show ?thesis
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_comma_cs_intros cat_cs_intros
)
qed
qed
(
use assms in
‹
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
›
)+
qed
lemma cf_comma_proj_left_is_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔄' = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⇩C⇩F⨅ ℌ : 𝔄' ↦↦⇩C⇘α⇙ 𝔄"
using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_left_is_functor)
lemma cf_comma_proj_right_is_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "𝔊 ⨅⇩C⇩F ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇘α⇙ 𝔅"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
from assms interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
show ?thesis
proof(rule is_functorI')
show "vfsequence (𝔊 ⨅⇩C⇩F ℌ)"
unfolding cf_comma_proj_right_def by auto
show "vcard (𝔊 ⨅⇩C⇩F ℌ) = 4⇩ℕ"
unfolding cf_comma_proj_right_def by (simp add: nat_omega_simps)
from assms show "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈) ⊆⇩∘ 𝔅⦇Obj⦈"
by (rule cf_comma_proj_right_ObjMap_vrange)
show "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇F⦈ : 𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇A⦈ ↦⇘𝔅⇙ 𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇B⦈"
if "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" for A B F
proof-
from assms that obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and h: "h : b ↦⇘𝔅⇙ b'"
by auto
from that h show
"𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇F⦈ : 𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇A⦈ ↦⇘𝔅⇙ 𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇B⦈"
unfolding F_def A_def B_def
by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
qed
show
"𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇G ∘⇩A⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ F⦈ =
𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇G⦈ ∘⇩A⇘𝔅⇙ 𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇F⦈"
if "G : B ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ C" and "F : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" for B C G A F
proof-
from assms that(2) obtain a b f a' b' f' g h
where F_def: "F = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [cat_cs_simps]: "f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
with that(1) assms obtain a'' b'' f'' g' h'
where G_def: "G = [[a', b', f']⇩∘, [a'', b'', f'']⇩∘, [g', h']⇩∘]⇩∘"
and C_def: "C = [a'', b'', f'']⇩∘"
and g': "g' : a' ↦⇘𝔄⇙ a''"
and h': "h' : b' ↦⇘𝔅⇙ b''"
and f'': "f'' : 𝔊⦇ObjMap⦈⦇a''⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b''⦈"
and [cat_cs_simps]: "f'' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g'⦈ = ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'"
by auto
note [cat_cs_simps] =
category.cat_assoc_helper
[
where ℭ=ℭ
and h=f''
and g=‹𝔊⦇ArrMap⦈⦇g'⦈›
and q=‹ℌ⦇ArrMap⦈⦇h'⦈ ∘⇩A⇘ℭ⇙ f'›
]
category.cat_assoc_helper
[
where ℭ=ℭ
and h=f
and g=‹ℌ⦇ArrMap⦈⦇h⦈›
and q=‹f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈›
]
from assms that g g' h h' f f' f'' show ?thesis
unfolding F_def G_def A_def B_def C_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_comma_cs_intros cat_cs_intros
)
qed
show "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇CId⦈⦇A⦈⦈ = 𝔅⦇CId⦈⦇𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈⦇A⦈⦈"
if "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈" for A
proof-
from assms that obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by auto
from assms that this(2-4) show ?thesis
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_comma_cs_intros cat_cs_intros
)
qed
qed
(
use assms in
‹
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
›
)+
qed
lemma cf_comma_proj_right_is_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔄' = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⨅⇩C⇩F ℌ : 𝔄' ↦↦⇩C⇘α⇙ 𝔅"
using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_right_is_functor)
subsubsection‹Opposite projections for a comma category›
lemma op_cf_comma_proj_left:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "op_cf (𝔊 ⇩C⇩F⨅ ℌ) = (op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
show "op_cf (𝔊 ⇩C⇩F⨅ ℌ) = (op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ"
proof(rule cf_eqI)
show "op_cf (𝔊 ⇩C⇩F⨅ ℌ) : op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇘α⇙ op_cat 𝔄"
by
(
cs_concl cs_shallow
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
then have ObjMap_dom_lhs: "𝒟⇩∘ (op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ObjMap⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and ArrMap_dom_lhs: "𝒟⇩∘ (op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ArrMap⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cat_op_simps)+
show "(op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ :
op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇘α⇙ op_cat 𝔄"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros)
then have ObjMap_dom_rhs:
"𝒟⇩∘ (((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈) =
𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and ArrMap_dom_rhs:
"𝒟⇩∘ (((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈) =
𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_concl cs_simp: cat_cs_simps cat_op_simps)+
show
"op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ObjMap⦈ =
((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
fix A assume "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
with assms obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by auto
from a b f show
"op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ObjMap⦈⦇A⦈ =
((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
cs_concl cs_shallow
cs_simp: cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)+
show
"op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ArrMap⦈ =
((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix ABF assume "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain A B where "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [symmetric, cat_cs_simps]:
"f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
from g h f f' show "op_cf (𝔊 ⇩C⇩F⨅ ℌ)⦇ArrMap⦈⦇ABF⦈ =
((op_cf ℌ) ⨅⇩C⇩F (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈⦇ABF⦈"
unfolding ABF_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
cs_concl cs_shallow
cs_simp: cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)+
qed simp_all
qed
lemma op_cf_comma_proj_right:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
shows "op_cf (𝔊 ⨅⇩C⇩F ℌ) = (op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔊ℌ: category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
show "op_cf (𝔊 ⨅⇩C⇩F ℌ) = (op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ"
proof(rule cf_eqI)
show "op_cf (𝔊 ⨅⇩C⇩F ℌ) : op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇘α⇙ op_cat 𝔅"
by
(
cs_concl cs_shallow
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
then have ObjMap_dom_lhs: "𝒟⇩∘ (op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ObjMap⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and ArrMap_dom_lhs: "𝒟⇩∘ (op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ArrMap⦈) = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cat_op_simps)+
show "(op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ :
op_cat (𝔊 ⇩C⇩F↓⇩C⇩F ℌ) ↦↦⇩C⇘α⇙ op_cat 𝔅"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros)
then have ObjMap_dom_rhs:
"𝒟⇩∘ (((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈) =
𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
and ArrMap_dom_rhs:
"𝒟⇩∘ (((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈) =
𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by (cs_concl cs_simp: cat_cs_simps cat_op_simps)+
show
"op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ObjMap⦈ =
((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
fix A assume prems: "A ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
with assms obtain a b f
where A_def: "A = [a, b, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b: "b ∈⇩∘ 𝔅⦇Obj⦈"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
by auto
from a b f show
"op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ObjMap⦈⦇A⦈ =
((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
cs_concl cs_shallow
cs_simp: cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)+
show
"op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ArrMap⦈ =
((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix ABF assume prems: "ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
then obtain A B where ABF: "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b, f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b ↦⇘𝔅⇙ b'"
and f: "f : 𝔊⦇ObjMap⦈⦇a⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b⦈"
and f': "f' : 𝔊⦇ObjMap⦈⦇a'⦈ ↦⇘ℭ⇙ ℌ⦇ObjMap⦈⦇b'⦈"
and [symmetric, cat_cs_simps]:
"f' ∘⇩A⇘ℭ⇙ 𝔊⦇ArrMap⦈⦇g⦈ = ℌ⦇ArrMap⦈⦇h⦈ ∘⇩A⇘ℭ⇙ f"
by auto
from g h f f' show "op_cf (𝔊 ⨅⇩C⇩F ℌ)⦇ArrMap⦈⦇ABF⦈ =
((op_cf ℌ) ⇩C⇩F⨅ (op_cf 𝔊) ∘⇩C⇩F op_cf_comma 𝔊 ℌ)⦇ArrMap⦈⦇ABF⦈"
unfolding ABF_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
cs_concl cs_shallow
cs_simp: cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)+
qed simp_all
qed
subsubsection‹Projections for a tiny comma category›
lemma cf_comma_proj_left_is_tm_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
shows "𝔊 ⇩C⇩F⨅ ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
proof(intro is_tm_functorI')
interpret 𝔊: is_tm_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_tm_functor α 𝔅 ℭ ℌ by (rule assms(2))
show Π_𝔊ℌ: "𝔊 ⇩C⇩F⨅ ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
interpret Π_𝔊ℌ: is_functor α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔄 ‹𝔊 ⇩C⇩F⨅ ℌ›
by (rule Π_𝔊ℌ)
interpret 𝔊ℌ: tiny_category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_small_cs_intros cat_comma_cs_intros)
show "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ ∈⇩∘ Vset α"
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈) ∈⇩∘ Vset α"
proof-
note Π_𝔊ℌ.cf_ObjMap_vrange
moreover have "𝔄⦇Obj⦈ ∈⇩∘ Vset α" by (intro cat_small_cs_intros)
ultimately show ?thesis by auto
qed
qed (auto simp: cf_comma_proj_left_components intro: cat_small_cs_intros)
show "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ ∈⇩∘ Vset α"
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈) ∈⇩∘ Vset α"
proof-
note Π_𝔊ℌ.cf_ArrMap_vrange
moreover have "𝔄⦇Arr⦈ ∈⇩∘ Vset α" by (intro cat_small_cs_intros)
ultimately show ?thesis by auto
qed
qed (auto simp: cf_comma_proj_left_components intro: cat_small_cs_intros)
qed
lemma cf_comma_proj_left_is_tm_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
and "𝔊ℌ = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⇩C⇩F⨅ ℌ : 𝔊ℌ ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_left_is_tm_functor)
lemma cf_comma_proj_right_is_tm_functor:
assumes "𝔊 : 𝔄 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ" and "ℌ : 𝔅 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
shows "𝔊 ⨅⇩C⇩F ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
proof(intro is_tm_functorI')
interpret 𝔊: is_tm_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_tm_functor α 𝔅 ℭ ℌ by (rule assms(2))
show Π_𝔊ℌ: "𝔊 ⨅⇩C⇩F ℌ : 𝔊 ⇩C⇩F↓⇩C⇩F ℌ ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
interpret Π_𝔊ℌ: is_functor α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔅 ‹𝔊 ⨅⇩C⇩F ℌ›
by (rule Π_𝔊ℌ)
interpret 𝔊ℌ: tiny_category α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ›
by (cs_concl cs_shallow cs_intro: cat_small_cs_intros cat_comma_cs_intros)
show "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ ∈⇩∘ Vset α"
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈) ∈⇩∘ Vset α"
proof-
note Π_𝔊ℌ.cf_ObjMap_vrange
moreover have "𝔅⦇Obj⦈ ∈⇩∘ Vset α" by (intro cat_small_cs_intros)
ultimately show ?thesis by auto
qed
qed (auto simp: cf_comma_proj_right_components intro: cat_small_cs_intros)
show "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ ∈⇩∘ Vset α"
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈) ∈⇩∘ Vset α"
proof-
note Π_𝔊ℌ.cf_ArrMap_vrange
moreover have "𝔅⦇Arr⦈ ∈⇩∘ Vset α" by (intro cat_small_cs_intros)
ultimately show ?thesis by auto
qed
qed (auto simp: cf_comma_proj_right_components intro: cat_small_cs_intros)
qed
lemma cf_comma_proj_right_is_tm_functor'[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ ℭ"
and "𝔊ℌ = 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⨅⇩C⇩F ℌ : 𝔊ℌ ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_right_is_tm_functor)
lemma cf_comp_cf_comma_proj_left_is_tm_functor[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⇩C⇩F⨅ ℌ ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
proof(intro is_tm_functorI')
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔉: is_tm_functor α 𝔍 ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔉 by (rule assms(3))
interpret 𝔊ℌ: is_functor α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔄 ‹𝔊 ⇩C⇩F⨅ ℌ›
by (rule cf_comma_proj_left_is_functor[OF assms(1-2)])
show "𝔊 ⇩C⇩F⨅ ℌ ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
show "(𝔊 ⇩C⇩F⨅ ℌ ∘⇩C⇩F 𝔉)⦇ObjMap⦈ ∈⇩∘ Vset α"
unfolding dghm_comp_components
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "vbrelation (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈)" by auto
show "Limit α" by auto
show "𝒟⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈) ∈⇩∘ Vset α"
by
(
cs_concl
cs_simp: V_cs_simps cat_cs_simps
cs_intro: 𝔉.cf_ObjMap_vrange cat_small_cs_intros
)
show "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈) ∈⇩∘ Vset α"
unfolding vrange_vcomp
proof-
have "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈) ⊆⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈))))"
proof(intro vsubsetI)
fix A assume prems: "A ∈⇩∘ 𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
then obtain abf
where abf_in_𝔉: "abf ∈⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
and 𝔊ℌ_abf: "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈⦇abf⦈ = A"
by auto
with 𝔉.ObjMap.vrange_atD obtain j
where "j ∈⇩∘ 𝔍⦇Obj⦈" and 𝔉j: "𝔉⦇ObjMap⦈⦇j⦈ = abf"
by (force simp: 𝔉.cf_ObjMap_vdomain)
from abf_in_𝔉 𝔉.cf_ObjMap_vrange have abf_in_𝔊ℌ:
"abf ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
then obtain a b f where abf_def: "abf = [a, b, f]⇩∘"
by (elim cat_comma_ObjE[OF _ assms(1,2)])
have "a ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈))))"
proof(intro VUnionI)
from abf_in_𝔉 show "[a, b, f]⇩∘ ∈⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
unfolding abf_def by auto
show "⟨0, a⟩ ∈⇩∘ [a, b, f]⇩∘" by auto
show "set {0, a} ∈⇩∘ ⟨0, a⟩" unfolding vpair_def by simp
qed auto
with abf_in_𝔊ℌ show "A ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘(𝔉⦇ObjMap⦈))))"
unfolding 𝔊ℌ_abf[symmetric] abf_def
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
qed
moreover have "⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈)))) ∈⇩∘ Vset α"
by (intro VUnion_in_VsetI vrange_in_VsetI[OF 𝔉.tm_cf_ObjMap_in_Vset])
ultimately show "𝔊 ⇩C⇩F⨅ ℌ⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈) ∈⇩∘ Vset α" by auto
qed
qed
show "(𝔊 ⇩C⇩F⨅ ℌ ∘⇩C⇩F 𝔉)⦇ArrMap⦈ ∈⇩∘ Vset α"
unfolding dghm_comp_components
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "vbrelation (𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈)" by auto
show "Limit α" by auto
show "𝒟⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈) ∈⇩∘ Vset α"
by
(
cs_concl
cs_simp: V_cs_simps cat_cs_simps
cs_intro: 𝔉.cf_ArrMap_vrange cat_small_cs_intros
)
show "ℛ⇩∘ (𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈) ∈⇩∘ Vset α"
unfolding vrange_vcomp
proof-
have
"𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈) ⊆⇩∘
⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
proof(intro vsubsetI)
fix F assume prems: "F ∈⇩∘ 𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
then obtain ABF
where ABF_in_𝔉: "ABF ∈⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
and 𝔊ℌ_ABF: "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈⦇ABF⦈ = F"
by auto
with 𝔉.ArrMap.vrange_atD obtain k
where "k ∈⇩∘ 𝔍⦇Arr⦈" and 𝔉j: "𝔉⦇ArrMap⦈⦇k⦈ = ABF"
by (force simp: 𝔉.cf_ArrMap_vdomain)
then obtain i j where "k : i ↦⇘𝔍⇙ j" by auto
from ABF_in_𝔉 𝔉.cf_ArrMap_vrange have ABF_in_𝔊ℌ:
"ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by auto
then obtain A B where "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
by (elim cat_comma_is_arrE[OF _ assms(1,2)])
have "g ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
proof(intro VUnionI)
from ABF_in_𝔉 show
"[[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘ ∈⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
unfolding ABF_def by auto
show "⟨2⇩ℕ, [g, h]⇩∘⟩ ∈⇩∘ [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
by (auto simp: nat_omega_simps)
show "set {2⇩ℕ, [g, h]⇩∘} ∈⇩∘ ⟨2⇩ℕ, [g, h]⇩∘⟩"
unfolding vpair_def by auto
show "[g, h]⇩∘ ∈⇩∘ set {2⇩ℕ, [g, h]⇩∘}" by simp
show "⟨0, g⟩ ∈⇩∘ [g, h]⇩∘" by auto
show "set {0, g} ∈⇩∘ ⟨0, g⟩" unfolding vpair_def by auto
qed auto
with ABF_in_𝔊ℌ show "F ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
unfolding 𝔊ℌ_ABF[symmetric] ABF_def
by (cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps)
qed
moreover have "⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈))))))) ∈⇩∘ Vset α"
by (intro VUnion_in_VsetI vrange_in_VsetI[OF 𝔉.tm_cf_ArrMap_in_Vset])
ultimately show "𝔊 ⇩C⇩F⨅ ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈) ∈⇩∘ Vset α" by auto
qed
qed
qed
lemma cf_comp_cf_comma_proj_right_is_tm_functor[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ"
shows "𝔊 ⨅⇩C⇩F ℌ ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
proof(intro is_tm_functorI')
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(2))
interpret 𝔉: is_tm_functor α 𝔍 ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔉 by (rule assms(3))
interpret 𝔊ℌ: is_functor α ‹𝔊 ⇩C⇩F↓⇩C⇩F ℌ› 𝔅 ‹𝔊 ⨅⇩C⇩F ℌ›
by (rule cf_comma_proj_right_is_functor[OF assms(1-2)])
show "𝔊 ⨅⇩C⇩F ℌ ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
show "(𝔊 ⨅⇩C⇩F ℌ ∘⇩C⇩F 𝔉)⦇ObjMap⦈ ∈⇩∘ Vset α"
unfolding dghm_comp_components
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "vbrelation (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈)" by auto
show "Limit α" by auto
show "𝒟⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈) ∈⇩∘ Vset α"
by
(
cs_concl
cs_simp: V_cs_simps cat_cs_simps
cs_intro: 𝔉.cf_ObjMap_vrange cat_small_cs_intros
)
show "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ ∘⇩∘ 𝔉⦇ObjMap⦈) ∈⇩∘ Vset α"
unfolding vrange_vcomp
proof-
have "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈) ⊆⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈))))"
proof(intro vsubsetI)
fix A assume prems: "A ∈⇩∘ (𝔊 ⨅⇩C⇩F ℌ)⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
then obtain abf
where abf_in_𝔉: "abf ∈⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
and 𝔊ℌ_abf: "(𝔊 ⨅⇩C⇩F ℌ)⦇ObjMap⦈⦇abf⦈ = A"
by (auto simp: cf_comma_proj_right_ObjMap_vsv)
with 𝔉.ObjMap.vrange_atD obtain j
where "j ∈⇩∘ 𝔍⦇Obj⦈" and 𝔉j: "𝔉⦇ObjMap⦈⦇j⦈ = abf"
by (force simp: 𝔉.cf_ObjMap_vdomain)
from abf_in_𝔉 𝔉.cf_ObjMap_vrange have abf_in_𝔊ℌ:
"abf ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Obj⦈"
by auto
then obtain a b f where abf_def: "abf = [a, b, f]⇩∘"
by (elim cat_comma_ObjE[OF _ assms(1,2)])
have "b ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈))))"
proof(intro VUnionI)
from abf_in_𝔉 show "[a, b, f]⇩∘ ∈⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈)"
unfolding abf_def by auto
show "⟨1⇩ℕ, b⟩ ∈⇩∘ [a, b, f]⇩∘" by (auto simp: nat_omega_simps)
show "set {1⇩ℕ, b} ∈⇩∘ ⟨1⇩ℕ, b⟩" unfolding vpair_def by simp
qed auto
with abf_in_𝔊ℌ show "A ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘(𝔉⦇ObjMap⦈))))"
unfolding 𝔊ℌ_abf[symmetric] abf_def
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
qed
moreover have "⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ObjMap⦈)))) ∈⇩∘ Vset α"
by (intro VUnion_in_VsetI vrange_in_VsetI[OF 𝔉.tm_cf_ObjMap_in_Vset])
ultimately show "𝔊 ⨅⇩C⇩F ℌ⦇ObjMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ObjMap⦈) ∈⇩∘ Vset α" by auto
qed
qed
show "(𝔊 ⨅⇩C⇩F ℌ ∘⇩C⇩F 𝔉)⦇ArrMap⦈ ∈⇩∘ Vset α"
unfolding dghm_comp_components
proof(rule vbrelation.vbrelation_Limit_in_VsetI)
show "vbrelation (𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈)" by auto
show "Limit α" by auto
show "𝒟⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈) ∈⇩∘ Vset α"
by
(
cs_concl
cs_simp: V_cs_simps cat_cs_simps
cs_intro: 𝔉.cf_ArrMap_vrange cat_small_cs_intros
)
show "ℛ⇩∘ (𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ ∘⇩∘ 𝔉⦇ArrMap⦈) ∈⇩∘ Vset α"
unfolding vrange_vcomp
proof-
have
"𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈) ⊆⇩∘
⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
proof(intro vsubsetI)
fix F assume prems: "F ∈⇩∘ 𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
then obtain ABF
where ABF_in_𝔉: "ABF ∈⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
and 𝔊ℌ_ABF: "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈⦇ABF⦈ = F"
by (auto simp: cf_comma_proj_right_ArrMap_vsv)
with 𝔉.ArrMap.vrange_atD obtain k
where "k ∈⇩∘ 𝔍⦇Arr⦈" and 𝔉j: "𝔉⦇ArrMap⦈⦇k⦈ = ABF"
by (force simp: 𝔉.cf_ArrMap_vdomain)
then obtain i j where "k : i ↦⇘𝔍⇙ j" by auto
from ABF_in_𝔉 𝔉.cf_ArrMap_vrange have ABF_in_𝔊ℌ:
"ABF ∈⇩∘ 𝔊 ⇩C⇩F↓⇩C⇩F ℌ⦇Arr⦈"
by auto
then obtain A B where "ABF : A ↦⇘𝔊 ⇩C⇩F↓⇩C⇩F ℌ⇙ B" by auto
with assms obtain a b f a' b' f' g h
where ABF_def: "ABF = [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
by (elim cat_comma_is_arrE[OF _ assms(1,2)])
have "h ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
proof(intro VUnionI)
from ABF_in_𝔉 show
"[[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘ ∈⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈)"
unfolding ABF_def by auto
show "⟨2⇩ℕ, [g, h]⇩∘⟩ ∈⇩∘ [[a, b, f]⇩∘, [a', b', f']⇩∘, [g, h]⇩∘]⇩∘"
by (auto simp: nat_omega_simps)
show "set {2⇩ℕ, [g, h]⇩∘} ∈⇩∘ ⟨2⇩ℕ, [g, h]⇩∘⟩"
unfolding vpair_def by auto
show "[g, h]⇩∘ ∈⇩∘ set {2⇩ℕ, [g, h]⇩∘}" by simp
show "⟨1⇩ℕ, h⟩ ∈⇩∘ [g, h]⇩∘" by (auto simp: nat_omega_simps)
show "set {1⇩ℕ, h} ∈⇩∘ ⟨1⇩ℕ, h⟩" unfolding vpair_def by auto
qed auto
with ABF_in_𝔊ℌ show "F ∈⇩∘ ⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈)))))))"
unfolding 𝔊ℌ_ABF[symmetric] ABF_def
by (cs_concl cs_shallow cs_simp: cat_cs_simps cat_comma_cs_simps)
qed
moreover have "⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(⋃⇩∘(ℛ⇩∘ (𝔉⦇ArrMap⦈))))))) ∈⇩∘ Vset α"
by (intro VUnion_in_VsetI vrange_in_VsetI[OF 𝔉.tm_cf_ArrMap_in_Vset])
ultimately show "𝔊 ⨅⇩C⇩F ℌ⦇ArrMap⦈ `⇩∘ ℛ⇩∘ (𝔉⦇ArrMap⦈) ∈⇩∘ Vset α" by auto
qed
qed
qed
subsection‹Comma categories constructed from a functor and an object›
subsubsection‹Definitions and elementary properties›
text‹See Chapter II-6 in \<^cite>‹"mac_lane_categories_2010"›.›
definition cat_cf_obj_comma :: "V ⇒ V ⇒ V" (‹(_ ⇩C⇩F↓ _)› [1000, 1000] 999)
where "𝔉 ⇩C⇩F↓ b ≡ 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b)"
definition cat_obj_cf_comma :: "V ⇒ V ⇒ V" (‹(_ ↓⇩C⇩F _)› [1000, 1000] 999)
where "b ↓⇩C⇩F 𝔉 ≡ (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b) ⇩C⇩F↓⇩C⇩F 𝔉"
text‹Alternative forms of the definitions.›
lemma (in is_functor) cat_cf_obj_comma_def:
"𝔉 ⇩C⇩F↓ b = 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) 𝔅 b)"
unfolding cat_cf_obj_comma_def cf_HomCod ..
lemma (in is_functor) cat_obj_cf_comma_def:
"b ↓⇩C⇩F 𝔉 = (cf_const (cat_1 0 0) 𝔅 b) ⇩C⇩F↓⇩C⇩F 𝔉"
unfolding cat_obj_cf_comma_def cf_HomCod ..
text‹Size.›
lemma small_cat_cf_obj_comma_Obj[simp]:
"small {[a, 0, f]⇩∘ | a f. a ∈⇩∘ 𝔄⦇Obj⦈ ∧ f : x ↦⇘ℭ⇙ 𝔊⦇ObjMap⦈⦇a⦈}"
(is ‹small ?afs›)
proof-
define Q where
"Q i = (if i = 0 then 𝔄⦇Obj⦈ else if i = 1⇩ℕ then set {0} else ℭ⦇Arr⦈)"
for i
have "?afs ⊆ elts (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
unfolding Q_def
proof
(
intro subsetI,
unfold mem_Collect_eq,
elim exE conjE,
intro vproductI;
simp only:
)
fix a f show "𝒟⇩∘ [a, 0, f]⇩∘ = set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: three nat_omega_simps)
qed (force simp: nat_omega_simps)+
then show "small ?afs" by (rule down)
qed
lemma small_cat_obj_cf_comma_Obj[simp]:
"small {[0, b, f]⇩∘ | b f. b ∈⇩∘ 𝔅⦇Obj⦈ ∧ f : x ↦⇘ℭ⇙ 𝔊⦇ObjMap⦈⦇b⦈}"
(is ‹small ?bfs›)
proof-
define Q where
"Q i = (if i = 0 then set {0} else if i = 1⇩ℕ then 𝔅⦇Obj⦈ else ℭ⦇Arr⦈)"
for i
have "?bfs ⊆ elts (∏⇩∘i∈⇩∘ set {0, 1⇩ℕ, 2⇩ℕ}. Q i)"
unfolding Q_def
proof
(
intro subsetI,
unfold mem_Collect_eq,
elim exE conjE,
intro vproductI;
simp only:
)
fix a b f show "𝒟⇩∘ [0, b, f]⇩∘ = set {0, 1⇩ℕ, 2⇩ℕ}"
by (simp add: three nat_omega_simps)
qed (force simp: nat_omega_simps)+
then show "small ?bfs" by (rule down)
qed
subsubsection‹Objects›
lemma (in is_functor) cat_cf_obj_comma_ObjI[cat_comma_cs_intros]:
assumes "A = [a, 0, f]⇩∘" and "a ∈⇩∘ 𝔄⦇Obj⦈" and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
shows "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
using assms(2,3)
unfolding assms(1)
by
(
cs_concl
cs_simp: cat_cs_simps cat_cf_obj_comma_def
cs_intro: cat_cs_intros vempty_is_zet cat_comma_ObjI
)
lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_ObjI
lemma (in is_functor) cat_obj_cf_comma_ObjI[cat_comma_cs_intros]:
assumes "A = [0, a, f]⇩∘" and "a ∈⇩∘ 𝔄⦇Obj⦈" and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
shows "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
using assms(2,3)
unfolding assms(1)
by
(
cs_concl
cs_simp: cat_cs_simps cat_obj_cf_comma_def
cs_intro: vempty_is_zet cat_cs_intros cat_comma_ObjI
)
lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_ObjI
lemma (in is_functor) cat_cf_obj_comma_ObjD[dest]:
assumes "[a, b', f]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "a ∈⇩∘ 𝔄⦇Obj⦈" and "b' = 0" and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
proof-
from assms(2) have "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
note obj = cat_comma_ObjD[
OF assms(1)[unfolded cat_cf_obj_comma_def] is_functor_axioms this
]
from obj[unfolded cat_1_components] have [cat_cs_simps]: "b' = 0" by simp
moreover have "cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ = b"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
ultimately show "a ∈⇩∘ 𝔄⦇Obj⦈" "b' = 0" "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
using obj by auto
qed
lemmas [dest] = is_functor.cat_cf_obj_comma_ObjD[rotated 1]
lemma (in is_functor) cat_obj_cf_comma_ObjD[dest]:
assumes "[b', a, f]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "a ∈⇩∘ 𝔄⦇Obj⦈" and "b' = 0" and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
proof-
from assms(2) have "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
note obj = cat_comma_ObjD[
OF assms(1)[unfolded cat_obj_cf_comma_def] this is_functor_axioms
]
from obj[unfolded cat_1_components] have [cat_cs_simps]: "b' = 0" by simp
moreover have "cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ = b"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
ultimately show "a ∈⇩∘ 𝔄⦇Obj⦈" "b' = 0" "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
using obj by auto
qed
lemmas [dest] = is_functor.cat_obj_cf_comma_ObjD[rotated 1]
lemma (in is_functor) cat_cf_obj_comma_ObjE[elim]:
assumes "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains a f
where "A = [a, 0, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
proof-
from assms(2) have "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(1)[unfolded cat_cf_obj_comma_def] is_functor_axioms this
obtain a b' f
where "A = [a, b', f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b': "b' ∈⇩∘ cat_1 0 0⦇Obj⦈"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈"
by auto
moreover from b' have [cat_cs_simps]: "b' = 0"
unfolding cat_1_components by auto
moreover have "cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ = b"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
ultimately show ?thesis using that by auto
qed
lemmas [elim] = is_functor.cat_cf_obj_comma_ObjE[rotated 1]
lemma (in is_functor) cat_obj_cf_comma_ObjE[elim]:
assumes "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains a f
where "A = [0, a, f]⇩∘"
and "a ∈⇩∘ 𝔄⦇Obj⦈"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
proof-
from assms(2) have "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(1)[unfolded cat_obj_cf_comma_def] is_functor_axioms this
obtain a b' f
where A_def: "A = [b', a, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and b': "b' ∈⇩∘ cat_1 0 0⦇Obj⦈"
and f: "f : cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
by auto
moreover from b' have [cat_cs_simps]: "b' = 0"
unfolding cat_1_components by auto
moreover have "cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ = b"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
ultimately show ?thesis using that by auto
qed
lemmas [elim] = is_functor.cat_obj_cf_comma_ObjE[rotated 1]
subsubsection‹Arrows›
lemma (in is_functor) cat_cf_obj_comma_ArrI[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "F = [A, B, [g, 0]⇩∘]⇩∘"
and "A = [a, 0, f]⇩∘"
and "B = [a', 0, f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
shows "F ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
unfolding cat_cf_obj_comma_def
proof(intro cat_comma_ArrI cat_comma_HomI)
show "𝔉 : 𝔄 ↦↦⇩C⇘α⇙ 𝔅" by (cs_concl cs_shallow cs_intro: cat_cs_intros)
from assms(1) show const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from vempty_is_zet show 0: "0 : 0 ↦⇘cat_1 0 0⇙ 0"
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_1_CId_app cs_intro: cat_cs_intros
)
from assms(6) show
"f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇0⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from assms(7) show
"f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇0⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from 0 assms(6) show
"f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = cf_const (cat_1 0 0) 𝔅 b⦇ArrMap⦈⦇0⦈ ∘⇩A⇘𝔅⇙ f"
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps assms(8) cs_intro: cat_cs_intros
)
from const assms(5,6) show "A ∈⇩∘ 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) 𝔅 b)⦇Obj⦈"
by (fold cat_cf_obj_comma_def)
(cs_concl cs_simp: assms(3) cs_intro: cat_cs_intros cat_comma_cs_intros)
from const assms(5,7) show "B ∈⇩∘ 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) 𝔅 b)⦇Obj⦈"
by (fold cat_cf_obj_comma_def)
(
cs_concl cs_shallow
cs_simp: assms(4) cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed (intro assms)+
lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_ArrI
lemma (in is_functor) cat_obj_cf_comma_ArrI[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "F = [A, B, [0, g]⇩∘]⇩∘"
and "A = [0, a, f]⇩∘"
and "B = [0, a', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈ "
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
shows "F ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
unfolding cat_obj_cf_comma_def
proof(intro cat_comma_ArrI cat_comma_HomI)
show "𝔉 : 𝔄 ↦↦⇩C⇘α⇙ 𝔅" by (cs_concl cs_shallow cs_intro: cat_cs_intros)
from assms(1) show const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from vempty_is_zet show 0: "0 : 0 ↦⇘cat_1 0 0⇙ 0"
by (cs_concl cs_shallow cs_simp: cat_1_CId_app cs_intro: cat_cs_intros)
from assms(6) show
"f : cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇0⦈ ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from assms(7) show
"f' : cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇0⦈ ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from 0 assms(7) show
"f' ∘⇩A⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ArrMap⦈⦇0⦈ = 𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f"
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps assms(8) cs_intro: cat_cs_intros
)
from const assms(5,6) show "A ∈⇩∘ (cf_const (cat_1 0 0) 𝔅 b) ⇩C⇩F↓⇩C⇩F 𝔉⦇Obj⦈"
by (fold cat_obj_cf_comma_def)
(cs_concl cs_simp: assms(3) cs_intro: cat_cs_intros cat_comma_cs_intros)
from const assms(5,7) show "B ∈⇩∘ (cf_const (cat_1 0 0) 𝔅 b) ⇩C⇩F↓⇩C⇩F 𝔉⦇Obj⦈"
by (fold cat_obj_cf_comma_def)
(
cs_concl cs_shallow
cs_simp: assms(4) cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed (intro assms)+
lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_ArrI
lemma (in is_functor) cat_cf_obj_comma_ArrE[elim]:
assumes "F ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains A B a f a' f' g
where "F = [A, B, [g, 0]⇩∘]⇩∘"
and "A = [a, 0, f]⇩∘"
and "B = [a', 0, f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
and "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
and "B ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
proof-
from cat_comma_ArrE[OF assms(1)[unfolded cat_cf_obj_comma_def]]
obtain A B
where F: "F ∈⇩∘ cat_comma_Hom 𝔉 (cf_const (cat_1 0 0) 𝔅 b) A B"
and A: "A ∈⇩∘ 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) 𝔅 b)⦇Obj⦈"
and B: "B ∈⇩∘ 𝔉 ⇩C⇩F↓⇩C⇩F (cf_const (cat_1 0 0) 𝔅 b)⦇Obj⦈"
by auto
from assms(2) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from F obtain a b'' f a' b' f' g h
where F_def: "F = [A, B, [g, h]⇩∘]⇩∘"
and A_def: "A = [a, b'', f]⇩∘"
and B_def: "B = [a', b', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and h: "h : b'' ↦⇘cat_1 0 0⇙ b'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b''⦈"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈"
and f_def:
"f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = cf_const (cat_1 0 0) 𝔅 b⦇ArrMap⦈⦇h⦈ ∘⇩A⇘𝔅⇙ f"
by (elim cat_comma_HomE[OF _ is_functor_axioms const]) blast
note hb'b'' = cat_1_is_arrD[OF h]
from F_def have F_def: "F = [A, B, [g, 0]⇩∘]⇩∘"
unfolding hb'b'' by simp
from A_def have A_def: "A = [a, 0, f]⇩∘"
unfolding hb'b'' by simp
from B_def have B_def: "B = [a', 0, f']⇩∘"
unfolding hb'b'' by simp
from f have f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from f' have f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from f_def f f' g h have f_def: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from
that F_def A_def B_def g f f' f_def
B[folded cat_cf_obj_comma_def] A[folded cat_cf_obj_comma_def]
show ?thesis
by blast
qed
lemmas [elim] = is_functor.cat_cf_obj_comma_ArrE[rotated 1]
lemma (in is_functor) cat_obj_cf_comma_ArrE[elim]:
assumes "F ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains A B a f a' f' g
where "F = [A, B, [0, g]⇩∘]⇩∘"
and "A = [0, a, f]⇩∘"
and "B = [0, a', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
and "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
and "B ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
proof-
from cat_comma_ArrE[OF assms(1)[unfolded cat_obj_cf_comma_def]]
obtain A B
where F: "F ∈⇩∘ cat_comma_Hom (cf_const (cat_1 0 0) 𝔅 b) 𝔉 A B"
and A: "A ∈⇩∘ (cf_const (cat_1 0 0) 𝔅 b) ⇩C⇩F↓⇩C⇩F 𝔉⦇Obj⦈"
and B: "B ∈⇩∘ (cf_const (cat_1 0 0) 𝔅 b) ⇩C⇩F↓⇩C⇩F 𝔉⦇Obj⦈"
by auto
from assms(2) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from F obtain a b'' f a' b' f' h g
where F_def: "F = [A, B, [h, g]⇩∘]⇩∘"
and A_def: "A = [b', a, f]⇩∘"
and B_def: "B = [b'', a', f']⇩∘"
and h: "h : b' ↦⇘cat_1 0 0⇙ b''"
and g: "g : a ↦⇘𝔄⇙ a'"
and f: "f : cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b'⦈ ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and f': "f' : cf_const (cat_1 0 0) 𝔅 b⦇ObjMap⦈⦇b''⦈ ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and f'_def:
"f' ∘⇩A⇘𝔅⇙ cf_const (cat_1 0 0) 𝔅 b⦇ArrMap⦈⦇h⦈ = 𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f"
by (elim cat_comma_HomE[OF _ const is_functor_axioms]) blast
note hb'b'' = cat_1_is_arrD[OF h]
from F_def have F_def: "F = [A, B, [0, g]⇩∘]⇩∘"
unfolding hb'b'' by simp
from A_def have A_def: "A = [0, a, f]⇩∘" unfolding hb'b'' by simp
from B_def have B_def: "B = [0, a', f']⇩∘" unfolding hb'b'' by simp
from f have f: "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from f' have f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from f'_def f f' g h have f'_def[symmetric]: "f' = 𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f"
unfolding hb'b''
by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
from
that F_def A_def B_def g f f' f'_def
A[folded cat_obj_cf_comma_def] B[folded cat_obj_cf_comma_def]
show ?thesis
by blast
qed
lemmas [elim] = is_functor.cat_obj_cf_comma_ArrE
lemma (in is_functor) cat_cf_obj_comma_ArrD[dest]:
assumes "[[a, b', f]⇩∘, [a', b'', f']⇩∘, [g, h]⇩∘]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b' = 0"
and "b'' = 0"
and "h = 0"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
and "[a, b', f]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
and "[a', b'', f']⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
using cat_cf_obj_comma_ArrE[OF assms] by auto
lemmas [dest] = is_functor.cat_cf_obj_comma_ArrD[rotated 1]
lemma (in is_functor) cat_obj_cf_comma_ArrD[dest]:
assumes "[[b', a, f]⇩∘, [b'', a', f']⇩∘, [h, g]⇩∘]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b' = 0"
and "b'' = 0"
and "h = 0"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
and "[b', a, f]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
and "[b'', a', f']⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
using cat_obj_cf_comma_ArrE[OF assms] by auto
lemmas [dest] = is_functor.cat_obj_cf_comma_ArrD
subsubsection‹Domain›
lemma cat_cf_obj_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (𝔉 ⇩C⇩F↓ b⦇Dom⦈)"
unfolding cat_cf_obj_comma_def cat_comma_components by simp
lemma cat_cf_obj_comma_Dom_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔉 ⇩C⇩F↓ b⦇Dom⦈) = 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
unfolding cat_cf_obj_comma_def cat_comma_components by simp
lemma cat_cf_obj_comma_Dom_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
shows "𝔉 ⇩C⇩F↓ b⦇Dom⦈⦇ABF⦈ = A"
using assms(2)
unfolding assms(1) cat_cf_obj_comma_def cat_comma_components
by simp
lemma (in is_functor) cat_cf_obj_comma_Dom_vrange:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "ℛ⇩∘ (𝔉 ⇩C⇩F↓ b⦇Dom⦈) ⊆⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Dom_vrange[
OF is_functor_axioms const, folded cat_cf_obj_comma_def
]
)
qed
lemma cat_obj_cf_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (b ↓⇩C⇩F 𝔉⦇Dom⦈)"
unfolding cat_obj_cf_comma_def cat_comma_components by simp
lemma cat_obj_cf_comma_Dom_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (b ↓⇩C⇩F 𝔉⦇Dom⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
unfolding cat_obj_cf_comma_def cat_comma_components by simp
lemma cat_obj_cf_comma_Dom_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
shows "b ↓⇩C⇩F 𝔉⦇Dom⦈⦇ABF⦈ = A"
using assms(2)
unfolding assms(1) cat_obj_cf_comma_def cat_comma_components
by simp
lemma (in is_functor) cat_obj_cf_comma_Dom_vrange:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "ℛ⇩∘ (b ↓⇩C⇩F 𝔉⦇Dom⦈) ⊆⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Dom_vrange[
OF const is_functor_axioms, folded cat_obj_cf_comma_def
]
)
qed
subsubsection‹Codomain›
lemma cat_cf_obj_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (𝔉 ⇩C⇩F↓ b⦇Cod⦈)"
unfolding cat_cf_obj_comma_def cat_comma_components by simp
lemma cat_cf_obj_comma_Cod_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔉 ⇩C⇩F↓ b⦇Cod⦈) = 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
unfolding cat_cf_obj_comma_def cat_comma_components by simp
lemma cat_cf_obj_comma_Cod_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
shows "𝔉 ⇩C⇩F↓ b⦇Cod⦈⦇ABF⦈ = B"
using assms(2)
unfolding assms(1) cat_cf_obj_comma_def cat_comma_components
by (simp add: nat_omega_simps)
lemma (in is_functor) cat_cf_obj_comma_Cod_vrange:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "ℛ⇩∘ (𝔉 ⇩C⇩F↓ b⦇Cod⦈) ⊆⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Cod_vrange[
OF is_functor_axioms const, folded cat_cf_obj_comma_def
]
)
qed
lemma cat_obj_cf_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (b ↓⇩C⇩F 𝔉⦇Cod⦈)"
unfolding cat_obj_cf_comma_def cat_comma_components by simp
lemma cat_obj_cf_comma_Cod_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (b ↓⇩C⇩F 𝔉⦇Cod⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
unfolding cat_obj_cf_comma_def cat_comma_components by simp
lemma cat_obj_cf_comma_Cod_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, F]⇩∘" and "ABF ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
shows "b ↓⇩C⇩F 𝔉⦇Cod⦈⦇ABF⦈ = B"
using assms(2)
unfolding assms(1) cat_obj_cf_comma_def cat_comma_components
by (simp add: nat_omega_simps)
lemma (in is_functor) cat_obj_cf_comma_Cod_vrange:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "ℛ⇩∘ (b ↓⇩C⇩F 𝔉⦇Dom⦈) ⊆⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Dom_vrange[
OF const is_functor_axioms, folded cat_obj_cf_comma_def
]
)
qed
subsubsection‹Arrow with a domain and a codomain›
lemma (in is_functor) cat_cf_obj_comma_is_arrI[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "ABF = [A, B, F]⇩∘"
and "A = [a, 0, f]⇩∘"
and "B = [a', 0, f']⇩∘"
and "F = [g, 0]⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
shows "ABF : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ B"
proof(intro is_arrI)
from assms(1,6,7,8) show "ABF ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
by
(
cs_concl cs_shallow
cs_simp: assms(2,3,4,5,9) cs_intro: cat_comma_cs_intros
)
with assms(2) show "𝔉 ⇩C⇩F↓ b⦇Dom⦈⦇ABF⦈ = A" "𝔉 ⇩C⇩F↓ b⦇Cod⦈⦇ABF⦈ = B"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)+
qed
lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_is_arrI
lemma (in is_functor) cat_obj_cf_comma_is_arrI[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "ABF = [A, B, F]⇩∘"
and "A = [0, a, f]⇩∘"
and "B = [0, a', f']⇩∘"
and "F = [0, g]⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
shows "ABF : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B"
proof(intro is_arrI)
from assms(1,6,7,8) show "ABF ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
by
(
cs_concl cs_shallow
cs_simp: assms(2,3,4,5,9) cs_intro: cat_comma_cs_intros
)
with assms(2) show "b ↓⇩C⇩F 𝔉⦇Dom⦈⦇ABF⦈ = A" "b ↓⇩C⇩F 𝔉⦇Cod⦈⦇ABF⦈ = B"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)+
qed
lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_is_arrI
lemma (in is_functor) cat_cf_obj_comma_is_arrD[dest]:
assumes "[[a, b', f]⇩∘, [a', b'', f']⇩∘, [g, h]⇩∘]⇩∘ :
[a, b', f]⇩∘ ↦⇘𝔉 ⇩C⇩F↓ b⇙ [a', b'', f']⇩∘"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b' = 0"
and "b'' = 0"
and "h = 0"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
and "[a, b', f]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
and "[a', b'', f']⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
by (intro cat_cf_obj_comma_ArrD[OF is_arrD(1)[OF assms(1)] assms(2)])+
lemma (in is_functor) cat_obj_cf_comma_is_arrD[dest]:
assumes "[[b', a, f]⇩∘, [b'', a', f']⇩∘, [h, g]⇩∘]⇩∘ :
[b', a, f]⇩∘ ↦⇘b ↓⇩C⇩F 𝔉⇙ [b'', a', f']⇩∘"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b' = 0"
and "b'' = 0"
and "h = 0"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
and "[b', a, f]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
and "[b'', a', f']⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (intro cat_obj_cf_comma_ArrD[OF is_arrD(1)[OF assms(1)] assms(2)])+
lemmas [dest] = is_functor.cat_obj_cf_comma_is_arrD
lemma (in is_functor) cat_cf_obj_comma_is_arrE[elim]:
assumes "ABF : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ B" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains a f a' f' g
where "ABF = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [g, 0]⇩∘]⇩∘"
and "A = [a, 0, f]⇩∘"
and "B = [a', 0, f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
and "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
and "B ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
proof-
note ABF = is_arrD[OF assms(1)]
from ABF(1) obtain C D a f a' f' g
where ABF_def: "ABF = [C, D, [g, 0]⇩∘]⇩∘"
and C_def: "C = [a, 0, f]⇩∘"
and D_def: "D = [a', 0, f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and f_def: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
and C: "C ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
and D: "D ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
by (elim cat_cf_obj_comma_ArrE[OF _ assms(2)])
from ABF(2) assms(2) C_def D_def g f f' f_def have [simp]: "C = A"
unfolding ABF_def
by
(
cs_prems cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
from ABF(3) assms(2) C_def D_def g f f' f_def have [simp]: "D = B"
unfolding ABF_def
by
(
cs_prems cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
from that ABF_def C_def D_def g f f' f_def C D show ?thesis by auto
qed
lemmas [elim] = is_functor.cat_cf_obj_comma_is_arrE
lemma (in is_functor) cat_obj_cf_comma_is_arrE[elim]:
assumes "ABF : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B" and "b ∈⇩∘ 𝔅⦇Obj⦈"
obtains a f a' f' g
where "ABF = [[0, a, f]⇩∘, [0, a', f']⇩∘, [0, g]⇩∘]⇩∘"
and "A = [0, a, f]⇩∘"
and "B = [0, a', f']⇩∘"
and "g : a ↦⇘𝔄⇙ a'"
and "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
and "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
and "B ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
proof-
note ABF = is_arrD[OF assms(1)]
from ABF(1) obtain C D a f a' f' g
where ABF_def: "ABF = [C, D, [0, g]⇩∘]⇩∘"
and C_def: "C = [0, a, f]⇩∘"
and D_def: "D = [0, a', f']⇩∘"
and g: "g : a ↦⇘𝔄⇙ a'"
and f: "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and f'_def: "𝔉⦇ArrMap⦈⦇g⦈ ∘⇩A⇘𝔅⇙ f = f'"
and C: "C ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
and D: "D ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (elim cat_obj_cf_comma_ArrE[OF _ assms(2)])
from ABF(2) assms(2) C_def D_def g f f' f'_def have [simp]: "C = A"
unfolding ABF_def
by
(
cs_prems cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
from ABF(3) assms(2) C_def D_def g f f' f'_def have [simp]: "D = B"
unfolding ABF_def
by
(
cs_prems cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
from that ABF_def C_def D_def g f f' f'_def C D show ?thesis by auto
qed
lemmas [elim] = is_functor.cat_obj_cf_comma_is_arrE
subsubsection‹Composition›
lemma cat_cf_obj_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (𝔉 ⇩C⇩F↓ b⦇Comp⦈)"
unfolding cat_cf_obj_comma_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma cat_obj_cf_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (b ↓⇩C⇩F 𝔉⦇Comp⦈)"
unfolding cat_obj_cf_comma_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma (in is_functor) cat_cf_obj_comma_Comp_app[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "BCG = [B, C, [g', h']⇩∘]⇩∘"
and "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "BCG : B ↦⇘𝔉 ⇩C⇩F↓ b⇙ C"
and "ABF : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ B"
shows "BCG ∘⇩A⇘𝔉 ⇩C⇩F↓ b⇙ ABF = [A, C, [g' ∘⇩A⇘𝔄⇙ g, 0]⇩∘]⇩∘"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(4) obtain a f a' f' g
where BCG_def: "BCG = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [g, 0]⇩∘]⇩∘"
by (elim cat_cf_obj_comma_is_arrE[OF _ assms(1)])
from assms(5) obtain a f a' f' g
where ABF_def: "ABF = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [g, 0]⇩∘]⇩∘"
by (elim cat_cf_obj_comma_is_arrE[OF _ assms(1)])
from assms(2)[unfolded BCG_def] assms(3)[unfolded ABF_def] have [cat_cs_simps]:
"h' = 0" "h = 0"
by simp_all
have "h' ∘⇩A⇘cat_1 0 0⇙ h = 0" by (cs_concl cs_shallow cs_simp: cat_cs_simps)
show ?thesis
by
(
rule cat_comma_Comp_app
[
OF
is_functor_axioms
const
assms(2,3)
assms(4)[unfolded cat_cf_obj_comma_def]
assms(5)[unfolded cat_cf_obj_comma_def],
folded cat_cf_obj_comma_def,
unfolded cat_cs_simps
]
)
qed
lemma (in is_functor) cat_obj_cf_comma_Comp_app[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "BCG = [B, C, [h', g']⇩∘]⇩∘"
and "ABF = [A, B, [h, g]⇩∘]⇩∘"
and "BCG : B ↦⇘b ↓⇩C⇩F 𝔉⇙ C"
and "ABF : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B"
shows "BCG ∘⇩A⇘b ↓⇩C⇩F 𝔉⇙ ABF = [A, C, [0, g' ∘⇩A⇘𝔄⇙ g]⇩∘]⇩∘"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(4) obtain a f a' f' g
where BCG_def: "BCG = [[0, a, f]⇩∘, [0, a', f']⇩∘, [0, g]⇩∘]⇩∘"
by (elim cat_obj_cf_comma_is_arrE[OF _ assms(1)])
from assms(5) obtain a f a' f' g
where ABF_def: "ABF = [[0, a, f]⇩∘, [0, a', f']⇩∘, [0, g]⇩∘]⇩∘"
by (elim cat_obj_cf_comma_is_arrE[OF _ assms(1)])
from assms(2)[unfolded BCG_def] assms(3)[unfolded ABF_def] have [cat_cs_simps]:
"h' = 0" "h = 0"
by simp_all
have "h' ∘⇩A⇘cat_1 0 0⇙ h = 0" by (cs_concl cs_shallow cs_simp: cat_cs_simps)
show ?thesis
by
(
rule cat_comma_Comp_app
[
OF
const
is_functor_axioms
assms(2,3)
assms(4)[unfolded cat_obj_cf_comma_def]
assms(5)[unfolded cat_obj_cf_comma_def],
folded cat_obj_cf_comma_def,
unfolded cat_cs_simps
]
)
qed
lemma (in is_functor) cat_cf_obj_comma_Comp_is_arr[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "BCG : B ↦⇘𝔉 ⇩C⇩F↓ b⇙ C"
and "ABF : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ B"
shows "BCG ∘⇩A⇘𝔉 ⇩C⇩F↓ b⇙ ABF : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ C"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Comp_is_arr
[
OF
is_functor_axioms
const
assms(2)[unfolded cat_cf_obj_comma_def]
assms(3)[unfolded cat_cf_obj_comma_def],
folded cat_cf_obj_comma_def
]
)
qed
lemma (in is_functor) cat_obj_cf_comma_Comp_is_arr[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "BCG : B ↦⇘b ↓⇩C⇩F 𝔉⇙ C"
and "ABF : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B"
shows "BCG ∘⇩A⇘b ↓⇩C⇩F 𝔉⇙ ABF : A ↦⇘b ↓⇩C⇩F 𝔉⇙ C"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_Comp_is_arr
[
OF
const
is_functor_axioms
assms(2)[unfolded cat_obj_cf_comma_def]
assms(3)[unfolded cat_obj_cf_comma_def],
folded cat_obj_cf_comma_def
]
)
qed
subsubsection‹Identity›
lemma cat_cf_obj_comma_CId_vsv[cat_comma_cs_intros]: "vsv (𝔉 ⇩C⇩F↓ b⦇CId⦈)"
unfolding cat_cf_obj_comma_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma cat_obj_cf_comma_CId_vsv[cat_comma_cs_intros]: "vsv (b ↓⇩C⇩F 𝔉⦇CId⦈)"
unfolding cat_obj_cf_comma_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma (in is_functor) cat_cf_obj_comma_CId_vdomain[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "𝒟⇩∘ (𝔉 ⇩C⇩F↓ b⦇CId⦈) = 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule cat_comma_CId_vdomain[
OF is_functor_axioms const, folded cat_cf_obj_comma_def
]
)
qed
lemma (in is_functor) cat_obj_cf_comma_CId_vdomain[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "𝒟⇩∘ (b ↓⇩C⇩F 𝔉⦇CId⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show "𝒟⇩∘ (b ↓⇩C⇩F 𝔉⦇CId⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by
(
rule cat_comma_CId_vdomain[
OF const is_functor_axioms, folded cat_obj_cf_comma_def
]
)
qed
lemma (in is_functor) cat_cf_obj_comma_CId_app[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "A = [a, b', f]⇩∘" and "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
shows "𝔉 ⇩C⇩F↓ b⦇CId⦈⦇A⦈ = [A, A, [𝔄⦇CId⦈⦇a⦈, 0]⇩∘]⇩∘"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(3,2) have b'_def: "b' = 0"
by (auto elim: cat_cf_obj_comma_ObjE[OF _ assms(1)])
have [cat_cs_simps]: "cat_1 0 0⦇CId⦈⦇b'⦈ = 0"
unfolding cat_1_components b'_def by simp
show ?thesis
by
(
rule cat_comma_CId_app
[
OF
is_functor_axioms
const
assms(2,3)[unfolded cat_cf_obj_comma_def],
unfolded cat_cf_obj_comma_def[symmetric] cat_cs_simps
]
)
qed
lemma (in is_functor) cat_obj_cf_comma_CId_app[cat_comma_cs_simps]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "A = [b', a, f]⇩∘" and "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
shows "b ↓⇩C⇩F 𝔉⦇CId⦈⦇A⦈ = [A, A, [0, 𝔄⦇CId⦈⦇a⦈]⇩∘]⇩∘"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
from assms(3,2) have b'_def: "b' = 0"
by (auto elim: cat_obj_cf_comma_ObjE[OF _ assms(1)])
have [cat_cs_simps]: "cat_1 0 0⦇CId⦈⦇b'⦈ = 0"
unfolding cat_1_components b'_def by simp
show ?thesis
by
(
rule cat_comma_CId_app
[
OF
const
is_functor_axioms
assms(2,3)[unfolded cat_obj_cf_comma_def],
unfolded cat_obj_cf_comma_def[symmetric] cat_cs_simps
]
)
qed
subsubsection‹
Comma categories constructed from a functor and an object are categories
›
lemma (in is_functor) category_cat_cf_obj_comma[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "category α (𝔉 ⇩C⇩F↓ b)"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule category_cat_comma[
OF is_functor_axioms const, folded cat_cf_obj_comma_def
]
)
qed
lemmas [cat_comma_cs_intros] = is_functor.category_cat_cf_obj_comma
lemma (in is_functor) category_cat_obj_cf_comma[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "category α (b ↓⇩C⇩F 𝔉)"
proof-
from assms(1) have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_cs_intros)
show ?thesis
by
(
rule category_cat_comma[
OF const is_functor_axioms, folded cat_obj_cf_comma_def
]
)
qed
lemmas [cat_comma_cs_intros] = is_functor.category_cat_obj_cf_comma
subsubsection‹Tiny comma categories constructed from a functor and an object›
lemma (in is_tm_functor) tiny_category_cat_cf_obj_comma[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "tiny_category α (𝔉 ⇩C⇩F↓ b)"
proof-
from assms(1) have const:
"cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_small_cs_intros cat_cs_intros)
show ?thesis
by
(
rule tiny_category_cat_comma[
OF is_tm_functor_axioms const, folded cat_cf_obj_comma_def
]
)
qed
lemma (in is_tm_functor) tiny_category_cat_obj_cf_comma[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "tiny_category α (b ↓⇩C⇩F 𝔉)"
proof-
from assms(1) have const:
"cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
by (cs_concl cs_intro: vempty_is_zet cat_small_cs_intros cat_cs_intros)
show ?thesis
by
(
rule tiny_category_cat_comma[
OF const is_tm_functor_axioms, folded cat_obj_cf_comma_def
]
)
qed
subsection‹
Opposite comma category functors for the comma categories
constructed from a functor and an object
›
subsubsection‹Definitions and elementary properties›
definition op_cf_obj_comma :: "V ⇒ V ⇒ V"
where "op_cf_obj_comma 𝔉 b =
op_cf_comma 𝔉 (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b)"
definition op_obj_cf_comma :: "V ⇒ V ⇒ V"
where "op_obj_cf_comma b 𝔉 =
op_cf_comma (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b) 𝔉"
text‹Alternative forms of the definitions.›
lemma (in is_functor) op_cf_obj_comma_def:
"op_cf_obj_comma 𝔉 b = op_cf_comma 𝔉 (cf_const (cat_1 0 0) 𝔅 b)"
unfolding op_cf_obj_comma_def cat_cs_simps by simp
lemma (in is_functor) op_obj_cf_comma_def:
"op_obj_cf_comma b 𝔉 = op_cf_comma (cf_const (cat_1 0 0) 𝔅 b) 𝔉"
unfolding op_obj_cf_comma_def cat_cs_simps by simp
subsubsection‹Object map›
lemma op_cf_obj_comma_ObjMap_vsv[cat_comma_cs_intros]:
"vsv (op_cf_obj_comma 𝔉 b⦇ObjMap⦈)"
unfolding op_cf_obj_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
lemma op_obj_cf_comma_ObjMap_vsv[cat_comma_cs_intros]:
"vsv (op_obj_cf_comma b 𝔉⦇ObjMap⦈)"
unfolding op_obj_cf_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
lemma (in is_functor) op_cf_obj_comma_ObjMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (op_cf_obj_comma 𝔉 b⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
unfolding op_cf_obj_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_cf_obj_comma_def[symmetric]
)
lemma (in is_functor) op_obj_cf_comma_ObjMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (op_obj_cf_comma b 𝔉⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
unfolding op_obj_cf_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_obj_cf_comma_def[symmetric]
)
lemma (in is_functor) op_cf_obj_comma_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a, 0, f]⇩∘" and "b ∈⇩∘ 𝔅⦇Obj⦈" and "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
shows "op_cf_obj_comma 𝔉 b⦇ObjMap⦈⦇A⦈ = [0, a, f]⇩∘"
proof-
have a: "a ∈⇩∘ 𝔄⦇Obj⦈" and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
by (intro cat_cf_obj_comma_ObjD[OF assms(3)[unfolded assms(1)] assms(2)])+
from assms(2) a f show ?thesis
using assms(2)
unfolding assms(1) op_cf_obj_comma_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros
)
qed
lemma (in is_functor) op_obj_cf_comma_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [0, a, f]⇩∘" and "b ∈⇩∘ 𝔅⦇Obj⦈" and "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
shows "op_obj_cf_comma b 𝔉 ⦇ObjMap⦈⦇A⦈ = [a, 0, f]⇩∘"
proof-
have a: "a ∈⇩∘ 𝔄⦇Obj⦈" and f: "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
by (intro cat_obj_cf_comma_ObjD[OF assms(3)[unfolded assms(1)] assms(2)])+
from assms(2) a f show ?thesis
using assms(2)
unfolding assms(1) op_obj_cf_comma_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros
)
qed
subsubsection‹Arrow map›
lemma op_cf_obj_comma_ArrMap_vsv[cat_comma_cs_intros]:
"vsv (op_cf_obj_comma 𝔉 b⦇ArrMap⦈)"
unfolding op_cf_obj_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
lemma op_obj_cf_comma_ArrMap_vsv[cat_comma_cs_intros]:
"vsv (op_obj_cf_comma b 𝔉⦇ArrMap⦈)"
unfolding op_obj_cf_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros
)
lemma (in is_functor) op_cf_obj_comma_ArrMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (op_cf_obj_comma 𝔉 b⦇ArrMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
unfolding op_cf_obj_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_cf_obj_comma_def[symmetric]
)
lemmas [cat_comma_cs_simps] = is_functor.op_cf_obj_comma_ArrMap_vdomain
lemma (in is_functor) op_obj_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (op_obj_cf_comma b 𝔉⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
unfolding op_obj_cf_comma_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_obj_cf_comma_def[symmetric]
)
lemmas [cat_comma_cs_simps] = is_functor.op_obj_cf_comma_ArrMap_vdomain
lemma (in is_functor) op_cf_obj_comma_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [g, 0]⇩∘]⇩∘"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "ABF ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
shows "op_cf_obj_comma 𝔉 b⦇ArrMap⦈⦇ABF⦈ = [[0, a', f']⇩∘, [0, a, f]⇩∘, [0, g]⇩∘]⇩∘"
proof-
from assms(3) have g: "g : a ↦⇘𝔄⇙ a'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ b"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and [cat_comma_cs_simps]: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇g⦈ = f"
by (intro cat_cf_obj_comma_ArrD[OF assms(3)[unfolded assms(1)] assms(2)])+
from assms(2) g f f' show ?thesis
unfolding assms(1) op_cf_obj_comma_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_1_CId_app
cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_1_is_arrI
)
qed
lemmas [cat_comma_cs_simps] = is_functor.op_cf_obj_comma_ArrMap_app
lemma (in is_functor) op_obj_cf_comma_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [[0, a, f]⇩∘, [0, a', f']⇩∘, [0, h]⇩∘]⇩∘"
and "b ∈⇩∘ 𝔅⦇Obj⦈"
and "ABF ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
shows "op_obj_cf_comma b 𝔉⦇ArrMap⦈⦇ABF⦈ = [[a', 0, f']⇩∘, [a, 0, f]⇩∘, [h, 0]⇩∘]⇩∘"
proof-
from assms(3) have h: "h : a ↦⇘𝔄⇙ a'"
and f: "f : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
and [cat_comma_cs_simps]: "𝔉⦇ArrMap⦈⦇h⦈ ∘⇩A⇘𝔅⇙ f = f'"
by (intro cat_obj_cf_comma_ArrD[OF assms(3)[unfolded assms(1)] assms(2)])+
from assms(2) h f f' show ?thesis
unfolding assms(1) op_obj_cf_comma_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps cat_1_CId_app
cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_1_is_arrI
)
qed
lemmas [cat_comma_cs_simps] = is_functor.op_obj_cf_comma_ArrMap_app
subsubsection‹
Opposite comma category functors for the comma categories
constructed from a functor and an object are isomorphisms of categories
›
lemma (in is_functor) op_cf_obj_comma_is_iso_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_cf_obj_comma 𝔉 b : op_cat (𝔉 ⇩C⇩F↓ b) ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ b ↓⇩C⇩F (op_cf 𝔉)"
proof-
from assms have cf_const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
note cat_obj_cf_comma_def =
is_functor.cat_obj_cf_comma_def[
OF is_functor_op, unfolded cat_op_simps
]
show ?thesis
by
(
rule op_cf_comma_is_iso_functor
[
OF is_functor_axioms cf_const,
folded cat_cf_obj_comma_def op_cf_obj_comma_def,
unfolded cat_op_simps,
folded cat_obj_cf_comma_def
]
)
qed
lemma (in is_functor) op_cf_obj_comma_is_iso_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "𝔄' = op_cat (𝔉 ⇩C⇩F↓ b)"
and "𝔅' = b ↓⇩C⇩F (op_cf 𝔉)"
shows "op_cf_obj_comma 𝔉 b : 𝔄' ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule op_cf_obj_comma_is_iso_functor)
lemmas [cat_comma_cs_intros] = is_functor.op_cf_obj_comma_is_iso_functor'
lemma (in is_functor) op_cf_obj_comma_is_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_cf_obj_comma 𝔉 b : op_cat (𝔉 ⇩C⇩F↓ b) ↦↦⇩C⇘α⇙ b ↓⇩C⇩F (op_cf 𝔉)"
by (rule is_iso_functorD(1)[OF op_cf_obj_comma_is_iso_functor[OF assms]])
lemma (in is_functor) op_cf_obj_comma_is_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "𝔄' = op_cat (𝔉 ⇩C⇩F↓ b)"
and "𝔅' = b ↓⇩C⇩F (op_cf 𝔉)"
shows "op_cf_obj_comma 𝔉 b : 𝔄' ↦↦⇩C⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule op_cf_obj_comma_is_functor)
lemmas [cat_comma_cs_intros] = is_functor.op_cf_obj_comma_is_functor'
lemma (in is_functor) op_obj_cf_comma_is_iso_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_obj_cf_comma b 𝔉 : op_cat (b ↓⇩C⇩F 𝔉) ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ (op_cf 𝔉) ⇩C⇩F↓ b"
proof-
from assms have cf_const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
note cat_cf_obj_comma_def =
is_functor.cat_cf_obj_comma_def[
OF is_functor_op, unfolded cat_op_simps
]
show ?thesis
by
(
rule op_cf_comma_is_iso_functor
[
OF cf_const is_functor_axioms,
folded cat_obj_cf_comma_def op_obj_cf_comma_def,
unfolded cat_op_simps,
folded cat_cf_obj_comma_def
]
)
qed
lemma (in is_functor) op_obj_cf_comma_is_iso_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "𝔄' = op_cat (b ↓⇩C⇩F 𝔉)"
and "𝔅' = (op_cf 𝔉) ⇩C⇩F↓ b"
shows "op_obj_cf_comma b 𝔉 : 𝔄' ↦↦⇩C⇩.⇩i⇩s⇩o⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule op_obj_cf_comma_is_iso_functor)
lemmas [cat_comma_cs_intros] = is_functor.op_obj_cf_comma_is_iso_functor'
lemma (in is_functor) op_obj_cf_comma_is_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_obj_cf_comma b 𝔉 : op_cat (b ↓⇩C⇩F 𝔉) ↦↦⇩C⇘α⇙ (op_cf 𝔉) ⇩C⇩F↓ b"
by (rule is_iso_functorD(1)[OF op_obj_cf_comma_is_iso_functor[OF assms]])
lemma (in is_functor) op_obj_cf_comma_is_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
and "𝔄' = op_cat (b ↓⇩C⇩F 𝔉)"
and "𝔅' = (op_cf 𝔉) ⇩C⇩F↓ b"
shows "op_obj_cf_comma b 𝔉 : 𝔄' ↦↦⇩C⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule op_obj_cf_comma_is_functor)
subsection‹
Projections for comma categories constructed from a functor and an object
›
subsubsection‹Definitions and elementary properties›
definition cf_cf_obj_comma_proj :: "V ⇒ V ⇒ V" (‹(_ ⇩C⇩F⨅⇩O _)› [1000, 1000] 999)
where "𝔉 ⇩C⇩F⨅⇩O b ≡ 𝔉 ⇩C⇩F⨅ (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b)"
definition cf_obj_cf_comma_proj :: "V ⇒ V ⇒ V" (‹(_ ⇩O⨅⇩C⇩F _)› [1000, 1000] 999)
where "b ⇩O⨅⇩C⇩F 𝔉 ≡ (cf_const (cat_1 0 0) (𝔉⦇HomCod⦈) b) ⨅⇩C⇩F 𝔉"
text‹Alternative forms of the definitions.›
lemma (in is_functor) cf_cf_obj_comma_proj_def:
"𝔉 ⇩C⇩F⨅⇩O b = 𝔉 ⇩C⇩F⨅ (cf_const (cat_1 0 0) 𝔅 b)"
unfolding cf_cf_obj_comma_proj_def cf_HomCod..
lemma (in is_functor) cf_obj_cf_comma_proj_def:
"b ⇩O⨅⇩C⇩F 𝔉 = (cf_const (cat_1 0 0) 𝔅 b) ⨅⇩C⇩F 𝔉"
unfolding cf_obj_cf_comma_proj_def cf_HomCod..
text‹Components.›
lemma (in is_functor) cf_cf_obj_comma_proj_components[cat_comma_cs_simps]:
shows "𝔉 ⇩C⇩F⨅⇩O b⦇HomDom⦈ = 𝔉 ⇩C⇩F↓ b"
and "𝔉 ⇩C⇩F⨅⇩O b⦇HomCod⦈ = 𝔄"
unfolding
cf_cf_obj_comma_proj_def
cf_comma_proj_left_components
cat_cf_obj_comma_def[symmetric]
cat_cs_simps
by simp_all
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_components
lemma (in is_functor) cf_obj_cf_comma_proj_components[cat_comma_cs_simps]:
shows "b ⇩O⨅⇩C⇩F 𝔉⦇HomDom⦈ = b ↓⇩C⇩F 𝔉"
and "b ⇩O⨅⇩C⇩F 𝔉⦇HomCod⦈ = 𝔄"
unfolding
cf_obj_cf_comma_proj_def
cf_comma_proj_right_components
cat_obj_cf_comma_def[symmetric]
cat_cs_simps
by simp_all
lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_components
subsubsection‹Object map›
lemma cf_cf_obj_comma_proj_ObjMap_vsv[cat_comma_cs_intros]:
"vsv (𝔉 ⇩C⇩F⨅⇩O b⦇ObjMap⦈)"
unfolding cf_cf_obj_comma_proj_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma cf_obj_cf_comma_proj_ObjMap_vsv[cat_comma_cs_intros]:
"vsv (b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈)"
unfolding cf_obj_cf_comma_proj_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma (in is_functor) cf_cf_obj_comma_proj_ObjMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔉 ⇩C⇩F⨅⇩O b⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
unfolding cf_cf_obj_comma_proj_def cf_comma_proj_left_ObjMap_vdomain
unfolding
cf_cf_obj_comma_proj_def[symmetric]
cf_comma_proj_left_components[symmetric]
cat_comma_cs_simps
by simp
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_vdomain
lemma (in is_functor) cf_obj_cf_comma_proj_ObjMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
unfolding cf_obj_cf_comma_proj_def cf_comma_proj_right_ObjMap_vdomain
unfolding
cf_obj_cf_comma_proj_def[symmetric]
cf_comma_proj_right_components[symmetric]
cat_comma_cs_simps
by simp
lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ObjMap_vdomain
lemma (in is_functor) cf_cf_obj_comma_proj_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a, b', f]⇩∘" and "[a, b', f]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
shows "𝔉 ⇩C⇩F⨅⇩O b⦇ObjMap⦈⦇A⦈ = a"
by
(
rule cf_comma_proj_left_ObjMap_app[
OF assms(1) assms(2)[unfolded cat_cf_obj_comma_def],
folded cf_cf_obj_comma_proj_def
]
)
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_app
lemma (in is_functor) cf_obj_cf_comma_proj_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [b', a, f]⇩∘" and "[b', a, f]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
shows "b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈ = a"
by
(
rule cf_comma_proj_right_ObjMap_app[
OF assms(1) assms(2)[unfolded cat_obj_cf_comma_def],
folded cf_obj_cf_comma_proj_def
]
)
lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ObjMap_app
subsubsection‹Arrow map›
lemma cf_cf_obj_comma_proj_ArrMap_vsv[cat_comma_cs_intros]:
"vsv (𝔉 ⇩C⇩F⨅⇩O b⦇ArrMap⦈)"
unfolding cf_cf_obj_comma_proj_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma cf_obj_cf_comma_proj_ArrMap_vsv[cat_comma_cs_intros]:
"vsv (b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈)"
unfolding cf_obj_cf_comma_proj_def
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma (in is_functor) cf_cf_obj_comma_proj_ArrMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (𝔉 ⇩C⇩F⨅⇩O b⦇ArrMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
unfolding cf_cf_obj_comma_proj_def cf_comma_proj_left_ArrMap_vdomain
unfolding
cf_cf_obj_comma_proj_def[symmetric]
cf_comma_proj_left_components[symmetric]
cat_comma_cs_simps
by simp
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_vdomain
lemma (in is_functor) cf_obj_cf_comma_proj_ArrMap_vdomain[cat_comma_cs_simps]:
"𝒟⇩∘ (b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
unfolding cf_obj_cf_comma_proj_def cf_comma_proj_right_ArrMap_vdomain
unfolding
cf_obj_cf_comma_proj_def[symmetric]
cf_comma_proj_right_components[symmetric]
cat_comma_cs_simps
by simp
lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ArrMap_vdomain
lemma (in is_functor) cf_cf_obj_comma_proj_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "[A, B, [g, h]⇩∘]⇩∘ ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
shows "𝔉 ⇩C⇩F⨅⇩O b⦇ArrMap⦈⦇ABF⦈ = g"
by
(
rule cf_comma_proj_left_ArrMap_app[
OF assms(1) assms(2)[unfolded cat_cf_obj_comma_def],
folded cf_cf_obj_comma_proj_def
]
)
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ArrMap_app
lemma (in is_functor) cf_obj_cf_comma_proj_ArrMap_app[cat_comma_cs_simps]:
assumes "ABF = [A, B, [g, h]⇩∘]⇩∘"
and "[A, B, [g, h]⇩∘]⇩∘ ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
shows "b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈⦇ABF⦈ = h"
by
(
rule cf_comma_proj_right_ArrMap_app[
OF assms(1) assms(2)[unfolded cat_obj_cf_comma_def],
folded cf_obj_cf_comma_proj_def
]
)
lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ArrMap_app
subsubsection‹Projections for a comma category are functors›
lemma (in is_functor) cf_cf_obj_comma_proj_is_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "𝔉 ⇩C⇩F⨅⇩O b : 𝔉 ⇩C⇩F↓ b ↦↦⇩C⇘α⇙ 𝔄"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comma_proj_left_is_functor[
OF is_functor_axioms const,
folded cf_cf_obj_comma_proj_def cat_cf_obj_comma_def
]
)
qed
lemma (in is_functor) cf_cf_obj_comma_proj_is_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "𝔄' = 𝔉 ⇩C⇩F↓ b"
shows "𝔉 ⇩C⇩F⨅⇩O b : 𝔄' ↦↦⇩C⇘α⇙ 𝔄"
using assms(1) unfolding assms(2) by (rule cf_cf_obj_comma_proj_is_functor)
lemmas [cat_comma_cs_intros] = is_functor.cf_cf_obj_comma_proj_is_functor'
lemma (in is_functor) cf_obj_cf_comma_proj_is_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b ⇩O⨅⇩C⇩F 𝔉 : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ 𝔄"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comma_proj_right_is_functor[
OF const is_functor_axioms,
folded cf_obj_cf_comma_proj_def cat_obj_cf_comma_def
]
)
qed
lemma (in is_functor) cf_obj_cf_comma_proj_is_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "𝔄' = b ↓⇩C⇩F 𝔉"
shows "b ⇩O⨅⇩C⇩F 𝔉 : 𝔄' ↦↦⇩C⇘α⇙ 𝔄"
using assms(1) unfolding assms(2) by (rule cf_obj_cf_comma_proj_is_functor)
lemmas [cat_comma_cs_intros] = is_functor.cf_obj_cf_comma_proj_is_functor'
subsubsection‹
Opposite projections for comma categories constructed from a functor
and an object
›
lemma (in is_functor) op_cf_cf_obj_comma_proj:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_cf (𝔉 ⇩C⇩F⨅⇩O b) = b ⇩O⨅⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 b"
proof-
from assms have cf_const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule op_cf_comma_proj_left
[
OF is_functor_axioms cf_const,
unfolded cat_op_simps,
folded
cf_cf_obj_comma_proj_def
op_cf_obj_comma_def
is_functor.cf_obj_cf_comma_proj_def[
OF is_functor_op, unfolded cat_op_simps
]
]
)
qed
lemma (in is_functor) op_cf_obj_cf_comma_proj:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "op_cf (b ⇩O⨅⇩C⇩F 𝔉) = (op_cf 𝔉) ⇩C⇩F⨅⇩O b ∘⇩C⇩F op_obj_cf_comma b 𝔉"
proof-
from assms have cf_const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇘α⇙ 𝔅"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule op_cf_comma_proj_right
[
OF cf_const is_functor_axioms,
unfolded cat_op_simps,
folded
cf_obj_cf_comma_proj_def
op_obj_cf_comma_def
is_functor.cf_cf_obj_comma_proj_def[
OF is_functor_op, unfolded cat_op_simps
]
]
)
qed
subsubsection‹Projections for a tiny comma category›
lemma (in is_tm_functor) cf_cf_obj_comma_proj_is_tm_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "𝔉 ⇩C⇩F⨅⇩O b : 𝔉 ⇩C⇩F↓ b ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
by (cs_concl cs_intro: V_cs_intros cat_small_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comma_proj_left_is_tm_functor[
OF is_tm_functor_axioms const,
folded cf_cf_obj_comma_proj_def cat_cf_obj_comma_def
]
)
qed
lemma (in is_tm_functor) cf_cf_obj_comma_proj_is_tm_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "𝔉b = 𝔉 ⇩C⇩F↓ b"
shows "𝔉 ⇩C⇩F⨅⇩O b : 𝔉b ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
using assms(1) unfolding assms(2) by (rule cf_cf_obj_comma_proj_is_tm_functor)
lemmas [cat_comma_cs_intros] = is_tm_functor.cf_cf_obj_comma_proj_is_tm_functor'
lemma (in is_tm_functor) cf_obj_cf_comma_proj_is_tm_functor:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "b ⇩O⨅⇩C⇩F 𝔉 : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
proof-
from assms have const: "cf_const (cat_1 0 0) 𝔅 b : cat_1 0 0 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
by (cs_concl cs_intro: V_cs_intros cat_small_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comma_proj_right_is_tm_functor[
OF const is_tm_functor_axioms,
folded cf_obj_cf_comma_proj_def cat_obj_cf_comma_def
]
)
qed
lemma (in is_tm_functor) cf_obj_cf_comma_proj_is_tm_functor'[cat_comma_cs_intros]:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈" and "𝔄' = b ↓⇩C⇩F 𝔉"
shows "b ⇩O⨅⇩C⇩F 𝔉 : 𝔄' ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
using assms(1) unfolding assms(2) by (rule cf_obj_cf_comma_proj_is_tm_functor)
lemmas [cat_comma_cs_intros] = is_tm_functor.cf_obj_cf_comma_proj_is_tm_functor'
lemma cf_comp_cf_cf_obj_comma_proj_is_tm_functor[cat_comma_cs_intros]:
assumes "𝔊 : 𝔄 ↦↦⇩C⇘α⇙ ℭ"
and "𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔊 ⇩C⇩F↓ c"
and "c ∈⇩∘ ℭ⦇Obj⦈"
shows "𝔊 ⇩C⇩F⨅⇩O c ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔄"
proof-
interpret 𝔊: is_functor α 𝔄 ℭ 𝔊 by (rule assms(1))
from assms(3) have cf_const: "cf_const (cat_1 0 0) ℭ c : cat_1 0 0 ↦↦⇩C⇘α⇙ ℭ"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comp_cf_comma_proj_left_is_tm_functor
[
OF assms(1) _ assms(2)[unfolded cat_cf_obj_comma_def],
unfolded cat_cs_simps,
OF cf_const,
folded 𝔊.cf_cf_obj_comma_proj_def
]
)
qed
lemma cf_comp_cf_obj_cf_comma_proj_is_tm_functor[cat_comma_cs_intros]:
assumes "ℌ : 𝔅 ↦↦⇩C⇘α⇙ ℭ"
and "𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ c ↓⇩C⇩F ℌ"
and "c ∈⇩∘ ℭ⦇Obj⦈"
shows "c ⇩O⨅⇩C⇩F ℌ ∘⇩C⇩F 𝔉 : 𝔍 ↦↦⇩C⇩.⇩t⇩m⇘α⇙ 𝔅"
proof-
interpret ℌ: is_functor α 𝔅 ℭ ℌ by (rule assms(1))
from assms(3) have cf_const: "cf_const (cat_1 0 0) ℭ c : cat_1 0 0 ↦↦⇩C⇘α⇙ ℭ"
by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros)
show ?thesis
by
(
rule cf_comp_cf_comma_proj_right_is_tm_functor
[
OF _ assms(1) assms(2)[unfolded cat_obj_cf_comma_def],
unfolded cat_cs_simps,
OF cf_const,
folded ℌ.cf_obj_cf_comma_proj_def
]
)
qed
subsection‹Comma functors›
subsubsection‹Definition and elementary properties›
text‹See Theorem 1 in Chapter X-3 in \<^cite>‹"mac_lane_categories_2010"›.›
definition cf_arr_cf_comma :: "V ⇒ V ⇒ V"
(‹(_ ⇩A↓⇩C⇩F _)› [1000, 1000] 999)
where "g ⇩A↓⇩C⇩F 𝔉 =
[
(λA∈⇩∘(𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉⦇Obj⦈. [0, A⦇1⇩ℕ⦈, A⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘),
(
λF∈⇩∘(𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉⦇Arr⦈.
[
[0, F⦇0⦈⦇1⇩ℕ⦈, F⦇0⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘,
[0, F⦇1⇩ℕ⦈⦇1⇩ℕ⦈, F⦇1⇩ℕ⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
),
(𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉,
(𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈) ↓⇩C⇩F 𝔉
]⇩∘"
definition cf_cf_arr_comma :: "V ⇒ V ⇒ V"
(‹(_ ⇩C⇩F↓⇩A _)› [1000, 1000] 999)
where "𝔉 ⇩C⇩F↓⇩A g =
[
(λA∈⇩∘𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈)⦇Obj⦈. [A⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ A⦇2⇩ℕ⦈]⇩∘),
(
λF∈⇩∘𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈)⦇Arr⦈.
[
[F⦇0⦈⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ F⦇0⦈⦇2⇩ℕ⦈]⇩∘,
[F⦇1⇩ℕ⦈⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ F⦇1⇩ℕ⦈⦇2⇩ℕ⦈]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
),
𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈),
𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈)
]⇩∘"
text‹Components.›
lemma cf_arr_cf_comma_components:
shows "g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈ =
(λA∈⇩∘(𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉⦇Obj⦈. [0, A⦇1⇩ℕ⦈, A⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘)"
and "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈ =
(
λF∈⇩∘(𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉⦇Arr⦈.
[
[0, F⦇0⦈⦇1⇩ℕ⦈, F⦇0⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘,
[0, F⦇1⇩ℕ⦈⦇1⇩ℕ⦈, F⦇1⇩ℕ⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔉⦇HomCod⦈⇙ g]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
)"
and "g ⇩A↓⇩C⇩F 𝔉⦇HomDom⦈ = (𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈) ↓⇩C⇩F 𝔉"
and "g ⇩A↓⇩C⇩F 𝔉⦇HomCod⦈ = (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈) ↓⇩C⇩F 𝔉"
unfolding cf_arr_cf_comma_def dghm_field_simps
by (simp_all add: nat_omega_simps)
lemma cf_cf_arr_comma_components:
shows "𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈ =
(λA∈⇩∘𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈)⦇Obj⦈. [A⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ A⦇2⇩ℕ⦈]⇩∘)"
and "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈ =
(
λF∈⇩∘𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈)⦇Arr⦈.
[
[F⦇0⦈⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ F⦇0⦈⦇2⇩ℕ⦈]⇩∘,
[F⦇1⇩ℕ⦈⦇0⦈, 0, g ∘⇩A⇘𝔉⦇HomCod⦈⇙ F⦇1⇩ℕ⦈⦇2⇩ℕ⦈]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
)"
and "𝔉 ⇩C⇩F↓⇩A g⦇HomDom⦈ = 𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Dom⦈⦇g⦈)"
and "𝔉 ⇩C⇩F↓⇩A g⦇HomCod⦈ = 𝔉 ⇩C⇩F↓ (𝔉⦇HomCod⦈⦇Cod⦈⦇g⦈)"
unfolding cf_cf_arr_comma_def dghm_field_simps
by (simp_all add: nat_omega_simps)
context is_functor
begin
lemma cf_arr_cf_comma_components':
assumes "g : c ↦⇘𝔅⇙ c'"
shows "g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈ = (λA∈⇩∘c' ↓⇩C⇩F 𝔉⦇Obj⦈. [0, A⦇1⇩ℕ⦈, A⦇2⇩ℕ⦈ ∘⇩A⇘𝔅⇙ g]⇩∘)"
and "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈ =
(
λF∈⇩∘c' ↓⇩C⇩F 𝔉⦇Arr⦈.
[
[0, F⦇0⦈⦇1⇩ℕ⦈, F⦇0⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔅⇙ g]⇩∘,
[0, F⦇1⇩ℕ⦈⦇1⇩ℕ⦈, F⦇1⇩ℕ⦈⦇2⇩ℕ⦈ ∘⇩A⇘𝔅⇙ g]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
)"
and [cat_comma_cs_simps]: "g ⇩A↓⇩C⇩F 𝔉⦇HomDom⦈ = c' ↓⇩C⇩F 𝔉"
and [cat_comma_cs_simps]: "g ⇩A↓⇩C⇩F 𝔉⦇HomCod⦈ = c ↓⇩C⇩F 𝔉"
using assms
unfolding cf_arr_cf_comma_components
by (simp_all add: cat_cs_simps)
lemma cf_cf_arr_comma_components':
assumes "g : c ↦⇘𝔅⇙ c'"
shows "𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈ = (λA∈⇩∘𝔉 ⇩C⇩F↓ c⦇Obj⦈. [A⦇0⦈, 0, g ∘⇩A⇘𝔅⇙ A⦇2⇩ℕ⦈]⇩∘)"
and "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈ =
(
λF∈⇩∘𝔉 ⇩C⇩F↓ c⦇Arr⦈.
[
[F⦇0⦈⦇0⦈, 0, g ∘⇩A⇘𝔅⇙ F⦇0⦈⦇2⇩ℕ⦈]⇩∘,
[F⦇1⇩ℕ⦈⦇0⦈, 0, g ∘⇩A⇘𝔅⇙ F⦇1⇩ℕ⦈⦇2⇩ℕ⦈]⇩∘,
F⦇2⇩ℕ⦈
]⇩∘
)"
and [cat_comma_cs_simps]: "𝔉 ⇩C⇩F↓⇩A g⦇HomDom⦈ = 𝔉 ⇩C⇩F↓ c"
and [cat_comma_cs_simps]: "𝔉 ⇩C⇩F↓⇩A g⦇HomCod⦈ = 𝔉 ⇩C⇩F↓ c'"
using assms
unfolding cf_cf_arr_comma_components
by (simp_all add: cat_cs_simps)
end
lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_components'(3,4)
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_components'(3,4)
subsubsection‹Object map›
mk_VLambda cf_arr_cf_comma_components(1)[unfolded VLambda_vid_on[symmetric]]
|vsv cf_arr_cf_comma_ObjMap_vsv[cat_comma_cs_intros]|
mk_VLambda cf_cf_arr_comma_components(1)[unfolded VLambda_vid_on[symmetric]]
|vsv cf_cf_arr_comma_ObjMap_vsv[cat_comma_cs_intros]|
context is_functor
begin
context
fixes g c c'
assumes g: "g : c ↦⇘𝔅⇙ c'"
begin
mk_VLambda
cf_arr_cf_comma_components'(1)[OF g, unfolded VLambda_vid_on[symmetric]]
|vdomain cf_arr_cf_comma_ObjMap_vdomain[cat_comma_cs_simps]|
mk_VLambda
cf_cf_arr_comma_components'(1)[OF g, unfolded VLambda_vid_on[symmetric]]
|vdomain cf_cf_arr_comma_ObjMap_vdomain[cat_comma_cs_simps]|
end
end
lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ObjMap_vdomain
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ObjMap_vdomain
lemma (in is_functor) cf_arr_cf_comma_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a', b', f']⇩∘" and "A ∈⇩∘ c' ↓⇩C⇩F 𝔉⦇Obj⦈" and "g : c ↦⇘𝔅⇙ c'"
shows "g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈ = [a', b', f' ∘⇩A⇘𝔅⇙ g]⇩∘"
proof-
from assms have b': "b' ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f' : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
and a'_def: "a' = 0"
by auto
from assms(2) show ?thesis
unfolding cf_arr_cf_comma_components'[OF assms(3)] assms(1)
by (simp add: nat_omega_simps a'_def)
qed
lemma (in is_functor) cf_cf_arr_comma_ObjMap_app[cat_comma_cs_simps]:
assumes "A = [a', b', f']⇩∘" and "A ∈⇩∘ 𝔉 ⇩C⇩F↓ c⦇Obj⦈" and "g : c ↦⇘𝔅⇙ c'"
shows "𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈⦇A⦈ = [a', b', g ∘⇩A⇘𝔅⇙ f']⇩∘"
proof-
from assms have b'_def: "b' = 0"
and f: "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ c"
and a': "a' ∈⇩∘ 𝔄⦇Obj⦈"
by auto
from assms(2) show ?thesis
unfolding cf_cf_arr_comma_components'[OF assms(3)] assms(1)
by (simp add: nat_omega_simps b'_def)
qed
lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ObjMap_app
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ObjMap_app
lemma (in is_functor) cf_arr_cf_comma_ObjMap_vrange:
assumes "g : c ↦⇘𝔅⇙ c'"
shows "ℛ⇩∘ (g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈) ⊆⇩∘ c ↓⇩C⇩F 𝔉⦇Obj⦈"
proof
(
rule vsv.vsv_vrange_vsubset,
unfold cf_arr_cf_comma_ObjMap_vdomain[OF assms]
)
fix A assume "A ∈⇩∘ c' ↓⇩C⇩F 𝔉⦇Obj⦈"
with assms is_functor_axioms obtain a f
where A_def: "A = [0, a, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
by auto
from assms a f show "g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈ ∈⇩∘ c ↓⇩C⇩F 𝔉⦇Obj⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps A_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
lemma (in is_functor) cf_cf_arr_comma_ObjMap_vrange:
assumes "g : c ↦⇘𝔅⇙ c'"
shows "ℛ⇩∘ (𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈) ⊆⇩∘ 𝔉 ⇩C⇩F↓ c'⦇Obj⦈"
proof
(
rule vsv.vsv_vrange_vsubset,
unfold cf_cf_arr_comma_ObjMap_vdomain[OF assms]
)
fix A assume "A ∈⇩∘ 𝔉 ⇩C⇩F↓ c⦇Obj⦈"
with assms is_functor_axioms obtain a f
where A_def: "A = [a, 0, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
by auto
from assms a f show "𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈⦇A⦈ ∈⇩∘ 𝔉 ⇩C⇩F↓ c'⦇Obj⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps A_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
subsubsection‹Arrow map›
mk_VLambda cf_arr_cf_comma_components(2)
|vsv cf_arr_cf_comma_ArrMap_vsv[cat_comma_cs_intros]|
mk_VLambda cf_cf_arr_comma_components(2)
|vsv cf_cf_arr_comma_ArrMap_vsv[cat_comma_cs_intros]|
context is_functor
begin
context
fixes g c c'
assumes g: "g : c ↦⇘𝔅⇙ c'"
begin
mk_VLambda
cf_arr_cf_comma_components'(2)[OF g, unfolded VLambda_vid_on[symmetric]]
|vdomain cf_arr_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]|
mk_VLambda
cf_cf_arr_comma_components'(2)[OF g, unfolded VLambda_vid_on[symmetric]]
|vdomain cf_cf_arr_comma_ArrMap_vdomain[cat_comma_cs_simps]|
end
end
lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ArrMap_vdomain
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ArrMap_vdomain
lemma (in is_functor) cf_arr_cf_comma_ArrMap_app[cat_comma_cs_simps]:
assumes "A = [[a, b, f]⇩∘, [a', b', f']⇩∘, [h, k]⇩∘]⇩∘"
and "[[a, b, f]⇩∘, [a', b', f']⇩∘, [h, k]⇩∘]⇩∘ :
[a, b, f]⇩∘ ↦⇘c' ↓⇩C⇩F 𝔉⇙ [a', b', f']⇩∘"
and "g : c ↦⇘𝔅⇙ c'"
shows "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇A⦈ =
[[a, b, f ∘⇩A⇘𝔅⇙ g]⇩∘, [a', b', f' ∘⇩A⇘𝔅⇙ g]⇩∘, [h, k]⇩∘]⇩∘"
proof-
from assms(3) have c': "c' ∈⇩∘ 𝔅⦇Obj⦈" by auto
from
cat_obj_cf_comma_is_arrD(1,2)[OF assms(2)[unfolded cat_comma_cs_simps] c']
is_arrD(1)[OF assms(2)]
show ?thesis
unfolding assms(1) cf_arr_cf_comma_components'[OF assms(3)]
by (simp_all add: nat_omega_simps)
qed
lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ArrMap_app
lemma (in is_functor) cf_cf_arr_comma_ArrMap_app[cat_comma_cs_simps]:
assumes "A = [[a, b, f]⇩∘, [a', b', f']⇩∘, [h, k]⇩∘]⇩∘"
and "[[a, b, f]⇩∘, [a', b', f']⇩∘, [h, k]⇩∘]⇩∘ :
[a, b, f]⇩∘ ↦⇘𝔉 ⇩C⇩F↓ c⇙ [a', b', f']⇩∘"
and "g : c ↦⇘𝔅⇙ c'"
shows "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇A⦈ =
[[a, b, g ∘⇩A⇘𝔅⇙ f]⇩∘, [a', b', g ∘⇩A⇘𝔅⇙ f']⇩∘, [h, k]⇩∘]⇩∘"
proof-
from assms(3) have c: "c ∈⇩∘ 𝔅⦇Obj⦈" by auto
from
cat_cf_obj_comma_is_arrD(1,2)[OF assms(2)[unfolded cat_comma_cs_simps] c]
is_arrD(1)[OF assms(2)]
show ?thesis
unfolding assms(1) cf_cf_arr_comma_components'[OF assms(3)]
by (simp_all add: nat_omega_simps)
qed
lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ArrMap_app
subsubsection‹Comma functors are functors›
lemma (in is_functor) cf_arr_cf_comma_is_functor:
assumes "g : c ↦⇘𝔅⇙ c'"
shows "g ⇩A↓⇩C⇩F 𝔉 : c' ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ c ↓⇩C⇩F 𝔉"
proof(rule is_functorI')
show "vfsequence (g ⇩A↓⇩C⇩F 𝔉)" unfolding cf_arr_cf_comma_def by simp
from assms show "category α (c' ↓⇩C⇩F 𝔉)"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms show "category α (c ↓⇩C⇩F 𝔉)"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
show "vcard (g ⇩A↓⇩C⇩F 𝔉) = 4⇩ℕ"
unfolding cf_arr_cf_comma_def by (simp_all add: nat_omega_simps)
from assms show "ℛ⇩∘ (g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈) ⊆⇩∘ c ↓⇩C⇩F 𝔉⦇Obj⦈"
by (intro cf_arr_cf_comma_ObjMap_vrange)
show "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇F⦈ :
g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈ ↦⇘c ↓⇩C⇩F 𝔉⇙ g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇B⦈"
if "F : A ↦⇘c' ↓⇩C⇩F 𝔉⇙ B" for A B F
proof-
from assms that obtain b f b' f' k
where F_def: "F = [[0, b, f]⇩∘, [0, b', f']⇩∘, [0, k]⇩∘]⇩∘"
and A_def: "A = [0, b, f]⇩∘"
and B_def: "B = [0, b', f']⇩∘"
and k: "k : b ↦⇘𝔄⇙ b'"
and f: "f : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b⦈"
and f': "f' : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
and f'_def: "𝔉⦇ArrMap⦈⦇k⦈ ∘⇩A⇘𝔅⇙ f = f'"
by auto
from assms that k f f' show ?thesis
unfolding F_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps f'_def[symmetric]
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
show "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇G ∘⇩A⇘c' ↓⇩C⇩F 𝔉⇙ F⦈ =
g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇G⦈ ∘⇩A⇘c ↓⇩C⇩F 𝔉⇙ g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇F⦈"
if "G : B ↦⇘c' ↓⇩C⇩F 𝔉⇙ C" and "F : A ↦⇘c' ↓⇩C⇩F 𝔉⇙ B" for B C G A F
proof-
from that(2) assms obtain b f b' f' k
where F_def: "F = [[0, b, f]⇩∘, [0, b', f']⇩∘, [0, k]⇩∘]⇩∘"
and A_def: "A = [0, b, f]⇩∘"
and B_def: "B = [0, b', f']⇩∘"
and k: "k : b ↦⇘𝔄⇙ b'"
and f: "f : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b⦈"
and f': "f' : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
and f'_def: "𝔉⦇ArrMap⦈⦇k⦈ ∘⇩A⇘𝔅⇙ f = f'"
by auto
with that(1) assms obtain b'' f'' k'
where G_def: "G = [[0, b', f']⇩∘, [0, b'', f'']⇩∘, [0, k']⇩∘]⇩∘"
and C_def: "C = [0, b'', f'']⇩∘"
and k': "k' : b' ↦⇘𝔄⇙ b''"
and f'': "f'' : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b''⦈"
and f''_def: "𝔉⦇ArrMap⦈⦇k'⦈ ∘⇩A⇘𝔅⇙ f' = f''"
by auto
from assms that k f f' f'' k' show ?thesis
unfolding F_def G_def A_def B_def C_def
by
(
cs_concl
cs_simp:
cat_cs_simps cat_comma_cs_simps
f''_def[symmetric] f'_def[symmetric]
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
show "g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇c' ↓⇩C⇩F 𝔉⦇CId⦈⦇C⦈⦈ = c ↓⇩C⇩F 𝔉⦇CId⦈⦇g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇C⦈⦈"
if "C ∈⇩∘ c' ↓⇩C⇩F 𝔉⦇Obj⦈" for C
proof-
from that assms obtain a f
where C_def: "C = [0, a, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : c' ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a⦈"
by auto
from a assms f show
"g ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇c' ↓⇩C⇩F 𝔉⦇CId⦈⦇C⦈⦈ = c ↓⇩C⇩F 𝔉⦇CId⦈⦇g ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇C⦈⦈"
unfolding C_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
qed
(
use assms in
‹
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
›
)+
lemma (in is_functor) cf_cf_arr_comma_is_functor:
assumes "g : c ↦⇘𝔅⇙ c'"
shows "𝔉 ⇩C⇩F↓⇩A g : 𝔉 ⇩C⇩F↓ c ↦↦⇩C⇘α⇙ 𝔉 ⇩C⇩F↓ c'"
proof(rule is_functorI')
from assms have c: "c ∈⇩∘ 𝔅⦇Obj⦈" by auto
show "vfsequence (𝔉 ⇩C⇩F↓⇩A g)" unfolding cf_cf_arr_comma_def by simp
from assms show "category α (𝔉 ⇩C⇩F↓ c')"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms show "category α (𝔉 ⇩C⇩F↓ c)"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
show "vcard (𝔉 ⇩C⇩F↓⇩A g) = 4⇩ℕ"
unfolding cf_cf_arr_comma_def by (simp_all add: nat_omega_simps)
from assms show "ℛ⇩∘ (𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈) ⊆⇩∘ 𝔉 ⇩C⇩F↓ c'⦇Obj⦈"
by (intro cf_cf_arr_comma_ObjMap_vrange)
show "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇F⦈ :
𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈⦇A⦈ ↦⇘𝔉 ⇩C⇩F↓ c'⇙ 𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈⦇B⦈"
if "F : A ↦⇘𝔉 ⇩C⇩F↓ c⇙ B" for A B F
proof-
from assms that obtain a f a' f' h
where F_def: "F = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [h, 0]⇩∘]⇩∘"
and A_def: "A = [a, 0, f]⇩∘"
and B_def: "B = [a', 0, f']⇩∘"
and h: "h : a ↦⇘𝔄⇙ a'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ c"
and f'_def: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇h⦈ = f"
by auto
from assms that h f f' show ?thesis
unfolding F_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps f'_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
show "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇G ∘⇩A⇘𝔉 ⇩C⇩F↓ c⇙ F⦈ =
𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇G⦈ ∘⇩A⇘𝔉 ⇩C⇩F↓ c'⇙ 𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇F⦈"
if "G : B ↦⇘𝔉 ⇩C⇩F↓ c⇙ C" and "F : A ↦⇘𝔉 ⇩C⇩F↓ c⇙ B" for B C G A F
proof-
from that(2) assms obtain a f a' f' h
where F_def: "F = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [h, 0]⇩∘]⇩∘"
and A_def: "A = [a, 0, f]⇩∘"
and B_def: "B = [a', 0, f']⇩∘"
and h: "h : a ↦⇘𝔄⇙ a'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ c"
and [cat_cs_simps]: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇h⦈ = f"
by auto
with that(1) assms obtain a'' f'' h'
where G_def: "G = [[a', 0, f']⇩∘, [a'', 0, f'']⇩∘, [h', 0]⇩∘]⇩∘"
and C_def: "C = [a'', 0, f'']⇩∘"
and h': "h' : a' ↦⇘𝔄⇙ a''"
and f'': "f'' : 𝔉⦇ObjMap⦈⦇a''⦈ ↦⇘𝔅⇙ c"
and [cat_cs_simps]: "f'' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇h'⦈ = f'"
by auto
note [cat_cs_simps] = category.cat_assoc_helper[
where ℭ=𝔅, where h=f'' and g=‹𝔉⦇ArrMap⦈⦇h'⦈› and q=f'
]
from assms that c h f f' f'' h' show ?thesis
unfolding F_def G_def A_def B_def C_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
show "𝔉 ⇩C⇩F↓⇩A g⦇ArrMap⦈⦇𝔉 ⇩C⇩F↓ c⦇CId⦈⦇C⦈⦈ = 𝔉 ⇩C⇩F↓ c'⦇CId⦈⦇𝔉 ⇩C⇩F↓⇩A g⦇ObjMap⦈⦇C⦈⦈"
if "C ∈⇩∘ 𝔉 ⇩C⇩F↓ c⦇Obj⦈" for C
proof-
from that assms obtain a f
where C_def: "C = [a, 0, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
by auto
from a c assms f show ?thesis
unfolding C_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
qed
(
use assms in
‹
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
›
)+
lemma (in is_functor) cf_arr_cf_comma_is_functor'[cat_comma_cs_intros]:
assumes "g : c ↦⇘𝔅⇙ c'" and "𝔄' = c' ↓⇩C⇩F 𝔉" and "𝔅' = c ↓⇩C⇩F 𝔉"
shows "g ⇩A↓⇩C⇩F 𝔉 : 𝔄' ↦↦⇩C⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule cf_arr_cf_comma_is_functor(1))
lemmas [cat_comma_cs_intros] = is_functor.cf_arr_cf_comma_is_functor'
lemma (in is_functor) cf_cf_arr_comma_is_functor'[cat_comma_cs_intros]:
assumes "g : c ↦⇘𝔅⇙ c'" and "𝔄' = 𝔉 ⇩C⇩F↓ c" and "𝔅' = 𝔉 ⇩C⇩F↓ c'"
shows "𝔉 ⇩C⇩F↓⇩A g : 𝔄' ↦↦⇩C⇘α⇙ 𝔅'"
using assms(1) unfolding assms(2,3) by (rule cf_cf_arr_comma_is_functor(1))
lemmas [cat_comma_cs_intros] = is_functor.cf_cf_arr_comma_is_functor'
lemma (in is_functor) cf_arr_cf_comma_CId:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉 = cf_id (b ↓⇩C⇩F 𝔉)"
proof(rule cf_eqI)
from vempty_is_zet assms show "cf_id (b ↓⇩C⇩F 𝔉) : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ b ↓⇩C⇩F 𝔉"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
from vempty_is_zet assms show "(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉 : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ b ↓⇩C⇩F 𝔉"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms have ObjMap_dom_lhs:
"𝒟⇩∘ ((𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
from assms have ObjMap_dom_rhs:
"𝒟⇩∘ (cf_id (b ↓⇩C⇩F 𝔉)⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
show "(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈ = cf_id (b ↓⇩C⇩F 𝔉)⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
fix A assume prems: "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
with assms obtain a' f'
where A_def: "A = [0, a', f']⇩∘"
and a': "a' ∈⇩∘ 𝔄⦇Obj⦈"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇a'⦈"
by auto
from prems assms vempty_is_zet a' f' show
"(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈ = cf_id (b ↓⇩C⇩F 𝔉)⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros
)
qed (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)+
from assms have ArrMap_dom_lhs:
"𝒟⇩∘ ((𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
from assms have ArrMap_dom_rhs:
"𝒟⇩∘ (cf_id (b ↓⇩C⇩F 𝔉)⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
show "(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈ = cf_id (b ↓⇩C⇩F 𝔉)⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix F assume prems: "F ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B" by (auto dest: is_arrI)
from assms F obtain b' f' b'' f'' h
where F_def: "F = [[0, b', f']⇩∘, [0, b'', f'']⇩∘, [0, h]⇩∘]⇩∘"
and A_def: "A = [0, b', f']⇩∘"
and B_def: "B = [0, b'', f'']⇩∘"
and h: "h : b' ↦⇘𝔄⇙ b''"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
and f'': "f'' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b''⦈"
and "𝔉⦇ArrMap⦈⦇h⦈ ∘⇩A⇘𝔅⇙ f' = f''"
by auto
from assms prems F h f' f'' show
"(𝔅⦇CId⦈⦇b⦈) ⇩A↓⇩C⇩F 𝔉⦇ArrMap⦈⦇F⦈ = cf_id (b ↓⇩C⇩F 𝔉)⦇ArrMap⦈⦇F⦈"
unfolding F_def A_def B_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros
)
qed (cs_concl cs_shallow cs_intro: cat_comma_cs_intros cat_cs_intros)+
qed simp_all
lemma (in is_functor) cf_cf_arr_comma_CId:
assumes "b ∈⇩∘ 𝔅⦇Obj⦈"
shows "𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈) = cf_id (𝔉 ⇩C⇩F↓ b)"
proof(rule cf_eqI)
from vempty_is_zet assms show "cf_id (𝔉 ⇩C⇩F↓ b) : 𝔉 ⇩C⇩F↓ b ↦↦⇩C⇘α⇙ 𝔉 ⇩C⇩F↓ b"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
from vempty_is_zet assms show "𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈) : 𝔉 ⇩C⇩F↓ b ↦↦⇩C⇘α⇙ 𝔉 ⇩C⇩F↓ b"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms have ObjMap_dom_lhs:
"𝒟⇩∘ (𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
from assms have ObjMap_dom_rhs:
"𝒟⇩∘ (cf_id (𝔉 ⇩C⇩F↓ b)⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
show "𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ObjMap⦈ = cf_id (𝔉 ⇩C⇩F↓ b)⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
fix A assume prems: "A ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Obj⦈"
with assms obtain a' f'
where A_def: "A = [a', 0, f']⇩∘"
and a': "a' ∈⇩∘ 𝔄⦇Obj⦈"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
by auto
from prems assms vempty_is_zet a' f' show
"𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ObjMap⦈⦇A⦈ = cf_id (𝔉 ⇩C⇩F↓ b)⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros
)
qed (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)+
from assms have ArrMap_dom_lhs:
"𝒟⇩∘ (𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ArrMap⦈) = 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros)
from assms have ArrMap_dom_rhs:
"𝒟⇩∘ (cf_id (𝔉 ↓⇩C⇩F b)⦇ArrMap⦈) = 𝔉 ↓⇩C⇩F b⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
show "𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ArrMap⦈ = cf_id (𝔉 ⇩C⇩F↓ b)⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix F assume prems: "F ∈⇩∘ 𝔉 ⇩C⇩F↓ b⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘𝔉 ⇩C⇩F↓ b⇙ B" by (auto dest: is_arrI)
from assms F obtain a' f' a'' f'' k
where F_def: "F = [[a', 0, f']⇩∘, [a'', 0, f'']⇩∘, [k, 0]⇩∘]⇩∘"
and A_def: "A = [a', 0, f']⇩∘"
and B_def: "B = [a'', 0, f'']⇩∘"
and k: "k : a' ↦⇘𝔄⇙ a''"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ b"
and f'': "f'' : 𝔉⦇ObjMap⦈⦇a''⦈ ↦⇘𝔅⇙ b"
and [cat_cs_simps]: "f'' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇k⦈ = f'"
by auto
from assms prems F k f' f'' show
"𝔉 ⇩C⇩F↓⇩A (𝔅⦇CId⦈⦇b⦈)⦇ArrMap⦈⦇F⦈ = cf_id (𝔉 ⇩C⇩F↓ b)⦇ArrMap⦈⦇F⦈"
unfolding F_def A_def B_def
by
(
cs_concl cs_shallow
cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros
)
qed
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros
)+
qed simp_all
subsubsection‹Comma functors and projections›
lemma (in is_functor)
cf_cf_comp_cf_obj_cf_comma_proj_cf_arr_cf_comma[cat_comma_cs_simps]:
assumes "f : a ↦⇘𝔅⇙ b"
shows "a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉 = b ⇩O⨅⇩C⇩F 𝔉"
proof(rule cf_eqI)
from assms vempty_is_zet show "b ⇩O⨅⇩C⇩F 𝔉 : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms show "a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉 : b ↓⇩C⇩F 𝔉 ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms have ObjMap_dom_lhs:
"𝒟⇩∘ ((a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by
(
cs_concl
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms have ObjMap_dom_rhs: "𝒟⇩∘ (b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
show "(a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ObjMap⦈ = b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
from assms show "vsv (b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈)"
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
fix A assume prems: "A ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Obj⦈"
with assms obtain b' f'
where A_def: "A = [0, b', f']⇩∘"
and b': "b' ∈⇩∘ 𝔄⦇Obj⦈"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
by auto
from prems assms b' f' show
"(a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ObjMap⦈⦇A⦈ = b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
(
use assms vempty_is_zet in
‹cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros›
)
from assms have ArrMap_dom_lhs:
"𝒟⇩∘ ((a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by
(
cs_concl
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms vempty_is_zet have ArrMap_dom_rhs:
"𝒟⇩∘ (b ⇩O⨅⇩C⇩F 𝔉⦇ObjMap⦈) = b ↓⇩C⇩F 𝔉⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
from assms vempty_is_zet have ArrMap_dom_lhs:
"𝒟⇩∘ ((a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
by
(
cs_concl
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms have ArrMap_dom_rhs: "𝒟⇩∘ (b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈) = b ↓⇩C⇩F 𝔉⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
show "(a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ArrMap⦈ = b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix F assume "F ∈⇩∘ b ↓⇩C⇩F 𝔉⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘b ↓⇩C⇩F 𝔉⇙ B"
by (auto dest: is_arrI)
with assms obtain b' f' b'' f'' h
where F_def: "F = [[0, b', f']⇩∘, [0, b'', f'']⇩∘, [0, h]⇩∘]⇩∘"
and A_def: "A = [0, b', f']⇩∘"
and B_def: "B = [0, b'', f'']⇩∘"
and h: "h : b' ↦⇘𝔄⇙ b''"
and f': "f' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b'⦈"
and f'': "f'' : b ↦⇘𝔅⇙ 𝔉⦇ObjMap⦈⦇b''⦈"
and f''_def: "𝔉⦇ArrMap⦈⦇h⦈ ∘⇩A⇘𝔅⇙ f' = f''"
by auto
from vempty_is_zet h assms f' f'' F show
"(a ⇩O⨅⇩C⇩F 𝔉 ∘⇩C⇩F f ⇩A↓⇩C⇩F 𝔉)⦇ArrMap⦈⦇F⦈ = b ⇩O⨅⇩C⇩F 𝔉⦇ArrMap⦈⦇F⦈"
unfolding F_def A_def B_def
by
(
cs_concl
cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric]
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
qed
(
use assms vempty_is_zet in
‹cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros›
)
qed simp_all
lemma (in is_functor)
cf_cf_comp_cf_cf_obj_comma_proj_cf_cf_arr_comma[cat_comma_cs_simps]:
assumes "f : a ↦⇘𝔅⇙ b"
shows "𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f = 𝔉 ⇩C⇩F⨅⇩O a"
proof(rule cf_eqI)
from assms vempty_is_zet show "𝔉 ⇩C⇩F⨅⇩O a : 𝔉 ⇩C⇩F↓ a ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms show "𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f : 𝔉 ⇩C⇩F↓ a ↦↦⇩C⇘α⇙ 𝔄"
by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)
from assms have ObjMap_dom_lhs:
"𝒟⇩∘ ((𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ a⦇Obj⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms have ObjMap_dom_rhs: "𝒟⇩∘ (𝔉 ⇩C⇩F⨅⇩O a⦇ObjMap⦈) = 𝔉 ⇩C⇩F↓ a⦇Obj⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
show "(𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ObjMap⦈ = 𝔉 ⇩C⇩F⨅⇩O a⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs)
from assms show "vsv (𝔉 ⇩C⇩F⨅⇩O a⦇ObjMap⦈)"
by (cs_concl cs_shallow cs_intro: cat_comma_cs_intros)
fix A assume prems: "A ∈⇩∘ 𝔉 ⇩C⇩F↓ a⦇Obj⦈"
with assms obtain a' f'
where A_def: "A = [a', 0, f']⇩∘"
and b': "a' ∈⇩∘ 𝔄⦇Obj⦈"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ a"
by auto
from prems assms b' f' show
"(𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ObjMap⦈⦇A⦈ = 𝔉 ⇩C⇩F⨅⇩O a⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps
cs_intro: cat_cs_intros cat_comma_cs_intros
)
qed
(
use assms vempty_is_zet in
‹cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros›
)
from assms vempty_is_zet have ArrMap_dom_lhs:
"𝒟⇩∘ ((𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ArrMap⦈) = 𝔉 ⇩C⇩F↓ a⦇Arr⦈"
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms have ArrMap_dom_rhs: "𝒟⇩∘ (𝔉 ⇩C⇩F⨅⇩O a⦇ArrMap⦈) = 𝔉 ⇩C⇩F↓ a⦇Arr⦈"
by (cs_concl cs_shallow cs_simp: cat_comma_cs_simps)
show "(𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ArrMap⦈ = 𝔉 ⇩C⇩F⨅⇩O a⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs)
fix F assume "F ∈⇩∘ 𝔉 ⇩C⇩F↓ a⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘𝔉 ⇩C⇩F↓ a⇙ B" by (auto dest: is_arrI)
with assms obtain a' f' a'' f'' k
where F_def: "F = [[a', 0, f']⇩∘, [a'', 0, f'']⇩∘, [k, 0]⇩∘]⇩∘"
and A_def: "A = [a', 0, f']⇩∘"
and B_def: "B = [a'', 0, f'']⇩∘"
and k: "k : a' ↦⇘𝔄⇙ a''"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ a"
and f'': "f'' : 𝔉⦇ObjMap⦈⦇a''⦈ ↦⇘𝔅⇙ a"
and f'_def: "f'' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇k⦈ = f'"
by auto
from vempty_is_zet k assms f' f'' F show
"(𝔉 ⇩C⇩F⨅⇩O b ∘⇩C⇩F 𝔉 ⇩C⇩F↓⇩A f)⦇ArrMap⦈⦇F⦈ = 𝔉 ⇩C⇩F⨅⇩O a⦇ArrMap⦈⦇F⦈"
unfolding F_def A_def B_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps f'_def
cs_intro: cat_cs_intros cat_comma_cs_intros
)+
qed
(
use assms vempty_is_zet in
‹cs_concl cs_shallow cs_intro: cat_cs_intros cat_comma_cs_intros›
)
qed simp_all
subsubsection‹Opposite comma functors›
lemma (in is_functor) cf_op_cf_obj_comma_cf_arr_cf_comma:
assumes "g : c ↦⇘𝔅⇙ c'"
shows
"op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g) =
g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c"
proof(rule cf_eqI)
from assms interpret 𝔉c: category α ‹𝔉 ⇩C⇩F↓ c›
by
(
cs_concl
cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros
)
from assms have c: "c ∈⇩∘ 𝔅⦇Obj⦈" by auto
from assms show "op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g) :
op_cat (𝔉 ⇩C⇩F↓ c) ↦↦⇩C⇘α⇙ c' ↓⇩C⇩F (op_cf 𝔉)"
by
(
cs_concl
cs_simp: cat_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
then have ObjMap_dom_lhs:
"𝒟⇩∘ ((op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ObjMap⦈) =
(op_cat (𝔉 ⇩C⇩F↓ c))⦇Obj⦈"
and ArrMap_dom_lhs:
"𝒟⇩∘ ((op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ArrMap⦈) =
(op_cat (𝔉 ⇩C⇩F↓ c))⦇Arr⦈"
by (cs_concl cs_simp: cat_cs_simps)+
from assms show
"g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c :
op_cat (𝔉 ⇩C⇩F↓ c) ↦↦⇩C⇘α⇙ c' ↓⇩C⇩F (op_cf 𝔉)"
by
(
cs_concl
cs_simp: cat_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
then have ObjMap_dom_rhs:
"𝒟⇩∘ ((g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ObjMap⦈) =
(op_cat (𝔉 ⇩C⇩F↓ c))⦇Obj⦈"
and ArrMap_dom_rhs:
"𝒟⇩∘ ((g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ArrMap⦈) =
(op_cat (𝔉 ⇩C⇩F↓ c))⦇Arr⦈"
by (cs_concl cs_simp: cat_cs_simps)+
show
"(op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ObjMap⦈ =
(g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ObjMap⦈"
proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_op_simps)
fix A assume "A ∈⇩∘ 𝔉 ⇩C⇩F↓ c⦇Obj⦈"
with assms obtain a f
where A_def: "A = [a, 0, f]⇩∘"
and a: "a ∈⇩∘ 𝔄⦇Obj⦈"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
by auto
from assms a f show
"(op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ObjMap⦈⦇A⦈ =
(g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ObjMap⦈⦇A⦈"
unfolding A_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
use assms in
‹cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros›
)+
show
"(op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ArrMap⦈ =
(g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ArrMap⦈"
proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs cat_op_simps)
fix F assume "F ∈⇩∘ 𝔉 ⇩C⇩F↓ c⦇Arr⦈"
then obtain A B where F: "F : A ↦⇘𝔉 ⇩C⇩F↓ c⇙ B" by auto
with assms c obtain a f a' f' h
where F_def: "F = [[a, 0, f]⇩∘, [a', 0, f']⇩∘, [h, 0]⇩∘]⇩∘"
and A_def: "A = [a, 0, f]⇩∘"
and B_def: "B = [a', 0, f']⇩∘"
and h: "h : a ↦⇘𝔄⇙ a'"
and f: "f : 𝔉⦇ObjMap⦈⦇a⦈ ↦⇘𝔅⇙ c"
and f': "f' : 𝔉⦇ObjMap⦈⦇a'⦈ ↦⇘𝔅⇙ c"
and [cat_comma_cs_simps]: "f' ∘⇩A⇘𝔅⇙ 𝔉⦇ArrMap⦈⦇h⦈ = f"
by auto
from F assms h f f' c show
"(op_cf_obj_comma 𝔉 c' ∘⇩C⇩F op_cf (𝔉 ⇩C⇩F↓⇩A g))⦇ArrMap⦈⦇F⦈ =
(g ⇩A↓⇩C⇩F (op_cf 𝔉) ∘⇩C⇩F op_cf_obj_comma 𝔉 c)⦇ArrMap⦈⦇F⦈"
unfolding F_def A_def B_def
by
(
cs_concl cs_shallow
cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps
cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros
)
qed
(
use assms in
‹cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros›
)+
qed simp_all
text‹\newpage›
end