diff --git a/thys/CZH_Elementary_Categories/ROOT b/thys/CZH_Elementary_Categories/ROOT --- a/thys/CZH_Elementary_Categories/ROOT +++ b/thys/CZH_Elementary_Categories/ROOT @@ -1,15 +1,15 @@ chapter AFP session CZH_Elementary_Categories (AFP) = CZH_Foundations + - options [timeout = 1200] + options [timeout = 7200] directories czh_ecategories theories CZH_ECAT_Conclusions document_files "iman.sty" "extra.sty" "isar.sty" "style.sty" "root.tex" "root.bib" \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Discrete.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Discrete.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Discrete.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Discrete.thy @@ -1,722 +1,722 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Discrete category\ theory CZH_ECAT_Discrete imports CZH_ECAT_Simple CZH_ECAT_Small_Functor begin subsection\Abstract discrete category\ named_theorems cat_discrete_cs_simps named_theorems cat_discrete_cs_intros subsubsection\Definition and elementary properties\ text\See Chapter I-2 in \cite{mac_lane_categories_2010}.\ locale cat_discrete = category \ \ for \ \ + assumes cat_discrete_Arr: "f \\<^sub>\ \\Arr\ \ f \\<^sub>\ \\<^sub>\ (\\CId\)" text\Rules.\ lemma (in cat_discrete) assumes "\' = \" "\' = \" shows "cat_discrete \' \'" unfolding assms by (rule cat_discrete_axioms) mk_ide rf cat_discrete_def[unfolded cat_discrete_axioms_def] |intro cat_discreteI| |dest cat_discreteD[dest]| |elim cat_discreteE[elim]| lemmas [cat_discrete_cs_intros] = cat_discreteD(1) text\Elementary properties.\ lemma (in cat_discrete) cat_discrete_is_arrD[dest]: assumes "f : a \\<^bsub>\\<^esub> b" shows "b = a" and "f = \\CId\\a\" proof- from assms cat_discrete_Arr have "f \\<^sub>\ \\<^sub>\ (\\CId\)" by (auto simp: cat_cs_simps) with cat_CId_vdomain obtain a' where f_def: "f = \\CId\\a'\" and "a' \\<^sub>\ \\Obj\" by (blast dest: CId.vrange_atD) then have "f : a' \\<^bsub>\\<^esub> a'" by (auto intro: cat_CId_is_arr') with assms have "a = a'" and "b = a'" by blast+ with f_def show "b = a" and "f = \\CId\\a\" by auto qed lemma (in cat_discrete) cat_discrete_is_arrE[elim]: assumes "f : b \\<^bsub>\\<^esub> c" obtains a where "f : a \\<^bsub>\\<^esub> a" and "f = \\CId\\a\" using assms by auto subsection\The discrete category\ text\ As explained in Chapter I-2 in \cite{mac_lane_categories_2010}, every discrete category is identified with its set of objects. In this work, it is assumed that the set of objects and the set of arrows in the canonical discrete category coincide; the domain and the codomain functions are identities. \ subsubsection\Definition and elementary properties\ definition the_cat_discrete :: "V \ V" (\:\<^sub>C\) where ":\<^sub>C I = [I, I, vid_on I, vid_on I, (\fg\\<^sub>\fid_on I. fg\0\), vid_on I]\<^sub>\" text\Components.\ lemma the_cat_discrete_components: shows ":\<^sub>C I\Obj\ = I" and ":\<^sub>C I\Arr\ = I" and ":\<^sub>C I\Dom\ = vid_on I" and ":\<^sub>C I\Cod\ = vid_on I" and ":\<^sub>C I\Comp\ = (\fg\\<^sub>\fid_on I. fg\0\)" and ":\<^sub>C I\CId\ = vid_on I" unfolding the_cat_discrete_def dg_field_simps by (simp_all add: nat_omega_simps) subsubsection\Domain\ mk_VLambda the_cat_discrete_components(3)[folded VLambda_vid_on] |vsv the_cat_discrete_Dom_vsv[cat_discrete_cs_intros]| |vdomain the_cat_discrete_Dom_vdomain[cat_discrete_cs_simps]| |app the_cat_discrete_Dom_app[cat_discrete_cs_simps]| subsubsection\Codomain\ mk_VLambda the_cat_discrete_components(4)[folded VLambda_vid_on] |vsv the_cat_discrete_Cod_vsv[cat_discrete_cs_intros]| |vdomain the_cat_discrete_Cod_vdomain[cat_discrete_cs_simps]| |app the_cat_discrete_Cod_app[cat_discrete_cs_simps]| subsubsection\Composition\ lemma the_cat_discrete_Comp_vsv[cat_discrete_cs_intros]: "vsv (:\<^sub>C I\Comp\)" unfolding the_cat_discrete_components by simp lemma the_cat_discrete_Comp_vdomain: "\\<^sub>\ (:\<^sub>C I\Comp\) = fid_on I" unfolding the_cat_discrete_components by simp lemma the_cat_discrete_Comp_vrange: "\\<^sub>\ (:\<^sub>C I\Comp\) = I" proof(intro vsubset_antisym vsubsetI) fix f assume "f \\<^sub>\ \\<^sub>\ (:\<^sub>C I\Comp\)" then obtain gg where f_def: "f = :\<^sub>C I\Comp\\gg\" and gg: "gg \\<^sub>\ fid_on I" unfolding the_cat_discrete_components by auto from gg show "f \\<^sub>\ I" unfolding f_def the_cat_discrete_components by clarsimp next fix f assume "f \\<^sub>\ I" then have "[f, f]\<^sub>\ \\<^sub>\ fid_on I" by clarsimp moreover then have "f = :\<^sub>C I\Comp\\f, f\\<^sub>\" unfolding the_cat_discrete_components by simp ultimately show "f \\<^sub>\ \\<^sub>\ (:\<^sub>C I\Comp\)" unfolding the_cat_discrete_components by (metis rel_VLambda.vsv_vimageI2 vdomain_VLambda) qed lemma the_cat_discrete_Comp_app[cat_discrete_cs_simps]: assumes "i \\<^sub>\ I" shows "i \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> i = i" proof- from assms have "[i, i]\<^sub>\ \\<^sub>\ fid_on I" by clarsimp then show ?thesis unfolding the_cat_discrete_components by simp qed subsubsection\Identity\ mk_VLambda the_cat_discrete_components(6)[folded VLambda_vid_on] |vsv the_cat_discrete_CId_vsv[cat_discrete_cs_intros]| |vdomain the_cat_discrete_CId_vdomain[cat_discrete_cs_simps]| |app the_cat_discrete_CId_app[cat_discrete_cs_simps]| subsubsection\Arrow with a domain and a codomain\ lemma the_cat_discrete_is_arrI: assumes "i \\<^sub>\ I" shows "i : i \\<^bsub>:\<^sub>C I\<^esub> i" using assms unfolding is_arr_def the_cat_discrete_components by simp lemma the_cat_discrete_is_arrI'[cat_discrete_cs_intros]: assumes "i \\<^sub>\ I" and "a = i" and "b = i" shows "i : a \\<^bsub>:\<^sub>C I\<^esub> b" using assms(1) unfolding assms(2,3) by (rule the_cat_discrete_is_arrI) lemma the_cat_discrete_is_arrD: assumes "f : a \\<^bsub>:\<^sub>C I\<^esub> b" shows "f : f \\<^bsub>:\<^sub>C I\<^esub> f" and "a : a \\<^bsub>:\<^sub>C I\<^esub> a" and "b : b \\<^bsub>:\<^sub>C I\<^esub> b" and "f \\<^sub>\ I" and "a \\<^sub>\ I" and "b \\<^sub>\ I" and "f = a" and "f = b" and "b = a" using assms unfolding is_arr_def the_cat_discrete_components by force+ subsubsection\The discrete category is a discrete category\ lemma (in \) cat_discrete_the_cat_discrete: assumes "I \\<^sub>\ Vset \" shows "cat_discrete \ (:\<^sub>C I)" proof(intro cat_discreteI categoryI') show "vfsequence (:\<^sub>C I)" unfolding the_cat_discrete_def by simp show "vcard (:\<^sub>C I) = 6\<^sub>\" unfolding the_cat_discrete_def by (simp add: nat_omega_simps) show "gf \\<^sub>\ \\<^sub>\ (:\<^sub>C I\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>:\<^sub>C I\<^esub> c \ f : a \\<^bsub>:\<^sub>C I\<^esub> b)" for gf unfolding the_cat_discrete_Comp_vdomain proof assume "gf \\<^sub>\ fid_on I" then obtain a where "gf = [a, a]\<^sub>\" and "a \\<^sub>\ I" by clarsimp moreover then have "a : a \\<^bsub>:\<^sub>C I\<^esub> a" by (auto intro: the_cat_discrete_is_arrI) ultimately show "\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>:\<^sub>C I\<^esub> c \ f : a \\<^bsub>:\<^sub>C I\<^esub> b" by auto next assume "\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>:\<^sub>C I\<^esub> c \ f : a \\<^bsub>:\<^sub>C I\<^esub> b" then obtain g f b c a where gf_def: "gf = [g, f]\<^sub>\" and g: "g : b \\<^bsub>:\<^sub>C I\<^esub> c" and f: "f : a \\<^bsub>:\<^sub>C I\<^esub> b" by clarsimp then have "g = f" by (metis is_arrE the_cat_discrete_is_arrD(1)) with the_cat_discrete_is_arrD(4)[OF f] show "gf \\<^sub>\ fid_on I" unfolding gf_def by clarsimp qed show "g \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> f : a \\<^bsub>:\<^sub>C I\<^esub> c" if "g : b \\<^bsub>:\<^sub>C I\<^esub> c" and "f : a \\<^bsub>:\<^sub>C I\<^esub> b" for g b c f a proof- from that have fba: "f = a" "b = a" and a: "a \\<^sub>\ I" unfolding the_cat_discrete_is_arrD[OF that(2)] by (simp_all add: \a \\<^sub>\ I\) from that have gcb: "g = b" "c = b" unfolding the_cat_discrete_is_arrD[OF that(1)] by simp_all from a show ?thesis unfolding fba gcb by ( cs_concl cs_simp: cat_discrete_cs_simps cs_intro: cat_discrete_cs_intros ) qed show "h \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> g \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> f = h \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> (g \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> f)" if "h : c \\<^bsub>:\<^sub>C I\<^esub> d" and "g : b \\<^bsub>:\<^sub>C I\<^esub> c" and "f : a \\<^bsub>:\<^sub>C I\<^esub> b" for h c d g b f a proof- from that have fba: "f = a" "b = a" and a: "a \\<^sub>\ I" unfolding the_cat_discrete_is_arrD[OF that(3)] by (simp_all add: \a \\<^sub>\ I\) from that have gcb: "g = b" "c = b" unfolding the_cat_discrete_is_arrD[OF that(2)] by simp_all from that have hcd: "h = c" "d = c" unfolding the_cat_discrete_is_arrD[OF that(1)] by simp_all from a show ?thesis unfolding fba gcb hcd by (cs_concl cs_simp: cat_discrete_cs_simps) qed show ":\<^sub>C I\CId\\b\ \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> f = f" if "f : a \\<^bsub>:\<^sub>C I\<^esub> b" for f a b proof- from that have fba: "f = a" "b = a" and a: "a \\<^sub>\ I" unfolding the_cat_discrete_is_arrD[OF that] by (simp_all add: \a \\<^sub>\ I\) from a show ?thesis by (cs_concl cs_simp: cat_discrete_cs_simps fba) qed show "f \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> :\<^sub>C I\CId\\b\ = f" if "f : b \\<^bsub>:\<^sub>C I\<^esub> c" for f b c proof- from that have fba: "f = b" "c = b" and b: "b \\<^sub>\ I" unfolding the_cat_discrete_is_arrD[OF that] by (simp_all add: \b \\<^sub>\ I\) from b show ?thesis by (cs_concl cs_simp: cat_discrete_cs_simps fba) qed show ":\<^sub>C I\CId\\a\ : a \\<^bsub>:\<^sub>C I\<^esub> a" if "a \\<^sub>\ :\<^sub>C I\Obj\" for a using that by (auto simp: the_cat_discrete_components intro: cat_discrete_cs_intros) show "\\<^sub>\((\a\\<^sub>\A. \\<^sub>\(VLambda B (Hom (:\<^sub>C I) a) `\<^sub>\ B)) `\<^sub>\ A) \\<^sub>\ Vset \" if "A \\<^sub>\ :\<^sub>C I\Obj\" and "B \\<^sub>\ :\<^sub>C I\Obj\" and "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" for A B proof- have "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (:\<^sub>C I) a b) \\<^sub>\ A \\<^sub>\ B" proof(intro vsubsetI, elim vifunionE, unfold in_Hom_iff) fix i j f assume prems: "i \\<^sub>\ A" "j \\<^sub>\ B" "f : i \\<^bsub>:\<^sub>C I\<^esub> j" then show "f \\<^sub>\ A \\<^sub>\ B" unfolding the_cat_discrete_is_arrD[OF prems(3)] by simp qed moreover have "A \\<^sub>\ B \\<^sub>\ Vset \" by (simp add: that(3,4) vunion_in_VsetI) ultimately show "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (:\<^sub>C I) a b) \\<^sub>\ Vset \" by (auto simp: vsubset_in_VsetI) qed qed (auto simp: assms the_cat_discrete_components intro: cat_cs_intros) lemmas [cat_discrete_cs_intros] = \.cat_discrete_the_cat_discrete subsubsection\Opposite discrete category\ lemma (in \) the_cat_discrete_op[cat_op_simps]: assumes "I \\<^sub>\ Vset \" shows "op_cat (:\<^sub>C I) = :\<^sub>C I" proof(rule cat_eqI[of \]) from assms show dI: "category \ (:\<^sub>C I)" by (cs_concl cs_intro: cat_discrete_the_cat_discrete cat_discrete_cs_intros) then show op_dI: "category \ (op_cat (:\<^sub>C I))" by (cs_concl cs_intro: cat_op_intros) interpret category \ \op_cat (:\<^sub>C I)\ by (rule op_dI) show "op_cat (:\<^sub>C I)\Comp\ = :\<^sub>C I\Comp\" proof(rule vsv_eqI) show "\\<^sub>\ (op_cat (:\<^sub>C I)\Comp\) = \\<^sub>\ (:\<^sub>C I\Comp\)" by (simp add: the_cat_discrete_components op_cat_components) fix gf assume "gf \\<^sub>\ \\<^sub>\ (op_cat (:\<^sub>C I)\Comp\)" then have "gf \\<^sub>\ fid_on I" by (simp add: the_cat_discrete_components op_cat_components) then obtain h where gf_def: "gf = [h, h]\<^sub>\" and h: "h \\<^sub>\ I" by clarsimp from dI h show "op_cat (:\<^sub>C I)\Comp\\gf\ = :\<^sub>C I\Comp\\gf\" by ( cs_concl cs_simp: cat_op_simps gf_def cs_intro: cat_discrete_cs_intros ) qed (auto intro: cat_discrete_cs_intros) qed (unfold the_cat_discrete_components op_cat_components, simp_all) subsection\Discrete functor\ subsubsection\Local assumptions for the discrete functor\ text\See Chapter III in \cite{mac_lane_categories_2010}).\ locale cf_discrete = category \ \ for \ I F \ + assumes cf_discrete_selector_vrange[cat_discrete_cs_intros]: "i \\<^sub>\ I \ F i \\<^sub>\ \\Obj\" and cf_discrete_vdomain_vsubset_Vset: "I \\<^sub>\ Vset \" lemmas (in cf_discrete) cf_discrete_category = category_axioms lemmas [cat_discrete_cs_intros] = cf_discrete.cf_discrete_category text\Rules.\ lemma (in cf_discrete) cf_discrete_axioms'[cat_discrete_cs_intros]: assumes "\' = \" and "I' = I" and "F' = F" shows "cf_discrete \' I' F' \" unfolding assms by (rule cf_discrete_axioms) mk_ide rf cf_discrete_def[unfolded cf_discrete_axioms_def] |intro cf_discreteI| |dest cf_discreteD[dest]| |elim cf_discreteE[elim]| text\Elementary properties.\ lemma (in cf_discrete) cf_discrete_is_functor_cf_CId_selector_is_arr: assumes "i \\<^sub>\ I" shows "\\CId\\F i\ : F i \\<^bsub>\\<^esub> F i" using assms by (meson cat_CId_is_arr' cf_discreteD(2) cf_discrete_axioms) lemma (in cf_discrete) cf_discrete_is_functor_cf_CId_selector_is_arr'[cat_discrete_cs_intros]: assumes "i \\<^sub>\ I" and "a = F i" and "b = F i" shows "\\CId\\F i\ : a \\<^bsub>\\<^esub> b" using assms(1) unfolding assms(2,3) by (rule cf_discrete_is_functor_cf_CId_selector_is_arr) subsubsection\Definition and elementary properties\ definition the_cf_discrete :: "V \ (V \ V) \ V \ V" (\:\:\) where ":\: I F \ = [VLambda I F, (\i\\<^sub>\I. \\CId\\F i\), :\<^sub>C I, \]\<^sub>\" text\Components.\ lemma the_cf_discrete_components: shows ":\: I F \\ObjMap\ = (\i\\<^sub>\I. F i)" and ":\: I F \\ArrMap\ = (\i\\<^sub>\I. \\CId\\F i\)" and [cat_discrete_cs_simps]: ":\: I F \\HomDom\ = :\<^sub>C I" and [cat_discrete_cs_simps]: ":\: I F \\HomCod\ = \" unfolding the_cf_discrete_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda the_cf_discrete_components(1) |vsv the_cf_discrete_ObjMap_vsv[cat_discrete_cs_intros]| |vdomain the_cf_discrete_ObjMap_vdomain[cat_discrete_cs_simps]| |app the_cf_discrete_ObjMap_app[cat_discrete_cs_simps]| lemma (in cf_discrete) cf_discrete_the_cf_discrete_ObjMap_vrange: "\\<^sub>\ (:\: I F \\ObjMap\) \\<^sub>\ \\Obj\" using cf_discrete_is_functor_cf_CId_selector_is_arr unfolding the_cf_discrete_components by (intro vrange_VLambda_vsubset) auto subsubsection\Arrow map\ mk_VLambda the_cf_discrete_components(2) |vsv the_cf_discrete_ArrMap_vsv[cat_discrete_cs_intros]| |vdomain the_cf_discrete_ArrMap_vdomain[cat_discrete_cs_simps]| |app the_cf_discrete_ArrMap_app[cat_discrete_cs_simps]| lemma (in cf_discrete) cf_discrete_the_cf_discrete_ArrMap_vrange: "\\<^sub>\ (:\: I F \\ArrMap\) \\<^sub>\ \\Arr\" using cf_discrete_is_functor_cf_CId_selector_is_arr unfolding the_cf_discrete_components by (intro vrange_VLambda_vsubset) (auto simp: cf_discrete_selector_vrange) subsubsection\Discrete functor is a functor\ lemma (in cf_discrete) cf_discrete_the_cf_discrete_is_functor: ":\: I F \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_functorI') show "vfsequence (:\: I F \)" unfolding the_cf_discrete_def by simp show "category \ (:\<^sub>C I)" by ( simp add: cat_discrete_the_cat_discrete cf_discrete_vdomain_vsubset_Vset cat_discrete.axioms(1) ) show "vcard (:\: I F \) = 4\<^sub>\" unfolding the_cf_discrete_def by (simp add: nat_omega_simps) show ":\: I F \\ArrMap\\f\ : :\: I F \\ObjMap\\a\ \\<^bsub>\\<^esub> :\: I F \\ObjMap\\b\" if "f : a \\<^bsub>:\<^sub>C I\<^esub> b" for f a b proof- from that have fba: "f = a" "b = a" and a: "a \\<^sub>\ I" unfolding the_cat_discrete_is_arrD[OF that] by (simp_all add: \a \\<^sub>\ I\) from that \a \\<^sub>\ I\ show ?thesis by ( cs_concl cs_simp: cat_discrete_cs_simps fba cs_intro: cat_discrete_cs_intros ) qed show ":\: I F \\ArrMap\\g \\<^sub>A\<^bsub>:\<^sub>C I\<^esub> f\ = :\: I F \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> :\: I F \\ArrMap\\f\" if "g : b \\<^bsub>:\<^sub>C I\<^esub> c" and "f : a \\<^bsub>:\<^sub>C I\<^esub> b" for g b c f a proof- from that have gfacb: "f = a" "a = b" "g = b" "c = b" and b: "b \\<^sub>\ I" by ( simp_all add: the_cat_discrete_is_arrD(8-9)[OF that(1)] the_cat_discrete_is_arrD(5-9)[OF that(2)] ) have "F b \\<^sub>\ \\Obj\" by (simp add: b cf_discrete_selector_vrange) from b category_axioms this show ?thesis using that unfolding gfacb by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps cs_intro: cat_cs_intros ) qed show ":\: I F \\ArrMap\\:\<^sub>C I\CId\\c\\ = \\CId\\:\: I F \\ObjMap\\c\\" if "c \\<^sub>\ :\<^sub>C I\Obj\" for c using that unfolding the_cat_discrete_components(1) by (cs_concl cs_simp: cat_discrete_cs_simps cs_intro: cat_cs_intros) qed ( auto simp: the_cf_discrete_components the_cat_discrete_components cat_cs_intros cat_discrete_cs_intros ) lemma (in cf_discrete) cf_discrete_the_cf_discrete_is_functor': assumes "\' = :\<^sub>C I" and "\' = \" shows ":\: I F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule cf_discrete_the_cf_discrete_is_functor) lemmas [cat_discrete_cs_intros] = cf_discrete.cf_discrete_the_cf_discrete_is_functor' subsubsection\Uniqueness of the discrete category\ lemma (in cat_discrete) cat_discrete_iso_the_cat_discrete: assumes "I \\<^sub>\ Vset \" and "I \\<^sub>\ \\Obj\" obtains F where ":\: I F \ : :\<^sub>C I \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof- from assms obtain F where v11_f: "v11 F" and dr[simp]: "\\<^sub>\ F = I" "\\<^sub>\ F = \\Obj\" by auto let ?F = "\i. F\i\" interpret F: v11 F by (rule v11_f) from assms(1) interpret \: cf_discrete \ I ?F \ apply(intro cf_discreteI) unfolding dr[symmetric] by (cs_concl cs_intro: V_cs_intros cat_cs_intros)+ have ":\: I ?F \ : :\<^sub>C I \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_functorI') from \.cf_discrete_selector_vrange show ":\: I ?F \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" by (intro cf_discrete.cf_discrete_the_cf_discrete_is_functor cf_discreteI) (auto simp: category_axioms assms(1)) show "v11 (:\: I ?F \\ArrMap\)" proof(rule vsv.vsv_valeq_v11I, unfold the_cf_discrete_ArrMap_vdomain) fix i j assume prems: "i \\<^sub>\ I" "j \\<^sub>\ I" ":\: I ?F \\ArrMap\\i\ = :\: I ?F \\ArrMap\\j\" from prems(3) have "\\CId\\?F i\ = \\CId\\?F j\" unfolding the_cf_discrete_ArrMap_app[OF prems(1)] the_cf_discrete_ArrMap_app[OF prems(2)]. then have "?F i = ?F j" by ( metis \.cf_discrete_is_functor_cf_CId_selector_is_arr prems(1,2) cat_is_arrD(4) ) with F.v11_eq_iff prems show "i = j" by simp qed (simp add: the_cf_discrete_components) show "\\<^sub>\ (:\: I ?F \\ArrMap\) = \\Arr\" proof(intro vsubset_antisym vsubsetI) fix f assume "f \\<^sub>\ \\<^sub>\ (:\: I ?F \\ArrMap\)" with \.cf_discrete_the_cf_discrete_ArrMap_vrange show "f \\<^sub>\ \\Arr\" by auto next fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then obtain a where f_def: "f = \\CId\\a\" and a: "a \\<^sub>\ \\Obj\" by auto from a F.vrange_atD dr obtain i where a_def: "a = ?F i" and i: "i \\<^sub>\ I" by blast from a i show "f \\<^sub>\ \\<^sub>\ (:\: I ?F \\ArrMap\)" unfolding a_def f_def the_cf_discrete_components by auto qed qed (auto simp: v11_f the_cf_discrete_components) with that show ?thesis by simp qed subsubsection\Opposite discrete functor\ lemma (in cf_discrete) cf_discrete_the_cf_discrete_op[cat_op_simps]: "op_cf (:\: I F \) = :\: I F (op_cat \)" proof(rule cf_eqI) from cf_discrete_vdomain_vsubset_Vset show "op_cf (:\: I F \) : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by ( cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros cat_discrete_cs_intros ) show ":\: I F (op_cat \) : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" proof(intro cf_discrete.cf_discrete_the_cf_discrete_is_functor cf_discreteI) fix i assume "i \\<^sub>\ I" then show "F i \\<^sub>\ op_cat \\Obj\" by (simp add: cat_op_simps cf_discrete_selector_vrange) qed (intro cf_discrete_vdomain_vsubset_Vset cat_cs_intros)+ qed (unfold cat_op_simps the_cf_discrete_components, simp_all) lemmas [cat_op_simps] = cf_discrete.cf_discrete_the_cf_discrete_op lemma (in cf_discrete) cf_discrete_op[cat_op_intros]: "cf_discrete \ I F (op_cat \)" proof(intro cf_discreteI) show "category \ (op_cat \)" by (cs_concl cs_intro: cat_cs_intros) fix i assume "i \\<^sub>\ I" then show "F i \\<^sub>\ op_cat \\Obj\" by (cs_concl cs_simp: cat_op_simps cs_intro: cat_discrete_cs_intros) qed (intro cf_discrete_vdomain_vsubset_Vset) lemmas [cat_op_intros] = cf_discrete.cf_discrete_op subsection\Tiny discrete category\ subsubsection\Background\ named_theorems cat_small_discrete_cs_simps named_theorems cat_small_discrete_cs_intros lemmas [cat_small_discrete_cs_simps] = cat_discrete_cs_simps lemmas [cat_small_discrete_cs_intros] = cat_discrete_cs_intros subsubsection\Definition and elementary properties\ locale tiny_cat_discrete = cat_discrete \ \ + tiny_category \ \ for \ \ text\Rules.\ lemma (in tiny_cat_discrete) tiny_cat_discrete_axioms'[cat_discrete_cs_intros]: assumes "\' = \" and "\' = \" shows "tiny_cat_discrete \' \'" unfolding assms by (rule tiny_cat_discrete_axioms) mk_ide rf tiny_cat_discrete_def |intro tiny_cat_discreteI| |dest tiny_cat_discreteD[dest]| |elim tiny_cat_discreteE[elim]| lemmas [cat_small_discrete_cs_intros] = tiny_cat_discreteD lemma tiny_cat_discreteI': assumes "tiny_category \ \" and "\f. f \\<^sub>\ \\Arr\ \ f \\<^sub>\ \\<^sub>\ (\\CId\)" shows "tiny_cat_discrete \ \" proof(intro tiny_cat_discreteI cat_discreteI) interpret tiny_category \ \ by (rule assms(1)) - show "category \ \" by (auto intro: tiny_dg_category) + show "category \ \" by (auto intro: tiny_cat_category) show "f \\<^sub>\ \\<^sub>\ (\\CId\)" if "f \\<^sub>\ \\Arr\" for f using that by (rule assms(2)) qed (auto intro: assms(1)) subsubsection\The discrete category is a tiny category\ lemma (in \) tiny_cat_discrete_the_cat_discrete[cat_small_discrete_cs_intros]: assumes "I \\<^sub>\ Vset \" shows "tiny_cat_discrete \ (:\<^sub>C I)" proof(intro tiny_cat_discreteI cat_discrete_the_cat_discrete) from assms show "I \\<^sub>\ Vset \" by auto then interpret cat_discrete \ \:\<^sub>C I\ by (intro cat_discrete_the_cat_discrete) show "tiny_category \ (:\<^sub>C I)" by (intro tiny_categoryI', unfold the_cat_discrete_components) (auto intro: cat_cs_intros assms) qed lemmas [cat_small_discrete_cs_intros] = \.cat_discrete_the_cat_discrete subsection\Discrete functor with tiny maps\ subsubsection\Definition and elementary properties\ locale tm_cf_discrete = category \ \ for \ I F \ + assumes tm_cf_discrete_selector_vrange[cat_small_discrete_cs_intros]: "i \\<^sub>\ I \ F i \\<^sub>\ \\Obj\" and tm_cf_discrete_ObjMap_in_Vset: "VLambda I F \\<^sub>\ Vset \" and tm_cf_discrete_ArrMap_in_Vset: "(\i\\<^sub>\I. \\CId\\F i\) \\<^sub>\ Vset \" text\Rules.\ lemma (in tm_cf_discrete) tm_cf_discrete_axioms'[cat_small_discrete_cs_intros]: assumes "\' = \" and "I' = I" and "F' = F" shows "tm_cf_discrete \' I' F' \" unfolding assms by (rule tm_cf_discrete_axioms) mk_ide rf tm_cf_discrete_def[unfolded tm_cf_discrete_axioms_def] |intro tm_cf_discreteI| |dest tm_cf_discreteD[dest]| |elim tm_cf_discreteE[elim]| lemma tm_cf_discreteI': assumes "cf_discrete \ I F \" and "(\i\\<^sub>\I. F i) \\<^sub>\ Vset \" and "(\i\\<^sub>\I. \\CId\\F i\) \\<^sub>\ Vset \" shows "tm_cf_discrete \ I F \" proof- interpret cf_discrete \ I F \ by (rule assms(1)) show ?thesis by (intro tm_cf_discreteI) (auto intro: assms cf_discrete_selector_vrange cat_cs_intros) qed text\Elementary properties.\ sublocale tm_cf_discrete \ cf_discrete proof(intro cf_discreteI) from tm_cf_discrete_ObjMap_in_Vset have "\\<^sub>\ (\i\\<^sub>\I. F i) \\<^sub>\ Vset \" by (cs_concl cs_intro: vdomain_in_VsetI) then show "I \\<^sub>\ Vset \" by auto qed (auto intro: cat_cs_intros tm_cf_discrete_selector_vrange) lemmas (in tm_cf_discrete) tm_cf_discrete_is_cf_discrete_axioms = cf_discrete_axioms lemmas [cat_small_discrete_cs_intros] = tm_cf_discrete.tm_cf_discrete_is_cf_discrete_axioms lemma (in tm_cf_discrete) tm_cf_discrete_index_in_Vset[cat_small_discrete_cs_intros]: "I \\<^sub>\ Vset \" proof- from tm_cf_discrete_ObjMap_in_Vset have "\\<^sub>\ (\i\\<^sub>\I. F i) \\<^sub>\ Vset \" by (cs_concl cs_intro: vdomain_in_VsetI) then show ?thesis by simp qed subsubsection\Opposite discrete functor with tiny maps\ lemma (in tm_cf_discrete) tm_cf_discrete_op[cat_op_intros]: "tm_cf_discrete \ I F (op_cat \)" using tm_cf_discrete_ObjMap_in_Vset tm_cf_discrete_ArrMap_in_Vset by (intro tm_cf_discreteI' cf_discrete_op) (auto simp: cat_op_simps) lemmas [cat_op_intros] = tm_cf_discrete.tm_cf_discrete_op subsubsection\Discrete functor with tiny maps is a functor with tiny maps\ lemma (in tm_cf_discrete) tm_cf_discrete_the_cf_discrete_is_tm_functor: ":\: I F \ : :\<^sub>C I \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by (intro is_tm_functorI' cf_discrete_the_cf_discrete_is_functor) ( auto simp: the_cf_discrete_components tm_cf_discrete_ObjMap_in_Vset tm_cf_discrete_ArrMap_in_Vset ) lemma (in tm_cf_discrete) tm_cf_discrete_the_cf_discrete_is_tm_functor': assumes "\' = :\<^sub>C I" and "\' = \" shows ":\: I F \ : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \'" unfolding assms by (rule tm_cf_discrete_the_cf_discrete_is_tm_functor) lemmas [cat_discrete_cs_intros] = tm_cf_discrete.tm_cf_discrete_the_cf_discrete_is_tm_functor' text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_FUNCT.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_FUNCT.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_FUNCT.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_FUNCT.thy @@ -1,2883 +1,3012 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\FUNCT\ and \Funct\\ theory CZH_ECAT_FUNCT imports CZH_SMC_FUNCT CZH_ECAT_Subcategory CZH_ECAT_NTCF begin subsection\Background\ text\ The subsection presents the theory of the categories of \\\-functors between two \\\-categories. It continues the development that was initiated in sections \ref{sec:dg_FUNCT} and \ref{sec:smc_FUNCT}. A general reference for this section is Chapter II-4 in \cite{mac_lane_categories_2010}. \ named_theorems cat_FUNCT_cs_simps named_theorems cat_FUNCT_cs_intros lemmas [cat_FUNCT_cs_simps] = cat_map_cs_simps lemmas [cat_FUNCT_cs_intros] = cat_map_cs_intros subsection\\FUNCT\\ subsubsection\Definition and elementary properties\ definition cat_FUNCT :: "V \ V \ V \ V" where "cat_FUNCT \ \ \ = [ cf_maps \ \ \, ntcf_arrows \ \ \, (\\\\<^sub>\ntcf_arrows \ \ \. \\NTDom\), (\\\\<^sub>\ntcf_arrows \ \ \. \\NTCod\), (\\\\\<^sub>\composable_arrs (dg_FUNCT \ \ \). \\\0\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^bsub>\,\\<^esub> \\\1\<^sub>\\), (\\\\<^sub>\cf_maps \ \ \. ntcf_arrow_id \ \ \) ]\<^sub>\" text\Components.\ lemma cat_FUNCT_components: shows [cat_FUNCT_cs_simps]: "cat_FUNCT \ \ \\Obj\ = cf_maps \ \ \" and "cat_FUNCT \ \ \\Arr\ = ntcf_arrows \ \ \" and "cat_FUNCT \ \ \\Dom\ = (\\\\<^sub>\ntcf_arrows \ \ \. \\NTDom\)" and "cat_FUNCT \ \ \\Cod\ = (\\\\<^sub>\ntcf_arrows \ \ \. \\NTCod\)" and "cat_FUNCT \ \ \\Comp\ = (\\\\\<^sub>\composable_arrs (dg_FUNCT \ \ \). \\\0\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^bsub>\,\\<^esub> \\\1\<^sub>\\)" and "cat_FUNCT \ \ \\CId\ = (\\\\<^sub>\cf_maps \ \ \. ntcf_arrow_id \ \ \)" unfolding cat_FUNCT_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_FUNCT: "cat_smc (cat_FUNCT \ \ \) = smc_FUNCT \ \ \" proof(rule vsv_eqI) show "vsv (cat_smc (cat_FUNCT \ \ \))" unfolding cat_smc_def by auto show "vsv (smc_FUNCT \ \ \)" unfolding smc_FUNCT_def by auto have dom_lhs: "\\<^sub>\ (cat_smc (cat_FUNCT \ \ \)) = 5\<^sub>\" unfolding cat_smc_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smc_FUNCT \ \ \) = 5\<^sub>\" unfolding smc_FUNCT_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cat_smc (cat_FUNCT \ \ \)) = \\<^sub>\ (smc_FUNCT \ \ \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (cat_smc (cat_FUNCT \ \ \)) \ cat_smc (cat_FUNCT \ \ \)\a\ = smc_FUNCT \ \ \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cat_smc_def dg_field_simps cat_FUNCT_def smc_FUNCT_def ) (auto simp: nat_omega_simps) qed context is_ntcf begin lemmas_with [folded cat_smc_FUNCT, unfolded slicing_simps]: cat_FUNCT_Dom_app = smc_FUNCT_Dom_app and cat_FUNCT_Cod_app = smc_FUNCT_Cod_app end lemmas [smc_FUNCT_cs_simps] = is_ntcf.cat_FUNCT_Dom_app is_ntcf.cat_FUNCT_Cod_app lemmas_with [folded cat_smc_FUNCT, unfolded slicing_simps]: cat_FUNCT_Dom_vsv[intro] = smc_FUNCT_Dom_vsv and cat_FUNCT_Dom_vdomain[cat_FUNCT_cs_simps] = smc_FUNCT_Dom_vdomain and cat_FUNCT_Cod_vsv[intro] = smc_FUNCT_Cod_vsv and cat_FUNCT_Cod_vdomain[cat_FUNCT_cs_simps] = smc_FUNCT_Cod_vdomain and cat_FUNCT_Dom_vrange = smc_FUNCT_Dom_vrange and cat_FUNCT_Cod_vrange = smc_FUNCT_Cod_vrange and cat_FUNCT_is_arrI = smc_FUNCT_is_arrI and cat_FUNCT_is_arrI'[cat_FUNCT_cs_intros] = smc_FUNCT_is_arrI' and cat_FUNCT_is_arrD = smc_FUNCT_is_arrD and cat_FUNCT_is_arrE[elim] = smc_FUNCT_is_arrE lemmas_with [folded cat_smc_FUNCT, unfolded slicing_simps]: cat_FUNCT_Comp_app[cat_FUNCT_cs_simps] = smc_FUNCT_Comp_app subsubsection\Identity\ mk_VLambda cat_FUNCT_components(6) |vsv cat_FUNCT_CId_vsv[cat_FUNCT_cs_intros]| |vdomain cat_FUNCT_CId_vdomain[cat_FUNCT_cs_simps]| |app cat_FUNCT_CId_app[cat_FUNCT_cs_simps]| lemma smc_FUNCT_CId_vrange: "\\<^sub>\ (cat_FUNCT \ \ \\CId\) \\<^sub>\ ntcf_arrows \ \ \" unfolding cat_FUNCT_components proof(rule vrange_VLambda_vsubset) fix x assume "x \\<^sub>\ cf_maps \ \ \" then obtain \ where x_def: "x = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp then show "ntcf_arrow_id \ \ x \\<^sub>\ ntcf_arrows \ \ \" unfolding x_def by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed subsubsection\ The conversion of a natural transformation arrow to a natural transformation is a bijection \ lemma bij_betw_ntcf_of_ntcf_arrow: "bij_betw (ntcf_of_ntcf_arrow \ \) (elts (ntcf_arrows \ \ \)) (elts (ntcfs \ \ \))" proof(intro bij_betw_imageI inj_onI subset_antisym subsetI) fix \ \ assume prems: "\ \\<^sub>\ ntcf_arrows \ \ \" "\ \\<^sub>\ ntcf_arrows \ \ \" "ntcf_of_ntcf_arrow \ \ \ = ntcf_of_ntcf_arrow \ \ \" from prems(1) obtain \' \ \ where \_def: "\ = ntcf_arrow \'" and \': "\' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from prems(2) obtain \' \' \' where \_def: "\ = ntcf_arrow \'" and \': "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from prems(3) have "\' = \'" unfolding \_def \_def is_ntcf.ntcf_of_ntcf_arrow[OF \'] is_ntcf.ntcf_of_ntcf_arrow[OF \'] by simp then show "\ = \" unfolding \_def \_def by auto next fix \ assume "\ \ ntcf_of_ntcf_arrow \ \ ` elts (ntcf_arrows \ \ \)" then obtain \' where \': "\' \\<^sub>\ ntcf_arrows \ \ \" and \_def: "\ = ntcf_of_ntcf_arrow \ \ \'" by auto from \' obtain \'' \ \ where \'_def: "\' = ntcf_arrow \''" and \'': "\'' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \'' show "\ \\<^sub>\ ntcfs \ \ \" unfolding \_def \'_def is_ntcf.ntcf_of_ntcf_arrow[OF \''] by auto next fix \ assume "\ \\<^sub>\ ntcfs \ \ \" then obtain \ \ where \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp then have "\ = ntcf_of_ntcf_arrow \ \ (ntcf_arrow \)" by (cs_concl cs_simp: cat_FUNCT_cs_simps) moreover from \ have "ntcf_arrow \ \\<^sub>\ ntcf_arrows \ \ \" by (cs_concl cs_intro: cat_FUNCT_cs_intros) ultimately show "\ \ ntcf_of_ntcf_arrow \ \ ` elts (ntcf_arrows \ \ \)" by simp qed lemma bij_betw_ntcf_of_ntcf_arrow_Hom: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bij_betw (ntcf_of_ntcf_arrow \ \) (elts (Hom (cat_FUNCT \ \ \) (cf_map \) (cf_map \))) (elts (these_ntcfs \ \ \ \ \))" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms have [cat_cs_simps]: "cf_of_cf_map \ \ (cf_map \) = \" "cf_of_cf_map \ \ (cf_map \) = \" by (cs_concl cs_simp: cat_FUNCT_cs_simps)+ show ?thesis proof ( rule bij_betw_subset[OF bij_betw_ntcf_of_ntcf_arrow]; (intro subset_antisym subsetI)?; (unfold in_Hom_iff)? ) fix \ assume prems: "\ : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" note \ = cat_FUNCT_is_arrD[OF prems, unfolded cat_cs_simps] from \(1) show "\ \\<^sub>\ ntcf_arrows \ \ \" by (subst \(2)) (cs_concl cs_intro: cat_FUNCT_cs_intros) next fix \ assume "\ \ ntcf_of_ntcf_arrow \ \ ` elts (Hom (cat_FUNCT \ \ \) (cf_map \) (cf_map \))" then obtain \' where \': "\' \\<^sub>\ Hom (cat_FUNCT \ \ \) (cf_map \) (cf_map \)" and \_def: "\ = ntcf_of_ntcf_arrow \ \ \'" by auto note \' = cat_FUNCT_is_arrD[ OF \'[unfolded cat_cs_simps], unfolded cat_cs_simps ] from \'(1) show "\ \\<^sub>\ these_ntcfs \ \ \ \ \" unfolding \_def by simp next fix \ assume "\ \\<^sub>\ these_ntcfs \ \ \ \ \" then have \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by simp then have "\ = ntcf_of_ntcf_arrow \ \ (ntcf_arrow \)" by (cs_concl cs_simp: cat_FUNCT_cs_simps) moreover from \ have "ntcf_arrow \ \\<^sub>\ Hom (cat_FUNCT \ \ \) (cf_map \) (cf_map \)" unfolding in_Hom_iff by (cs_concl cs_intro: cat_FUNCT_cs_intros) ultimately show "\ \ ntcf_of_ntcf_arrow \ \ ` elts (Hom (cat_FUNCT \ \ \) (cf_map \) (cf_map \))" by simp qed qed subsubsection\\FUNCT\ is a category\ lemma (in \) tiny_category_cat_FUNCT[cat_FUNCT_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" shows "tiny_category \ (cat_FUNCT \ \ \)" (is \tiny_category \ ?FUNCT\) proof(intro tiny_categoryI) show "vfsequence ?FUNCT" unfolding cat_FUNCT_def by auto show "vcard ?FUNCT = 6\<^sub>\" unfolding cat_FUNCT_def by (simp add: nat_omega_simps) from assms show "tiny_semicategory \ (cat_smc ?FUNCT)" unfolding cat_smc_FUNCT by (auto simp add: tiny_semicategory_smc_FUNCT) show CId_a: "?FUNCT\CId\\\'\ : \' \\<^bsub>?FUNCT\<^esub> \'" if "\' \\<^sub>\ ?FUNCT\Obj\" for \' proof- from that obtain \ where \'_def: "\' = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding cat_FUNCT_components by clarsimp show ?thesis using that \ unfolding cat_FUNCT_components(1) \'_def by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "?FUNCT\CId\\\\ \\<^sub>A\<^bsub>?FUNCT\<^esub> \ = \" if "\ : \ \\<^bsub>?FUNCT\<^esub> \" for \ \ \ proof- from that obtain \' \' \' where \': "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \_def[cat_FUNCT_cs_simps]: "\ = ntcf_arrow \'" and \_def[cat_FUNCT_cs_simps]: "\ = cf_map \'" and \_def[cat_FUNCT_cs_simps]: "\ = cf_map \'" by auto from \' show "cat_FUNCT \ \ \\CId\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> \ = \" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "\ \\<^sub>A\<^bsub>?FUNCT\<^esub> ?FUNCT\CId\\\\ = \" if "\ : \ \\<^bsub>?FUNCT\<^esub> \" for \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that] from \(1) show "\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> cat_FUNCT \ \ \\CId\\\\ = \" by (subst (1 2) \(2), subst \(3), remdups) ( cs_concl cs_simp: cat_FUNCT_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed qed (simp_all add: assms cat_FUNCT_components) lemmas (in \) [cat_FUNCT_cs_intros] = tiny_category_cat_FUNCT subsubsection\Isomorphism\ lemma (in \) cat_FUNCT_is_arr_isomorphismI: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_arrow \ : cf_map \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" proof(intro is_arr_isomorphismI is_inverseI) interpret \: is_iso_ntcf \ \ \ \ \ \ by (rule assms) show is_arr_\: "ntcf_arrow \ : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" by (simp add: assms cat_FUNCT_is_arrI is_iso_ntcf.axioms(1)) interpret inv_\: is_iso_ntcf \ \ \ \ \ \inv_ntcf \\ using CZH_ECAT_NTCF.iso_ntcf_is_arr_isomorphism(1)[OF assms] by simp from assms show is_arr_inv_\: "ntcf_arrow (inv_ntcf \) : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" by ( cs_concl cs_intro: ntcf_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from assms show "ntcf_arrow \ : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms show "ntcf_arrow (inv_ntcf \) \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> ntcf_arrow \ = cat_FUNCT \ \ \\CId\\cf_map \\" "ntcf_arrow \ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> ntcf_arrow (inv_ntcf \) = cat_FUNCT \ \ \\CId\\cf_map \\" by ( cs_concl cs_simp: iso_ntcf_is_arr_isomorphism(2,3) cat_FUNCT_cs_simps cs_intro: ntcf_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma (in \) cat_FUNCT_is_arr_isomorphismI': assumes "\' = ntcf_arrow \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = cf_map \" and "\' = cf_map \" shows "\' : \' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" using assms(2) unfolding assms(1,3,4) by (rule cat_FUNCT_is_arr_isomorphismI) lemmas [cat_FUNCT_cs_intros] = \.cat_FUNCT_is_arr_isomorphismI'[rotated 2] lemma (in \) cat_FUNCT_is_arr_isomorphismD: assumes "\ : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_FUNCT \ \ \\<^esub> \" (is \\ : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>?FUNCT\<^esub> \\) shows "ntcf_of_ntcf_arrow \ \ \ : cf_of_cf_map \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_of_cf_map \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = ntcf_arrow (ntcf_of_ntcf_arrow \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" proof- define \ where "\ = \ + \" have \\: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_\_\\ \.intro \_Limit_\\ \_\_\\ \_def) interpret FUNCT: tiny_category \ ?FUNCT by (rule tiny_category_cat_FUNCT[OF \\ \\, of \ \]) have inv_\: "\\\<^sub>C\<^bsub>?FUNCT\<^esub> : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>?FUNCT\<^esub> \" and inv_\_\: "\\\<^sub>C\<^bsub>?FUNCT\<^esub> \\<^sub>A\<^bsub>?FUNCT\<^esub> \ = ?FUNCT\CId\\\\" and \_inv_\: "\ \\<^sub>A\<^bsub>?FUNCT\<^esub> \\\<^sub>C\<^bsub>?FUNCT\<^esub> = ?FUNCT\CId\\\\" by ( intro FUNCT.cat_the_inverse_is_arr_isomorphism[OF assms] FUNCT.cat_the_inverse_Comp_CId[OF assms] )+ from assms is_arr_isomorphismD inv_\ have \_is_arr: "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" and inv_\_is_arr: "\\\<^sub>C\<^bsub>?FUNCT\<^esub> : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" by auto note \_is_arr = cat_FUNCT_is_arrD[OF \_is_arr] note inv_\_is_arr = cat_FUNCT_is_arrD[OF inv_\_is_arr] let ?\ = \ntcf_of_ntcf_arrow \ \ \\ and ?inv_\ = \ntcf_of_ntcf_arrow \ \ (\\\<^sub>C\<^bsub>cat_FUNCT \ \ \\<^esub>)\ from inv_\_\ \_is_arr(1) inv_\_is_arr(1) have inv_\_\: "?inv_\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\ = ntcf_id (cf_of_cf_map \ \ \)" by ( subst (asm) inv_\_is_arr(2), use nothing in \subst (asm) (2) \_is_arr(2), subst (asm) \_is_arr(3)\ ) ( cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_cs_intros ) from \_inv_\ inv_\_is_arr(1) \_is_arr(1) have \_inv_\: "?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?inv_\ = ntcf_id (cf_of_cf_map \ \ \)" by ( subst (asm) inv_\_is_arr(2), use nothing in \subst (asm) \_is_arr(2), subst (asm) \_is_arr(4)\ ) ( cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_cs_intros ) show "ntcf_of_ntcf_arrow \ \ \ : cf_of_cf_map \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_of_cf_map \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( rule CZH_ECAT_NTCF.is_arr_isomorphism_is_iso_ntcf[ OF \_is_arr(1) inv_\_is_arr(1) \_inv_\ inv_\_\ ] ) show "\ = ntcf_arrow (ntcf_of_ntcf_arrow \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" by (intro \_is_arr(2-4))+ qed subsection\\Funct\\ subsubsection\Definition and elementary properties\ definition cat_Funct :: "V \ V \ V \ V" where "cat_Funct \ \ \ = [ tm_cf_maps \ \ \, tm_ntcf_arrows \ \ \, (\\\\<^sub>\tm_ntcf_arrows \ \ \. \\NTDom\), (\\\\<^sub>\tm_ntcf_arrows \ \ \. \\NTCod\), (\\\\\<^sub>\composable_arrs (dg_Funct \ \ \). \\\0\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^bsub>\,\\<^esub> \\\1\<^sub>\\), (\\\\<^sub>\tm_cf_maps \ \ \. ntcf_arrow_id \ \ \) ]\<^sub>\" text\Components.\ lemma cat_Funct_components: shows "cat_Funct \ \ \\Obj\ = tm_cf_maps \ \ \" and "cat_Funct \ \ \\Arr\ = tm_ntcf_arrows \ \ \" and "cat_Funct \ \ \\Dom\ = (\\\\<^sub>\tm_ntcf_arrows \ \ \. \\NTDom\)" and "cat_Funct \ \ \\Cod\ = (\\\\<^sub>\tm_ntcf_arrows \ \ \. \\NTCod\)" and "cat_Funct \ \ \\Comp\ = (\\\\\<^sub>\composable_arrs (dg_Funct \ \ \). \\\0\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^bsub>\,\\<^esub> \\\1\<^sub>\\)" and "cat_Funct \ \ \\CId\ = (\\\\<^sub>\tm_cf_maps \ \ \. ntcf_arrow_id \ \ \)" unfolding cat_Funct_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_Funct: "cat_smc (cat_Funct \ \ \) = smc_Funct \ \ \" proof(rule vsv_eqI) show "vsv (cat_smc (cat_Funct \ \ \))" unfolding cat_smc_def by auto show "vsv (smc_Funct \ \ \)" unfolding smc_Funct_def by auto have dom_lhs: "\\<^sub>\ (cat_smc (cat_Funct \ \ \)) = 5\<^sub>\" unfolding cat_smc_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smc_Funct \ \ \) = 5\<^sub>\" unfolding smc_Funct_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cat_smc (cat_Funct \ \ \)) = \\<^sub>\ (smc_Funct \ \ \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (cat_smc (cat_Funct \ \ \)) \ cat_smc (cat_Funct \ \ \)\a\ = smc_Funct \ \ \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cat_smc_def dg_field_simps cat_Funct_def smc_Funct_def ) (auto simp: nat_omega_simps) qed context is_tm_ntcf begin lemmas_with [folded cat_smc_Funct, unfolded slicing_simps]: cat_Funct_Dom_app = smc_Funct_Dom_app and cat_Funct_Cod_app = smc_Funct_Cod_app end lemmas [cat_FUNCT_cs_simps] = is_tm_ntcf.cat_Funct_Dom_app is_tm_ntcf.cat_Funct_Cod_app lemmas_with [folded cat_smc_Funct, unfolded slicing_simps]: cat_Funct_Dom_vsv[cat_FUNCT_cs_intros] = smc_Funct_Dom_vsv and cat_Funct_Dom_vdomain[cat_FUNCT_cs_simps] = smc_Funct_Dom_vdomain and cat_Funct_Cod_vsv[cat_FUNCT_cs_intros] = smc_Funct_Cod_vsv and cat_Funct_Cod_vdomain[cat_FUNCT_cs_simps] = smc_Funct_Cod_vdomain and cat_Funct_Dom_vrange = smc_Funct_Dom_vrange and cat_Funct_Cod_vrange = smc_Funct_Cod_vrange and cat_Funct_is_arrI = smc_Funct_is_arrI and cat_Funct_is_arrI'[cat_FUNCT_cs_intros] = smc_Funct_is_arrI' and cat_Funct_is_arrD = smc_Funct_is_arrD and cat_Funct_is_arrE[elim] = smc_Funct_is_arrE lemmas_with [folded cat_smc_Funct, unfolded slicing_simps]: cat_Funct_Comp_app[cat_FUNCT_cs_simps] = smc_Funct_Comp_app subsubsection\Identity\ mk_VLambda cat_Funct_components(6) |vsv cat_Funct_CId_vsv[intro]| |vdomain cat_Funct_CId_vdomain[cat_FUNCT_cs_simps]| |app cat_Funct_CId_app[cat_FUNCT_cs_simps]| lemma smc_Funct_CId_vrange: "\\<^sub>\ (cat_Funct \ \ \\CId\) \\<^sub>\ ntcf_arrows \ \ \" unfolding cat_Funct_components proof(rule vrange_VLambda_vsubset) fix \' assume "\' \\<^sub>\ tm_cf_maps \ \ \" then obtain \ where \'_def: "\' = cf_map \" and \: "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by clarsimp then show "ntcf_arrow_id \ \ \' \\<^sub>\ ntcf_arrows \ \ \" by ( cs_concl cs_simp: cat_FUNCT_cs_simps \'_def cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) qed subsubsection\\Funct\ is a category\ lemma category_cat_Funct: assumes "tiny_category \ \" and "category \ \" shows "category \ (cat_Funct \ \ \)" (is \category \ ?Funct\) proof- interpret tiny_category \ \ by (rule assms(1)) show ?thesis proof(intro categoryI) show "vfsequence ?Funct" by (simp add: cat_Funct_def) show "vcard ?Funct = 6\<^sub>\" unfolding cat_Funct_def by (simp add: nat_omega_simps) from assms show "semicategory \ (cat_smc (cat_Funct \ \ \))" unfolding cat_smc_Funct by (rule semicategory_smc_Funct) show "\\<^sub>\ (cat_Funct \ \ \\CId\) = cat_Funct \ \ \\Obj\" by (cs_concl cs_simp: cat_Funct_components cat_FUNCT_cs_simps) show "cat_Funct \ \ \\CId\\\\ : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" if "\ \\<^sub>\ cat_Funct \ \ \\Obj\" for \ proof- from that have "\ \\<^sub>\ tm_cf_maps \ \ \" unfolding cat_Funct_components by simp then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by auto from assms \' show "cat_Funct \ \ \\CId\\\\ : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" by ( cs_concl cs_simp: cat_FUNCT_cs_simps \_def cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) qed show "cat_Funct \ \ \\CId\\\\ \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> \ = \" if "\ : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" for \ \ \ proof- note \ = cat_Funct_is_arrD[OF that] from assms \(1) show "cat_Funct \ \ \\CId\\\\ \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> \ = \" by (subst (1 2) \(2), use nothing in \subst \(4)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) qed show "\ \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> cat_Funct \ \ \\CId\\\\ = \" if "\ : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" for \ \ \ proof- note \ = cat_Funct_is_arrD[OF that] from assms \(1) show "\ \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> cat_Funct \ \ \\CId\\\\ = \" by (subst (1 2) \(2), use nothing in \subst \(3)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) qed qed auto qed lemma category_cat_Funct'[cat_FUNCT_cs_intros]: assumes "tiny_category \ \" and "category \ \" and "\ = \" shows "category \ (cat_Funct \ \ \)" using assms(1,2) unfolding assms(3) by (rule category_cat_Funct) subsubsection\\Funct\ is a subcategory of \FUNCT\\ lemma subcategory_cat_Funct_cat_FUNCT: assumes "\ \" and "\ \\<^sub>\ \" and "tiny_category \ \" and "category \ \" shows "cat_Funct \ \ \ \\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" proof ( intro subcategoryI, unfold cat_smc_FUNCT cat_smc_Funct cat_Funct_components(1) ) interpret category \ \ by (rule assms(4)) interpret \\: category \ \cat_Funct \ \ \\ by (rule category_cat_Funct[OF assms(3,4)]) show "category \ (cat_Funct \ \ \)" by (rule category.cat_category_if_ge_Limit[OF _ assms(1,2)]) (auto intro: cat_cs_intros) from assms show "category \ (cat_FUNCT \ \ \)" by (cs_concl cs_intro: tiny_category_cat_FUNCT cat_small_cs_intros) show "smc_Funct \ \ \ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_FUNCT \ \ \" by (rule subsemicategory_smc_Funct_smc_FUNCT[OF assms]) show "cat_Funct \ \ \\CId\\\\ = cat_FUNCT \ \ \\CId\\\\" if \\ \\<^sub>\ tm_cf_maps \ \ \\ for \ proof- from that obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by auto from that show ?thesis by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros tm_cf_maps_in_cf_maps ) qed qed subsubsection\Isomorphism\ lemma (in is_tm_iso_ntcf) cat_Funct_is_arr_isomorphismI: assumes "category \ \" shows "ntcf_arrow \ : cf_map \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" proof(intro is_arr_isomorphismI is_inverseI) from is_tm_iso_ntcf_axioms show "ntcf_arrow \ : cf_map \ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" by (cs_concl cs_intro: ntcf_cs_intros cat_FUNCT_cs_intros) interpret inv_\: is_tm_iso_ntcf \ \ \ \ \ \inv_ntcf \\ by (rule iso_tm_ntcf_is_arr_isomorphism(1)[OF assms is_tm_iso_ntcf_axioms]) from inv_\.is_tm_iso_ntcf_axioms show "ntcf_arrow (inv_ntcf \) : cf_map \ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" by (cs_concl cs_intro: ntcf_cs_intros cat_FUNCT_cs_intros) from is_tm_iso_ntcf_axioms show "ntcf_arrow \ : cf_map \ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" by (cs_concl cs_intro: ntcf_cs_intros cat_FUNCT_cs_intros) from assms is_tm_iso_ntcf_axioms show "ntcf_arrow (inv_ntcf \) \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> ntcf_arrow \ = cat_Funct \ \ \\CId\\cf_map \\" "ntcf_arrow \ \\<^sub>A\<^bsub>cat_Funct \ \ \\<^esub> ntcf_arrow (inv_ntcf \) = cat_Funct \ \ \\CId\\cf_map \\" by ( cs_concl cs_simp: iso_tm_ntcf_is_arr_isomorphism(2,3) cat_FUNCT_cs_simps cs_intro: ntcf_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros )+ qed lemma (in is_tm_iso_ntcf) cat_Funct_is_arr_isomorphismI': assumes "category \ \" and "\' = ntcf_arrow \" and "\' = cf_map \" and "\' = cf_map \" shows "\' : \' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" using assms(1) unfolding assms(2-4) by (rule cat_Funct_is_arr_isomorphismI) lemmas [cat_FUNCT_cs_intros] = is_tm_iso_ntcf.cat_Funct_is_arr_isomorphismI'[rotated 2] lemma (in \) cat_Funct_is_arr_isomorphismD: assumes "tiny_category \ \" and "category \ \" and "\ : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Funct \ \ \\<^esub> \" (is \\ : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>?Funct\<^esub> \\) shows "ntcf_of_ntcf_arrow \ \ \ : cf_of_cf_map \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_of_cf_map \ \ \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ = ntcf_arrow (ntcf_of_ntcf_arrow \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" proof- interpret Funct: category \ ?Funct by (rule category_cat_Funct[OF assms(1,2)]) have inv_\: "\\\<^sub>C\<^bsub>?Funct\<^esub> : \ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>?Funct\<^esub> \" and inv_\_\: "\\\<^sub>C\<^bsub>?Funct\<^esub> \\<^sub>A\<^bsub>?Funct\<^esub> \ = ?Funct\CId\\\\" and \_inv_\: "\ \\<^sub>A\<^bsub>?Funct\<^esub> \\\<^sub>C\<^bsub>?Funct\<^esub> = ?Funct\CId\\\\" by ( intro Funct.cat_the_inverse_is_arr_isomorphism[OF assms(3)] Funct.cat_the_inverse_Comp_CId[OF assms(3)] )+ from assms is_arr_isomorphismD inv_\ have \_is_arr: "\ : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" and inv_\_is_arr: "\\\<^sub>C\<^bsub>?Funct\<^esub> : \ \\<^bsub>cat_Funct \ \ \\<^esub> \" by auto note \_is_arr = cat_Funct_is_arrD[OF \_is_arr] note inv_\_is_arr = cat_Funct_is_arrD[OF inv_\_is_arr] let ?\ = \ntcf_of_ntcf_arrow \ \ \\ and ?inv_\ = \ntcf_of_ntcf_arrow \ \ (\\\<^sub>C\<^bsub>cat_Funct \ \ \\<^esub>)\ from inv_\_\ \_is_arr(1) inv_\_is_arr(1) have inv_\_\: "?inv_\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\ = ntcf_id (cf_of_cf_map \ \ \)" by ( subst (asm) inv_\_is_arr(2), use nothing in \subst (asm) (2) \_is_arr(2), subst (asm) \_is_arr(3)\ ) ( cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) from \_inv_\ inv_\_is_arr(1) \_is_arr(1) have \_inv_\: "?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?inv_\ = ntcf_id (cf_of_cf_map \ \ \)" by ( subst (asm) inv_\_is_arr(2), use nothing in \subst (asm) \_is_arr(2), subst (asm) \_is_arr(4)\ ) ( cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros ) show "ntcf_of_ntcf_arrow \ \ \ : cf_of_cf_map \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_of_cf_map \ \ \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by ( rule is_arr_isomorphism_is_tm_iso_ntcf[ OF \_is_arr(1) inv_\_is_arr(1) \_inv_\ inv_\_\ ] ) show "\ = ntcf_arrow (ntcf_of_ntcf_arrow \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" and "\ = cf_map (cf_of_cf_map \ \ \)" by (intro \_is_arr(2-4))+ qed subsection\Diagonal functor\ subsubsection\Definition and elementary properties\ text\See Chapter III-3 in \cite{mac_lane_categories_2010}.\ -definition cf_diagonal :: "V \ V \ V \ V" (\\\<^sub>C\) - where "\\<^sub>C \ \ \ = +definition cf_diagonal :: "V \ V \ V \ V" (\\\<^sub>C\<^sub>F\) + where "\\<^sub>C\<^sub>F \ \ \ = [ (\a\\<^sub>\\\Obj\. cf_map (cf_const \ \ a)), (\f\\<^sub>\\\Arr\. ntcf_arrow (ntcf_const \ \ f)), \, - cat_Funct \ \ \ + cat_FUNCT \ \ \ ]\<^sub>\" text\Components.\ lemma cf_diagonal_components: - shows "\\<^sub>C \ \ \\ObjMap\ = (\a\\<^sub>\\\Obj\. cf_map (cf_const \ \ a))" - and "\\<^sub>C \ \ \\ArrMap\ = (\f\\<^sub>\\\Arr\. ntcf_arrow (ntcf_const \ \ f))" - and "\\<^sub>C \ \ \\HomDom\ = \" - and "\\<^sub>C \ \ \\HomCod\ = cat_Funct \ \ \" + shows "\\<^sub>C\<^sub>F \ \ \\ObjMap\ = (\a\\<^sub>\\\Obj\. cf_map (cf_const \ \ a))" + and "\\<^sub>C\<^sub>F \ \ \\ArrMap\ = (\f\\<^sub>\\\Arr\. ntcf_arrow (ntcf_const \ \ f))" + and "\\<^sub>C\<^sub>F \ \ \\HomDom\ = \" + and "\\<^sub>C\<^sub>F \ \ \\HomCod\ = cat_FUNCT \ \ \" unfolding cf_diagonal_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda cf_diagonal_components(1) |vsv cf_diagonal_ObjMap_vsv[cat_cs_intros]| |vdomain cf_diagonal_ObjMap_vdomain[cat_cs_simps]| |app cf_diagonal_ObjMap_app[cat_cs_simps]| lemma cf_diagonal_ObjMap_vrange: - assumes "tiny_category \ \" and "category \ \" - shows "\\<^sub>\ (\\<^sub>C \ \ \\ObjMap\) \\<^sub>\ cat_Funct \ \ \\Obj\" + assumes "\ \"and "\ \\<^sub>\ \" and "category \ \" and "category \ \" + shows "\\<^sub>\ (\\<^sub>C\<^sub>F \ \ \\ObjMap\) \\<^sub>\ cat_FUNCT \ \ \\Obj\" unfolding cf_diagonal_components proof(rule vrange_VLambda_vsubset) - fix x assume "x \\<^sub>\ \\Obj\" - with assms category_cat_Funct[OF assms] show - "cf_map (cf_const \ \ x) \\<^sub>\ cat_Funct \ \ \\Obj\" - unfolding cat_Funct_components(1) - by (cs_concl cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros) + interpret \: \ \ by (rule assms(1)) + interpret category \ \ by (rule assms(3)) + interpret FUNCT: tiny_category \ \(cat_FUNCT \ \ \)\ + by (rule \.tiny_category_cat_FUNCT[OF \_axioms assms(1,2)]) + fix x assume prems: "x \\<^sub>\ \\Obj\" + from prems assms show "cf_map (cf_const \ \ x) \\<^sub>\ cat_FUNCT \ \ \\Obj\" + unfolding cat_FUNCT_components(1) + by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed subsubsection\Arrow map\ mk_VLambda cf_diagonal_components(2) |vsv cf_diagonal_ArrMap_vsv[cat_cs_intros]| |vdomain cf_diagonal_ArrMap_vdomain[cat_cs_simps]| |app cf_diagonal_ArrMap_app[cat_cs_simps]| subsubsection\Diagonal functor is a functor\ lemma cf_diagonal_is_functor[cat_cs_intros]: + assumes "\ \" and "\ \\<^sub>\ \" and "category \ \" and "category \ \" + shows "\\<^sub>C\<^sub>F \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" (is \?\ : \ \\\<^sub>C\<^bsub>\\<^esub> ?FUNCT\) +proof- + + interpret \: \ \ by (rule assms(1)) + interpret \: category \ \ by (rule assms(3)) + interpret \: category \ \ by (rule assms(4)) + interpret FUNCT: tiny_category \ \(cat_FUNCT \ \ \)\ + by (rule \.tiny_category_cat_FUNCT[OF \.\_axioms assms(1,2)]) + + show ?thesis + proof(intro is_functorI') + show "vfsequence ?\" + unfolding cf_diagonal_def by (simp add: nat_omega_simps) + show "category \ \" by (rule \.cat_category_if_ge_Limit[OF assms(1,2)]) + from assms show "category \ (cat_FUNCT \ \ \)" + by (cs_concl cs_intro: cat_cs_intros) + show "vcard ?\ = 4\<^sub>\" + unfolding cf_diagonal_def by (simp add: nat_omega_simps) + show "vsv (?\\ObjMap\)" unfolding cf_diagonal_components by simp + from assms show "\\<^sub>\ (?\\ObjMap\) \\<^sub>\ ?FUNCT\Obj\" + by (rule cf_diagonal_ObjMap_vrange) + show "?\\ArrMap\\f\ : ?\\ObjMap\\a\ \\<^bsub>?FUNCT\<^esub> ?\\ObjMap\\b\" + if "f : a \\<^bsub>\\<^esub> b" for f a b + using that + by + ( + cs_concl + cs_simp: cat_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros + ) + show "?\\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = ?\\ArrMap\\g\ \\<^sub>A\<^bsub>?FUNCT\<^esub> ?\\ArrMap\\f\" + if "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" for g b c f a + using that \.category_axioms \.category_axioms + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros + ) + fix c assume "c \\<^sub>\ \\Obj\" + with \.category_axioms \.category_axioms show + "?\\ArrMap\\\\CId\\c\\ = ?FUNCT\CId\\?\\ObjMap\\c\\" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros + ) + qed (auto simp: cf_diagonal_components cat_smc_FUNCT) + +qed + +lemma cf_diagonal_is_functor'[cat_cs_intros]: + assumes "\ \" + and "\ \\<^sub>\ \" + and "category \ \" + and "category \ \" + and "\' = \" + and "\' = cat_FUNCT \ \ \" + shows "\\<^sub>C\<^sub>F \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + using assms(1-4) unfolding assms(5-6) by (rule cf_diagonal_is_functor) + + + +(*TODO: functor codomain substitution*) +subsection\Diagonal functor for functors with tiny maps\ + + +subsubsection\Definition and elementary properties\ + + +text\See Chapter III-3 in \cite{mac_lane_categories_2010}.\ + +definition tm_cf_diagonal :: "V \ V \ V \ V" (\\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m\) + where "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \ = + [ + (\a\\<^sub>\\\Obj\. cf_map (cf_const \ \ a)), + (\f\\<^sub>\\\Arr\. ntcf_arrow (ntcf_const \ \ f)), + \, + cat_Funct \ \ \ + ]\<^sub>\" + + +text\Components.\ + +lemma tm_cf_diagonal_components: + shows "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \\ObjMap\ = (\a\\<^sub>\\\Obj\. cf_map (cf_const \ \ a))" + and "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \\ArrMap\ = (\f\\<^sub>\\\Arr\. ntcf_arrow (ntcf_const \ \ f))" + and "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \\HomDom\ = \" + and "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \\HomCod\ = cat_Funct \ \ \" + unfolding tm_cf_diagonal_def dghm_field_simps by (simp_all add: nat_omega_simps) + + +subsubsection\Object map\ + +mk_VLambda tm_cf_diagonal_components(1) + |vsv tm_cf_diagonal_ObjMap_vsv[cat_cs_intros]| + |vdomain tm_cf_diagonal_ObjMap_vdomain[cat_cs_simps]| + |app tm_cf_diagonal_ObjMap_app[cat_cs_simps]| + +lemma tm_cf_diagonal_ObjMap_vrange: assumes "tiny_category \ \" and "category \ \" - shows "\\<^sub>C \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Funct \ \ \" (is \?\ : \ \\\<^sub>C\<^bsub>\\<^esub> ?Funct\) + shows "\\<^sub>\ (\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \\ObjMap\) \\<^sub>\ cat_Funct \ \ \\Obj\" + unfolding tm_cf_diagonal_components +proof(rule vrange_VLambda_vsubset) + fix x assume "x \\<^sub>\ \\Obj\" + with assms category_cat_Funct[OF assms] show + "cf_map (cf_const \ \ x) \\<^sub>\ cat_Funct \ \ \\Obj\" + unfolding cat_Funct_components(1) + by (cs_concl cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros) +qed + + +subsubsection\Arrow map\ + +mk_VLambda tm_cf_diagonal_components(2) + |vsv tm_cf_diagonal_ArrMap_vsv[cat_cs_intros]| + |vdomain tm_cf_diagonal_ArrMap_vdomain[cat_cs_simps]| + |app tm_cf_diagonal_ArrMap_app[cat_cs_simps]| + + +subsubsection\Diagonal functor for functors with tiny maps is a functor\ + +lemma tm_cf_diagonal_is_functor[cat_cs_intros]: + assumes "tiny_category \ \" and "category \ \" + shows "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Funct \ \ \" + (is \?\ : \ \\\<^sub>C\<^bsub>\\<^esub> ?Funct\) proof- interpret \: tiny_category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) show ?thesis proof(intro is_functorI') show "vfsequence ?\" - unfolding cf_diagonal_def by (simp add: nat_omega_simps) + unfolding tm_cf_diagonal_def by (simp add: nat_omega_simps) from assms(2) show "category \ \" by (cs_concl cs_intro: cat_cs_intros) from assms show "category \ ?Funct" by (cs_concl cs_intro: cat_cs_intros category_cat_Funct) show "vcard ?\ = 4\<^sub>\" - unfolding cf_diagonal_def by (simp add: nat_omega_simps) - show "vsv (?\\ObjMap\)" unfolding cf_diagonal_components by simp + unfolding tm_cf_diagonal_def by (simp add: nat_omega_simps) + show "vsv (?\\ObjMap\)" unfolding tm_cf_diagonal_components by simp from assms show "\\<^sub>\ (?\\ObjMap\) \\<^sub>\ ?Funct\Obj\" - by (rule cf_diagonal_ObjMap_vrange) + by (rule tm_cf_diagonal_ObjMap_vrange) show "?\\ArrMap\\f\ : ?\\ObjMap\\a\ \\<^bsub>?Funct\<^esub> ?\\ObjMap\\b\" if "f : a \\<^bsub>\\<^esub> b" for f a b using that by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros ) show "?\\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = ?\\ArrMap\\g\ \\<^sub>A\<^bsub>?Funct\<^esub> ?\\ArrMap\\f\" if "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" for g b c f a using that \.category_axioms \.category_axioms by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) fix c assume "c \\<^sub>\ \\Obj\" with \.category_axioms \.category_axioms show "?\\ArrMap\\\\CId\\c\\ = ?Funct\CId\\?\\ObjMap\\c\\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) - qed (auto simp: cf_diagonal_components cat_smc_FUNCT) + qed (auto simp: tm_cf_diagonal_components cat_smc_FUNCT) qed -lemma cf_diagonal_is_functor'[cat_cs_intros]: +lemma tm_cf_diagonal_is_functor'[cat_cs_intros]: assumes "tiny_category \ \" and "category \ \" and "\' = \" and "\ = \" and "\ = cat_Funct \ \ \" - shows "\\<^sub>C \ \ \ : \ \\\<^sub>C\<^bsub>\'\<^esub> \" - using assms(1-2) unfolding assms(3-5) by (rule cf_diagonal_is_functor) + shows "\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ \ \ : \ \\\<^sub>C\<^bsub>\'\<^esub> \" + using assms(1-2) unfolding assms(3-5) by (rule tm_cf_diagonal_is_functor) subsection\Functor raised to the power of a category\ subsubsection\Definition and elementary properties\ text\ Most of the definitions and the results presented in this and the remaining subsections can be found in \cite{mac_lane_categories_2010} and \cite{riehl_category_2016} (e.g., see Chapter X-3 in \cite{mac_lane_categories_2010}). \ definition exp_cf_cat :: "V \ V \ V \ V" where "exp_cf_cat \ \ \ = [ ( \\\\<^sub>\cat_FUNCT \ \ (\\HomDom\)\Obj\. cf_map (\ \\<^sub>C\<^sub>F cf_of_cf_map \ (\\HomDom\) \) ), ( \\\\<^sub>\cat_FUNCT \ \ (\\HomDom\)\Arr\. ntcf_arrow (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_of_ntcf_arrow \ (\\HomDom\) \) ), cat_FUNCT \ \ (\\HomDom\), cat_FUNCT \ \ (\\HomCod\) ]\<^sub>\" text\Components.\ lemma exp_cf_cat_components: shows "exp_cf_cat \ \ \\ObjMap\ = ( \\\\<^sub>\cat_FUNCT \ \ (\\HomDom\)\Obj\. cf_map (\ \\<^sub>C\<^sub>F cf_of_cf_map \ (\\HomDom\) \) )" and "exp_cf_cat \ \ \\ArrMap\ = ( \\\\<^sub>\cat_FUNCT \ \ (\\HomDom\)\Arr\. ntcf_arrow (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_of_ntcf_arrow \ (\\HomDom\) \)) )" and "exp_cf_cat \ \ \\HomDom\ = cat_FUNCT \ \ (\\HomDom\)" and "exp_cf_cat \ \ \\HomCod\ = cat_FUNCT \ \ (\\HomCod\)" unfolding exp_cf_cat_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda exp_cf_cat_components(1) |vsv exp_cf_cat_components_ObjMap_vsv[cat_FUNCT_cs_intros]| context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda exp_cf_cat_components(1)[where \=\ and \=\, unfolded cat_cs_simps] |vdomain exp_cf_cat_components_ObjMap_vdomain[cat_FUNCT_cs_simps]| |app exp_cf_cat_components_ObjMap_app[cat_FUNCT_cs_simps]| end subsubsection\Arrow map\ mk_VLambda exp_cf_cat_components(2) |vsv exp_cf_cat_components_ArrMap_vsv[cat_FUNCT_cs_intros]| context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda exp_cf_cat_components(2)[where \=\ and \=\, unfolded cat_cs_simps] |vdomain exp_cf_cat_components_ArrMap_vdomain[cat_FUNCT_cs_simps]| |app exp_cf_cat_components_ArrMap_app[cat_FUNCT_cs_simps]| end subsubsection\Domain and codomain\ context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) lemmas exp_cf_cat_HomDom[cat_FUNCT_cs_simps] = exp_cf_cat_components(3)[where \=\ and \=\, unfolded cat_cs_simps] and exp_cf_cat_HomCod[cat_FUNCT_cs_simps] = exp_cf_cat_components(4)[where \=\ and \=\, unfolded cat_cs_simps] end subsubsection\Functor raised to the power of a category is a functor\ lemma exp_cf_cat_is_tiny_functor: assumes "\ \" and "\ \\<^sub>\ \" and "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" proof- interpret \: \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(3)) interpret \: is_functor \ \ \ \ by (rule assms(4)) from assms(2-4) interpret \\: tiny_category \ \cat_FUNCT \ \ \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2-4) interpret \\: tiny_category \ \cat_FUNCT \ \ \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) show ?thesis proof(intro is_tiny_functorI' is_functorI') show "vfsequence (exp_cf_cat \ \ \)" unfolding exp_cf_cat_def by simp show "vcard (exp_cf_cat \ \ \) = 4\<^sub>\" unfolding exp_cf_cat_def by (simp add: nat_omega_simps) show "\\<^sub>\ (exp_cf_cat \ \ \\ObjMap\) \\<^sub>\ cat_FUNCT \ \ \\Obj\" proof ( unfold cat_FUNCT_components exp_cf_cat_components, intro vrange_VLambda_vsubset, unfold cat_cs_simps ) fix \ assume "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from assms(2-4) \' show "cf_map (\ \\<^sub>C\<^sub>F cf_of_cf_map \ \ \) \\<^sub>\ cf_maps \ \ \" by (cs_concl cs_simp: \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed show "exp_cf_cat \ \ \\ArrMap\\\\ : exp_cf_cat \ \ \\ObjMap\\\\ \\<^bsub>cat_FUNCT \ \ \\<^esub> exp_cf_cat \ \ \\ObjMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" for \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that] from \(1,3,4) assms(2-4) show ?thesis by (subst \(2), use nothing in \subst \(3), subst \(4)\) ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "exp_cf_cat \ \ \\ArrMap\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> \\ = exp_cf_cat \ \ \\ArrMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> exp_cf_cat \ \ \\ArrMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" and "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" for \ \ \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that(1)] and \ = cat_FUNCT_is_arrD[OF that(2)] from \(1,3,4) \(1,3,4) assms(2-4) show ?thesis by (subst (1 2) \(2), use nothing in \subst (1 2) \(2)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cf_ntcf_comp_ntcf_vcomp cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "exp_cf_cat \ \ \\ArrMap\\cat_FUNCT \ \ \\CId\\\\\ = cat_FUNCT \ \ \\CId\\exp_cf_cat \ \ \\ObjMap\\\\\" if "\ \\<^sub>\ cat_FUNCT \ \ \\Obj\" for \ proof- from that[unfolded cat_FUNCT_components] obtain \ where \_def: "\ = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \ show "exp_cf_cat \ \ \\ArrMap\\cat_FUNCT \ \ \\CId\\\\\ = cat_FUNCT \ \ \\CId\\exp_cf_cat \ \ \\ObjMap\\\\\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed qed ( use assms(1,2) in \ cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros \ )+ qed lemma exp_cf_cat_is_tiny_functor'[cat_FUNCT_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = cat_FUNCT \ \ \" and "\' = cat_FUNCT \ \ \" shows "exp_cf_cat \ \ \ : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> \'" using assms(1-4) unfolding assms(5,6) by (rule exp_cf_cat_is_tiny_functor) subsubsection\Further properties\ lemma exp_cf_cat_cf_comp: assumes "category \ \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
= exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
" proof(rule cf_eqI) interpret \
: category \ \
by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \
.\_Limit_\\ \
.\_\_\\ \_def \
.\_\_\\) then interpret \: \ \ by simp from \\ show "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
: cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ show "exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
: cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_lhs: "\\<^sub>\ (exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ObjMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ObjMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ObjMap\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ObjMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) show "vsv (exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ObjMap\)" by (cs_concl cs_intro: cat_FUNCT_cs_intros) from \\ show "vsv ((exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ObjMap\)" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \
\\Obj\" then have "\ \\<^sub>\ cf_maps \ \
\" unfolding cat_FUNCT_components by simp then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by auto from assms \\ \' show "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ObjMap\\\\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ObjMap\\\\" by (subst (1 2) \_def) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed simp from \\ have dom_lhs: "\\<^sub>\ (exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ArrMap\) = cat_FUNCT \ \
\\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ArrMap\) = cat_FUNCT \ \
\\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ArrMap\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ArrMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) show "vsv (exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ArrMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_FUNCT_cs_intros) from \\ show "vsv ((exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ArrMap\)" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \
\\Arr\" then obtain \ \' where \: "\ : \ \\<^bsub>cat_FUNCT \ \
\\<^esub> \'" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from \\ assms \(1,3,4) show "exp_cf_cat \ (\ \\<^sub>C\<^sub>F \) \
\ArrMap\\\\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cf_comp_cf_ntcf_comp_assoc cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed simp qed simp_all lemma exp_cf_cat_cf_id_cat: assumes "category \ \" and "category \ \
" shows "exp_cf_cat \ (cf_id \) \
= cf_id (cat_FUNCT \ \
\)" proof(rule cf_eqI) interpret \: category \ \ by (rule assms) interpret \
: category \ \
by (rule assms) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp from \\ show "cf_id (cat_FUNCT \ \
\) : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ show "exp_cf_cat \ (cf_id \) \
: cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ have ObjMap_dom_lhs: "\\<^sub>\ (exp_cf_cat \ (cf_id \) \
\ObjMap\) = cat_FUNCT \ \
\\Obj\" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) from \\ have ObjMap_dom_rhs: "\\<^sub>\ (cf_id (cat_FUNCT \ \
\)\ObjMap\) = cat_FUNCT \ \
\\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "exp_cf_cat \ (cf_id \) \
\ObjMap\ = cf_id (cat_FUNCT \ \
\)\ObjMap\" proof ( rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_FUNCT_components(1) ) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \
\" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from prems \' show "exp_cf_cat \ (cf_id \) \
\ObjMap\\\\ = cf_id (cat_FUNCT \ \
\)\ObjMap\\\\" by (subst (1 2) \_def) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros)+ from \\ have ArrMap_dom_lhs: "\\<^sub>\ (cf_id (cat_FUNCT \ \
\)\ArrMap\) = cat_FUNCT \ \
\\Arr\" by (cs_concl cs_simp: cat_cs_simps) from \\ have ArrMap_dom_rhs: "\\<^sub>\ (exp_cf_cat \ (cf_id \) \
\ArrMap\) = cat_FUNCT \ \
\\Arr\" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) show "exp_cf_cat \ (cf_id \) \
\ArrMap\ = cf_id (cat_FUNCT \ \
\)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \
\\Arr\" then obtain \ \ where \: "\ : \ \\<^bsub>cat_FUNCT \ \
\\<^esub> \" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from \(1,3,4) \\ show "exp_cf_cat \ (cf_id \) \
\ArrMap\\\\ = cf_id (cat_FUNCT \ \
\)\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed simp_all lemma cf_comp_exp_cf_cat_exp_cf_cat_cf_id[cat_FUNCT_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \ = exp_cf_cat \ \ \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule cf_eqI) from assms \\ \ show \\: "exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by (cs_concl cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros) with assms \\ show "exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms \\ have ObjMap_dom_lhs: "\\<^sub>\ ((exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ObjMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms have ObjMap_dom_rhs: "\\<^sub>\ (exp_cf_cat \ \ \\ObjMap\) = cat_FUNCT \ \ \\Obj\" by (cs_concl cs_simp: cat_FUNCT_cs_simps) from assms \\ have ArrMap_dom_lhs: "\\<^sub>\ ((exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ArrMap\) = cat_FUNCT \ \ \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms have ArrMap_dom_rhs: "\\<^sub>\ (exp_cf_cat \ \ \\ArrMap\) = cat_FUNCT \ \ \\Arr\" by (cs_concl cs_simp: cat_FUNCT_cs_simps) show "(exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ObjMap\ = exp_cf_cat \ \ \\ObjMap\" proof ( rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_FUNCT_components(1) ) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from prems \' assms \\ \\ show "(exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ObjMap\\\\ = exp_cf_cat \ \ \\ObjMap\\\\" unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros ) qed ( use assms \\ \\ in \ cs_concl cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros \ ) show "(exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ArrMap\ = exp_cf_cat \ \ \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \ \\Arr\" then obtain \' \' where \: "\ : \' \\<^bsub>cat_FUNCT \ \ \\<^esub> \'" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from \(1) assms \\ \\ show "(exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ (cf_id \) \)\ArrMap\\\\ = exp_cf_cat \ \ \\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros ) qed ( use assms \\ in \ cs_concl cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros \ ) qed simp_all qed lemma cf_comp_exp_cf_cat_cf_id_exp_cf_cat[cat_FUNCT_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ = exp_cf_cat \ \ \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule cf_eqI) from assms \\ \ show \\: "exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by (cs_concl cs_simp: cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros) with assms \\ show "exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms \\ have ObjMap_dom_lhs: "\\<^sub>\ ((exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ObjMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms have ObjMap_dom_rhs: "\\<^sub>\ (exp_cf_cat \ \ \\ObjMap\) = cat_FUNCT \ \ \\Obj\" by (cs_concl cs_simp: cat_FUNCT_cs_simps) from assms \\ have ArrMap_dom_lhs: "\\<^sub>\ ((exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ArrMap\) = cat_FUNCT \ \ \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms have ArrMap_dom_rhs: "\\<^sub>\ (exp_cf_cat \ \ \\ArrMap\) = cat_FUNCT \ \ \\Arr\" by (cs_concl cs_simp: cat_FUNCT_cs_simps) show "(exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ObjMap\ = exp_cf_cat \ \ \\ObjMap\" proof ( rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_FUNCT_components(1) ) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from prems \' assms \\ \\ show "(exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ObjMap\\\\ = exp_cf_cat \ \ \\ObjMap\\\\" unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros ) qed ( use assms \\ \\ in \ cs_concl cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros \ ) show "(exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ArrMap\ = exp_cf_cat \ \ \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \ \\Arr\" then obtain \' \' where \: "\ : \' \\<^bsub>cat_FUNCT \ \ \\<^esub> \'" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from \(1) assms \\ \\ show "(exp_cf_cat \ (cf_id \) \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \)\ArrMap\\\\ = exp_cf_cat \ \ \\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros ) qed ( use assms \\ in \ cs_concl cs_intro: cat_FUNCT_cs_intros cat_small_cs_intros cat_cs_intros \ ) qed simp_all qed subsection\Category raised to the power of a functor\ subsubsection\Definition and elementary properties\ definition exp_cat_cf :: "V \ V \ V \ V" where "exp_cat_cf \ \ \ = [ ( \\\\<^sub>\cat_FUNCT \ (\\HomCod\) \\Obj\. cf_map (cf_of_cf_map (\\HomCod\) \ \ \\<^sub>C\<^sub>F \) ), ( \\\\<^sub>\cat_FUNCT \ (\\HomCod\) \\Arr\. ntcf_arrow (ntcf_of_ntcf_arrow (\\HomCod\) \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) ), cat_FUNCT \ (\\HomCod\) \, cat_FUNCT \ (\\HomDom\) \ ]\<^sub>\" text\Components.\ lemma exp_cat_cf_components: shows "exp_cat_cf \ \ \\ObjMap\ = ( \\\\<^sub>\cat_FUNCT \ (\\HomCod\) \\Obj\. cf_map (cf_of_cf_map (\\HomCod\) \ \ \\<^sub>C\<^sub>F \) )" and "exp_cat_cf \ \ \\ArrMap\ = ( \\\\<^sub>\cat_FUNCT \ (\\HomCod\) \\Arr\. ntcf_arrow (ntcf_of_ntcf_arrow (\\HomCod\) \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) )" and "exp_cat_cf \ \ \\HomDom\ = cat_FUNCT \ (\\HomCod\) \" and "exp_cat_cf \ \ \\HomCod\ = cat_FUNCT \ (\\HomDom\) \" unfolding exp_cat_cf_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda exp_cat_cf_components(1)[where \=\ and \=\, unfolded cat_cs_simps] |vsv exp_cat_cf_components_ObjMap_vsv[cat_FUNCT_cs_intros]| |vdomain exp_cat_cf_components_ObjMap_vdomain[cat_FUNCT_cs_simps]| |app exp_cat_cf_components_ObjMap_app[cat_FUNCT_cs_simps]| end subsubsection\Arrow map\ context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda exp_cat_cf_components(2)[where \=\ and \=\, unfolded cat_cs_simps] |vsv exp_cat_cf_components_ArrMap_vsv[cat_FUNCT_cs_intros]| |vdomain exp_cat_cf_components_ArrMap_vdomain[cat_FUNCT_cs_simps]| |app exp_cat_cf_components_ArrMap_app[cat_FUNCT_cs_simps]| end subsubsection\Domain and codomain\ context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) lemmas exp_cat_cf_HomDom[cat_FUNCT_cs_simps] = exp_cat_cf_components(3)[where \=\ and \=\, unfolded cat_cs_simps] and exp_cat_cf_HomCod[cat_FUNCT_cs_simps] = exp_cat_cf_components(4)[where \=\ and \=\, unfolded cat_cs_simps] end subsubsection\Category raised to the power of a functor is a functor\ lemma exp_cat_cf_is_tiny_functor: assumes "\ \" and "\ \\<^sub>\ \" and "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" proof- interpret \: \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(3)) interpret \: is_functor \ \ \ \ by (rule assms(4)) from assms(2-4) interpret \\: tiny_category \ \cat_FUNCT \ \ \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2-4) interpret \\: tiny_category \ \cat_FUNCT \ \ \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) show ?thesis proof(intro is_tiny_functorI' is_functorI') show "vfsequence (exp_cat_cf \ \ \)" unfolding exp_cat_cf_def by auto show "vcard (exp_cat_cf \ \ \) = 4\<^sub>\" unfolding exp_cat_cf_def by (simp_all add: nat_omega_simps) show "\\<^sub>\ (exp_cat_cf \ \ \\ObjMap\) \\<^sub>\ cat_FUNCT \ \ \\Obj\" proof ( unfold cat_FUNCT_components exp_cat_cf_components, intro vrange_VLambda_vsubset, unfold cat_cs_simps ) fix \ assume "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from assms(2-4) \' show "cf_map (cf_of_cf_map \ \ \ \\<^sub>C\<^sub>F \) \\<^sub>\ cf_maps \ \ \" unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "exp_cat_cf \ \ \\ArrMap\\\\ : exp_cat_cf \ \ \\ObjMap\\\\ \\<^bsub>cat_FUNCT \ \ \\<^esub> exp_cat_cf \ \ \\ObjMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" for \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that] from \(1) assms(2-4) show ?thesis by (subst \(2), use nothing in \subst \(3), subst \(4)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "exp_cat_cf \ \ \\ArrMap\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> \\ = exp_cat_cf \ \ \\ArrMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> exp_cat_cf \ \ \\ArrMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" and "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" for \ \ \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that(1)] and \ = cat_FUNCT_is_arrD[OF that(2)] from \(1) \(1) assms(2-4) show ?thesis by (subst (1 2) \(2), use nothing in \subst (1 2) \(2)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "exp_cat_cf \ \ \\ArrMap\\cat_FUNCT \ \ \\CId\\\\\ = cat_FUNCT \ \ \\CId\\exp_cat_cf \ \ \\ObjMap\\\\\" if "\ \\<^sub>\ cat_FUNCT \ \ \\Obj\" for \ proof- from that have \: "\ \\<^sub>\ cf_maps \ \ \" unfolding cat_FUNCT_components by simp then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from assms(2-4) \ \' show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_FUNCT_components(1) \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed qed ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros cat_cs_intros )+ qed lemma exp_cat_cf_is_tiny_functor'[cat_FUNCT_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = cat_FUNCT \ \ \" and "\' = cat_FUNCT \ \ \" shows "exp_cat_cf \ \ \ : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> \'" using assms(1-4) unfolding assms(5,6) by (rule exp_cat_cf_is_tiny_functor) subsubsection\Further properties\ lemma exp_cat_cf_cat_cf_id: assumes "category \ \" and "category \ \" shows "exp_cat_cf \ \ (cf_id \) = cf_id (cat_FUNCT \ \ \)" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule cf_eqI) from \\ show "exp_cat_cf \ \ (cf_id \) : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ show "cf_id (cat_FUNCT \ \ \) : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have ObjMap_dom_lhs: "\\<^sub>\ (exp_cat_cf \ \ (cf_id \)\ObjMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ have ObjMap_dom_rhs: "\\<^sub>\ (cf_id (cat_FUNCT \ \ \)\ObjMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_cf \ \ (cf_id \)\ObjMap\ = cf_id (cat_FUNCT \ \ \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_FUNCT_components(1)) fix \ assume "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from \' show "exp_cat_cf \ \ (cf_id \)\ObjMap\\\\ = cf_id (cat_FUNCT \ \ \)\ObjMap\\\\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros)+ from \\ have ArrMap_dom_lhs: "\\<^sub>\ (exp_cat_cf \ \ (cf_id \)\ArrMap\) = cat_FUNCT \ \ \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ have ArrMap_dom_rhs: "\\<^sub>\ (cf_id (cat_FUNCT \ \ \)\ArrMap\) = cat_FUNCT \ \ \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_cf \ \ (cf_id \)\ArrMap\ = cf_id (cat_FUNCT \ \ \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \ \\Arr\" then obtain \ \' where \: "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \'" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from \(1) show "exp_cat_cf \ \ (cf_id \)\ArrMap\\\\ = cf_id (cat_FUNCT \ \ \)\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros)+ qed simp_all qed lemma exp_cat_cf_cf_comp: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \) = exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \ \
\ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule cf_eqI) from \ \\ show "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \) : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \ \\ show "exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \ : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \ \\ have ObjMap_dom_lhs: "\\<^sub>\ (exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ObjMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \ \\ have ObjMap_dom_rhs: "\\<^sub>\ ((exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ObjMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \ \\ have ArrMap_dom_lhs: "\\<^sub>\ (exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ArrMap\) = cat_FUNCT \ \
\\Arr\" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \ \\ have ArrMap_dom_rhs: "\\<^sub>\ ((exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ArrMap\) = cat_FUNCT \ \
\\Arr\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ObjMap\ = (exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ObjMap\" proof ( rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_FUNCT_components(1) ) fix \ assume "\ \\<^sub>\ cf_maps \ \
\" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from \ \\ \' assms show "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ObjMap\\\\ = (exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ObjMap\\\\" unfolding \_def (*slow*) by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed ( use \ \\ in \ cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros \ )+ show "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ArrMap\ = (exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix \ assume "\ \\<^sub>\ cat_FUNCT \ \
\\Arr\" then obtain \ \' where \: "\ : \ \\<^bsub>cat_FUNCT \ \
\\<^esub> \'" by (auto intro: is_arrI) note \ = cat_FUNCT_is_arrD[OF \] from assms \(1) \ \\ show "exp_cat_cf \ \ (\ \\<^sub>C\<^sub>F \)\ArrMap\\\\ = (exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \)\ArrMap\\\\" by (subst (1 2) \(2)) ( cs_concl cs_simp: cat_FUNCT_cs_simps cat_cs_simps ntcf_cf_comp_ntcf_cf_comp_assoc cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed ( use \ \\ in \ cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros \ )+ qed simp_all qed subsection\Natural transformation raised to the power of a category\ subsubsection\Definition and elementary properties\ definition exp_ntcf_cat :: "V \ V \ V \ V" where "exp_ntcf_cat \ \ \
= [ ( \\\\<^sub>\cat_FUNCT \ \
(\\NTDGDom\)\Obj\. ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_of_cf_map \
(\\NTDGDom\) \) ), exp_cf_cat \ (\\NTDom\) \
, exp_cf_cat \ (\\NTCod\) \
, cat_FUNCT \ \
(\\NTDGDom\), cat_FUNCT \ \
(\\NTDGCod\) ]\<^sub>\" text\Components.\ lemma exp_ntcf_cat_components: shows "exp_ntcf_cat \ \ \
\NTMap\ = ( \\\\<^sub>\cat_FUNCT \ \
(\\NTDGDom\)\Obj\. ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_of_cf_map \
(\\NTDGDom\) \) )" and "exp_ntcf_cat \ \ \
\NTDom\ = exp_cf_cat \ (\\NTDom\) \
" and "exp_ntcf_cat \ \ \
\NTCod\ = exp_cf_cat \ (\\NTCod\) \
" and "exp_ntcf_cat \ \ \
\NTDGDom\ = cat_FUNCT \ \
(\\NTDGDom\)" and "exp_ntcf_cat \ \ \
\NTDGCod\ = cat_FUNCT \ \
(\\NTDGCod\)" unfolding exp_ntcf_cat_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda exp_ntcf_cat_components(1) |vsv exp_ntcf_cat_components_NTMap_vsv[cat_FUNCT_cs_intros]| context is_ntcf begin lemmas exp_ntcf_cat_components' = exp_ntcf_cat_components[where \=\ and \=\, unfolded cat_cs_simps] lemmas [cat_FUNCT_cs_simps] = exp_ntcf_cat_components'(2-5) mk_VLambda exp_ntcf_cat_components(1)[where \=\, unfolded cat_cs_simps] |vdomain exp_ntcf_cat_components_NTMap_vdomain[cat_FUNCT_cs_simps]| |app exp_ntcf_cat_components_NTMap_app[cat_FUNCT_cs_simps]| end lemmas [cat_FUNCT_cs_simps] = is_ntcf.exp_ntcf_cat_components'(2-5) is_ntcf.exp_ntcf_cat_components_NTMap_vdomain is_ntcf.exp_ntcf_cat_components_NTMap_app subsubsection\ Natural transformation raised to the power of a category is a natural transformation \ lemma exp_ntcf_cat_is_tiny_ntcf: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" shows "exp_ntcf_cat \ \ \
: exp_cf_cat \ \ \
\\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y exp_cf_cat \ \ \
: cat_FUNCT \ \
\ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \
\" proof(rule is_tiny_ntcfI') interpret \: \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) interpret \
: category \ \
by (rule assms(4)) let ?exp_\ = \exp_ntcf_cat \ \ \
\ let ?exp_\ = \exp_cf_cat \ \ \
\ let ?exp_\ = \exp_cf_cat \ \ \
\ from assms(1,2) show "exp_cf_cat \ \ \
: cat_FUNCT \ \
\ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \
\" by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(1,2) show "exp_cf_cat \ \ \
: cat_FUNCT \ \
\ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \
\" by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) show "?exp_\ : ?exp_\ \\<^sub>C\<^sub>F ?exp_\ : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" proof(rule is_ntcfI') show "vfsequence (?exp_\)" unfolding exp_ntcf_cat_def by auto show "vcard (?exp_\) = 5\<^sub>\" unfolding exp_ntcf_cat_def by (simp add: nat_omega_simps) from assms(1,2) show "?exp_\ : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from assms(1,2) show "?exp_\ : cat_FUNCT \ \
\ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \
\" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "?exp_\\NTMap\\\\ : ?exp_\\ObjMap\\\\ \\<^bsub>cat_FUNCT \ \
\\<^esub> ?exp_\\ObjMap\\\\" if "\ \\<^sub>\ cat_FUNCT \ \
\\Obj\" for \ proof- from that[unfolded cat_FUNCT_cs_simps] have "\ \\<^sub>\ cf_maps \ \
\" by simp then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by auto from \' show ?thesis by ( cs_concl cs_simp: cat_FUNCT_cs_simps \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "?exp_\\NTMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \
\\<^esub> ?exp_\\ArrMap\\\\ = ?exp_\\ArrMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \
\\<^esub> ?exp_\\NTMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \
\\<^esub> \" for \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that] let ?\ = \cf_of_cf_map \
\ \\ and ?\ = \cf_of_cf_map \
\ \\ and ?\ = \ntcf_of_ntcf_arrow \
\ \\ have [cat_cs_simps]: "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) = (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\)" proof(rule ntcf_eqI) from \(1) show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) : \ \\<^sub>C\<^sub>F ?\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?\ : \
\\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from \(1) show "(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) : \ \\<^sub>C\<^sub>F ?\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?\ : \
\\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from \(1) have dom_lhs: "\\<^sub>\ (((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\))\NTMap\) = \
\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \(1) have dom_rhs: "\\<^sub>\ (((\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\))\NTMap\) = \
\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\))\NTMap\ = ((\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\))\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix d assume "d \\<^sub>\ \
\Obj\" with \(1) show "((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\))\NTMap\\d\ = ((\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F ?\))\NTMap\\d\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all from \(1,3,4) that show ?thesis by (subst (1 2) \(2), use nothing in \subst \(3), subst \(4)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed qed ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_ntcf_cat_is_tiny_ntcf'[cat_FUNCT_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" and "\' = exp_cf_cat \ \ \
" and "\' = exp_cf_cat \ \ \
" and "\' = cat_FUNCT \ \
\" and "\' = cat_FUNCT \ \
\" shows "exp_ntcf_cat \ \ \
: \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y \' : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> \'" using assms(1-4) unfolding assms(5-8) by (rule exp_ntcf_cat_is_tiny_ntcf) subsubsection\Further properties\ lemma exp_ntcf_cat_cf_ntcf_comp: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" shows "exp_ntcf_cat \ (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \
= exp_cf_cat \ \ \
\\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \
" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \
: category \ \
by (rule assms(3)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ have dom_lhs: "\\<^sub>\ (exp_ntcf_cat \ (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \
\NTMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cf_cat \ \ \
\\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \
)\NTMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_ntcf_cat \ (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \
\NTMap\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \
)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \
\" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: is_arrI) from \\ prems \' show "exp_ntcf_cat \ (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \
\NTMap\\\\ = (exp_cf_cat \ \ \
\\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \
)\NTMap\\\\" by ( cs_concl cs_simp: cf_ntcf_comp_ntcf_cf_comp_assoc cat_cs_simps cat_FUNCT_cs_simps \_def cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed ( cs_concl cs_simp: exp_cf_cat_cf_comp cat_cs_simps cat_FUNCT_cs_simps cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_ntcf_cat_ntcf_cf_comp: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" shows "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \
= exp_ntcf_cat \ \ \
\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cf_cat \ \ \
" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \
: category \ \
by (rule assms(3)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ have dom_lhs: "\\<^sub>\ (exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \
\NTMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_ntcf_cat \ \ \
\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\NTMap\) = cat_FUNCT \ \
\\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \
\NTMap\ = (exp_ntcf_cat \ \ \
\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \
\" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \
\\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: is_arrI) from \\ assms prems \' show "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \
\NTMap\\\\ = (exp_ntcf_cat \ \ \
\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cf_cat \ \ \
)\NTMap\\\\" by ( cs_concl cs_simp: ntcf_cf_comp_ntcf_cf_comp_assoc cat_cs_simps cat_FUNCT_cs_simps \_def cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed ( cs_concl cs_simp: exp_cf_cat_cf_comp cat_cs_simps cat_FUNCT_cs_simps cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_ntcf_cat_ntcf_vcomp: assumes "category \ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \ = exp_ntcf_cat \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ show "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \ : exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ show "exp_ntcf_cat \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \ : exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_lhs: "\\<^sub>\ ((exp_ntcf_cat \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \)\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) have dom_rhs: "\\<^sub>\ (exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\NTMap\) = cat_FUNCT \ \ \\Obj\" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) show "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\NTMap\ = (exp_ntcf_cat \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \' assume "\' \\<^sub>\ cf_maps \ \ \" then obtain \'' where \'_def: "\' = cf_map \''" and \'': "\'' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \'' \\ show "exp_ntcf_cat \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\NTMap\\\'\ = (exp_ntcf_cat \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_ntcf_cat \ \ \)\NTMap\\\'\" unfolding \'_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros)+ qed simp_all qed lemma ntcf_id_exp_cf_cat: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id (exp_cf_cat \ \ \) = exp_ntcf_cat \ (ntcf_id \) \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ show "exp_ntcf_cat \ (ntcf_id \) \ : exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ show "ntcf_id (exp_cf_cat \ \ \) : exp_cf_cat \ \ \ \\<^sub>C\<^sub>F exp_cf_cat \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ assms have dom_lhs: "\\<^sub>\ (ntcf_id (exp_cf_cat \ \ \)\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ assms have dom_rhs: "\\<^sub>\ (exp_ntcf_cat \ (ntcf_id \) \\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) show "ntcf_id (exp_cf_cat \ \ \)\NTMap\ = exp_ntcf_cat \ (ntcf_id \) \\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \' \\ show "ntcf_id (exp_cf_cat \ \ \)\NTMap\\\\ = exp_ntcf_cat \ (ntcf_id \) \\NTMap\\\\" unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed ( cs_concl cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed simp_all qed subsection\Category raised to the power of the natural transformation\ subsubsection\Definition and elementary properties\ definition exp_cat_ntcf :: "V \ V \ V \ V" where "exp_cat_ntcf \ \ \ = [ ( \\\\<^sub>\cat_FUNCT \ (\\NTDGCod\) \\Obj\. ntcf_arrow (cf_of_cf_map (\\NTDGCod\) \ \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) ), exp_cat_cf \ \ (\\NTDom\), exp_cat_cf \ \ (\\NTCod\), cat_FUNCT \ (\\NTDGCod\) \, cat_FUNCT \ (\\NTDGDom\) \ ]\<^sub>\" text\Components.\ lemma exp_cat_ntcf_components: shows "exp_cat_ntcf \ \ \\NTMap\ = ( \\\\<^sub>\cat_FUNCT \ (\\NTDGCod\) \\Obj\. ntcf_arrow (cf_of_cf_map (\\NTDGCod\) \ \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) )" and "exp_cat_ntcf \ \ \\NTDom\ = exp_cat_cf \ \ (\\NTDom\)" and "exp_cat_ntcf \ \ \\NTCod\ = exp_cat_cf \ \ (\\NTCod\)" and "exp_cat_ntcf \ \ \\NTDGDom\ = cat_FUNCT \ (\\NTDGCod\) \" and "exp_cat_ntcf \ \ \\NTDGCod\ = cat_FUNCT \ (\\NTDGDom\) \" unfolding exp_cat_ntcf_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda exp_cat_ntcf_components(1) |vsv exp_cat_ntcf_components_NTMap_vsv[cat_FUNCT_cs_intros]| context is_ntcf begin lemmas exp_cat_ntcf_components' = exp_cat_ntcf_components[where \=\ and \=\, unfolded cat_cs_simps] lemmas [cat_FUNCT_cs_simps] = exp_cat_ntcf_components'(2-5) mk_VLambda exp_cat_ntcf_components(1)[where \=\, unfolded cat_cs_simps] |vdomain exp_cat_ntcf_components_NTMap_vdomain[cat_FUNCT_cs_simps]| |app exp_cat_ntcf_components_NTMap_app[cat_FUNCT_cs_simps]| end lemmas exp_cat_ntcf_components' = is_ntcf.exp_cat_ntcf_components' lemmas [cat_FUNCT_cs_simps] = is_ntcf.exp_cat_ntcf_components'(2-5) is_ntcf.exp_cat_ntcf_components_NTMap_vdomain is_ntcf.exp_cat_ntcf_components_NTMap_app subsubsection\ Category raised to the power of a natural transformation is a natural transformation \ lemma exp_cat_ntcf_is_tiny_ntcf: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \" shows "exp_cat_ntcf \ \ \ : exp_cat_cf \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" proof(rule is_tiny_ntcfI') interpret \: \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) interpret \: category \ \ by (rule assms(4)) let ?exp_\ = \exp_cat_ntcf \ \ \\ let ?exp_\ = \exp_cat_cf \ \ \\ let ?exp_\ = \exp_cat_cf \ \ \\ from assms(1,2) show "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(1,2) show "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) show "?exp_\ : ?exp_\ \\<^sub>C\<^sub>F ?exp_\ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" proof(rule is_ntcfI') show "vfsequence (?exp_\)" unfolding exp_cat_ntcf_def by auto show "vcard (?exp_\) = 5\<^sub>\" unfolding exp_cat_ntcf_def by (simp add: nat_omega_simps) from assms(1,2) show "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from assms(1,2) show "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_ntcf \ \ \\NTMap\\\\ : exp_cat_cf \ \ \\ObjMap\\\\ \\<^bsub>cat_FUNCT \ \ \\<^esub> exp_cat_cf \ \ \\ObjMap\\\\" if "\ \\<^sub>\ cat_FUNCT \ \ \\Obj\" for \ proof- from that[unfolded cat_FUNCT_cs_simps] have "\ \\<^sub>\ cf_maps \ \ \" by simp then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \' show ?thesis unfolding \_def by ( cs_concl cs_simp: cat_FUNCT_cs_simps \_def cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "?exp_\\NTMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> ?exp_\\ArrMap\\\\ = ?exp_\\ArrMap\\\\ \\<^sub>A\<^bsub>cat_FUNCT \ \ \\<^esub> ?exp_\\NTMap\\\\" if "\ : \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \" for \ \ \ proof- note \ = cat_FUNCT_is_arrD[OF that] let ?\ = \cf_of_cf_map \ \ \\ and ?\ = \cf_of_cf_map \ \ \\ and ?\ = \ntcf_of_ntcf_arrow \ \ \\ have [cat_cs_simps]: "(?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) = (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" proof(rule ntcf_eqI) from \(1) show "(?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) : ?\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from \(1) show "(?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : ?\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from \(1) have dom_lhs: "\\<^sub>\ (((?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \(1) have dom_rhs: "\\<^sub>\ (((?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "((?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\ = ((?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with \(1) show "((?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\\a\ = ((?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\\a\" by ( cs_concl cs_simp: cat_cs_simps is_ntcf.ntcf_Comp_commute cs_intro: cat_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all from \(1,3,4) that show ?thesis by (subst (1 2) \(2), use nothing in \subst \(3), subst \(4)\) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed qed ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_cat_ntcf_is_tiny_ntcf'[cat_FUNCT_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \" and "\' = exp_cat_cf \ \ \" and "\' = exp_cat_cf \ \ \" and "\' = cat_FUNCT \ \ \" and "\' = cat_FUNCT \ \ \" shows "exp_cat_ntcf \ \ \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y \' : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> \'" using assms(1-4) unfolding assms(5-8) by (rule exp_cat_ntcf_is_tiny_ntcf) subsubsection\Further properties\ lemma ntcf_id_exp_cat_cf: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id (exp_cat_cf \ \ \) = exp_cat_ntcf \ \ (ntcf_id \)" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ show "exp_cat_ntcf \ \ (ntcf_id \) : exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from assms \ \\ show "ntcf_id (exp_cat_cf \ \ \) : exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ assms have dom_lhs: "\\<^sub>\ (exp_cat_ntcf \ \ (ntcf_id \)\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from \\ assms have dom_rhs: "\\<^sub>\ (ntcf_id (exp_cat_cf \ \ \)\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) show "ntcf_id (exp_cat_cf \ \ \)\NTMap\ = exp_cat_ntcf \ \ (ntcf_id \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume "\ \\<^sub>\ cf_maps \ \ \" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from \' \\ show "ntcf_id (exp_cat_cf \ \ \)\NTMap\\\\ = exp_cat_ntcf \ \ (ntcf_id \)\NTMap\\\\" unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed ( cs_concl cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed simp_all qed lemma exp_cat_ntcf_ntcf_cf_comp: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" shows "exp_cat_ntcf \ \
(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) = exp_cat_cf \ \
\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \
\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \
: category \ \
by (rule assms(3)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ have dom_lhs: "\\<^sub>\ (exp_cat_ntcf \ \
(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\) = cat_FUNCT \ \ \
\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cat_cf \ \
\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \
\)\NTMap\) = cat_FUNCT \ \ \
\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_ntcf \ \
(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\ = (exp_cat_cf \ \
\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \
\)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \ \
" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by (auto intro: is_arrI) from \\ assms prems \' show "exp_cat_ntcf \ \
(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\\\\ = (exp_cat_cf \ \
\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \
\)\NTMap\\\\" unfolding \_def by ( cs_concl cs_simp: cf_ntcf_comp_ntcf_cf_comp_assoc cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed ( cs_concl cs_simp: exp_cat_cf_cf_comp cat_cs_simps cat_FUNCT_cs_simps cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_cat_ntcf_cf_ntcf_comp: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \
" shows "exp_cat_ntcf \ \
(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = exp_cat_ntcf \ \
\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cat_cf \ \
\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \
: category \ \
by (rule assms(3)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \\ have dom_lhs: "\\<^sub>\ (exp_cat_ntcf \ \
(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = cat_FUNCT \ \ \
\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cat_ntcf \ \
\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cat_cf \ \
\)\NTMap\) = cat_FUNCT \ \ \
\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_ntcf \ \
(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = (exp_cat_ntcf \ \
\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cat_cf \ \
\)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \ assume prems: "\ \\<^sub>\ cf_maps \ \ \
" then obtain \' where \_def: "\ = cf_map \'" and \': "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by (auto intro: is_arrI) from assms \\ prems \' show "exp_cat_ntcf \ \
(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\\\ = (exp_cat_ntcf \ \
\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F exp_cat_cf \ \
\)\NTMap\\\\" by ( cs_concl cs_simp: cf_comp_cf_ntcf_comp_assoc cat_cs_simps cat_FUNCT_cs_simps \_def cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed ( cs_concl cs_simp: exp_cat_cf_cf_comp cat_cs_simps cat_FUNCT_cs_simps cs_intro: \\ cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemma exp_cat_ntcf_ntcf_vcomp: assumes "category \ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "exp_cat_ntcf \ \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = exp_cat_ntcf \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \ \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show ?thesis proof(rule ntcf_eqI) from \ \\ show "exp_cat_ntcf \ \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ show "exp_cat_ntcf \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \ \ : exp_cat_cf \ \ \ \\<^sub>C\<^sub>F exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_lhs: "\\<^sub>\ ((exp_cat_ntcf \ \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \\ have dom_rhs: "\\<^sub>\ ((exp_cat_ntcf \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \ \)\NTMap\) = cat_FUNCT \ \ \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "exp_cat_ntcf \ \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = (exp_cat_ntcf \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \ \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_FUNCT_components(1)) fix \' assume "\' \\<^sub>\ cf_maps \ \ \" then obtain \'' where \'_def: "\' = cf_map \''" and \'': "\'' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp from \'' \\ show "exp_cat_ntcf \ \ (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\\'\ = (exp_cat_ntcf \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F exp_cat_ntcf \ \ \)\NTMap\\\'\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cf_ntcf_comp_ntcf_vcomp \'_def cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros)+ qed simp_all qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Functor.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Functor.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Functor.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Functor.thy @@ -1,2041 +1,2042 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Functor\ theory CZH_ECAT_Functor imports CZH_ECAT_Category CZH_Foundations.CZH_SMC_Semifunctor begin subsection\Background\ named_theorems cf_cs_simps named_theorems cf_cs_intros named_theorems cat_cn_cs_simps named_theorems cat_cn_cs_intros lemmas [cat_cs_simps] = dg_shared_cs_simps lemmas [cat_cs_intros] = dg_shared_cs_intros subsubsection\Slicing\ definition cf_smcf :: "V \ V" where "cf_smcf \ = [\\ObjMap\, \\ArrMap\, cat_smc (\\HomDom\), cat_smc (\\HomCod\)]\<^sub>\" text\Components.\ lemma cf_smcf_components: shows [slicing_simps]: "cf_smcf \\ObjMap\ = \\ObjMap\" and [slicing_simps]: "cf_smcf \\ArrMap\ = \\ArrMap\" and [slicing_commute]: "cf_smcf \\HomDom\ = cat_smc (\\HomDom\)" and [slicing_commute]: "cf_smcf \\HomCod\ = cat_smc (\\HomCod\)" unfolding cf_smcf_def dghm_field_simps by (auto simp: nat_omega_simps) subsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_functor = \ \ + vfsequence \ + HomDom: category \ \ + HomCod: category \ \ for \ \ \ \ + assumes cf_length[cat_cs_simps]: "vcard \ = 4\<^sub>\" and cf_is_semifunctor[slicing_intros]: "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" and cf_HomDom[cat_cs_simps]: "\\HomDom\ = \" and cf_HomCod[cat_cs_simps]: "\\HomCod\ = \" and cf_ObjMap_CId[cat_cs_intros]: "c \\<^sub>\ \\Obj\ \ \\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\" syntax "_is_functor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>C\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_functor \ \ \ \" abbreviation (input) is_cn_cf :: "V \ V \ V \ V \ bool" where "is_cn_cf \ \ \ \ \ \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> \" syntax "_is_cn_cf" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \<^sub>C\\\ _)\ [51, 51, 51] 51) translations "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" \ "CONST is_cn_cf \ \ \ \" abbreviation all_cfs :: "V \ V" where "all_cfs \ \ set {\. \\ \. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" abbreviation cfs :: "V \ V \ V \ V" where "cfs \ \ \ \ set {\. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" lemmas [cat_cs_simps] = is_functor.cf_length is_functor.cf_HomDom is_functor.cf_HomCod is_functor.cf_ObjMap_CId lemma cn_cf_ObjMap_CId[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" shows "\\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\" proof- interpret is_functor \ \op_cat \\ \ \ by (rule assms(1)) from assms(2) have c: "c \\<^sub>\ op_cat \\Obj\" unfolding cat_op_simps by simp show ?thesis by (rule cf_ObjMap_CId[OF c, unfolded cat_op_simps]) qed lemma (in is_functor) cf_is_semifunctor': assumes "\' = cat_smc \" and "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule cf_is_semifunctor) lemmas [slicing_intros] = is_functor.cf_is_semifunctor' lemma cn_smcf_comp_is_semifunctor: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "cf_smcf \ : cat_smc \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub>cat_smc \" using assms unfolding slicing_simps slicing_commute by (rule is_functor.cf_is_semifunctor) lemma cn_smcf_comp_is_semifunctor'[slicing_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\' = op_smc (cat_smc \)" and "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \'" using assms(1) unfolding assms(2,3) by (rule cn_smcf_comp_is_semifunctor) text\Rules.\ lemma (in is_functor) is_functor_axioms'[cat_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_functor_axioms) mk_ide rf is_functor_def[unfolded is_functor_axioms_def] |intro is_functorI| |dest is_functorD[dest]| |elim is_functorE[elim]| lemmas [cat_cs_intros] = is_functorD(3,4) lemma is_functorI': assumes "\ \" and "vfsequence \" and "category \ \" and "category \ \" and "vcard \ = 4\<^sub>\" and "\\HomDom\ = \" and "\\HomCod\ = \" and "vsv (\\ObjMap\)" and "vsv (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) \\<^sub>\ \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "\b c g a f. \ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ \\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" and "(\c. c \\<^sub>\ \\Obj\ \ \\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\)" shows "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( intro is_functorI is_semifunctorI', unfold cf_smcf_components slicing_simps ) (simp_all add: assms cf_smcf_def nat_omega_simps category.cat_semicategory) lemma is_functorD': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \" and "vfsequence \" and "category \ \" and "category \ \" and "vcard \ = 4\<^sub>\" and "\\HomDom\ = \" and "\\HomCod\ = \" and "vsv (\\ObjMap\)" and "vsv (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) \\<^sub>\ \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "\b c g a f. \ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ \\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" and "(\c. c \\<^sub>\ \\Obj\ \ \\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\)" by ( simp_all add: is_functorD(2-9)[OF assms] is_semifunctorD'[OF is_functorD(6)[OF assms], unfolded slicing_simps] ) lemma is_functorE': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains "\ \" and "vfsequence \" and "category \ \" and "category \ \" and "vcard \ = 4\<^sub>\" and "\\HomDom\ = \" and "\\HomCod\ = \" and "vsv (\\ObjMap\)" and "vsv (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) \\<^sub>\ \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "\b c g a f. \ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ \\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" and "(\c. c \\<^sub>\ \\Obj\ \ \\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\)" using assms by (simp add: is_functorD') text\A functor is a semifunctor.\ context is_functor begin interpretation smcf: is_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule cf_is_semifunctor) sublocale ObjMap: vsv \\\ObjMap\\ by (rule smcf.ObjMap.vsv_axioms[unfolded slicing_simps]) sublocale ArrMap: vsv \\\ArrMap\\ by (rule smcf.ArrMap.vsv_axioms[unfolded slicing_simps]) lemmas_with [unfolded slicing_simps]: cf_ObjMap_vsv = smcf.smcf_ObjMap_vsv and cf_ArrMap_vsv = smcf.smcf_ArrMap_vsv and cf_ObjMap_vdomain[cat_cs_simps] = smcf.smcf_ObjMap_vdomain and cf_ObjMap_vrange = smcf.smcf_ObjMap_vrange and cf_ArrMap_vdomain[cat_cs_simps] = smcf.smcf_ArrMap_vdomain and cf_ArrMap_is_arr = smcf.smcf_ArrMap_is_arr and cf_ArrMap_is_arr''[cat_cs_intros] = smcf.smcf_ArrMap_is_arr'' and cf_ArrMap_is_arr'[cat_cs_intros] = smcf.smcf_ArrMap_is_arr' and cf_ObjMap_app_in_HomCod_Obj[cat_cs_intros] = smcf.smcf_ObjMap_app_in_HomCod_Obj and cf_ArrMap_vrange = smcf.smcf_ArrMap_vrange and cf_ArrMap_app_in_HomCod_Arr[cat_cs_intros] = smcf.smcf_ArrMap_app_in_HomCod_Arr and cf_ObjMap_vsubset_Vset = smcf.smcf_ObjMap_vsubset_Vset and cf_ArrMap_vsubset_Vset = smcf.smcf_ArrMap_vsubset_Vset and cf_ObjMap_in_Vset = smcf.smcf_ObjMap_in_Vset and cf_ArrMap_in_Vset = smcf.smcf_ArrMap_in_Vset and cf_is_semifunctor_if_ge_Limit = smcf.smcf_is_semifunctor_if_ge_Limit and cf_is_arr_HomCod = smcf.smcf_is_arr_HomCod and cf_vimage_dghm_ArrMap_vsubset_Hom = smcf.smcf_vimage_dghm_ArrMap_vsubset_Hom lemmas_with [unfolded slicing_simps]: cf_ArrMap_Comp = smcf.smcf_ArrMap_Comp end lemmas [cat_cs_simps] = is_functor.cf_ObjMap_vdomain is_functor.cf_ArrMap_vdomain is_functor.cf_ArrMap_Comp lemmas [cat_cs_intros] = is_functor.cf_ObjMap_app_in_HomCod_Obj is_functor.cf_ArrMap_app_in_HomCod_Arr is_functor.cf_ArrMap_is_arr' text\Elementary properties.\ lemma cn_cf_ArrMap_Comp[cat_cn_cs_simps]: assumes "category \ \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "g : c \\<^bsub>\\<^esub> b" and "f : b \\<^bsub>\\<^esub> a" shows "\\ArrMap\\f \\<^sub>A\<^bsub>\\<^esub> g\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" proof- interpret \: category \ \ by (rule assms(1)) interpret \: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule cn_smcf_ArrMap_Comp [ OF \.cat_semicategory \.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_simps, OF assms(3,4) ] ) qed lemma cf_eqI: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\\ObjMap\ = \\ObjMap\" and "\\ArrMap\ = \\ArrMap\" and "\ = \" and "\ = \
" shows "\ = \" proof(rule vsv_eqI) interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \
\ by (rule assms(2)) from assms(1) show "vsv \" by auto from assms(2) show "vsv \" by auto have dom: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: cat_cs_simps V_cs_simps) show "\\<^sub>\ \ = \\<^sub>\ \" by (cs_concl cs_simp: cat_cs_simps V_cs_simps) from assms(5,6) have sup: "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by (simp_all add: cat_cs_simps) show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = \\a\" for a by (unfold dom, elim_in_numeral, insert assms(3,4) sup) (auto simp: dghm_field_simps) qed lemma cf_smcf_eqI: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ = \" and "\ = \
" and "cf_smcf \ = cf_smcf \" shows "\ = \" proof(rule cf_eqI) from assms(5) have "cf_smcf \\ObjMap\ = cf_smcf \\ObjMap\" "cf_smcf \\ArrMap\ = cf_smcf \\ArrMap\" by simp_all then show "\\ObjMap\ = \\ObjMap\" "\\ArrMap\ = \\ArrMap\" unfolding slicing_simps by simp_all qed (auto intro: assms(1,2) simp: assms(3-5)) lemma (in is_functor) cf_def: "\ = [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: cat_cs_simps V_cs_simps) have dom_rhs: "\\<^sub>\ [\\Obj\, \\Arr\, \\Dom\, \\Cod\]\<^sub>\ = 4\<^sub>\" by (simp add: nat_omega_simps) then show "\\<^sub>\ \ = \\<^sub>\ [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\" unfolding dom_lhs dom_rhs by (simp add: nat_omega_simps) show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\\a\" for a by (unfold dom_lhs, elim_in_numeral, unfold dghm_field_simps) (simp_all add: nat_omega_simps) qed (auto simp: vsv_axioms) text\Size.\ lemma (in is_functor) cf_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "\ \\<^sub>\ Vset \" proof- interpret \: \ \ by (rule assms(1)) note [cat_cs_intros] = cf_ObjMap_in_Vset cf_ArrMap_in_Vset HomDom.cat_in_Vset HomCod.cat_in_Vset from assms(2) show ?thesis by (subst cf_def) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros) qed lemma (in is_functor) cf_is_functor_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule is_functorI) ( auto simp: cat_cs_simps assms vfsequence_axioms cf_is_semifunctor_if_ge_Limit HomDom.cat_category_if_ge_Limit HomCod.cat_category_if_ge_Limit intro: cat_cs_intros ) lemma small_all_cfs[simp]: "small {\. \\ \. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" proof(cases \\ \\) case True from is_functor.cf_in_Vset show ?thesis by (intro down[of _ \Vset (\ + \)\]) (auto simp: True \.\_Limit_\\ \.\_\_\\ \.intro \.\_\_\\) next case False then have "{\. \\ \. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \} = {}" by auto then show ?thesis by simp qed lemma (in is_functor) cf_in_Vset_7: "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof- note [folded VPow_iff, folded Vset_succ[OF Ord_\], cat_cs_intros] = cf_ObjMap_vsubset_Vset cf_ArrMap_vsubset_Vset from HomDom.cat_category_in_Vset_4 have [cat_cs_intros]: "\ \\<^sub>\ Vset (succ (succ (succ (succ \))))" by (succ_of_numeral) (cs_prems cs_simp: plus_V_succ_right V_cs_simps) from HomCod.cat_category_in_Vset_4 have [cat_cs_intros]: "\ \\<^sub>\ Vset (succ (succ (succ (succ \))))" by (succ_of_numeral) (cs_prems cs_simp: plus_V_succ_right V_cs_simps) show ?thesis by (subst cf_def, succ_of_numeral) ( cs_concl cs_simp: plus_V_succ_right V_cs_simps cat_cs_simps cs_intro: cat_cs_intros V_cs_intros ) qed lemma (in \) all_cfs_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "all_cfs \ \\<^sub>\ Vset \" proof(rule vsubset_in_VsetI) interpret \: \ \ by (rule assms(1)) show "all_cfs \ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof(intro vsubsetI) fix \ assume "\ \\<^sub>\ all_cfs \" then obtain \ \ where \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret is_functor \ \ \ \ using \ by simp show "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" by (rule cf_in_Vset_7) qed from assms(2) show "Vset (\ + 7\<^sub>\) \\<^sub>\ Vset \" by (cs_concl cs_intro: V_cs_intros Ord_cs_intros) qed lemma small_cfs[simp]: "small {\. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" by (rule down[of _ \set {\. \\ \. \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}\]) auto subsection\Opposite functor\ subsubsection\Definition and elementary properties\ text\See Chapter II-2 in \cite{mac_lane_categories_2010}.\ definition op_cf :: "V \ V" where "op_cf \ = [\\ObjMap\, \\ArrMap\, op_cat (\\HomDom\), op_cat (\\HomCod\)]\<^sub>\" text\Components.\ lemma op_cf_components[cat_op_simps]: shows "op_cf \\ObjMap\ = \\ObjMap\" and "op_cf \\ArrMap\ = \\ArrMap\" and "op_cf \\HomDom\ = op_cat (\\HomDom\)" and "op_cf \\HomCod\ = op_cat (\\HomCod\)" unfolding op_cf_def dghm_field_simps by (auto simp: nat_omega_simps) text\Slicing.\ lemma cf_smcf_op_cf[slicing_commute]: "op_smcf (cf_smcf \) = cf_smcf (op_cf \)" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (op_smcf (cf_smcf \)) = 4\<^sub>\" unfolding op_smcf_def by (auto simp: nat_omega_simps) have dom_rhs: "\\<^sub>\ (cf_smcf (op_cf \)) = 4\<^sub>\" unfolding cf_smcf_def by (auto simp: nat_omega_simps) show "\\<^sub>\ (op_smcf (cf_smcf \)) = \\<^sub>\ (cf_smcf (op_cf \))" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (op_smcf (cf_smcf \)) \ op_smcf (cf_smcf \)\a\ = cf_smcf (op_cf \)\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cf_smcf_def op_cf_def op_smcf_def dghm_field_simps ) (auto simp: nat_omega_simps slicing_commute) qed (auto simp: cf_smcf_def op_smcf_def) text\Elementary properties.\ lemma op_cf_vsv[cat_op_intros]: "vsv (op_cf \)" unfolding op_cf_def by auto subsubsection\Further properties\ lemma (in is_functor) is_functor_op: "op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" proof(intro is_functorI, unfold cat_op_simps) show "vfsequence (op_cf \)" unfolding op_cf_def by simp show "vcard (op_cf \) = 4\<^sub>\" unfolding op_cf_def by (auto simp: nat_omega_simps) fix c assume "c \\<^sub>\ \\Obj\" then show "\\ArrMap\\\\CId\\c\\ = \\CId\\\\ObjMap\\c\\" unfolding cat_op_simps by (auto intro: cat_cs_intros) qed ( auto simp: cat_cs_simps slicing_commute[symmetric] is_semifunctor.is_semifunctor_op cf_is_semifunctor HomCod.category_op HomDom.category_op ) lemma (in is_functor) is_functor_op'[cat_op_intros]: assumes "\' = op_cat \" and "\' = op_cat \" shows "op_cf \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms(1,2) by (rule is_functor_op) lemmas is_functor_op[cat_op_intros] = is_functor.is_functor_op' lemma (in is_functor) cf_op_cf_op_cf[cat_op_simps]: "op_cf (op_cf \) = \" proof(rule cf_eqI[of \ \ \ _ \ \], unfold cat_op_simps) show "op_cf (op_cf \) : \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( metis HomCod.cat_op_cat_op_cat HomDom.cat_op_cat_op_cat is_functor.is_functor_op is_functor_op ) qed (auto simp: cat_cs_intros) lemmas cf_op_cf_op_cf[cat_op_simps] = is_functor.cf_op_cf_op_cf lemma eq_op_cf_iff[cat_op_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "op_cf \ = op_cf \ \ \ = \" proof interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \
\ by (rule assms(2)) assume prems: "op_cf \ = op_cf \" show "\ = \" proof(rule cf_eqI[OF assms]) from prems R.cf_op_cf_op_cf L.cf_op_cf_op_cf show "\\ObjMap\ = \\ObjMap\" "\\ArrMap\ = \\ArrMap\" by metis+ from prems R.cf_op_cf_op_cf L.cf_op_cf_op_cf have "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by auto then show "\ = \" "\ = \
" by (simp_all add: cat_cs_simps) qed qed auto subsection\Composition of covariant functors\ subsubsection\Definition and elementary properties\ abbreviation (input) cf_comp :: "V \ V \ V" (infixl "\\<^sub>C\<^sub>F" 55) where "cf_comp \ dghm_comp" text\Slicing.\ lemma cf_smcf_smcf_comp[slicing_commute]: "cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ = cf_smcf (\ \\<^sub>C\<^sub>F \)" unfolding dghm_comp_def cf_smcf_def dghm_field_simps by (simp add: nat_omega_simps) subsubsection\Object map\ lemma cf_comp_ObjMap_vsv[cat_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>C\<^sub>F \)\ObjMap\)" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ObjMap_vsv [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>C\<^sub>F \)\ObjMap\) = \\Obj\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ObjMap_vdomain [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ObjMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>C\<^sub>F \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ObjMap_vrange [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ObjMap_app[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and [simp]: "a \\<^sub>\ \\Obj\" shows "(\ \\<^sub>C\<^sub>F \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ObjMap_app [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute, OF assms(3) ] ) qed subsubsection\Arrow map\ lemma cf_comp_ArrMap_vsv[cat_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>C\<^sub>F \)\ArrMap\)" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ArrMap_vsv [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>C\<^sub>F \)\ArrMap\) = \\Arr\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ArrMap_vdomain [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ArrMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>C\<^sub>F \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ArrMap_vrange [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute ] ) qed lemma cf_comp_ArrMap_app[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and [simp]: "f \\<^sub>\ \\Arr\" shows "(\ \\<^sub>C\<^sub>F \)\ArrMap\\f\ = \\ArrMap\\\\ArrMap\\f\\" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_comp_ArrMap_app [ OF L.cf_is_semifunctor R.cf_is_semifunctor, unfolded slicing_simps slicing_commute, OF assms(3) ] ) qed subsubsection\Further properties\ lemma cf_comp_is_functorI[cat_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: is_functor \ \ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_functorI, unfold dghm_comp_components(3,4)) show "vfsequence (\ \\<^sub>C\<^sub>F \)" by (simp add: dghm_comp_def) show "vcard (\ \\<^sub>C\<^sub>F \) = 4\<^sub>\" unfolding dghm_comp_def by (simp add: nat_omega_simps) show "cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" unfolding cf_smcf_smcf_comp[symmetric] by (cs_concl cs_intro: smc_cs_intros slicing_intros cat_cs_intros) fix c assume "c \\<^sub>\ \\Obj\" with assms show "(\ \\<^sub>C\<^sub>F \)\ArrMap\\\\CId\\c\\ = \\CId\\(\ \\<^sub>C\<^sub>F \)\ObjMap\\c\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: cat_cs_simps intro: cat_cs_intros) qed lemma cf_comp_assoc[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^sub>C\<^sub>F \) \\<^sub>C\<^sub>F \ = \ \\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F \)" proof(rule cf_eqI[of \ \ \
_ \ \
]) interpret \: is_functor \ \ \
\ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) from \.is_functor_axioms \.is_functor_axioms \.is_functor_axioms show "\ \\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F \) : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by (auto simp: cat_cs_simps intro: cat_cs_intros) qed (simp_all add: dghm_comp_components vcomp_assoc) text\The opposite of the covariant composition of functors.\ lemma op_cf_cf_comp[cat_op_simps]: "op_cf (\ \\<^sub>C\<^sub>F \) = op_cf \ \\<^sub>C\<^sub>F op_cf \" unfolding dghm_comp_def op_cf_def dghm_field_simps by (simp add: nat_omega_simps) subsection\Composition of contravariant functors\ subsubsection\Definition and elementary properties\ text\See section 1.2 in \cite{bodo_categories_1970}.\ definition cf_cn_comp :: "V \ V \ V" (infixl "\<^sub>C\<^sub>F\" 55) where "\ \<^sub>C\<^sub>F\ \ = [ \\ObjMap\ \\<^sub>\ \\ObjMap\, \\ArrMap\ \\<^sub>\ \\ArrMap\, op_cat (\\HomDom\), \\HomCod\ ]\<^sub>\" text\Components.\ lemma cf_cn_comp_components: shows "(\ \<^sub>C\<^sub>F\ \)\ObjMap\ = \\ObjMap\ \\<^sub>\ \\ObjMap\" and "(\ \<^sub>C\<^sub>F\ \)\ArrMap\ = \\ArrMap\ \\<^sub>\ \\ArrMap\" and [cat_cn_cs_simps]: "(\ \<^sub>C\<^sub>F\ \)\HomDom\ = op_cat (\\HomDom\)" and [cat_cn_cs_simps]: "(\ \<^sub>C\<^sub>F\ \)\HomCod\ = \\HomCod\" unfolding cf_cn_comp_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cf_smcf_cf_cn_comp[slicing_commute]: "cf_smcf \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ cf_smcf \ = cf_smcf (\ \<^sub>C\<^sub>F\ \)" unfolding smcf_cn_comp_def cf_cn_comp_def cf_smcf_def by (simp add: nat_omega_simps slicing_commute dghm_field_simps) subsubsection\Object map: two contravariant functors\ lemma cf_cn_comp_ObjMap_vsv[cat_cn_cs_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>C\<^sub>F\ \)\ObjMap\)" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ObjMap_vsv [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ObjMap_vdomain[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ObjMap\) = \\Obj\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ObjMap_vdomain [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ObjMap_vrange: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ObjMap_vrange [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ObjMap_app[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>C\<^sub>F\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ObjMap_app [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Arrow map: two contravariant functors\ lemma cf_cn_comp_ArrMap_vsv[cat_cn_cs_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>C\<^sub>F\ \)\ArrMap\)" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ArrMap_vsv [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ArrMap_vdomain[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ArrMap\) = \\Arr\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ArrMap_vdomain [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ArrMap_vrange: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ArrMap_vrange [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_comp_ArrMap_app[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Arr\" shows "(\ \<^sub>C\<^sub>F\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_comp_ArrMap_app [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Object map: contravariant and covariant functor\ lemma cf_cn_cov_comp_ObjMap_vsv[cat_cn_cs_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>C\<^sub>F\ \)\ObjMap\)" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ObjMap_vsv [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ObjMap_vdomain[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ObjMap\) = \\Obj\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ObjMap_vdomain [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ObjMap_vrange: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ObjMap_vrange [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ObjMap_app[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>C\<^sub>F\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ObjMap_app [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Arrow map: contravariant and covariant functors\ lemma cf_cn_cov_comp_ArrMap_vsv[cat_cn_cs_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>C\<^sub>F\ \)\ArrMap\)" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ArrMap_vsv [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ArrMap_vdomain[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ArrMap\) = \\Arr\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ArrMap_vdomain [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ArrMap_vrange: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>C\<^sub>F\ \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ArrMap_vrange [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps ] ) qed lemma cf_cn_cov_comp_ArrMap_app[cat_cn_cs_simps]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Arr\" shows "(\ \<^sub>C\<^sub>F\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule smcf_cn_cov_comp_ArrMap_app [ OF L.cf_is_semifunctor[unfolded slicing_commute[symmetric]] R.cf_is_semifunctor, unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Further properties\ lemma cf_cn_comp_is_functorI[cat_cn_cs_intros]: assumes "category \ \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\ \<^sub>C\<^sub>F\ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(2)) interpret R: is_functor \ \op_cat \\ \ \ by (rule assms(3)) interpret \: category \ \ by (rule assms(1)) show ?thesis proof(rule is_functorI, unfold cf_cn_comp_components(3,4) cat_op_simps) show "vfsequence (\ \<^sub>C\<^sub>F\ \)" unfolding cf_cn_comp_def by (simp add: nat_omega_simps) show "vcard (\ \<^sub>C\<^sub>F\ \) = 4\<^sub>\" unfolding cf_cn_comp_def by (simp add: nat_omega_simps) from assms(1) L.cf_is_semifunctor R.cf_is_semifunctor show "cf_smcf (\ \<^sub>C\<^sub>F\ \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" unfolding cf_smcf_cf_cn_comp[symmetric] by ( cs_concl cs_intro: cat_cs_intros slicing_intros smc_cn_cs_intros ) fix c assume "c \\<^sub>\ \\Obj\" with assms show "(\ \<^sub>C\<^sub>F\ \)\ArrMap\\\\CId\\c\\ = \\CId\\(\ \<^sub>C\<^sub>F\ \)\ObjMap\\c\\" by ( cs_concl cs_simp: cat_op_simps cat_cn_cs_simps cs_intro: cat_cs_intros ) qed (auto simp: cat_cs_simps cat_cs_intros cat_op_simps) qed text\See section 1.2 in \cite{bodo_categories_1970}).\ lemma cf_cn_cov_comp_is_functor[cat_cn_cs_intros]: assumes "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \<^sub>C\<^sub>F\ \ : \ \<^sub>C\\\<^bsub>\\<^esub> \" proof- interpret L: is_functor \ \op_cat \\ \ \ by (rule assms(1)) interpret R: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof ( rule is_functorI, unfold cf_cn_comp_components(3,4) cat_op_simps slicing_commute[symmetric] ) show "vfsequence (\ \<^sub>C\<^sub>F\ \)" unfolding cf_cn_comp_def by simp show "vcard (\ \<^sub>C\<^sub>F\ \) = 4\<^sub>\" unfolding cf_cn_comp_def by (auto simp: nat_omega_simps) from L.cf_is_semifunctor show "cf_smcf \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ cf_smcf \ : op_smc (cat_smc \) \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by (cs_concl cs_intro: cat_cs_intros slicing_intros smc_cs_intros) fix c assume "c \\<^sub>\ \\Obj\" with assms show "(\ \<^sub>C\<^sub>F\ \)\ArrMap\\\\CId\\c\\ = \\CId\\(\ \<^sub>C\<^sub>F\ \)\ObjMap\\c\\" by ( cs_concl cs_simp: cat_cs_simps cat_cn_cs_simps cs_intro: cat_cs_intros ) qed (auto simp: cat_cs_simps cat_cs_intros) qed text\See section 1.2 in \cite{bodo_categories_1970}.\ lemma cf_cov_cn_comp_is_functor[cat_cn_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \<^sub>C\\\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \<^sub>C\\\<^bsub>\\<^esub> \" using assms by (rule cf_comp_is_functorI) text\The opposite of the contravariant composition of functors.\ lemma op_cf_cf_cn_comp[cat_op_simps]: "op_cf (\ \<^sub>C\<^sub>F\ \) = op_cf \ \<^sub>C\<^sub>F\ op_cf \" unfolding op_cf_def cf_cn_comp_def dghm_field_simps by (auto simp: nat_omega_simps) subsection\Identity functor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ abbreviation (input) cf_id :: "V \ V" where "cf_id \ dghm_id" text\Slicing.\ lemma cf_smcf_cf_id[slicing_commute]: "smcf_id (cat_smc \) = cf_smcf (cf_id \)" unfolding dghm_id_def cat_smc_def cf_smcf_def dghm_field_simps dg_field_simps by (simp add: nat_omega_simps) context category begin interpretation smc: semicategory \ \cat_smc \\ by (rule cat_semicategory) lemmas_with [unfolded slicing_simps]: cat_smcf_id_is_semifunctor = smc.smc_smcf_id_is_semifunctor end subsubsection\Object map\ lemmas [cat_cs_simps] = dghm_id_ObjMap_app subsubsection\Arrow map\ lemmas [cat_cs_simps] = dghm_id_ArrMap_app subsubsection\Opposite of an identity functor.\ lemma op_cf_cf_id[cat_op_simps]: "op_cf (cf_id \) = cf_id (op_cat \)" unfolding dghm_id_def op_cat_def op_cf_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\An identity functor is a functor\ lemma (in category) cat_cf_id_is_functor: "cf_id \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_functorI, unfold dghm_id_components) from cat_smcf_id_is_semifunctor show "cf_smcf (cf_id \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by (simp add: slicing_commute) from cat_CId_is_arr show "c \\<^sub>\ \\Obj\ \ vid_on (\\Arr\)\\\CId\\c\\ = \\CId\\vid_on (\\Obj\)\c\\" for c by auto qed (auto simp: dghm_id_def nat_omega_simps cat_cs_intros) lemma (in category) cat_cf_id_is_functor': assumes "\ = \" and "\ = \" shows "cf_id \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding assms by (rule cat_cf_id_is_functor) lemmas [cat_cs_intros] = category.cat_cf_id_is_functor' subsubsection\Further properties\ lemma (in is_functor) cf_cf_comp_cf_id_left[cat_cs_simps]: "cf_id \ \\<^sub>C\<^sub>F \ = \" \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ by ( rule cf_eqI, unfold dghm_id_components dghm_comp_components dghm_id_components ) (auto intro: cat_cs_intros simp: cf_ArrMap_vrange cf_ObjMap_vrange) lemmas [cat_cs_simps] = is_functor.cf_cf_comp_cf_id_left lemma (in is_functor) cf_cf_comp_cf_id_right[cat_cs_simps]: "\ \\<^sub>C\<^sub>F cf_id \ = \" \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ by ( rule cf_eqI, unfold dghm_id_components dghm_comp_components dghm_id_components ) ( auto intro: cat_cs_intros simp: cat_cs_simps cf_ArrMap_vrange cf_ObjMap_vrange ) lemmas [cat_cs_simps] = is_functor.cf_cf_comp_cf_id_right subsection\Constant functor\ subsubsection\Definition and elementary properties\ text\See Chapter III-3 in \cite{mac_lane_categories_2010}.\ abbreviation cf_const :: "V \ V \ V \ V" where "cf_const \ \
a \ smcf_const \ \
a (\
\CId\\a\)" text\Slicing.\ lemma cf_smcf_cf_const[slicing_commute]: "smcf_const (cat_smc \) (cat_smc \
) a (\
\CId\\a\) = cf_smcf (cf_const \ \
a)" unfolding dghm_const_def cat_smc_def cf_smcf_def dghm_field_simps dg_field_simps by (simp add: nat_omega_simps) subsubsection\Object map and arrow map\ context fixes \
a :: V begin lemmas_with [where \
=\
and a=a and f=\\
\CId\\a\\, cat_cs_simps]: dghm_const_ObjMap_app dghm_const_ArrMap_app end subsubsection\Opposite constant functor\ lemma op_cf_cf_const[cat_op_simps]: "op_cf (cf_const \ \
a) = cf_const (op_cat \) (op_cat \
) a" unfolding dghm_const_def op_cat_def op_cf_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\A constant functor is a functor\ lemma cf_const_is_functor: assumes "category \ \" and "category \ \
" and "a \\<^sub>\ \
\Obj\" shows "cf_const \ \
a : \ \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: category \ \ by (rule assms(1)) interpret \
: category \ \
by (rule assms(2)) show ?thesis proof(intro is_functorI, tactic\distinct_subgoals_tac\) show "vfsequence (dghm_const \ \
a (\
\CId\\a\))" unfolding dghm_const_def by simp show "vcard (cf_const \ \
a) = 4\<^sub>\" unfolding dghm_const_def by (simp add: nat_omega_simps) from assms show "cf_smcf (cf_const \ \
a) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \
" by ( cs_concl cs_simp: cat_cs_simps slicing_simps slicing_commute[symmetric] cs_intro: smc_cs_intros cat_cs_intros slicing_intros ) fix c assume "c \\<^sub>\ \\Obj\" with assms show "cf_const \ \
a\ArrMap\\\\CId\\c\\ = \
\CId\\cf_const \ \
a\ObjMap\\c\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: dghm_const_components assms) qed lemma cf_const_is_functor'[cat_cs_intros]: assumes "category \ \" and "category \ \
" and "a \\<^sub>\ \
\Obj\" and "\ = \" and "\ = \
" - shows "cf_const \ \
a : \ \\\<^sub>C\<^bsub>\\<^esub> \" - using assms(1-3) unfolding assms(4,5) by (rule cf_const_is_functor) + and "f = (\
\CId\\a\)" + shows "dghm_const \ \
a f : \ \\\<^sub>C\<^bsub>\\<^esub> \" + using assms(1-3) unfolding assms(4-6) by (rule cf_const_is_functor) subsection\Faithful functor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_ft_functor = is_functor \ \ \ \ for \ \ \ \ + assumes ft_cf_is_ft_semifunctor[slicing_intros]: "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> cat_smc \" syntax "_is_ft_functor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_ft_functor \ \ \ \" lemma (in is_ft_functor) ft_cf_is_ft_functor': assumes "\' = cat_smc \" and "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule ft_cf_is_ft_semifunctor) lemmas [slicing_intros] = is_ft_functor.ft_cf_is_ft_functor' text\Rules.\ lemma (in is_ft_functor) is_ft_functor_axioms'[cf_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_ft_functor_axioms) mk_ide rf is_ft_functor_def[unfolded is_ft_functor_axioms_def] |intro is_ft_functorI| |dest is_ft_functorD[dest]| |elim is_ft_functorE[elim]| lemmas [cf_cs_intros] = is_ft_functorD(1) lemma is_ft_functorI': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" shows "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" using assms by (intro is_ft_functorI) ( simp_all add: assms(1) is_ft_semifunctorI'[OF is_functorD(6)[ OF assms(1)], unfolded slicing_simps ] ) lemma is_ft_functorD': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" shows "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" by ( simp_all add: is_ft_functorD[OF assms(1)] is_ft_semifunctorD'(2)[ OF is_ft_functorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_ft_functorE': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" obtains "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" using assms by (simp_all add: is_ft_functorD') text\Elementary properties.\ context is_ft_functor begin interpretation smcf: is_ft_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule ft_cf_is_ft_semifunctor) lemmas_with [unfolded slicing_simps]: ft_cf_v11_on_Hom = smcf.ft_smcf_v11_on_Hom end subsubsection\Opposite faithful functor.\ lemma (in is_ft_functor) is_ft_functor_op': "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> op_cat \" by (rule is_ft_functorI, unfold slicing_commute[symmetric]) ( simp_all add: is_functor_op is_ft_semifunctor.is_ft_semifunctor_op ft_cf_is_ft_semifunctor ) lemma (in is_ft_functor) is_ft_functor_op: assumes "\' = op_cat \" and "\' = op_cat \" shows "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> op_cat \" unfolding assms by (rule is_ft_functor_op') lemmas is_ft_functor_op[cat_op_intros] = is_ft_functor.is_ft_functor_op' subsubsection\The composition of faithful functors is a faithful functor\ lemma cf_comp_is_ft_functor[cf_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" proof(intro is_ft_functorI) interpret \: is_ft_functor \ \ \ \ by (simp add: assms(1)) interpret \: is_ft_functor \ \ \ \ by (simp add: assms(2)) from \.is_functor_axioms \.is_functor_axioms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) then interpret is_functor \ \ \ \\ \\<^sub>C\<^sub>F \\ . show "cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> cat_smc \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: cf_cs_intros smcf_cs_intros slicing_intros ) qed subsection\Full functor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_fl_functor = is_functor \ \ \ \ for \ \ \ \ + assumes fl_cf_is_fl_semifunctor: "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> cat_smc \" syntax "_is_fl_functor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_fl_functor \ \ \ \" lemma (in is_fl_functor) fl_cf_is_fl_functor'[slicing_intros]: assumes "\' = cat_smc \" and "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule fl_cf_is_fl_semifunctor) lemmas [slicing_intros] = is_fl_functor.fl_cf_is_fl_semifunctor text\Rules.\ lemma (in is_fl_functor) is_fl_functor_axioms'[cf_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_fl_functor_axioms) mk_ide rf is_fl_functor_def[unfolded is_fl_functor_axioms_def] |intro is_fl_functorI| |dest is_fl_functorD[dest]| |elim is_fl_functorE[elim]| lemmas [cf_cs_intros] = is_fl_functorD(1) lemma is_fl_functorI': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ \\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" shows "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" using assms by (intro is_fl_functorI) ( simp_all add: assms(1) is_fl_semifunctorI'[ OF is_functorD(6)[OF assms(1)], unfolded slicing_simps ] ) lemma is_fl_functorD': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" shows "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ \\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" by ( simp_all add: is_fl_functorD[OF assms(1)] is_fl_semifunctorD'(2)[ OF is_fl_functorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_fl_functorE': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" obtains "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ \\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" using assms by (simp_all add: is_fl_functorD') text\Elementary properties.\ context is_fl_functor begin interpretation smcf: is_fl_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule fl_cf_is_fl_semifunctor) lemmas_with [unfolded slicing_simps]: fl_cf_surj_on_Hom = smcf.fl_smcf_surj_on_Hom end subsubsection\Opposite full functor\ lemma (in is_fl_functor) is_fl_functor_op[cat_op_intros]: "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> op_cat \" by (rule is_fl_functorI, unfold slicing_commute[symmetric]) (simp_all add: cat_op_intros smc_op_intros slicing_intros) lemmas is_fl_functor_op[cat_op_intros] = is_fl_functor.is_fl_functor_op subsubsection\The composition of full functor is a full functor\ lemma cf_comp_is_fl_functor[cf_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" proof(intro is_fl_functorI) interpret \: is_fl_functor \ \ \ \ using assms(2) by simp interpret \: is_fl_functor \ \ \ \ using assms(1) by simp from \.is_functor_axioms \.is_functor_axioms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> cat_smc \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: cf_cs_intros smcf_cs_intros slicing_intros ) qed subsection\Fully faithful functor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_ff_functor = is_ft_functor \ \ \ \ + is_fl_functor \ \ \ \ for \ \ \ \ syntax "_is_ff_functor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" \ "CONST is_ff_functor \ \ \ \" text\Rules.\ mk_ide rf is_ff_functor_def |intro is_ff_functorI| |dest is_ff_functorD[dest]| |elim is_ff_functorE[elim]| lemmas [cf_cs_intros] = is_ff_functorD text\Elementary properties.\ lemma (in is_ff_functor) ff_cf_is_ff_semifunctor: "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> cat_smc \" by (rule is_ff_semifunctorI) (auto intro: slicing_intros) lemma (in is_ff_functor) ff_cf_is_ff_semifunctor'[slicing_intros]: assumes "\' = cat_smc \" and "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \'" unfolding assms by (rule ff_cf_is_ff_semifunctor) lemmas [slicing_intros] = is_ff_functor.ff_cf_is_ff_semifunctor' subsubsection\Opposite fully faithful functor\ lemma (in is_ff_functor) is_ff_functor_op: "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> op_cat \" by (rule is_ff_functorI) (auto simp: is_fl_functor_op is_ft_functor_op) lemma (in is_ff_functor) is_ff_functor_op'[cat_op_intros]: assumes "\' = op_cat \" and "\' = op_cat \" shows "op_cf \ : \' \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \'" unfolding assms by (rule is_ff_functor_op) lemmas is_ff_functor_op[cat_op_intros] = is_ff_functor.is_ff_functor_op subsubsection\ The composition of fully faithful functors is a fully faithful functor \ lemma cf_comp_is_ff_functor[cf_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" using assms by (intro is_ff_functorI, elim is_ff_functorE) (auto simp: cf_cs_intros) subsection\Isomorphism of categories\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_iso_functor = is_functor \ \ \ \ for \ \ \ \ + assumes iso_cf_is_iso_semifunctor: "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_smc \" syntax "_is_iso_functor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" \ "CONST is_iso_functor \ \ \ \" lemma (in is_iso_functor) iso_cf_is_iso_semifunctor'[slicing_intros]: assumes "\' = cat_smc \" "\' = cat_smc \" shows "cf_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" unfolding assms by (rule iso_cf_is_iso_semifunctor) lemmas [slicing_intros] = is_iso_semifunctor.iso_smcf_is_iso_dghm' text\Rules.\ lemma (in is_iso_functor) is_iso_functor_axioms'[cf_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_iso_functor_axioms) mk_ide rf is_iso_functor_def[unfolded is_iso_functor_axioms_def] |intro is_iso_functorI| |dest is_iso_functorD[dest]| |elim is_iso_functorE[elim]| lemma is_iso_functorI': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "v11 (\\ObjMap\)" and "v11 (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" shows "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms by (intro is_iso_functorI) ( simp_all add: assms(1) is_iso_semifunctorI'[ OF is_functorD(6)[OF assms(1)], unfolded slicing_simps ] ) lemma is_iso_functorD': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "v11 (\\ObjMap\)" and "v11 (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" by ( simp_all add: is_iso_functorD[OF assms(1)] is_iso_semifunctorD'(2-5)[ OF is_iso_functorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_iso_functorE': assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" obtains "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "v11 (\\ObjMap\)" and "v11 (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" using assms by (simp_all add: is_iso_functorD') text\Elementary properties.\ context is_iso_functor begin interpretation smcf: is_iso_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule iso_cf_is_iso_semifunctor) lemmas_with [unfolded slicing_simps]: iso_cf_ObjMap_vrange[simp] = smcf.iso_smcf_ObjMap_vrange and iso_cf_ArrMap_vrange[simp] = smcf.iso_smcf_ArrMap_vrange sublocale ObjMap: v11 \\\ObjMap\\ rewrites "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" by (rule smcf.ObjMap.v11_axioms[unfolded slicing_simps]) (simp_all add: cat_cs_simps cf_cs_simps) sublocale ArrMap: v11 \\\ArrMap\\ rewrites "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" by (rule smcf.ArrMap.v11_axioms[unfolded slicing_simps]) (simp_all add: cat_cs_simps smcf_cs_simps) lemmas_with [unfolded slicing_simps]: iso_cf_Obj_HomDom_if_Obj_HomCod[elim] = smcf.iso_smcf_Obj_HomDom_if_Obj_HomCod and iso_cf_Arr_HomDom_if_Arr_HomCod[elim] = smcf.iso_smcf_Arr_HomDom_if_Arr_HomCod and iso_cf_ObjMap_eqE[elim] = smcf.iso_smcf_ObjMap_eqE and iso_cf_ArrMap_eqE[elim] = smcf.iso_smcf_ArrMap_eqE end sublocale is_iso_functor \ is_ff_functor proof(intro is_ff_functorI) interpret is_iso_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule iso_cf_is_iso_semifunctor) show "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" by unfold_locales show "\ : \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" by unfold_locales qed lemmas (in is_iso_functor) iso_cf_is_ff_functor = is_ff_functor_axioms lemmas [cf_cs_intros] = is_iso_functor.iso_cf_is_ff_functor subsubsection\Opposite isomorphism of categories\ lemma (in is_iso_functor) is_iso_functor_op: "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> op_cat \" by (rule is_iso_functorI, unfold slicing_simps slicing_commute[symmetric]) (simp_all add: cat_op_intros smc_op_intros slicing_intros) lemma (in is_iso_functor) is_iso_functor_op': assumes "\' = op_cat \" and "\' = op_cat \" shows "op_cf \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> op_cat \" unfolding assms by (rule is_iso_functor_op) lemmas is_iso_functor_op[cat_op_intros] = is_iso_functor.is_iso_functor_op' subsubsection\ The composition of isomorphisms of categories is an isomorphism of categories \ lemma cf_comp_is_iso_functor[cf_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_functorI) interpret \: is_iso_functor \ \ \ \ using assms by auto interpret \: is_iso_functor \ \ \ \ using assms by auto from \.is_functor_axioms \.is_functor_axioms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_smc \" unfolding slicing_commute[symmetric] by (cs_concl cs_intro: smcf_cs_intros slicing_intros) qed subsection\Inverse functor\ abbreviation (input) inv_cf :: "V \ V" where "inv_cf \ inv_dghm" text\Slicing.\ lemma dghm_inv_semifunctor[slicing_commute]: "inv_smcf (cf_smcf \) = cf_smcf (inv_cf \)" unfolding cf_smcf_def inv_dghm_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_iso_functor begin interpretation smcf: is_iso_semifunctor \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ by (rule iso_cf_is_iso_semifunctor) lemmas_with [unfolded slicing_simps slicing_commute]: inv_cf_ObjMap_v11 = smcf.inv_smcf_ObjMap_v11 and inv_cf_ObjMap_vdomain = smcf.inv_smcf_ObjMap_vdomain and inv_cf_ObjMap_app = smcf.inv_smcf_ObjMap_app and inv_cf_ObjMap_vrange = smcf.inv_smcf_ObjMap_vrange and inv_cf_ArrMap_v11 = smcf.inv_smcf_ArrMap_v11 and inv_cf_ArrMap_vdomain = smcf.inv_smcf_ArrMap_vdomain and inv_cf_ArrMap_app = smcf.inv_smcf_ArrMap_app and inv_cf_ArrMap_vrange = smcf.inv_smcf_ArrMap_vrange and iso_cf_ObjMap_inv_cf_ObjMap_app = smcf.iso_smcf_ObjMap_inv_smcf_ObjMap_app and iso_cf_ArrMap_inv_cf_ArrMap_app = smcf.iso_smcf_ArrMap_inv_smcf_ArrMap_app and iso_cf_HomDom_is_arr_conv = smcf.iso_smcf_HomDom_is_arr_conv and iso_cf_HomCod_is_arr_conv = smcf.iso_smcf_HomCod_is_arr_conv end subsection\An isomorphism of categories is an isomorphism in the category \CAT\\ lemma is_arr_isomorphism_is_iso_functor: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>C\<^sub>F \ = cf_id \" and "\ \\<^sub>C\<^sub>F \ = cf_id \" shows "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_iso_functorI) have \\\: "cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ = smcf_id (cat_smc \)" by (simp add: assms(3) cf_smcf_cf_id cf_smcf_smcf_comp) have \\\: "cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ = smcf_id (cat_smc \)" by (simp add: assms(4) cf_smcf_cf_id cf_smcf_smcf_comp) from \.cf_is_semifunctor \.cf_is_semifunctor \\\ \\\ show "cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_smc \" by (rule is_arr_isomorphism_is_iso_semifunctor) qed (auto simp: cat_cs_intros) qed lemma is_iso_functor_is_arr_isomorphism: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows [cf_cs_intros]: "inv_cf \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "inv_cf \ \\<^sub>C\<^sub>F \ = cf_id \" and "\ \\<^sub>C\<^sub>F inv_cf \ = cf_id \" proof- let ?\ = "inv_cf \" interpret is_iso_functor \ \ \ \ by (rule assms(1)) show \: "?\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_functorI is_functorI, unfold inv_dghm_components) show "vfsequence ?\" by (simp add: inv_dghm_def) show "vcard ?\ = 4\<^sub>\" unfolding inv_dghm_def by (simp add: nat_omega_simps) show "cf_smcf ?\ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by ( metis dghm_inv_semifunctor iso_cf_is_iso_semifunctor is_iso_semifunctor_def is_iso_semifunctor_is_arr_isomorphism(1) ) show "cf_smcf ?\ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_smc \" by ( metis dghm_inv_semifunctor iso_cf_is_iso_semifunctor is_iso_semifunctor_is_arr_isomorphism(1) ) fix c assume prems: "c \\<^sub>\ \\Obj\" from prems show "(\\ArrMap\)\\<^sub>\\\\CId\\c\\ = \\CId\\(\\ObjMap\)\\<^sub>\\c\\" by (intro v11.v11_vconverse_app) ( cs_concl cs_intro: cat_cs_intros V_cs_intros cs_simp: V_cs_simps cat_cs_simps )+ qed (simp_all add: cat_cs_simps cat_cs_intros) show "?\ \\<^sub>C\<^sub>F \ = cf_id \" proof(rule cf_eqI, unfold dghm_comp_components inv_dghm_components) from \ is_functor_axioms show "?\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (blast intro: cat_cs_intros) qed ( simp_all add: HomDom.cat_cf_id_is_functor ObjMap.v11_vcomp_vconverse ArrMap.v11_vcomp_vconverse dghm_id_components ) show "\ \\<^sub>C\<^sub>F ?\ = cf_id \" proof(rule cf_eqI, unfold dghm_comp_components inv_dghm_components) from \ is_functor_axioms show "\ \\<^sub>C\<^sub>F ?\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (blast intro: cat_cs_intros) show "cf_id \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (simp add: HomCod.cat_cf_id_is_functor) qed ( simp_all add: ObjMap.v11_vcomp_vconverse' ArrMap.v11_vcomp_vconverse' dghm_id_components ) qed subsubsection\An identity functor is an isomorphism of categories\ lemma (in category) cat_cf_id_is_iso_functor: "cf_id \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (rule is_iso_functorI, unfold slicing_commute[symmetric]) ( simp_all add: cat_cf_id_is_functor semicategory.smc_smcf_id_is_iso_semifunctor cat_semicategory ) subsection\Isomorphic categories\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale iso_category = L: category \ \ + R: category \ \ for \ \ \ + assumes iso_cat_is_iso_functor: "\\. \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" notation iso_category (infixl "\\<^sub>C\" 50) text\Rules.\ lemma iso_categoryI: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^bsub>\\<^esub> \" using assms unfolding iso_category_def iso_category_axioms_def by auto lemma iso_categoryD[dest]: assumes "\ \\<^sub>C\<^bsub>\\<^esub> \" shows "\\. \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms unfolding iso_category_def iso_category_axioms_def by simp_all lemma iso_categoryE[elim]: assumes "\ \\<^sub>C\<^bsub>\\<^esub> \" obtains \ where "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms by auto text\Isomorphic categories are isomorphic semicategories.\ lemma (in iso_category) iso_cat_iso_semicategory: "cat_smc \ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" using iso_cat_is_iso_functor by (auto intro: slicing_intros iso_semicategoryI) subsubsection\A category isomorphism is an equivalence relation\ lemma iso_category_refl: assumes "category \ \" shows "\ \\<^sub>C\<^bsub>\\<^esub> \" proof(rule iso_categoryI[of _ _ _ \cf_id \\]) interpret category \ \ by (rule assms) show "cf_id \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (simp add: cat_cf_id_is_iso_functor) qed lemma iso_category_sym[sym]: assumes "\ \\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^bsub>\\<^esub> \" proof- interpret iso_category \ \ \ by (rule assms) from iso_cat_is_iso_functor obtain \ where "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by clarsimp from is_iso_functor_is_arr_isomorphism(1)[OF this] show ?thesis by (auto intro: iso_categoryI) qed lemma iso_category_trans[trans]: assumes "\ \\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: iso_category \ \ \ by (rule assms(1)) interpret R: iso_category \ \ \ by (rule assms(2)) from L.iso_cat_is_iso_functor R.iso_cat_is_iso_functor show ?thesis by (auto intro: iso_categoryI is_iso_functorI cf_comp_is_iso_functor) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy @@ -1,2095 +1,2095 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\Set\\ theory CZH_ECAT_Set imports CZH_Foundations.CZH_SMC_Set CZH_ECAT_Par CZH_ECAT_Subcategory CZH_ECAT_PCategory begin subsection\Background\ text\ The methodology chosen for the exposition of \Set\ as a category is analogous to the one used in \cite{milehins_category_2021} for the exposition of \Set\ as a semicategory. \ named_theorems cat_Set_cs_simps named_theorems cat_Set_cs_intros lemmas (in arr_Set) [cat_Set_cs_simps] = dg_Rel_shared_cs_simps lemmas [cat_Set_cs_simps] = dg_Rel_shared_cs_simps arr_Set.arr_Set_ArrVal_vdomain arr_Set_comp_Set_id_Set_left arr_Set_comp_Set_id_Set_right lemmas [cat_Set_cs_intros] = dg_Rel_shared_cs_intros arr_Set_comp_Set (* Certain lemmas are applicable to any of the categories among Rel, Par, Set. If these lemmas are included in general-purpose collections like cat_cs_simps/cat_cs_intros, then backtracking can become slow. The following collections were created to resolve such issues. *) named_theorems cat_rel_par_Set_cs_intros named_theorems cat_rel_par_Set_cs_simps named_theorems cat_rel_Par_set_cs_intros named_theorems cat_rel_Par_set_cs_simps named_theorems cat_Rel_par_set_cs_intros named_theorems cat_Rel_par_set_cs_simps subsection\\Set\ as a category\ subsubsection\Definition and elementary properties\ definition cat_Set :: "V \ V" where "cat_Set \ = [ Vset \, set {T. arr_Set \ T}, (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrDom\), (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrCod\), (\ST\\<^sub>\composable_arrs (dg_Set \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\), VLambda (Vset \) id_Set ]\<^sub>\" text\Components.\ lemma cat_Set_components: shows "cat_Set \\Obj\ = Vset \" and "cat_Set \\Arr\ = set {T. arr_Set \ T}" and "cat_Set \\Dom\ = (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrDom\)" and "cat_Set \\Cod\ = (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrCod\)" and "cat_Set \\Comp\ = (\ST\\<^sub>\composable_arrs (dg_Set \). ST\0\ \\<^sub>P\<^sub>a\<^sub>r ST\1\<^sub>\\)" and "cat_Set \\CId\ = VLambda (Vset \) id_Set" unfolding cat_Set_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_cat_Set: "cat_smc (cat_Set \) = smc_Set \" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (cat_smc (cat_Set \)) = 5\<^sub>\" unfolding cat_smc_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smc_Set \) = 5\<^sub>\" unfolding smc_Set_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cat_smc (cat_Set \)) = \\<^sub>\ (smc_Set \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (cat_smc (cat_Set \)) \ cat_smc (cat_Set \)\a\ = smc_Set \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cat_smc_def dg_field_simps cat_Set_def smc_Set_def ) (auto simp: nat_omega_simps) qed (auto simp: cat_smc_def smc_Set_def) lemmas_with [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_Obj_iff = smc_Set_Obj_iff and cat_Set_Arr_iff[cat_Set_cs_simps] = smc_Set_Arr_iff and cat_Set_Dom_vsv[intro] = smc_Set_Dom_vsv and cat_Set_Dom_vdomain[simp] = smc_Set_Dom_vdomain and cat_Set_Dom_vrange = smc_Set_Dom_vrange and cat_Set_Dom_app = smc_Set_Dom_app and cat_Set_Cod_vsv[intro] = smc_Set_Cod_vsv and cat_Set_Cod_vdomain[simp] = smc_Set_Cod_vdomain and cat_Set_Cod_vrange = smc_Set_Cod_vrange and cat_Set_Cod_app[cat_Set_cs_simps] = smc_Set_Cod_app and cat_Set_is_arrI = smc_Set_is_arrI and cat_Set_is_arrD = smc_Set_is_arrD and cat_Set_is_arrE = smc_Set_is_arrE and cat_Set_ArrVal_vdomain[cat_cs_simps] = smc_Set_ArrVal_vdomain and cat_Set_ArrVal_app_vrange[cat_Set_cs_intros] = smc_Set_ArrVal_app_vrange lemmas [cat_cs_simps] = cat_Set_is_arrD(2,3) lemmas [cat_Set_cs_intros] = cat_Set_is_arrI lemmas_with [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_composable_arrs_dg_Set = smc_Set_composable_arrs_dg_Set and cat_Set_Comp = smc_Set_Comp and cat_Set_Comp_app[cat_Set_cs_simps] = smc_Set_Comp_app and cat_Set_Comp_vdomain[cat_Set_cs_simps] = smc_Set_Comp_vdomain lemmas_with (in \) [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_Hom_vifunion_in_Vset = smc_Set_Hom_vifunion_in_Vset and cat_Set_incl_Set_is_arr = smc_Set_incl_Set_is_arr and cat_Set_incl_Set_is_arr'[cat_Set_cs_intros] = smc_Set_incl_Set_is_arr' and cat_Set_Comp_ArrVal = smc_Set_Comp_ArrVal and cat_Set_Comp_vrange = smc_Set_Comp_vrange and cat_Set_is_monic_arrI = smc_Set_is_monic_arrI and cat_Set_is_monic_arr = smc_Set_is_monic_arr and cat_Set_is_epic_arrI = smc_Set_is_epic_arrI and cat_Set_is_epic_arrD = smc_Set_is_epic_arrD and cat_Set_is_epic_arr = smc_Set_is_epic_arr and cat_Set_obj_terminal = smc_Set_obj_terminal and cat_Set_obj_initial = smc_Set_obj_initial and cat_Set_obj_null = smc_Set_obj_null and cat_Set_is_zero_arr = smc_Set_is_zero_arr lemmas [cat_Set_cs_intros] = \.cat_Set_incl_Set_is_arr' lemmas [cat_cs_simps] = \.cat_Set_Comp_ArrVal subsubsection\Identity\ lemma cat_Set_CId_app[cat_Set_cs_simps]: assumes "A \\<^sub>\ Vset \" shows "cat_Set \\CId\\A\ = id_Set A" using assms unfolding cat_Set_components by simp lemma id_Par_CId_app_app[cat_cs_simps]: assumes "A \\<^sub>\ Vset \" and "a \\<^sub>\ A" shows "cat_Set \\CId\\A\\ArrVal\\a\ = a" unfolding cat_Set_CId_app[OF assms(1)] id_Rel_ArrVal_app[OF assms(2)] by simp subsubsection\\Set\ is a category\ lemma (in \) category_cat_Set: "category \ (cat_Set \)" proof(rule categoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) interpret Set: semicategory \ \cat_smc (cat_Set \)\ unfolding cat_smc_cat_Set by (simp add: semicategory_smc_Set) show "vfsequence (cat_Set \)" unfolding cat_Set_def by simp show "vcard (cat_Set \) = 6\<^sub>\" unfolding cat_Set_def by (simp add: nat_omega_simps) show "semicategory \ (smc_Set \)" by (simp add: semicategory_smc_Set) show "cat_Set \\CId\\A\ : A \\<^bsub>cat_Set \\<^esub> A" if "A \\<^sub>\ cat_Set \\Obj\" for A using that unfolding cat_Set_Obj_iff by ( cs_concl cs_simp: cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) show "cat_Set \\CId\\B\ \\<^sub>A\<^bsub>cat_Set \\<^esub> F = F" if "F : A \\<^bsub>cat_Set \\<^esub> B" for F A B proof- from that have "arr_Set \ F" "B \\<^sub>\ Vset \" by (auto elim: cat_Set_is_arrE) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) qed show "F \\<^sub>A\<^bsub>cat_Set \\<^esub> cat_Set \\CId\\B\ = F" if "F : B \\<^bsub>cat_Set \\<^esub> C" for F B C proof- from that have "arr_Set \ F" "B \\<^sub>\ Vset \" by (auto elim: cat_Set_is_arrE) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) qed qed (auto simp: cat_Set_components) lemma (in \) category_cat_Set': assumes "\ = \" shows "category \ (cat_Set \)" unfolding assms by (rule category_cat_Set) lemmas [cat_cs_intros] = \.category_cat_Set' subsubsection\\Set\ is a wide replete subcategory of \Par\\ lemma (in \) wide_replete_subcategory_cat_Set_cat_Par: "cat_Set \ \\<^sub>C\<^sub>.\<^sub>w\<^sub>r\<^bsub>\\<^esub> cat_Par \" proof(intro wide_replete_subcategoryI) show wide_subcategory_cat_Set_cat_Par: "cat_Set \ \\<^sub>C\<^sub>.\<^sub>w\<^sub>i\<^sub>d\<^sub>e\<^bsub>\\<^esub> cat_Par \" proof(intro wide_subcategoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) interpret Par: category \ \cat_Par \\ by (rule category_cat_Par) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) interpret wide_subsemicategory \ \smc_Set \\ \smc_Par \\ by (simp add: wide_subsemicategory_smc_Set_smc_Par) show "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Par \" proof(intro subcategoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) show "smc_Set \ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Par \" by (simp add: subsemicategory_axioms) fix A assume "A \\<^sub>\ cat_Set \\Obj\" then show "cat_Set \\CId\\A\ = cat_Par \\CId\\A\" unfolding cat_Set_components cat_Par_components by simp qed ( auto simp: subsemicategory_axioms Par.category_axioms Set.category_axioms ) qed (rule wide_subsemicategory_smc_Set_smc_Par) show "cat_Set \ \\<^sub>C\<^sub>.\<^sub>r\<^sub>e\<^sub>p\<^bsub>\\<^esub> cat_Par \" proof(intro replete_subcategoryI) interpret wide_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_subcategory_cat_Set_cat_Par) show "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Par \" by (rule subcategory_axioms) fix A B F assume "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" note arr_Par = cat_Par_is_arr_isomorphismD[OF this] from arr_Par show "F : A \\<^bsub>cat_Set \\<^esub> B" by (intro cat_Set_is_arrI arr_Set_arr_ParI cat_Par_is_arrD[OF arr_Par(1)]) (auto simp: cat_Par_is_arrD(2)) qed qed subsubsection\\Set\ is a subcategory of \Set\\ lemma (in \) subcategory_cat_Set_cat_Set:(*TODO: generalize*) assumes "\ \" and "\ \\<^sub>\ \" shows "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: \ \ by (rule assms(1)) show ?thesis proof(intro subcategoryI') show "category \ (cat_Set \)" by (rule category.cat_category_if_ge_Limit, insert assms(2)) (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros)+ show "A \\<^sub>\ cat_Set \\Obj\" if "A \\<^sub>\ cat_Set \\Obj\" for A using that unfolding cat_Set_components(1) by (meson assms(2) Vset_in_mono \.Axiom_of_Extensionality(3)) show is_arr_if_is_arr: "F : A \\<^bsub>cat_Set \\<^esub> B" if "F : A \\<^bsub>cat_Set \\<^esub> B" for A B F proof- note f = cat_Set_is_arrD[OF that] interpret f: arr_Set \ F by (rule f(1)) show ?thesis proof(intro cat_Set_is_arrI arr_SetI) show "\\<^sub>\ (F\ArrVal\) \\<^sub>\ F\ArrCod\" by (auto simp: f.arr_Set_ArrVal_vrange) show "F\ArrDom\ \\<^sub>\ Vset \" by (auto intro!: f.arr_Set_ArrDom_in_Vset Vset_in_mono assms(2)) show "F\ArrCod\ \\<^sub>\ Vset \" by (auto intro!: f.arr_Set_ArrCod_in_Vset Vset_in_mono assms(2)) qed ( auto simp: f f.arr_Set_ArrVal_vdomain f.vfsequence_axioms f.arr_Set_length ) qed show "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F = G \\<^sub>A\<^bsub>cat_Set \\<^esub> F" if "G : B \\<^bsub>cat_Set \\<^esub> C" and "F : A \\<^bsub>cat_Set \\<^esub> B" for B C G A F proof- note g = cat_Set_is_arrD[OF that(1)] and f = cat_Set_is_arrD[OF that(2)] from that have \_gf_is_arr: "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F : A \\<^bsub>cat_Set \\<^esub> C" by (cs_concl cs_intro: cat_cs_intros is_arr_if_is_arr) from that have \_gf_is_arr: "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F : A \\<^bsub>cat_Set \\<^esub> C" by (cs_concl cs_intro: cat_cs_intros is_arr_if_is_arr) note \_gf = cat_Set_is_arrD[OF \_gf_is_arr] and \_gf = cat_Set_is_arrD[OF \_gf_is_arr] show ?thesis proof(rule arr_Set_eqI) show "arr_Set \ (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)" by (rule \_gf(1)) then interpret arr_Set_\_gf: arr_Set \ \(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ by simp from \_gf_is_arr have dom_lhs: "\\<^sub>\ ((G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) show "arr_Set \ (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)" by (rule \_gf(1)) then interpret arr_Set_\_gf: arr_Set \ \(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ by simp from \_gf_is_arr have dom_rhs: "\\<^sub>\ ((G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) show "(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\ = (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ A" from that this show "(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\ = (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros is_arr_if_is_arr ) qed auto qed (use \_gf_is_arr \_gf_is_arr in \cs_concl cs_simp: cat_cs_simps\)+ qed qed ( auto simp: assms(2) cat_Set_components Vset_trans Vset_in_mono cat_cs_intros ) qed subsection\Isomorphism\ lemma cat_Set_is_arr_isomorphismI[intro]: \\ See \cite{noauthor_nlab_nodate}\footnote{\url{ https://ncatlab.org/nlab/show/isomorphism }}). \ assumes "T : A \\<^bsub>cat_Set \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" shows "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" proof- interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF assms(1)]) note [cat_cs_intros] = cat_Par_is_arr_isomorphismI from wide_replete_subcategory_cat_Set_cat_Par assms have "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_cs_intros cat_sub_cs_intros cat_sub_fw_cs_intros) with wide_replete_subcategory_cat_Set_cat_Par assms show "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" by (cs_concl cs_simp: cat_sub_bw_cs_simps) qed lemma cat_Set_is_arr_isomorphismD[dest]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" shows "T : A \\<^bsub>cat_Set \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) from wide_replete_subcategory_cat_Set_cat_Par assms have T: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros) show "v11 (T\ArrVal\)" "\\<^sub>\ (T\ArrVal\) = A" "\\<^sub>\ (T\ArrVal\) = B" by (intro cat_Par_is_arr_isomorphismD[OF T])+ qed (rule is_arr_isomorphismD(1)[OF assms]) lemma cat_Set_is_arr_isomorphism: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B \ T : A \\<^bsub>cat_Set \\<^esub> B \ v11 (T\ArrVal\) \ \\<^sub>\ (T\ArrVal\) = A \ \\<^sub>\ (T\ArrVal\) = B" by auto subsection\The inverse arrow\ lemma cat_Set_ArrVal_app_is_arr[cat_cs_intros]: assumes "f : a \\<^bsub>\\<^esub> b" and "category \ \" (*the order of premises is important*) and "F : Hom \ a b \\<^bsub>cat_Set \\<^esub> Hom \ c d" shows "F\ArrVal\\f\ : c \\<^bsub>\\<^esub> d" proof- interpret \: category \ \ by (rule assms(2)) interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF assms(3)]) from assms have "F\ArrVal\\f\ \\<^sub>\ Hom \ c d" by (cs_concl cs_intro: cat_cs_intros cat_Set_cs_intros) then show ?thesis unfolding in_Hom_iff by simp qed abbreviation (input) converse_Set :: "V \ V" ("(_\\<^sub>S\<^sub>e\<^sub>t)" [1000] 999) where "a\\<^sub>S\<^sub>e\<^sub>t \ a\\<^sub>R\<^sub>e\<^sub>l" lemma cat_Set_the_inverse[cat_Set_cs_simps]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" shows "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>S\<^sub>e\<^sub>t" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) from wide_replete_subcategory_cat_Set_cat_Par assms have T: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros) from wide_replete_subcategory_cat_Set_cat_Par assms have [cat_cs_simps]: "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>C\<^bsub>cat_Par \\<^esub>" by ( cs_concl cs_full cs_simp: cat_sub_bw_cs_simps cs_intro: cat_sub_cs_intros ) from T show "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>R\<^sub>e\<^sub>l" by (cs_concl cs_simp: cat_Par_cs_simps cat_cs_simps cs_intro: \_\) qed lemma cat_Set_the_inverse_app[cat_cs_intros]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "a \\<^sub>\ A" and [cat_cs_simps]: "T\ArrVal\\a\ = b" shows "(T\\<^sub>C\<^bsub>cat_Set \\<^esub>)\ArrVal\\b\ = a" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) note T = cat_Set_is_arr_isomorphismD[OF assms(1)] interpret T: v11 \T\ArrVal\\ by (rule T(2)) from T.v11_axioms assms(1,2) show "T\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\ = a" by ( cs_concl cs_simp: converse_Rel_components V_cs_simps cat_Set_cs_simps cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) qed lemma cat_Set_ArrVal_app_the_inverse_is_arr[cat_cs_intros]: assumes "f : c \\<^bsub>\\<^esub> d" and "category \ \" (*the order of premises is important*) and "F : Hom \ a b \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> Hom \ c d" shows "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\f\ : a \\<^bsub>\\<^esub> b" proof- interpret \: category \ \ by (rule assms(2)) from cat_Set_is_arr_isomorphismD[OF assms(3)] interpret F: arr_Set \ F by (simp add: cat_Set_is_arrD) from assms have "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\f\ \\<^sub>\ Hom \ a b" by (cs_concl cs_intro: cat_cs_intros cat_arrow_cs_intros) then show ?thesis unfolding in_Hom_iff by simp qed lemma cat_Set_app_the_inverse_app[cat_cs_simps]: assumes "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "b \\<^sub>\ B" shows "F\ArrVal\\F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\\ = b" proof- note F = cat_Set_is_arr_isomorphismD[OF assms(1)] note F = F cat_Set_is_arrD[OF F(1)] interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF F(1)]) from assms have [cat_cs_simps]: "F \\<^sub>A\<^bsub>cat_Set \\<^esub> F\\<^sub>C\<^bsub>cat_Set \\<^esub> = cat_Set \\CId\\B\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have [cat_cs_simps]: "F\ArrVal\\F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\\ = (F \\<^sub>A\<^bsub>cat_Set \\<^esub> F\\<^sub>C\<^bsub>cat_Set \\<^esub>)\ArrVal\\b\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) from assms F.arr_Par_ArrCod_in_Vset[unfolded F] show ?thesis by (cs_concl cs_simp: cat_cs_simps) qed lemma cat_Set_the_inverse_app_app[cat_cs_simps]: assumes "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "a \\<^sub>\ A" shows "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\F\ArrVal\\a\\ = a" proof- note F = cat_Set_is_arr_isomorphismD[OF assms(1)] note F = F cat_Set_is_arrD[OF F(1)] interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF F(1)]) from assms have [cat_cs_simps]: "F\\<^sub>C\<^bsub>cat_Set \\<^esub> \\<^sub>A\<^bsub>cat_Set \\<^esub> F = cat_Set \\CId\\A\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have [cat_cs_simps]: "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\F\ArrVal\\a\\ = (F\\<^sub>C\<^bsub>cat_Set \\<^esub> \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) from assms F.arr_Par_ArrDom_in_Vset[unfolded F] show ?thesis by (cs_concl cs_simp: cat_cs_simps) qed subsection\Projection arrows for \vtimes\\ subsubsection\Definition and elementary properties\ definition vfst_arrow :: "V \ V \ V" where "vfst_arrow A B = [(\ab\\<^sub>\A \\<^sub>\ B. vfst ab), A \\<^sub>\ B, A]\<^sub>\" definition vsnd_arrow :: "V \ V \ V" where "vsnd_arrow A B = [(\ab\\<^sub>\A \\<^sub>\ B. vsnd ab), A \\<^sub>\ B, B]\<^sub>\" text\Components.\ lemma vfst_arrow_components: shows "vfst_arrow A B\ArrVal\ = (\ab\\<^sub>\A \\<^sub>\ B. vfst ab)" and [cat_cs_simps]: "vfst_arrow A B\ArrDom\ = A \\<^sub>\ B" and [cat_cs_simps]: "vfst_arrow A B\ArrCod\ = A" unfolding vfst_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) lemma vsnd_arrow_components: shows "vsnd_arrow A B\ArrVal\ = (\ab\\<^sub>\A \\<^sub>\ B. vsnd ab)" and [cat_cs_simps]: "vsnd_arrow A B\ArrDom\ = A \\<^sub>\ B" and [cat_cs_simps]: "vsnd_arrow A B\ArrCod\ = B" unfolding vsnd_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow value\ mk_VLambda vfst_arrow_components(1) |vsv vfst_arrow_ArrVal_vsv[cat_cs_intros]| |vdomain vfst_arrow_ArrVal_vdomain[cat_cs_simps]| |app vfst_arrow_ArrVal_app'| mk_VLambda vsnd_arrow_components(1) |vsv vsnd_arrow_ArrVal_vsv[cat_cs_intros]| |vdomain vsnd_arrow_ArrVal_vdomain[cat_cs_simps]| |app vsnd_arrow_ArrVal_app'| lemma vfst_arrow_ArrVal_app[cat_cs_simps]: assumes "ab = \a, b\" and "ab \\<^sub>\ A \\<^sub>\ B" shows "vfst_arrow A B\ArrVal\\ab\ = a" using assms(2) unfolding assms(1) by (simp add: vfst_arrow_ArrVal_app') lemma vfst_arrow_vrange: "\\<^sub>\ (vfst_arrow A B\ArrVal\) \\<^sub>\ A" unfolding vfst_arrow_components proof(intro vrange_VLambda_vsubset) fix ab assume "ab \\<^sub>\ A \\<^sub>\ B" then obtain a b where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ A" by clarsimp from a show "vfst ab \\<^sub>\ A" unfolding ab_def by simp qed lemma vsnd_arrow_ArrVal_app[cat_cs_simps]: assumes "ab = \a, b\" and "ab \\<^sub>\ A \\<^sub>\ B" shows "vsnd_arrow A B\ArrVal\\ab\ = b" using assms(2) unfolding assms(1) by (simp add: vsnd_arrow_ArrVal_app') lemma vsnd_arrow_vrange: "\\<^sub>\ (vsnd_arrow A B\ArrVal\) \\<^sub>\ B" unfolding vsnd_arrow_components proof(intro vrange_VLambda_vsubset) fix ab assume "ab \\<^sub>\ A \\<^sub>\ B" then obtain a b where ab_def: "ab = \a, b\" and b: "b \\<^sub>\ B" by clarsimp from b show "vsnd ab \\<^sub>\ B" unfolding ab_def by simp qed subsubsection\Projection arrows are arrows in the category \Set\\ lemma (in \) vfst_arrow_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A" proof(intro cat_Set_is_arrI arr_SetI, unfold cat_cs_simps) show "vfsequence (vfst_arrow A B)" unfolding vfst_arrow_def by simp show "vcard (vfst_arrow A B) = 3\<^sub>\" unfolding vfst_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (vfst_arrow A B\ArrVal\) \\<^sub>\ A" by (rule vfst_arrow_vrange) qed (use assms in \cs_concl cs_intro: V_cs_intros cat_cs_intros\)+ lemma (in \) vfst_arrow_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A" using assms unfolding cat_Set_components by (rule vfst_arrow_is_cat_Set_arr_Vset) lemma (in \) vfst_arrow_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Set \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.vfst_arrow_is_cat_Set_arr' lemma (in \) vsnd_arrow_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> B" proof(intro cat_Set_is_arrI arr_SetI , unfold cat_cs_simps) show "vfsequence (vsnd_arrow A B)" unfolding vsnd_arrow_def by simp show "vcard (vsnd_arrow A B) = 3\<^sub>\" unfolding vsnd_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (vsnd_arrow A B\ArrVal\) \\<^sub>\ B" by (rule vsnd_arrow_vrange) qed (use assms in \cs_concl cs_intro: V_cs_intros cat_cs_intros\)+ lemma (in \) vsnd_arrow_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> B" using assms unfolding cat_Set_components by (rule vsnd_arrow_is_cat_Set_arr_Vset) lemma (in \) vsnd_arrow_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Set \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.vsnd_arrow_is_cat_Set_arr' subsubsection\Projection arrows are arrows in the category \Par\\ lemma (in \) vfst_arrow_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Par \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD vfst_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vfst_arrow_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Par \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.vfst_arrow_is_cat_Par_arr' lemma (in \) vsnd_arrow_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Par \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD vsnd_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vsnd_arrow_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Par \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.vsnd_arrow_is_cat_Par_arr' subsubsection\Projection arrows are arrows in the category \Rel\\ lemma (in \) vfst_arrow_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD vfst_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vfst_arrow_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Rel \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.vfst_arrow_is_cat_Rel_arr' lemma (in \) vsnd_arrow_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD vsnd_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vsnd_arrow_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Rel \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.vsnd_arrow_is_cat_Rel_arr' subsubsection\Projection arrows are isomorphisms in the category \Set\\ lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism_Vset: assumes "A \\<^sub>\ Vset \" and "b \\<^sub>\ Vset \" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A" proof ( intro cat_Set_is_arr_isomorphismI arr_SetI vfst_arrow_is_cat_Set_arr_Vset assms, unfold cat_cs_simps ) show "v11 (vfst_arrow A (set {b})\ArrVal\)" proof(rule vsv.vsv_valeq_v11I, unfold cat_cs_simps) fix ab ab' assume prems: "ab \\<^sub>\ A \\<^sub>\ set {b}" "ab' \\<^sub>\ A \\<^sub>\ set {b}" "vfst_arrow A (set {b})\ArrVal\\ab\ = vfst_arrow A (set {b})\ArrVal\\ab'\" from prems obtain a where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ A" by clarsimp from prems obtain a' where ab'_def: "ab' = \a', b\" and a': "a' \\<^sub>\ A" by clarsimp from prems(3) a a' have "a = a'" unfolding ab_def ab'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: V_cs_intros) then show "ab = ab'" unfolding ab_def ab'_def by simp qed (cs_concl cs_intro: cat_cs_intros) show "\\<^sub>\ (vfst_arrow A (set {b})\ArrVal\) = A" proof(intro vsubset_antisym) show "A \\<^sub>\ \\<^sub>\ (vfst_arrow A (set {b})\ArrVal\)" proof(intro vsubsetI) fix a assume a: "a \\<^sub>\ A" then have a_def: "a = vfst_arrow A (set {b})\ArrVal\\\a, b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from a assms show "a \\<^sub>\ \\<^sub>\ (vfst_arrow A (set {b})\ArrVal\)" by (subst a_def, use nothing in \intro vsv.vsv_vimageI2\) (auto simp: cat_cs_simps cat_cs_intros) qed qed (rule vfst_arrow_vrange) qed (use assms in auto) lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism: assumes "A \\<^sub>\ cat_Set \\Obj\" and "b \\<^sub>\ cat_Set \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A" using assms unfolding cat_Set_components by (rule vfst_arrow_is_cat_Set_arr_isomorphism_Vset) lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "b \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Set \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.vfst_arrow_is_cat_Set_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism_Vset: assumes "a \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" proof ( intro cat_Set_is_arr_isomorphismI arr_SetI vsnd_arrow_is_cat_Set_arr_Vset assms, unfold cat_cs_simps ) show "v11 (vsnd_arrow (set {a}) B\ArrVal\)" proof(rule vsv.vsv_valeq_v11I, unfold cat_cs_simps) fix ab ab' assume prems: "ab \\<^sub>\ set {a} \\<^sub>\ B" "ab' \\<^sub>\ set {a} \\<^sub>\ B" "vsnd_arrow (set {a}) B\ArrVal\\ab\ = vsnd_arrow (set {a}) B\ArrVal\\ab'\" from prems obtain b where ab_def: "ab = \a, b\" and b: "b \\<^sub>\ B" by clarsimp from prems obtain b' where ab'_def: "ab' = \a, b'\" and b': "b' \\<^sub>\ B" by clarsimp from prems(3) b b' have "b = b'" unfolding ab_def ab'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: V_cs_intros) then show "ab = ab'" unfolding ab_def ab'_def by simp qed (cs_concl cs_intro: cat_cs_intros) show "\\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\) = B" proof(intro vsubset_antisym) show "B \\<^sub>\ \\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\)" proof(intro vsubsetI) fix b assume b: "b \\<^sub>\ B" then have b_def: "b = vsnd_arrow (set {a}) B\ArrVal\\\a, b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from b assms show "b \\<^sub>\ \\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\)" by (subst b_def, use nothing in \intro vsv.vsv_vimageI2\) (auto simp: cat_cs_simps cat_cs_intros) qed qed (rule vsnd_arrow_vrange) qed (use assms in auto) lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism: assumes "a \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" using assms unfolding cat_Set_components by (rule vsnd_arrow_is_cat_Set_arr_isomorphism_Vset) lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "a \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Set \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.vsnd_arrow_is_cat_Set_arr_isomorphism' subsubsection\Projection arrows are isomorphisms in the category \Par\\ lemma (in \) vfst_arrow_is_cat_Par_arr_isomorphism: assumes "A \\<^sub>\ cat_Par \\Obj\" and "b \\<^sub>\ cat_Par \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) show "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vfst_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) qed lemma (in \) vfst_arrow_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "b \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Par \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.vfst_arrow_is_cat_Par_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Par_arr_isomorphism: assumes "a \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) show "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vsnd_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) qed lemma (in \) vsnd_arrow_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "a \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Par \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.vsnd_arrow_is_cat_Par_arr_isomorphism' subsubsection\Projection arrows are isomorphisms in the category \Rel\\ lemma (in \) vfst_arrow_is_cat_Rel_arr_isomorphism: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "b \\<^sub>\ cat_Rel \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show ?thesis by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vfst_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) qed lemma (in \) vfst_arrow_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "b \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Rel \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.vfst_arrow_is_cat_Rel_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Rel_arr_isomorphism: assumes "a \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show ?thesis by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vsnd_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) qed lemma (in \) vsnd_arrow_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "a \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Rel \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.vsnd_arrow_is_cat_Rel_arr_isomorphism' subsection\Projection arrow for \vproduct\\ definition vprojection_arrow :: "V \ (V \ V) \ V \ V" where "vprojection_arrow I A i = [vprojection I A i, (\\<^sub>\i\\<^sub>\I. A i), A i]\<^sub>\" text\Components.\ lemma vprojection_arrow_components: shows "vprojection_arrow I A i\ArrVal\ = vprojection I A i" and "vprojection_arrow I A i\ArrDom\ = (\\<^sub>\i\\<^sub>\I. A i)" and "vprojection_arrow I A i\ArrCod\ = A i" unfolding vprojection_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Projection arrow value\ mk_VLambda vprojection_arrow_components(1)[unfolded vprojection_def] |vsv vprojection_arrow_vsv[cat_Set_cs_intros]| |vdomain vprojection_arrow_vdomain[cat_Set_cs_simps]| |app vprojection_arrow_app[cat_Set_cs_simps]| subsubsection\Projection arrow is an arrow in the category \Set\\ lemma (in \) arr_Set_vprojection_arrow: assumes "i \\<^sub>\ I" and "VLambda I A \\<^sub>\ Vset \" shows "arr_Set \ (vprojection_arrow I A i)" proof(intro arr_SetI) show "vfsequence (vprojection_arrow I A i)" unfolding vprojection_arrow_def by auto show "vcard (vprojection_arrow I A i) = 3\<^sub>\" unfolding vprojection_arrow_def by (simp add: nat_omega_simps) show "vprojection_arrow I A i\ArrCod\ \\<^sub>\ Vset \" unfolding vprojection_arrow_components proof- from assms(1) have "i \\<^sub>\ I" by simp then have "A i \\<^sub>\ \\<^sub>\ (VLambda I A)" by auto moreover from assms(2) have "\\<^sub>\ (VLambda I A) \\<^sub>\ Vset \" by (meson vrange_in_VsetI) ultimately show "A i \\<^sub>\ Vset \" by auto qed qed ( auto simp: vprojection_arrow_components intro!: assms vprojection_vrange_vsubset Limit_vproduct_in_Vset_if_VLambda_in_VsetI ) lemma (in \) vprojection_arrow_is_arr: assumes "i \\<^sub>\ I" and "VLambda I A \\<^sub>\ Vset \" shows "vprojection_arrow I A i : (\\<^sub>\i\\<^sub>\I. A i) \\<^bsub>cat_Set \\<^esub> A i" proof(intro cat_Set_is_arrI) from assms show "arr_Set \ (vprojection_arrow I A i)" by (rule arr_Set_vprojection_arrow) qed (simp_all add: vprojection_arrow_components) subsection\Product arrow value for \Rel\\ subsubsection\Definition and elementary properties\ definition prod_2_Rel_ArrVal :: "V \ V \ V" where "prod_2_Rel_ArrVal S T = set {\\a, b\, \c, d\\ | a b c d. \a, c\ \\<^sub>\ S \ \b, d\ \\<^sub>\ T}" lemma small_prod_2_Rel_ArrVal[simp]: "small {\\a, b\, \c, d\\ | a b c d. \a, c\ \\<^sub>\ S \ \b, d\ \\<^sub>\ T}" (is \small ?S\) proof(rule down) show "?S \ elts ((\\<^sub>\ S \\<^sub>\ \\<^sub>\ T) \\<^sub>\ (\\<^sub>\ S \\<^sub>\ \\<^sub>\ T))" by auto qed text\Rules.\ lemma prod_2_Rel_ArrValI: assumes "ab_cd = \\a, b\, \c, d\\" and "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" shows "ab_cd \\<^sub>\ prod_2_Rel_ArrVal S T" using assms unfolding prod_2_Rel_ArrVal_def by simp lemma prod_2_Rel_ArrValD[dest]: assumes "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" shows "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" using assms unfolding prod_2_Rel_ArrVal_def by auto lemma prod_2_Rel_ArrValE[elim]: assumes "ab_cd \\<^sub>\ prod_2_Rel_ArrVal S T" obtains a b c d where "ab_cd = \\a, b\, \c, d\\" and "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" using assms unfolding prod_2_Rel_ArrVal_def by auto text\Elementary properties\ lemma prod_2_Rel_ArrVal_vsubset_vprod: "prod_2_Rel_ArrVal S T \\<^sub>\ ((\\<^sub>\ S \\<^sub>\ \\<^sub>\ T) \\<^sub>\ (\\<^sub>\ S \\<^sub>\ \\<^sub>\ T))" by auto lemma prod_2_Rel_ArrVal_vbrelation: "vbrelation (prod_2_Rel_ArrVal S T)" using prod_2_Rel_ArrVal_vsubset_vprod by auto lemma prod_2_Rel_ArrVal_vdomain: "\\<^sub>\ (prod_2_Rel_ArrVal S T) = \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" proof(intro vsubset_antisym) show "\\<^sub>\ S \\<^sub>\ \\<^sub>\ T \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" proof(intro vsubsetI) fix ab assume "ab \\<^sub>\ \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" then obtain a b where ab_def: "ab = \a, b\" and "a \\<^sub>\ \\<^sub>\ S" and "b \\<^sub>\ \\<^sub>\ T" by auto then obtain c d where "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" by force then have "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" by (intro prod_2_Rel_ArrValI) auto then show "ab \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" unfolding ab_def by auto qed qed (use prod_2_Rel_ArrVal_vsubset_vprod in blast) lemma prod_2_Rel_ArrVal_vrange: "\\<^sub>\ (prod_2_Rel_ArrVal S T) = \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" proof(intro vsubset_antisym) show "\\<^sub>\ S \\<^sub>\ \\<^sub>\ T \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" proof(intro vsubsetI) fix cd assume "cd \\<^sub>\ \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" then obtain c d where cd_def: "cd = \c, d\" and "c \\<^sub>\ \\<^sub>\ S" and "d \\<^sub>\ \\<^sub>\ T" by auto then obtain a b where "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" by force then have "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" by (intro prod_2_Rel_ArrValI) auto then show "cd \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" unfolding cd_def by auto qed qed (use prod_2_Rel_ArrVal_vsubset_vprod in blast) subsubsection\Further properties\ lemma assumes "vsv g" and "vsv f" shows prod_2_Rel_ArrVal_vsv: "vsv (prod_2_Rel_ArrVal g f)" and prod_2_Rel_ArrVal_app: "\a b. \ a \\<^sub>\ \\<^sub>\ g; b \\<^sub>\ \\<^sub>\ f \ \ prod_2_Rel_ArrVal g f\\a,b\\ = \g\a\, f\b\\" proof- interpret g: vsv g by (rule assms(1)) interpret f: vsv f by (rule assms(2)) show vsv_gf: "vsv (prod_2_Rel_ArrVal g f)" by (intro vsvI; (elim prod_2_Rel_ArrValE)?; (unfold prod_2_Rel_ArrVal_def)?) (auto simp: g.vsv f.vsv) fix a b assume "a \\<^sub>\ \\<^sub>\ g" "b \\<^sub>\ \\<^sub>\ f" then have a_ga: "\a, g\a\\ \\<^sub>\ g" and b_fb: "\b, f\b\\ \\<^sub>\ f" by auto from a_ga b_fb show "prod_2_Rel_ArrVal g f\\a, b\\ = \g\a\, f\b\\" by (cs_concl cs_simp: vsv.vsv_appI[OF vsv_gf] cs_intro: prod_2_Rel_ArrValI) qed lemma prod_2_Rel_ArrVal_v11: assumes "v11 g" and "v11 f" shows "v11 (prod_2_Rel_ArrVal g f)" proof- interpret g: v11 g by (rule assms(1)) interpret f: v11 f by (rule assms(2)) show ?thesis proof ( intro vsv.vsv_valeq_v11I prod_2_Rel_ArrVal_vsv g.vsv_axioms f.vsv_axioms, unfold prod_2_Rel_ArrVal_vdomain ) fix ab cd assume prems: "ab \\<^sub>\ \\<^sub>\ g \\<^sub>\ \\<^sub>\ f" "cd \\<^sub>\ \\<^sub>\ g \\<^sub>\ \\<^sub>\ f" "prod_2_Rel_ArrVal g f\ab\ = prod_2_Rel_ArrVal g f\cd\" from prems(1) obtain a b where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ \\<^sub>\ g" and b: "b \\<^sub>\ \\<^sub>\ f" by auto from prems(2) obtain c d where cd_def: "cd = \c, d\" and c: "c \\<^sub>\ \\<^sub>\ g" and d: "d \\<^sub>\ \\<^sub>\ f" by auto from prems(3) a b c d have "\g\a\, f\b\\ = \g\c\, f\d\\" unfolding ab_def cd_def by (cs_prems cs_simp: prod_2_Rel_ArrVal_app cs_intro: V_cs_intros) then have "g\a\ = g\c\" and "f\b\ = f\d\" by simp_all then show "ab = cd" by (auto simp: ab_def cd_def a b c d f.v11_injective g.v11_injective) qed qed lemma prod_2_Rel_ArrVal_vcomp: "prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T = prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" proof- interpret ST': vbrelation \prod_2_Rel_ArrVal S' T'\ by (rule prod_2_Rel_ArrVal_vbrelation) interpret ST: vbrelation \prod_2_Rel_ArrVal S T\ by (rule prod_2_Rel_ArrVal_vbrelation) show ?thesis (*TODO: simplify proof*) proof(intro vsubset_antisym vsubsetI) fix aa'_cc' assume "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T" then obtain aa' bb' cc' where ac_def: "aa'_cc' = \aa', cc'\" and bc: "\bb', cc'\ \\<^sub>\ prod_2_Rel_ArrVal S' T'" and ab: "\aa', bb'\ \\<^sub>\ prod_2_Rel_ArrVal S T" by auto from bc obtain b b' c c' where bb'_cc'_def: "\bb', cc'\ = \\b, b'\, \c, c'\\" and bc: "\b, c\ \\<^sub>\ S'" and bc': "\b', c'\ \\<^sub>\ T'" by auto with ab obtain a a' where aa'_bb'_def: "\aa', bb'\ = \\a, a'\, \b, b'\\" and ab: "\a, b\ \\<^sub>\ S" and ab': "\a', b'\ \\<^sub>\ T" by auto from bb'_cc'_def have bb'_def: "bb' = \b, b'\" and cc'_def: "cc' = \c, c'\" by simp_all from aa'_bb'_def have aa'_def: "aa' = \a, a'\" and bb'_def: "bb' = \b, b'\" by simp_all from bc bc' ab ab' show "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" unfolding ac_def aa'_def cc'_def by (intro prod_2_Rel_ArrValI) (cs_concl cs_intro: prod_2_Rel_ArrValI vcompI)+ next fix aa'_cc' assume "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" then obtain a a' c c' where aa'_cc'_def: "aa'_cc' = \\a, a'\, \c, c'\\" and ac: "\a, c\ \\<^sub>\ S' \\<^sub>\ S" and ac': "\a', c'\ \\<^sub>\ T' \\<^sub>\ T" by blast from ac obtain b where ab: "\a, b\ \\<^sub>\ S" and bc: "\b, c\ \\<^sub>\ S'" by auto from ac' obtain b' where ab': "\a', b'\ \\<^sub>\ T" and bc': "\b', c'\ \\<^sub>\ T'" by auto from ab bc ab' bc' show "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T" unfolding aa'_cc'_def by (cs_concl cs_intro: vcompI prod_2_Rel_ArrValI) qed qed lemma prod_2_Rel_ArrVal_vid_on[cat_cs_simps]: "prod_2_Rel_ArrVal (vid_on A) (vid_on B) = vid_on (A \\<^sub>\ B)" unfolding prod_2_Rel_ArrVal_def by auto subsection\Product arrow for \Rel\\ subsubsection\Definition and elementary properties\ definition prod_2_Rel :: "V \ V \ V" where "prod_2_Rel S T = [ prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\), S\ArrDom\ \\<^sub>\ T\ArrDom\, S\ArrCod\ \\<^sub>\ T\ArrCod\ ]\<^sub>\" text\Components.\ lemma prod_2_Rel_components: shows "prod_2_Rel S T\ArrVal\ = prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)" and [cat_cs_simps]: "prod_2_Rel S T\ArrDom\ = S\ArrDom\ \\<^sub>\ T\ArrDom\" and [cat_cs_simps]: "prod_2_Rel S T\ArrCod\ = S\ArrCod\ \\<^sub>\ T\ArrCod\" unfolding prod_2_Rel_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Product arrow for \Rel\ is an arrow in \Rel\\ lemma prod_2_Rel_is_cat_Rel_arr: assumes "S : A \\<^bsub>cat_Rel \\<^esub> B" and "T : C \\<^bsub>cat_Rel \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^bsub>cat_Rel \\<^esub> B \\<^sub>\ D" proof- note S = cat_Rel_is_arrD[OF assms(1)] note T = cat_Rel_is_arrD[OF assms(2)] interpret S: arr_Rel \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Rel \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) show ?thesis proof(intro cat_Rel_is_arrI arr_RelI) show "vfsequence (prod_2_Rel S T)" unfolding prod_2_Rel_def by simp show "vcard (prod_2_Rel S T) = 3\<^sub>\" unfolding prod_2_Rel_def by (simp add: nat_omega_simps) from S have "\\<^sub>\ (S\ArrVal\) \\<^sub>\ A" and "\\<^sub>\ (S\ArrVal\) \\<^sub>\ B" by auto moreover from T have "\\<^sub>\ (T\ArrVal\) \\<^sub>\ C" and "\\<^sub>\ (T\ArrVal\) \\<^sub>\ D" by auto ultimately have "\\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\) \\<^sub>\ A \\<^sub>\ C" "\\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\) \\<^sub>\ B \\<^sub>\ D" by auto then show "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrDom\" "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrCod\" unfolding prod_2_Rel_components prod_2_Rel_ArrVal_vdomain prod_2_Rel_ArrVal_vrange by (force simp: prod_2_Rel_components)+ from S.arr_Rel_ArrDom_in_Vset T.arr_Rel_ArrDom_in_Vset S.arr_Rel_ArrCod_in_Vset T.arr_Rel_ArrCod_in_Vset show "prod_2_Rel S T\ArrDom\ \\<^sub>\ Vset \" "prod_2_Rel S T\ArrCod\ \\<^sub>\ Vset \" unfolding prod_2_Rel_components by (all\intro Limit_vtimes_in_VsetI\) auto qed (auto simp: prod_2_Rel_components intro: prod_2_Rel_ArrVal_vbrelation) qed lemma prod_2_Rel_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "S : A \\<^bsub>cat_Rel \\<^esub> B" and "T : C \\<^bsub>cat_Rel \\<^esub> D" and "A' = A \\<^sub>\ C" and "B' = B \\<^sub>\ D" and "\' = cat_Rel \" shows "prod_2_Rel S T : A' \\<^bsub>\'\<^esub> B'" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Rel_arr) subsubsection\Product arrow for \Rel\ is an arrow in \Set\\ lemma prod_2_Rel_app[cat_rel_par_Set_cs_simps]: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" and "a \\<^sub>\ A" and "c \\<^sub>\ C" and "ac = \a, c\" shows "prod_2_Rel S T\ArrVal\\ac\ = \S\ArrVal\\a\, T\ArrVal\\c\\" proof- note S = cat_Set_is_arrD[OF assms(1)] note T = cat_Set_is_arrD[OF assms(2)] interpret S: arr_Set \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Set \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) from assms(3,4) show ?thesis unfolding prod_2_Rel_components(1) assms(5) by ( cs_concl cs_simp: S.arr_Set_ArrVal_vdomain T.arr_Set_ArrVal_vdomain prod_2_Rel_ArrVal_app cs_intro: V_cs_intros ) qed lemma prod_2_Rel_is_cat_Set_arr: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> B \\<^sub>\ D" proof- note S = cat_Set_is_arrD[OF assms(1)] note T = cat_Set_is_arrD[OF assms(2)] interpret S: arr_Set \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Set \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) show ?thesis proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (prod_2_Rel S T)" unfolding prod_2_Rel_def by simp show "vcard (prod_2_Rel S T) = 3\<^sub>\" unfolding prod_2_Rel_def by (simp add: nat_omega_simps) from S.arr_Set_ArrVal_vrange T.arr_Set_ArrVal_vrange show "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrCod\" unfolding prod_2_Rel_components prod_2_Rel_ArrVal_vdomain prod_2_Rel_ArrVal_vrange by auto from assms S.arr_Par_ArrDom_in_Vset T.arr_Par_ArrDom_in_Vset show "prod_2_Rel S T\ArrDom\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from assms S.arr_Par_ArrCod_in_Vset T.arr_Par_ArrCod_in_Vset show "prod_2_Rel S T\ArrCod\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from assms show "prod_2_Rel S T\ArrDom\ = A \\<^sub>\ C" by (cs_concl cs_simp: cat_cs_simps) from assms show "prod_2_Rel S T\ArrCod\ = B \\<^sub>\ D" by (cs_concl cs_simp: cat_cs_simps) show "vsv (prod_2_Rel S T\ArrVal\)" unfolding prod_2_Rel_components by (intro prod_2_Rel_ArrVal_vsv S.ArrVal.vsv_axioms T.ArrVal.vsv_axioms) qed ( auto simp: cat_cs_simps cat_Set_cs_simps prod_2_Rel_ArrVal_vdomain prod_2_Rel_components(1) ) qed lemma prod_2_Rel_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" and "AC = A \\<^sub>\ C" and "BD = B \\<^sub>\ D" and "\' = cat_Set \" shows "prod_2_Rel S T : AC \\<^bsub>\'\<^esub> BD" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Set_arr) subsubsection\Product arrow for \Rel\ is an isomorphism in \Set\\ lemma prod_2_Rel_is_cat_Set_arr_isomorphism: assumes "S : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B \\<^sub>\ D" proof- note S = cat_Set_is_arr_isomorphismD[OF assms(1)] note T = cat_Set_is_arr_isomorphismD[OF assms(2)] show ?thesis proof ( intro cat_Set_is_arr_isomorphismI prod_2_Rel_is_cat_Set_arr[OF S(1) T(1)], unfold prod_2_Rel_components ) show "\\<^sub>\ (prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)) = A \\<^sub>\ C" unfolding prod_2_Rel_ArrVal_vdomain by (cs_concl cs_simp: S(3) T(3) cs_intro: cat_cs_intros) show "\\<^sub>\ (prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)) = B \\<^sub>\ D" unfolding prod_2_Rel_ArrVal_vrange by (cs_concl cs_simp: S(4) T(4) cs_intro: cat_cs_intros) qed (use S(2) T(2) in \cs_concl cs_intro: prod_2_Rel_ArrVal_v11\) qed lemma prod_2_Rel_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "S : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> D" and "AC = A \\<^sub>\ C" and "BD = B \\<^sub>\ D" and "\' = cat_Set \" shows "prod_2_Rel S T : AC \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> BD" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Set_arr_isomorphism) subsubsection\Further elementary properties\ lemma prod_2_Rel_Comp: assumes "G' : B' \\<^bsub>cat_Rel \\<^esub> B''" and "F' : A' \\<^bsub>cat_Rel \\<^esub> A''" and "G : B \\<^bsub>cat_Rel \\<^esub> B'" and "F : A \\<^bsub>cat_Rel \\<^esub> A'" shows "prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F = prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)" proof- from cat_Rel_is_arrD(1)[OF assms(1)] interpret \ \ by auto interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) note (*prefer cat_Rel*)[cat_cs_simps] = cat_Rel_is_arrD(2,3) from assms have GF'_GF: "prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F : B \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B'' \\<^sub>\ A''" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) from assms Rel.category_axioms have GG'_FF': "prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F) : B \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B'' \\<^sub>\ A''" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) show ?thesis proof(rule arr_Rel_eqI[of \]) from GF'_GF show arr_Rel_GF'_GF: "arr_Rel \ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)" by (auto dest: cat_Rel_is_arrD(1)) from GG'_FF' show arr_Rel_GG'_FF': "arr_Rel \ (prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F))" by (auto dest: cat_Rel_is_arrD(1)) show "(prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\ = prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" proof(intro vsubset_antisym vsubsetI) fix R assume "R \\<^sub>\ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\" from this assms have "R \\<^sub>\ prod_2_Rel_ArrVal (G'\ArrVal\) (F'\ArrVal\) \\<^sub>\ prod_2_Rel_ArrVal (G\ArrVal\) (F\ArrVal\)" by ( cs_prems cs_simp: prod_2_Rel_components(1) comp_Rel_components(1) cat_Rel_cs_simps cs_intro: cat_Rel_par_set_cs_intros ) from this[unfolded prod_2_Rel_ArrVal_vcomp] assms show "R \\<^sub>\ prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" by ( cs_concl cs_simp: prod_2_Rel_components comp_Rel_components(1) cat_Rel_cs_simps ) next fix R assume "R \\<^sub>\ prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" from this assms have "R \\<^sub>\ prod_2_Rel_ArrVal (G'\ArrVal\ \\<^sub>\ G\ArrVal\) (F'\ArrVal\ \\<^sub>\ F\ArrVal\)" by ( cs_prems cs_simp: comp_Rel_components prod_2_Rel_components cat_Rel_cs_simps ) from this[folded prod_2_Rel_ArrVal_vcomp] assms show "R \\<^sub>\ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\" by ( cs_concl cs_simp: prod_2_Rel_components comp_Rel_components(1) cat_Rel_cs_simps cs_intro: cat_Rel_par_set_cs_intros ) qed qed ( use GF'_GF assms in (*slow*) \ cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Rel_cs_intros \ )+ qed lemma (in \) prod_2_Rel_CId[cat_cs_simps]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\) = cat_Rel \\CId\\A \\<^sub>\ B\" proof- interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) from assms have A_B: "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\) : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ B" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) from assms Rel.category_axioms have AB: "cat_Rel \\CId\\A \\<^sub>\ B\ : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ B" by ( cs_concl cs_simp: cat_Rel_components(1) cs_intro: V_cs_intros cat_cs_intros ) show ?thesis proof(rule arr_Rel_eqI) from A_B show arr_Rel_GF'_GF: "arr_Rel \ (prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\))" by (auto dest: cat_Rel_is_arrD(1)) from AB show arr_Rel_GG'_FF': "arr_Rel \ (cat_Rel \\CId\\A \\<^sub>\ B\)" by (auto dest: cat_Rel_is_arrD(1)) from assms show "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\)\ArrVal\ = cat_Rel \\CId\\A \\<^sub>\ B\\ArrVal\" by ( cs_concl cs_simp: id_Rel_components prod_2_Rel_components cat_cs_simps cat_Rel_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed ( use A_B assms in \ cs_concl cs_simp: prod_2_Rel_components cat_Rel_cs_simps cs_intro: cat_cs_intros \ )+ qed subsection\Product functor for \Rel\\ definition cf_prod_2_Rel :: "V \ V" where "cf_prod_2_Rel \ = [ (\AB\\<^sub>\(\ \\<^sub>C \)\Obj\. AB\0\ \\<^sub>\ AB\1\<^sub>\\), (\ST\\<^sub>\(\ \\<^sub>C \)\Arr\. prod_2_Rel (ST\0\) (ST\1\<^sub>\\)), \ \\<^sub>C \, \ ]\<^sub>\" text\Components.\ lemma cf_prod_2_Rel_components: shows "cf_prod_2_Rel \\ObjMap\ = (\AB\\<^sub>\(\ \\<^sub>C \)\Obj\. AB\0\ \\<^sub>\ AB\1\<^sub>\\)" and "cf_prod_2_Rel \\ArrMap\ = (\ST\\<^sub>\(\ \\<^sub>C \)\Arr\. prod_2_Rel (ST\0\) (ST\1\<^sub>\\))" and [cat_cs_simps]: "cf_prod_2_Rel \\HomDom\ = \ \\<^sub>C \" and [cat_cs_simps]: "cf_prod_2_Rel \\HomCod\ = \" unfolding cf_prod_2_Rel_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda cf_prod_2_Rel_components(1) |vsv cf_prod_2_Rel_ObjMap_vsv[cat_cs_intros]| |vdomain cf_prod_2_Rel_ObjMap_vdomain[cat_cs_simps]| lemma cf_prod_2_Rel_ObjMap_app[cat_cs_simps]: assumes "AB = [A, B]\<^sub>\" and "AB \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "A \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>cf_prod_2_Rel \\<^esub> B = A \\<^sub>\ B" using assms(2) unfolding assms(1) cf_prod_2_Rel_components by (simp add: nat_omega_simps) lemma (in \) cf_prod_2_Rel_ObjMap_vrange: "\\<^sub>\ (cf_prod_2_Rel (cat_Rel \)\ObjMap\) \\<^sub>\ cat_Rel \\Obj\" proof- interpret Rel: category \ \cat_Rel \\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(rule vsv.vsv_vrange_vsubset, unfold cat_cs_simps) fix AB assume prems: "AB \\<^sub>\ (cat_Rel \ \\<^sub>C cat_Rel \)\Obj\" with Rel.category_axioms obtain A B where AB_def: "AB = [A, B]\<^sub>\" and A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) from prems A B show "cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\ \\<^sub>\ cat_Rel \\Obj\" unfolding AB_def cat_Rel_components(1) by (cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: V_cs_intros) qed (cs_concl cs_intro: cat_cs_intros) qed subsubsection\Arrow map\ mk_VLambda cf_prod_2_Rel_components(2) |vsv cf_prod_2_Rel_ArrMap_vsv[cat_cs_intros]| |vdomain cf_prod_2_Rel_ArrMap_vdomain[cat_cs_simps]| lemma cf_prod_2_Rel_ArrMap_app[cat_cs_simps]: assumes "GF = [G, F]\<^sub>\" and "GF \\<^sub>\ (\ \\<^sub>C \)\Arr\" shows "G \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>cf_prod_2_Rel \\<^esub> F = prod_2_Rel G F" using assms(2) unfolding assms(1) cf_prod_2_Rel_components by (simp add: nat_omega_simps) subsubsection\Product functor for \Rel\ is a functor\ lemma (in \) cf_prod_2_Rel_is_functor: "cf_prod_2_Rel (cat_Rel \) : cat_Rel \ \\<^sub>C cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof- interpret Rel: category \ \cat_Rel \\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(rule is_functorI') show "vfsequence (cf_prod_2_Rel (cat_Rel \))" unfolding cf_prod_2_Rel_def by auto show "vcard (cf_prod_2_Rel (cat_Rel \)) = 4\<^sub>\" unfolding cf_prod_2_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_prod_2_Rel (cat_Rel \)\ObjMap\) \\<^sub>\ cat_Rel \\Obj\" by (rule cf_prod_2_Rel_ObjMap_vrange) show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF\ : cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\ \\<^bsub>cat_Rel \\<^esub> cf_prod_2_Rel (cat_Rel \)\ObjMap\\CD\" if "GF : AB \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> CD" for AB CD GF proof- from that obtain G F A B C D where GF_def: "GF = [G, F]\<^sub>\" and AB_def: "AB = [A, B]\<^sub>\" and CD_def: "CD = [C, D]\<^sub>\" and G: "G : A \\<^bsub>cat_Rel \\<^esub> C" and F: "F : B \\<^bsub>cat_Rel \\<^esub> D" by (elim cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms]) from that G F show ?thesis unfolding GF_def AB_def CD_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros cat_prod_cs_intros ) qed show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF' \\<^sub>A\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> GF\ = cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF'\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF\" if "GF' : AB' \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> AB''" and "GF : AB \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> AB'" for AB' AB'' GF' AB GF proof- from that(2) obtain G F A A' B B' where GF_def: "GF = [G, F]\<^sub>\" and AB_def: "AB = [A, B]\<^sub>\" and AB'_def: "AB' = [A', B']\<^sub>\" and G: "G : A \\<^bsub>cat_Rel \\<^esub> A'" and F: "F : B \\<^bsub>cat_Rel \\<^esub> B'" by (elim cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms]) with that(1) obtain G' F' A'' B'' where GF'_def: "GF' = [G', F']\<^sub>\" and AB''_def: "AB'' = [A'', B'']\<^sub>\" and G': "G' : A' \\<^bsub>cat_Rel \\<^esub> A''" and F': "F' : B' \\<^bsub>cat_Rel \\<^esub> B''" by ( auto elim: cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms] ) from that G F G' F' show ?thesis unfolding GF_def AB_def AB'_def GF'_def AB''_def by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps prod_2_Rel_Comp cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\(cat_Rel \ \\<^sub>C cat_Rel \)\CId\\AB\\ = cat_Rel \\CId\\cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\\" if "AB \\<^sub>\ (cat_Rel \ \\<^sub>C cat_Rel \)\Obj\" for AB proof- from that obtain A B where AB_def: "AB = [A, B]\<^sub>\" and A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" by (elim cat_prod_2_ObjE[OF Rel.category_axioms Rel.category_axioms]) from A B show ?thesis unfolding AB_def by ( cs_concl cs_simp: cf_prod_2_Rel_ObjMap_app cf_prod_2_Rel_ArrMap_app cat_cs_simps cat_prod_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_Rel_cs_intros cat_prod_cs_intros ) qed qed ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_cs_intros cat_Rel_cs_intros )+ qed lemma (in \) cf_prod_2_Rel_is_functor'[cat_cs_intros]: assumes "\' = cat_Rel \ \\<^sub>C cat_Rel \" and "\' = cat_Rel \" and "\' = \" shows "cf_prod_2_Rel (cat_Rel \) : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule cf_prod_2_Rel_is_functor) lemmas [cat_cs_intros] = \.cf_prod_2_Rel_is_functor' subsection\Product universal property arrow for \Set\\ subsubsection\Definition and elementary properties\ definition cat_Set_obj_prod_up :: "V \ (V \ V) \ V \ (V \ V) \ V" where "cat_Set_obj_prod_up I F A \ = [(\a\\<^sub>\A. (\i\\<^sub>\I. \ i\ArrVal\\a\)), A, (\\<^sub>\i\\<^sub>\I. F i)]\<^sub>\" text\Components.\ lemma cat_Set_obj_prod_up_components: shows "cat_Set_obj_prod_up I F A \\ArrVal\ = (\a\\<^sub>\A. (\i\\<^sub>\I. \ i\ArrVal\\a\))" and [cat_Set_cs_simps]: "cat_Set_obj_prod_up I F A \\ArrDom\ = A" and [cat_Set_cs_simps]: "cat_Set_obj_prod_up I F A \\ArrCod\ = (\\<^sub>\i\\<^sub>\I. F i)" unfolding cat_Set_obj_prod_up_def arr_field_simps by (simp_all add: nat_omega_simps) text\Arrow value.\ mk_VLambda cat_Set_obj_prod_up_components(1) |vsv cat_Set_obj_prod_up_ArrVal_vsv[cat_Set_cs_intros]| |vdomain cat_Set_obj_prod_up_ArrVal_vdomain[cat_Set_cs_simps]| |app cat_Set_obj_prod_up_ArrVal_app| lemma cat_Set_obj_prod_up_ArrVal_vrange: assumes "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" unfolding cat_Set_obj_prod_up_components proof(intro vrange_VLambda_vsubset vproductI) fix a assume prems: "a \\<^sub>\ A" show "\i\\<^sub>\I. (\i\\<^sub>\I. \ i\ArrVal\\a\)\i\ \\<^sub>\ F i" proof(intro ballI) fix i assume "i \\<^sub>\ I" with assms prems show "(\i\\<^sub>\I. \ i\ArrVal\\a\)\i\ \\<^sub>\ F i" by (cs_concl cs_simp: V_cs_simps cs_intro: cat_Set_cs_intros) qed qed auto lemma cat_Set_obj_prod_up_ArrVal_app_vdomain[cat_Set_cs_simps]: assumes "a \\<^sub>\ A" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\) = I" unfolding cat_Set_obj_prod_up_ArrVal_app[OF assms] by simp lemma cat_Set_obj_prod_up_ArrVal_app_component[cat_Set_cs_simps]: assumes "a \\<^sub>\ A" and "i \\<^sub>\ I" shows "cat_Set_obj_prod_up I F A \\ArrVal\\a\\i\ = \ i\ArrVal\\a\" using assms by (cs_concl cs_simp: cat_Set_obj_prod_up_ArrVal_app V_cs_simps) lemma cat_Set_obj_prod_up_ArrVal_app_vrange: assumes "a \\<^sub>\ A" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" proof(intro vsubsetI) fix b assume prems: "b \\<^sub>\ \\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\)" from assms(1) have "vsv (cat_Set_obj_prod_up I F A \\ArrVal\\a\)" by (auto simp: cat_Set_obj_prod_up_components) with prems obtain i where b_def: "b = cat_Set_obj_prod_up I F A \\ArrVal\\a\\i\" and i: "i \\<^sub>\ I" by ( auto elim: vsv.vrange_atE simp: cat_Set_obj_prod_up_ArrVal_app[OF assms(1)] ) from cat_Set_obj_prod_up_ArrVal_app_component[OF assms(1) i] b_def have b_def': "b = \ i\ArrVal\\a\" by simp from assms(1) assms(2)[OF i] have "b \\<^sub>\ F i" unfolding b_def' by (cs_concl cs_intro: cat_Set_cs_intros) with i show "b \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" by force qed subsubsection\Product universal property arrow for \Set\ is an arrow in \Set\\ lemma (in \) cat_Set_obj_prod_up_cat_Set_is_arr: assumes "A \\<^sub>\ Vset \" and "VLambda I F \\<^sub>\ Vset \" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "cat_Set_obj_prod_up I F A \ : A \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (cat_Set_obj_prod_up I F A \)" unfolding cat_Set_obj_prod_up_def by auto show "vcard (cat_Set_obj_prod_up I F A \) = 3\<^sub>\" unfolding cat_Set_obj_prod_up_def by (auto simp: nat_omega_simps) show "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\) \\<^sub>\ cat_Set_obj_prod_up I F A \\ArrCod\" unfolding cat_Set_obj_prod_up_components(3) by (rule cat_Set_obj_prod_up_ArrVal_vrange[OF assms(3)]) show "cat_Set_obj_prod_up I F A \\ArrCod\ \\<^sub>\ Vset \" unfolding cat_Set_cs_simps by (rule Limit_vproduct_in_Vset_if_VLambda_in_VsetI) (simp_all add: cat_Set_cs_simps assms) qed (auto simp: assms cat_Set_cs_simps intro: cat_Set_cs_intros) -lemma (in \) pdg_dghm_comp_dghm_proj_dghm_up: +lemma (in \) cat_Set_cf_comp_proj_obj_prod_up: assumes "A \\<^sub>\ Vset \" and "VLambda I F \\<^sub>\ Vset \" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" and "i \\<^sub>\ I" shows "\ i = vprojection_arrow I F i \\<^sub>A\<^bsub>cat_Set \\<^esub> cat_Set_obj_prod_up I F A \" (is \\ i = ?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\) proof(rule arr_Set_eqI[of \]) note \i = assms(3)[OF assms(4)] note \i = cat_Set_is_arrD[OF \i] \i have Fi: "?Fi : (\\<^sub>\i\\<^sub>\I. F i) \\<^bsub>cat_Set \\<^esub> F i" by (rule vprojection_arrow_is_arr[OF assms(4,2)]) from cat_Set_obj_prod_up_cat_Set_is_arr[OF assms(1,2,3)] have \: "cat_Set_obj_prod_up I F A \ : A \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" by simp show "arr_Set \ (\ i)" by (rule \i(1)) interpret \i: arr_Set \ \\ i\ by (rule \i(1)) from Fi \ have Fi_\: "?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ : A \\<^bsub>cat_Set \\<^esub> F i" by (cs_concl cs_intro: cat_cs_intros) then show arr_Set_Fi_\: "arr_Set \ (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)" by (auto simp: cat_Set_is_arrD(1)) interpret arr_Set \ \?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\ by (rule arr_Set_Fi_\) from \i have dom_lhs: "\\<^sub>\ (\ i\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) from Fi_\ have dom_rhs: "\\<^sub>\ ((?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\ i\ArrVal\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ A" from assms(4) prems \i(4) \ Fi show "\ i\ArrVal\\a\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\\a\" by ( cs_concl cs_simp: cat_Set_cs_simps cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros ) qed auto from Fi \ show "\ i\ArrDom\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrDom\" by (cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps \i(2)) from Fi \ show "\ i\ArrCod\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrCod\" by (cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps \i(3)) qed subsection\Equalizer object for the category \Set\\ text\ The definition of the (non-categorical concept of an) equalizer can be found in \cite{noauthor_wikipedia_2001}\footnote{ \url{https://en.wikipedia.org/wiki/Equaliser_(mathematics)} }\ definition vequalizer :: "V \ V \ V \ V" where "vequalizer X f g = set {x. x \\<^sub>\ X \ f\ArrVal\\x\ = g\ArrVal\\x\}" lemma small_vequalizer[simp]: "small {x. x \\<^sub>\ X \ f\ArrVal\\x\ = g\ArrVal\\x\}" by auto text\Rules.\ lemma vequalizerI: assumes "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" shows "x \\<^sub>\ vequalizer X f g" using assms unfolding vequalizer_def by auto lemma vequalizerD[dest]: assumes "x \\<^sub>\ vequalizer X f g" shows "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" using assms unfolding vequalizer_def by auto lemma vequalizerE[elim]: assumes "x \\<^sub>\ vequalizer X f g" obtains "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" using assms unfolding vequalizer_def by auto text\Elementary results.\ lemma vequalizer_vsubset_vdomain[cat_Set_cs_intros]: "vequalizer a g f \\<^sub>\ a" by auto lemma Limit_vequalizer_in_Vset[cat_Set_cs_intros]: assumes "Limit \" and "a \\<^sub>\ Vset \" shows "vequalizer a g f \\<^sub>\ Vset \" using assms by auto lemma vequalizer_flip: "vequalizer a f g = vequalizer a g f" unfolding vequalizer_def by auto lemma (in \) cat_Set_incl_Set_commute: assumes "\ : \ \\<^bsub>cat_Set \\<^esub> \" and "\ : \ \\<^bsub>cat_Set \\<^esub> \" shows "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer \ \ \) \ = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer \ \ \) \" (is \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\) proof- note \ = cat_Set_is_arrD[OF assms(1)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note \ = cat_Set_is_arrD[OF assms(2)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note [cat_Set_cs_intros] = \.arr_Set_ArrDom_in_Vset \.arr_Set_ArrCod_in_Vset from assms have \_incl: "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl : vequalizer \ \ \ \\<^bsub>cat_Set \\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros) then have dom_lhs: "\\<^sub>\ ((\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\) = vequalizer \ \ \" by (cs_concl cs_simp: cat_cs_simps)+ from assms have \_incl: "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl : vequalizer \ \ \ \\<^bsub>cat_Set \\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros) then have dom_rhs: "\\<^sub>\ ((\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\) = vequalizer \ \ \" by (cs_concl cs_simp: cat_cs_simps)+ show ?thesis proof(rule arr_Set_eqI) from \_incl show arr_Set_\_incl: "arr_Set \ (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by (auto dest: cat_Set_is_arrD(1)) interpret arr_Set_\_incl: arr_Set \ \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\ by (rule arr_Set_\_incl) from \_incl show arr_Set_\_incl: "arr_Set \ (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by (auto dest: cat_Set_is_arrD(1)) interpret arr_Set_\_incl: arr_Set \ \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\ by (rule arr_Set_\_incl) show "(\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\ = (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ vequalizer \ \ \" with assms show "(\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\\a\ = (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\\a\" by ( cs_concl cs_simp: vequalizerD(2) cat_Set_cs_simps cat_cs_simps cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros ) qed auto qed (use \_incl \_incl in \cs_concl cs_simp: cat_cs_simps\)+ qed subsection\Auxiliary\ text\ This subsection is reserved for insignificant helper lemmas and rules that are used in applied formalization elsewhere. \ lemma (in \) cat_Rel_CId_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" shows "cat_Rel \\CId\\A\ : A \\<^bsub>cat_Set \\<^esub> A" proof- from assms show ?thesis unfolding cat_Rel_components cat_Set_components(6)[symmetric] by (cs_concl cs_simp: cat_Set_components(1) cs_intro: cat_cs_intros) qed lemma (in \) cat_Rel_CId_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B' = A" and "C' = A" and "\' = cat_Set \" shows "cat_Rel \\CId\\A\ : B' \\<^bsub>\'\<^esub> C'" using assms(1) unfolding assms(2-4) by (rule cat_Rel_CId_is_cat_Set_arr) text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Small_Category.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Small_Category.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Small_Category.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Small_Category.thy @@ -1,306 +1,306 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Smallness for categories\ theory CZH_ECAT_Small_Category imports CZH_Foundations.CZH_SMC_Small_Semicategory CZH_ECAT_Category begin subsection\Background\ text\ An explanation of the methodology chosen for the exposition of all matters related to the size of the categories and associated entities is given in \cite{milehins_category_2021}. \ named_theorems cat_small_cs_simps named_theorems cat_small_cs_intros subsection\Tiny category\ subsubsection\Definition and elementary properties\ locale tiny_category = \ \ + vfsequence \ + CId: vsv \\\CId\\ for \ \ + assumes tiny_cat_length[cat_cs_simps]: "vcard \ = 6\<^sub>\" and tiny_cat_tiny_semicategory[slicing_intros]: "tiny_semicategory \ (cat_smc \)" and tiny_cat_CId_vdomain[cat_cs_simps]: "\\<^sub>\ (\\CId\) = \\Obj\" and tiny_cat_CId_is_arr[cat_cs_intros]: "a \\<^sub>\ \\Obj\ \ \\CId\\a\ : a \\<^bsub>\\<^esub> a" and tiny_cat_CId_left_left[cat_cs_simps]: "f : a \\<^bsub>\\<^esub> b \ \\CId\\b\ \\<^sub>A\<^bsub>\\<^esub> f = f" and tiny_cat_CId_right_left[cat_cs_simps]: "f : b \\<^bsub>\\<^esub> c \ f \\<^sub>A\<^bsub>\\<^esub> \\CId\\b\ = f" lemmas [slicing_intros] = tiny_category.tiny_cat_tiny_semicategory text\Rules.\ lemma (in tiny_category) tiny_category_axioms'[cat_small_cs_intros]: assumes "\' = \" shows "tiny_category \' \" unfolding assms by (rule tiny_category_axioms) mk_ide rf tiny_category_def[unfolded tiny_category_axioms_def] |intro tiny_categoryI| |dest tiny_categoryD[dest]| |elim tiny_categoryE[elim]| lemma tiny_categoryI': assumes "category \ \" and "\\Obj\ \\<^sub>\ Vset \" and "\\Arr\ \\<^sub>\ Vset \" shows "tiny_category \ \" proof- interpret category \ \ by (rule assms(1)) show ?thesis proof(intro tiny_categoryI) from assms show "tiny_semicategory \ (cat_smc \)" by (intro tiny_semicategoryI') (auto simp: slicing_simps) qed (auto simp: vfsequence_axioms cat_cs_simps cat_cs_intros) qed lemma tiny_categoryI'': assumes "\ \" and "vfsequence \" and "vcard \ = 6\<^sub>\" and "vsv (\\Dom\)" and "vsv (\\Cod\)" and "vsv (\\Comp\)" and "vsv (\\CId\)" and "\\<^sub>\ (\\Dom\) = \\Arr\" and "\\<^sub>\ (\\Dom\) \\<^sub>\ \\Obj\" and "\\<^sub>\ (\\Cod\) = \\Arr\" and "\\<^sub>\ (\\Cod\) \\<^sub>\ \\Obj\" and "\gf. gf \\<^sub>\ \\<^sub>\ (\\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>\\<^esub> c \ f : a \\<^bsub>\\<^esub> b)" and "\\<^sub>\ (\\CId\) = \\Obj\" and "\b c g a f. \ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" and "\c d h b g a f. \ h : c \\<^bsub>\\<^esub> d; g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ (h \\<^sub>A\<^bsub>\\<^esub> g) \\<^sub>A\<^bsub>\\<^esub> f = h \\<^sub>A\<^bsub>\\<^esub> (g \\<^sub>A\<^bsub>\\<^esub> f)" and "\a. a \\<^sub>\ \\Obj\ \ \\CId\\a\ : a \\<^bsub>\\<^esub> a" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\CId\\b\ \\<^sub>A\<^bsub>\\<^esub> f = f" and "\b c f. f : b \\<^bsub>\\<^esub> c \ f \\<^sub>A\<^bsub>\\<^esub> \\CId\\b\ = f" and "\\Obj\ \\<^sub>\ Vset \" and "\\Arr\ \\<^sub>\ Vset \" shows "tiny_category \ \" by (intro tiny_categoryI tiny_semicategoryI'', unfold slicing_simps) (simp_all add: cat_smc_def nat_omega_simps assms) text\Slicing.\ context tiny_category begin interpretation smc: tiny_semicategory \ \cat_smc \\ by (rule tiny_cat_tiny_semicategory) lemmas_with [unfolded slicing_simps]: tiny_cat_semicategory = smc.semicategory_axioms and tiny_cat_Obj_in_Vset[cat_small_cs_intros] = smc.tiny_smc_Obj_in_Vset and tiny_cat_Arr_in_Vset[cat_small_cs_intros] = smc.tiny_smc_Arr_in_Vset and tiny_cat_Dom_in_Vset[cat_small_cs_intros] = smc.tiny_smc_Dom_in_Vset and tiny_cat_Cod_in_Vset[cat_small_cs_intros] = smc.tiny_smc_Cod_in_Vset and tiny_cat_Comp_in_Vset[cat_small_cs_intros] = smc.tiny_smc_Comp_in_Vset end text\Elementary properties.\ sublocale tiny_category \ category by (rule categoryI) ( auto simp: vfsequence_axioms tiny_cat_semicategory cat_cs_intros cat_cs_simps ) -lemmas (in tiny_category) tiny_dg_category = category_axioms +lemmas (in tiny_category) tiny_cat_category = category_axioms -lemmas [cat_small_cs_intros] = tiny_category.tiny_dg_category +lemmas [cat_small_cs_intros] = tiny_category.tiny_cat_category text\Size.\ lemma (in tiny_category) tiny_cat_CId_in_Vset: "\\CId\ \\<^sub>\ Vset \" proof- from tiny_cat_Obj_in_Vset have "\\<^sub>\ (\\CId\) \\<^sub>\ Vset \" by (simp add: tiny_cat_Obj_in_Vset cat_cs_simps) moreover from tiny_cat_Arr_in_Vset cat_CId_vrange tiny_cat_Arr_in_Vset have "\\<^sub>\ (\\CId\) \\<^sub>\ Vset \" by auto ultimately show ?thesis by (blast intro: \_Limit_\\) qed lemma (in tiny_category) tiny_cat_in_Vset: "\ \\<^sub>\ Vset \" proof- note [cat_cs_intros] = tiny_cat_Obj_in_Vset tiny_cat_Arr_in_Vset tiny_cat_Dom_in_Vset tiny_cat_Cod_in_Vset tiny_cat_Comp_in_Vset tiny_cat_CId_in_Vset show ?thesis by (subst cat_def) (cs_concl cs_intro: cat_cs_intros V_cs_intros) qed lemma tiny_category[simp]: "small {\. tiny_category \ \}" proof(rule down) show "{\. tiny_category \ \} \ elts (set {\. category \ \})" by (auto intro: cat_small_cs_intros) qed lemma small_categories_vsubset_Vset: "set {\. tiny_category \ \} \\<^sub>\ Vset \" by (rule vsubsetI) (simp_all add: tiny_category.tiny_cat_in_Vset) lemma (in category) cat_tiny_category_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "tiny_category \ \" proof(intro tiny_categoryI) show "tiny_semicategory \ (cat_smc \)" by ( rule semicategory.smc_tiny_semicategory_if_ge_Limit, rule cat_semicategory; intro assms ) qed (auto simp: assms(1) cat_cs_simps cat_cs_intros vfsequence_axioms) subsubsection\Opposite tiny category\ lemma (in tiny_category) tiny_category_op: "tiny_category \ (op_cat \)" by (intro tiny_categoryI') (auto simp: cat_op_simps cat_cs_intros cat_small_cs_intros) lemmas tiny_category_op[cat_op_intros] = tiny_category.tiny_category_op subsection\Finite category\ subsubsection\Definition and elementary properties\ text\ A definition of a finite category can be found in nLab \cite{noauthor_nlab_nodate}\footnote{ \url{https://ncatlab.org/nlab/show/finite+category} }. \ (*TODO: implicit redundant assumption*) locale finite_category = \ \ + vfsequence \ + CId: vsv \\\CId\\ for \ \ + assumes fin_cat_length[cat_cs_simps]: "vcard \ = 6\<^sub>\" and fin_cat_finite_semicategory[slicing_intros]: "finite_semicategory \ (cat_smc \)" and fin_cat_CId_vdomain[cat_cs_simps]: "\\<^sub>\ (\\CId\) = \\Obj\" and fin_cat_CId_is_arr[cat_cs_intros]: "a \\<^sub>\ \\Obj\ \ \\CId\\a\ : a \\<^bsub>\\<^esub> a" and fin_cat_CId_left_left[cat_cs_simps]: "f : a \\<^bsub>\\<^esub> b \ \\CId\\b\ \\<^sub>A\<^bsub>\\<^esub> f = f" and fin_cat_CId_right_left[cat_cs_simps]: "f : b \\<^bsub>\\<^esub> c \ f \\<^sub>A\<^bsub>\\<^esub> \\CId\\b\ = f" lemmas [slicing_intros] = finite_category.fin_cat_finite_semicategory text\Rules.\ lemma (in finite_category) fin_category_axioms'[cat_small_cs_intros]: assumes "\' = \" shows "finite_category \' \" unfolding assms by (rule finite_category_axioms) mk_ide rf finite_category_def[unfolded finite_category_axioms_def] |intro finite_categoryI| |dest finite_categoryD[dest]| |elim finite_categoryE[elim]| lemma finite_categoryI': assumes "category \ \" and "vfinite (\\Obj\)" and "vfinite (\\Arr\)" shows "finite_category \ \" proof- interpret category \ \ by (rule assms(1)) show ?thesis proof(intro finite_categoryI) from assms show "finite_semicategory \ (cat_smc \)" by (intro finite_semicategoryI') (auto simp: slicing_simps) qed (auto simp: vfsequence_axioms cat_cs_simps cat_cs_intros) qed lemma finite_categoryI'': assumes "tiny_category \ \" and "vfinite (\\Obj\)" and "vfinite (\\Arr\)" shows "finite_category \ \" using assms by (intro finite_categoryI') (auto intro: cat_small_cs_intros) text\Slicing.\ context finite_category begin interpretation smc: finite_semicategory \ \cat_smc \\ by (rule fin_cat_finite_semicategory) lemmas_with [unfolded slicing_simps]: fin_cat_tiny_semicategory = smc.tiny_semicategory_axioms and fin_smc_Obj_vfinite[cat_small_cs_intros] = smc.fin_smc_Obj_vfinite and fin_smc_Arr_vfinite[cat_small_cs_intros] = smc.fin_smc_Arr_vfinite end text\Elementary properties.\ sublocale finite_category \ tiny_category by (rule tiny_categoryI) ( auto simp: vfsequence_axioms intro: cat_cs_intros cat_cs_simps cat_small_cs_intros finite_category.fin_cat_tiny_semicategory ) lemmas (in finite_category) fin_cat_tiny_category = tiny_category_axioms lemmas [cat_small_cs_intros] = finite_category.fin_cat_tiny_category lemma (in finite_category) fin_cat_in_Vset: "\ \\<^sub>\ Vset \" by (rule tiny_cat_in_Vset) text\Size.\ lemma small_finite_categories[simp]: "small {\. finite_category \ \}" proof(rule down) show "{\. finite_category \ \} \ elts (set {\. tiny_category \ \})" by (auto intro: cat_small_cs_intros) qed lemma small_finite_categories_vsubset_Vset: "set {\. finite_category \ \} \\<^sub>\ Vset \" by (rule vsubsetI) (simp_all add: finite_category.fin_cat_in_Vset) subsubsection\Opposite finite category\ lemma (in finite_category) finite_category_op: "finite_category \ (op_cat \)" by (intro finite_categoryI', unfold cat_op_simps) (auto simp: cat_cs_intros cat_small_cs_intros) lemmas finite_category_op[cat_op_intros] = finite_category.finite_category_op text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/ROOT b/thys/CZH_Foundations/ROOT --- a/thys/CZH_Foundations/ROOT +++ b/thys/CZH_Foundations/ROOT @@ -1,31 +1,31 @@ chapter AFP session CZH_Foundations (AFP) = ZFC_in_HOL + - options [timeout = 1200] + options [timeout = 2400] sessions "HOL-Library" "HOL-ex" Intro_Dest_Elim Conditional_Simplification "HOL-Eisbach" directories czh_introduction czh_sets "czh_sets/ex" czh_digraphs czh_semicategories theories [document = false] CZH_Sets_MIF CZH_Utilities theories CZH_Introduction CZH_Sets_Conclusions CZH_DG_Conclusions CZH_SMC_Conclusions document_files "iman.sty" "extra.sty" "isar.sty" "style.sty" "root.tex" "root.bib" diff --git a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_GRPH.thy b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_GRPH.thy --- a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_GRPH.thy +++ b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_GRPH.thy @@ -1,469 +1,468 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\GRPH\ as a semicategory\ theory CZH_SMC_GRPH imports CZH_DG_Simple CZH_DG_GRPH CZH_SMC_Small_Semicategory begin subsection\Background\ text\ -The methodology for the exposition -of \GRPH\ as a semicategory is analogous to the -one used in the previous chapter -for the exposition of \GRPH\ as a digraph. +The methodology for the exposition of \GRPH\ as a semicategory is analogous +to the one used in the previous chapter for the exposition of \GRPH\ +as a digraph. \ named_theorems smc_GRPH_cs_simps named_theorems smc_GRPH_cs_intros subsection\Definition and elementary properties\ definition smc_GRPH :: "V \ V" where "smc_GRPH \ = [ set {\. digraph \ \}, all_dghms \, (\\\\<^sub>\all_dghms \. \\HomDom\), (\\\\<^sub>\all_dghms \. \\HomCod\), (\\\\\<^sub>\composable_arrs (dg_GRPH \). \\\0\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \\\1\<^sub>\\) ]\<^sub>\" text\Components.\ lemma smc_GRPH_components: shows "smc_GRPH \\Obj\ = set {\. digraph \ \}" and "smc_GRPH \\Arr\ = all_dghms \" and "smc_GRPH \\Dom\ = (\\\\<^sub>\all_dghms \. \\HomDom\)" and "smc_GRPH \\Cod\ = (\\\\<^sub>\all_dghms \. \\HomCod\)" and "smc_GRPH \\Comp\ = (\\\\\<^sub>\composable_arrs (dg_GRPH \). \\\0\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \\\1\<^sub>\\)" unfolding smc_GRPH_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smc_dg_GRPH: "smc_dg (smc_GRPH \) = dg_GRPH \" proof(rule vsv_eqI) show "vsv (smc_dg (smc_GRPH \))" unfolding smc_dg_def by auto show "vsv (dg_GRPH \)" unfolding dg_GRPH_def by auto have dom_lhs: "\\<^sub>\ (smc_dg (smc_GRPH \)) = 4\<^sub>\" unfolding smc_dg_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (dg_GRPH \) = 4\<^sub>\" unfolding dg_GRPH_def by (simp add: nat_omega_simps) show "\\<^sub>\ (smc_dg (smc_GRPH \)) = \\<^sub>\ (dg_GRPH \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (smc_dg (smc_GRPH \)) \ smc_dg (smc_GRPH \)\a\ = dg_GRPH \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold smc_dg_def dg_field_simps smc_GRPH_def dg_GRPH_def ) (auto simp: nat_omega_simps) qed lemmas_with [folded smc_dg_GRPH, unfolded slicing_simps]: smc_GRPH_ObjI = dg_GRPH_ObjI and smc_GRPH_ObjD = dg_GRPH_ObjD and smc_GRPH_ObjE = dg_GRPH_ObjE and smc_GRPH_Obj_iff[smc_GRPH_cs_simps] = dg_GRPH_Obj_iff and smc_GRPH_Dom_app[smc_GRPH_cs_simps] = dg_GRPH_Dom_app and smc_GRPH_Cod_app[smc_GRPH_cs_simps] = dg_GRPH_Cod_app and smc_GRPH_is_arrI = dg_GRPH_is_arrI and smc_GRPH_is_arrD = dg_GRPH_is_arrD and smc_GRPH_is_arrE = dg_GRPH_is_arrE and smc_GRPH_is_arr_iff[smc_GRPH_cs_simps] = dg_GRPH_is_arr_iff subsection\Composable arrows\ lemma smc_GRPH_composable_arrs_dg_GRPH: "composable_arrs (dg_GRPH \) = composable_arrs (smc_GRPH \)" unfolding composable_arrs_def smc_dg_GRPH[symmetric] slicing_simps by auto lemma smc_GRPH_Comp: "smc_GRPH \\Comp\ = (\\\\\<^sub>\composable_arrs (smc_GRPH \). \\\0\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \\\1\<^sub>\\)" unfolding smc_GRPH_components smc_GRPH_composable_arrs_dg_GRPH .. subsection\Composition\ lemma smc_GRPH_Comp_app: assumes "\ : \ \\<^bsub>smc_GRPH \\<^esub> \" and "\ : \ \\<^bsub>smc_GRPH \\<^esub> \" shows "\ \\<^sub>A\<^bsub>smc_GRPH \\<^esub> \ = \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \" proof- from assms have "[\, \]\<^sub>\ \\<^sub>\ composable_arrs (smc_GRPH \)" by (auto intro: smc_cs_intros) then show "\ \\<^sub>A\<^bsub>smc_GRPH \\<^esub> \ = \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \" unfolding smc_GRPH_Comp by (simp add: nat_omega_simps) qed lemma smc_GRPH_Comp_vdomain: "\\<^sub>\ (smc_GRPH \\Comp\) = composable_arrs (smc_GRPH \)" unfolding smc_GRPH_Comp by auto subsection\\GRPH\ is a semicategory\ lemma (in \) tiny_semicategory_smc_GRPH: assumes "\ \" and "\ \\<^sub>\ \" shows "tiny_semicategory \ (smc_GRPH \)" proof(intro tiny_semicategoryI, unfold smc_GRPH_is_arr_iff) show "vfsequence (smc_GRPH \)" unfolding smc_GRPH_def by auto show "vcard (smc_GRPH \) = 5\<^sub>\" unfolding smc_GRPH_def by (simp add: nat_omega_simps) show "(gf \\<^sub>\ \\<^sub>\ (smc_GRPH \\Comp\)) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c \ f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> b)" for gf unfolding smc_GRPH_Comp_vdomain proof show "gf \\<^sub>\ composable_arrs (smc_GRPH \) \ \g f b c a. gf = [g, f]\<^sub>\ \ g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c \ f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> b" by (elim composable_arrsE) (auto simp: smc_GRPH_is_arr_iff) next assume "\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c \ f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> b" with smc_GRPH_is_arr_iff show "gf \\<^sub>\ composable_arrs (smc_GRPH \)" unfolding smc_GRPH_Comp_vdomain by (auto intro: smc_cs_intros) qed show "\ g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c; f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> b \ \ g \\<^sub>A\<^bsub>smc_GRPH \\<^esub> f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c" for g b c f a by (auto simp: smc_GRPH_Comp_app dghm_comp_is_dghm smc_GRPH_cs_simps) fix h c d g b f a assume "h : c \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> d" "g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c" "f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> b" moreover then have "g \\<^sub>D\<^sub>G\<^sub>H\<^sub>M f : a \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> c" "h \\<^sub>D\<^sub>G\<^sub>H\<^sub>M g : b \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> d" by (auto simp: dghm_comp_is_dghm smc_GRPH_cs_simps) ultimately show "h \\<^sub>A\<^bsub>smc_GRPH \\<^esub> g \\<^sub>A\<^bsub>smc_GRPH \\<^esub> f = h \\<^sub>A\<^bsub>smc_GRPH \\<^esub> (g \\<^sub>A\<^bsub>smc_GRPH \\<^esub> f)" by (simp add: smc_GRPH_is_arr_iff smc_GRPH_Comp_app dghm_comp_assoc) qed (simp_all add: assms smc_dg_GRPH tiny_digraph_dg_GRPH smc_GRPH_components) subsection\Initial object\ lemma (in \) smc_GRPH_obj_initialI: "obj_initial (smc_GRPH \) dg_0" unfolding obj_initial_def proof ( intro obj_terminalI, unfold smc_op_simps smc_GRPH_is_arr_iff smc_GRPH_Obj_iff ) show "digraph \ dg_0" by (intro digraph_dg_0) fix \ assume "digraph \ \" then interpret digraph \ \ . show "\!f. f : dg_0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof show dghm_0: "dghm_0 \ : dg_0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" by (simp add: dghm_0_is_dghm digraph_axioms is_ft_dghm.axioms(1)) fix \ assume prems: "\ : dg_0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" then interpret \: is_dghm \ dg_0 \ \ . show "\ = dghm_0 \" proof(rule dghm_eqI) from dghm_0 show "dghm_0 \ : dg_0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" unfolding smc_GRPH_is_arr_iff by simp have [simp]: "\\<^sub>\ (\\ObjMap\) = 0" by (simp add: dg_cs_simps dg_0_components) with \.ObjMap.vdomain_vrange_is_vempty show "\\ObjMap\ = dghm_0 \\ObjMap\" by ( auto intro: \.ObjMap.vsv_vrange_vempty simp: dg_0_components dghm_0_components ) from \.dghm_ObjMap_vdomain have "\\<^sub>\ (\\ArrMap\) = 0" by ( auto simp: \.dghm_ArrMap_vdomain intro: \.HomDom.dg_Arr_vempty_if_Obj_vempty ) then show "\\ArrMap\ = dghm_0 \\ArrMap\" by ( metis \.ArrMap.vsv_axioms dghm_0_components(2) vsv.vdomain_vrange_is_vempty vsv.vsv_vrange_vempty ) qed (auto simp: dghm_0_components prems) qed qed lemma (in \) smc_GRPH_obj_initialD: assumes "obj_initial (smc_GRPH \) \" shows "\ = dg_0" using assms unfolding obj_initial_def proof ( elim obj_terminalE, unfold smc_op_simps smc_GRPH_is_arr_iff smc_GRPH_Obj_iff ) assume prems: "digraph \ \" "digraph \ \ \ \!\. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" for \ from prems(2)[OF digraph_dg_0] obtain \ where \: "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_0" by meson interpret \: is_dghm \ \ dg_0 \ by (rule \) have "\\<^sub>\ (\\ObjMap\) \\<^sub>\ 0" unfolding dg_0_components(1)[symmetric] by (simp add: \.dghm_ObjMap_vrange) then have "\\ObjMap\ = 0" by (auto intro: \.ObjMap.vsv_vrange_vempty) with \.dghm_ObjMap_vdomain have Obj[simp]: "\\Obj\ = 0" by auto have "\\<^sub>\ (\\ArrMap\) \\<^sub>\ 0" unfolding dg_0_components(2)[symmetric] by (simp add: \.dghm_ArrMap_vrange) then have "\\ArrMap\ = 0" by (auto intro: \.ArrMap.vsv_vrange_vempty) with \.dghm_ArrMap_vdomain have Arr[simp]: "\\Arr\ = 0" by auto from Arr \.HomDom.dg_Dom_vempty_if_Arr_vempty have [simp]: "\\Dom\ = []\<^sub>\" by auto from Arr \.HomDom.dg_Cod_vempty_if_Arr_vempty have [simp]: "\\Cod\ = []\<^sub>\" by auto show "\ = dg_0" by (rule dg_eqI[of \]) (simp_all add: prems(1) dg_0_components digraph_dg_0) qed lemma (in \) smc_GRPH_obj_initialE: assumes "obj_initial (smc_GRPH \) \" obtains "\ = dg_0" using assms by (auto dest: smc_GRPH_obj_initialD) lemma (in \) smc_GRPH_obj_initial_iff[smc_GRPH_cs_simps]: "obj_initial (smc_GRPH \) \ \ \ = dg_0" using smc_GRPH_obj_initialI smc_GRPH_obj_initialD by auto subsection\Terminal object\ lemma (in \) smc_GRPH_obj_terminalI[smc_GRPH_cs_intros]: assumes "a \\<^sub>\ Vset \" and "f \\<^sub>\ Vset \" shows "obj_terminal (smc_GRPH \) (dg_1 a f)" proof ( intro obj_terminalI, unfold smc_op_simps smc_GRPH_is_arr_iff smc_GRPH_Obj_iff ) fix \ assume "digraph \ \" then interpret digraph \ \ . show "\!\'. \' : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_1 a f" proof show dghm_1: "dghm_const \ (dg_1 a f) a f : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_1 a f" by ( auto simp: assms dg_1_is_arr_iff dghm_const_is_dghm digraph_axioms' digraph_dg_1 ) fix \' assume prems: "\' : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_1 a f" then interpret \': is_dghm \ \ \dg_1 a f\ \' . show "\' = dghm_const \ (dg_1 a f) a f" proof(rule dghm_eqI, unfold dghm_const_components) show "dghm_const \ (dg_1 a f) a f : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_1 a f" by (rule dghm_1) show "\'\ObjMap\ = vconst_on (\\Obj\) a" proof(cases\\\Obj\ = 0\) case True then have "\'\ObjMap\ = 0" by ( simp add: \'.ObjMap.vdomain_vrange_is_vempty \'.dghm_ObjMap_vsv vsv.vsv_vrange_vempty ) with True show ?thesis by simp next case False then have "\\<^sub>\ (\'\ObjMap\) \ 0" by (auto simp: \'.dghm_ObjMap_vdomain) with False have "\\<^sub>\ (\'\ObjMap\) \ 0" by fastforce moreover from \'.dghm_ObjMap_vrange have "\\<^sub>\ (\'\ObjMap\) \\<^sub>\ set {a}" by (simp add: dg_1_components) ultimately have "\\<^sub>\ (\'\ObjMap\) = set {a}" by auto with \'.dghm_ObjMap_vdomain show ?thesis by (intro vsv.vsv_is_vconst_onI) blast+ qed show "\'\ArrMap\ = vconst_on (\\Arr\) f" proof(cases\\\Arr\ = 0\) case True then have "\'\ArrMap\ = 0" by ( simp add: \'.ArrMap.vdomain_vrange_is_vempty \'.dghm_ArrMap_vsv vsv.vsv_vrange_vempty ) with True show ?thesis by simp next case False then have "\\<^sub>\ (\'\ArrMap\) \ 0" by (auto simp: \'.dghm_ArrMap_vdomain) with False have "\\<^sub>\ (\'\ArrMap\) \ 0" by (force simp: \'.ArrMap.vdomain_vrange_is_vempty) moreover from \'.dghm_ArrMap_vrange have "\\<^sub>\ (\'\ArrMap\) \\<^sub>\ set {f}" by (simp add: dg_1_components) ultimately have "\\<^sub>\ (\'\ArrMap\) = set {f}" by auto then show ?thesis by (intro vsv.vsv_is_vconst_onI) (auto simp: \'.dghm_ArrMap_vdomain) qed qed (auto intro: prems) qed qed (simp add: assms digraph_dg_1) lemma (in \) smc_GRPH_obj_terminalE: assumes "obj_terminal (smc_GRPH \) \" obtains a f where "a \\<^sub>\ Vset \" and "f \\<^sub>\ Vset \" and "\ = dg_1 a f" using assms proof ( elim obj_terminalE; unfold smc_op_simps smc_GRPH_is_arr_iff smc_GRPH_Obj_iff ) assume prems: "digraph \ \" "digraph \ \ \ \!\. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" for \ then interpret \: digraph \ \ by simp obtain a where \_Obj: "\\Obj\ = set {a}" and a: "a \\<^sub>\ Vset \" proof- have dg_10: "digraph \ (dg_10 0)" by (rule digraph_dg_10) auto from prems(2)[OF dg_10] obtain \ where \: "\ : dg_10 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and \\: "\ : dg_10 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \ \ \ = \" for \ by fastforce interpret \: is_dghm \ \dg_10 0\ \ \ by (rule \) have "\\<^sub>\ (\\ObjMap\) = set {0}" by (simp add: dg_cs_simps dg_10_components) then obtain a where vrange_\[simp]: "\\<^sub>\ (\\ObjMap\) = set {a}" by ( auto simp: dg_cs_simps intro: \.ObjMap.vsv_vdomain_vsingleton_vrange_vsingleton ) with \.dg_Obj_vsubset_Vset \.dghm_ObjMap_vrange have [simp]: "a \\<^sub>\ Vset \" by auto from \.dghm_ObjMap_vrange have "set {a} \\<^sub>\ \\Obj\" by simp moreover have "\\Obj\ \\<^sub>\ set {a}" proof(rule ccontr) assume "\\\Obj\ \\<^sub>\ set {a}" then obtain b where ba: "b \ a" and b: "b \\<^sub>\ \\Obj\" by force define \ where "\ = [set {\0, b\}, 0, dg_10 0, \]\<^sub>\" have \_components: "\\ObjMap\ = set {\0, b\}" "\\ArrMap\ = 0" "\\HomDom\ = dg_10 0" "\\HomCod\ = \" unfolding \_def dghm_field_simps by (simp_all add: nat_omega_simps) have \: "\ : dg_10 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" by (rule is_dghmI, unfold \_components dg_10_components) ( auto simp: dg_cs_intros nat_omega_simps digraph_dg_10 \_def dg_10_is_arr_iff b vsubset_vsingleton_leftI ) then have \_def: "\ = \" by (rule \\) have "\\<^sub>\ (\\ObjMap\) = set {b}" unfolding \_components by simp with vrange_\ ba show False unfolding \_def by simp qed ultimately have "\\Obj\ = set {a}" by simp with that show ?thesis by simp qed obtain f where \_Arr: "\\Arr\ = set {f}" and f: "f \\<^sub>\ Vset \" proof- from prems(2)[OF digraph_dg_1, of 0 0] obtain \ where \: "\ : dg_1 0 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and \\: "\ : dg_1 0 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \ \ \ = \" for \ by fastforce interpret \: is_dghm \ \dg_1 0 0\ \ \ by (rule \) have "\\<^sub>\ (\\ObjMap\) = set {0}" by (simp add: dg_cs_simps dg_1_components) then obtain a' where "\\<^sub>\ (\\ObjMap\) = set {a'}" by ( auto simp: dg_cs_simps intro: \.ObjMap.vsv_vdomain_vsingleton_vrange_vsingleton ) with \_Obj \.dghm_ObjMap_vrange have "\\<^sub>\ (\\ObjMap\) = set {a}" by auto have "\\<^sub>\ (\\ArrMap\) = set {0}" by (simp add: dg_cs_simps dg_1_components) then obtain f where vrange_\[simp]: "\\<^sub>\ (\\ArrMap\) = set {f}" by ( auto simp: dg_cs_simps intro: \.ArrMap.vsv_vdomain_vsingleton_vrange_vsingleton ) with \.dg_Arr_vsubset_Vset \.dghm_ArrMap_vrange have [simp]: "f \\<^sub>\ Vset \" by auto from \.dghm_ArrMap_vrange have "set {f} \\<^sub>\ \\Arr\" by simp moreover have "\\Arr\ \\<^sub>\ set {f}" proof(rule ccontr) assume "\\\Arr\ \\<^sub>\ set {f}" then obtain g where gf: "g \ f" and g: "g \\<^sub>\ \\Arr\" by force have g: "g : a \\<^bsub>\\<^esub> a" proof(intro is_arrI) from g \_Obj show "\\Dom\\g\ = a" by (metis \.dg_is_arrD(2) is_arr_def vsingleton_iff) from g \_Obj show "\\Cod\\g\ = a" by (metis \.dg_is_arrD(3) is_arr_def vsingleton_iff) qed (auto simp: g) define \ where "\ = [set {\0, a\}, set {\0, g\}, dg_1 0 0, \]\<^sub>\" have \_components: "\\ObjMap\ = set {\0, a\}" "\\ArrMap\ = set {\0, g\}" "\\HomDom\ = dg_1 0 0" "\\HomCod\ = \" unfolding \_def dghm_field_simps by (simp_all add: nat_omega_simps) have \: "\ : dg_1 0 0 \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" by (rule is_dghmI, unfold \_components dg_1_components) ( auto simp: dg_cs_intros nat_omega_simps \_def dg_1_is_arr_iff \_Obj g ) then have \_def: "\ = \" by (rule \\) have "\\<^sub>\ (\\ArrMap\) = set {g}" unfolding \_components by simp with vrange_\ gf show False unfolding \_def by simp qed ultimately have "\\Arr\ = set {f}" by simp with that show ?thesis by simp qed have "\ = dg_1 a f" proof(rule dg_eqI[of \], unfold dg_1_components) show "\\Obj\ = set {a}" by (simp add: \_Obj) moreover show "\\Arr\ = set {f}" by (simp add: \_Arr) ultimately have "\\Dom\\f\ = a" "\\Cod\\f\ = a" by (metis \.dg_is_arrE is_arr_def vsingleton_iff)+ have "\\<^sub>\ (\\Dom\) = set {f}" by (simp add: dg_cs_simps \_Arr) moreover from \.Dom.vsv_vrange_vempty \.dg_Dom_vdomain \.dg_Dom_vrange have "\\<^sub>\ (\\Dom\) = set {a}" by (fastforce simp: \_Arr \_Obj) ultimately show "\\Dom\ = set {\f, a\}" using \.Dom.vsv_vdomain_vrange_vsingleton by simp have "\\<^sub>\ (\\Cod\) = set {f}" by (simp add: dg_cs_simps \_Arr) moreover from \.Cod.vsv_vrange_vempty \.dg_Cod_vdomain \.dg_Cod_vrange have "\\<^sub>\ (\\Cod\) = set {a}" by (fastforce simp: \_Arr \_Obj) ultimately show "\\Cod\ = set {\f, a\}" using assms \.Cod.vsv_vdomain_vrange_vsingleton by simp qed (auto simp: dg_cs_intros \_Obj digraph_dg_1 a f) with a f that show ?thesis by auto qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Introduction.thy b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Introduction.thy --- a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Introduction.thy +++ b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Introduction.thy @@ -1,43 +1,43 @@ (* Copyright 2021 (C) Mihails Milehins *) chapter\Semicategories\ section\Introduction\ theory CZH_SMC_Introduction imports CZH_DG_Introduction begin subsection\Background\ text\ Many concepts that are normally associated with category theory can be generalized to semicategories. It is the goal of -this chapter to expose these generalized concepts and provide the -relevant foundations for the development of the notion of a category -in the next chapter. +this chapter to expose these generalized concepts and provide a +foundation for the development of the notion of a category +in \cite{milehins_category_2021}. \ subsection\Preliminaries\ named_theorems smc_op_simps named_theorems smc_op_intros named_theorems smc_cs_simps named_theorems smc_cs_intros named_theorems smc_arrow_cs_intros subsection\CS setup for foundations\ lemmas (in \) [smc_cs_intros] = \_\ text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Small_Semicategory.thy b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Small_Semicategory.thy --- a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Small_Semicategory.thy +++ b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Small_Semicategory.thy @@ -1,322 +1,322 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Smallness for semicategories\ theory CZH_SMC_Small_Semicategory imports CZH_DG_Small_Digraph CZH_SMC_Semicategory begin subsection\Background\ text\ An explanation of the methodology chosen for the exposition of all matters related to the size of the semicategories and associated entities is given in the previous chapter. \ named_theorems smc_small_cs_simps named_theorems smc_small_cs_intros subsection\Tiny semicategory\ subsubsection\Definition and elementary properties\ locale tiny_semicategory = \ \ + vfsequence \ + Comp: vsv \\\Comp\\ for \ \ + assumes tiny_smc_length[smc_cs_simps]: "vcard \ = 5\<^sub>\" and tiny_smc_tiny_digraph[slicing_intros]: "tiny_digraph \ (smc_dg \)" and tiny_smc_Comp_vdomain: "gf \\<^sub>\ \\<^sub>\ (\\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>\\<^esub> c \ f : a \\<^bsub>\\<^esub> b)" and tiny_smc_Comp_is_arr[smc_cs_intros]: "\ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" and tiny_smc_assoc[smc_cs_simps]: "\ h : c \\<^bsub>\\<^esub> d; g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ (h \\<^sub>A\<^bsub>\\<^esub> g) \\<^sub>A\<^bsub>\\<^esub> f = h \\<^sub>A\<^bsub>\\<^esub> (g \\<^sub>A\<^bsub>\\<^esub> f)" lemmas [smc_cs_simps] = tiny_semicategory.tiny_smc_length tiny_semicategory.tiny_smc_assoc lemmas [slicing_intros] = tiny_semicategory.tiny_smc_Comp_is_arr text\Rules.\ lemma (in tiny_semicategory) tiny_semicategory_axioms'[smc_small_cs_intros]: assumes "\' = \" shows "tiny_semicategory \' \" unfolding assms by (rule tiny_semicategory_axioms) mk_ide rf tiny_semicategory_def[unfolded tiny_semicategory_axioms_def] |intro tiny_semicategoryI| |dest tiny_semicategoryD[dest]| |elim tiny_semicategoryE[elim]| lemma tiny_semicategoryI': assumes "semicategory \ \" and "\\Obj\ \\<^sub>\ Vset \" and "\\Arr\ \\<^sub>\ Vset \" shows "tiny_semicategory \ \" proof- interpret semicategory \ \ by (rule assms(1)) show ?thesis proof(intro tiny_semicategoryI) show "vfsequence \" by (simp add: vfsequence_axioms) from assms show "tiny_digraph \ (smc_dg \)" by (intro tiny_digraphI') (auto simp: slicing_simps) qed (auto simp: smc_cs_simps intro: smc_cs_intros) qed lemma tiny_semicategoryI'': assumes "\ \" and "vfsequence \" and "vsv (\\Comp\)" and "vcard \ = 5\<^sub>\" and "vsv (\\Dom\)" and "vsv (\\Cod\)" and "\\<^sub>\ (\\Dom\) = \\Arr\" and "\\<^sub>\ (\\Dom\) \\<^sub>\ \\Obj\" and "\\<^sub>\ (\\Cod\) = \\Arr\" and "\\<^sub>\ (\\Cod\) \\<^sub>\ \\Obj\" and "\gf. gf \\<^sub>\ \\<^sub>\ (\\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>\\<^esub> c \ f : a \\<^bsub>\\<^esub> b)" and "\b c g a f. \ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" and "\c d h b g a f. \ h : c \\<^bsub>\\<^esub> d; g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ (h \\<^sub>A\<^bsub>\\<^esub> g) \\<^sub>A\<^bsub>\\<^esub> f = h \\<^sub>A\<^bsub>\\<^esub> (g \\<^sub>A\<^bsub>\\<^esub> f)" and "\\Obj\ \\<^sub>\ Vset \" and "\\Arr\ \\<^sub>\ Vset \" shows "tiny_semicategory \ \" by (intro tiny_semicategoryI tiny_digraphI, unfold slicing_simps) (simp_all add: smc_dg_def nat_omega_simps assms) text\Slicing.\ context tiny_semicategory begin interpretation dg: tiny_digraph \ \smc_dg \\ by (rule tiny_smc_tiny_digraph) lemmas_with [unfolded slicing_simps]: tiny_smc_Obj_in_Vset[smc_small_cs_intros] = dg.tiny_dg_Obj_in_Vset and tiny_smc_Arr_in_Vset[smc_small_cs_intros] = dg.tiny_dg_Arr_in_Vset and tiny_smc_Dom_in_Vset[smc_small_cs_intros] = dg.tiny_dg_Dom_in_Vset and tiny_smc_Cod_in_Vset[smc_small_cs_intros] = dg.tiny_dg_Cod_in_Vset end text\Elementary properties.\ sublocale tiny_semicategory \ semicategory by (rule semicategoryI) ( auto simp: vfsequence_axioms tiny_digraph.tiny_dg_digraph tiny_smc_tiny_digraph tiny_smc_Comp_vdomain intro: smc_cs_intros smc_cs_simps ) -lemmas (in tiny_semicategory) tiny_dg_semicategory = semicategory_axioms +lemmas (in tiny_semicategory) tiny_smc_semicategory = semicategory_axioms -lemmas [smc_small_cs_intros] = tiny_semicategory.tiny_dg_semicategory +lemmas [smc_small_cs_intros] = tiny_semicategory.tiny_smc_semicategory text\Size.\ lemma (in tiny_semicategory) tiny_smc_Comp_in_Vset: "\\Comp\ \\<^sub>\ Vset \" proof- have "\\Arr\ \\<^sub>\ Vset \" by (simp add: tiny_smc_Arr_in_Vset) with Axiom_of_Infinity have "\\Arr\ ^\<^sub>\ 2\<^sub>\ \\<^sub>\ Vset \" by (intro Limit_vcpower_in_VsetI) auto with Comp.pnop_vdomain have D: "\\<^sub>\ (\\Comp\) \\<^sub>\ Vset \" by auto moreover from tiny_smc_Arr_in_Vset smc_Comp_vrange have "\\<^sub>\ (\\Comp\) \\<^sub>\ Vset \" by auto ultimately show ?thesis by (simp add: Comp.vbrelation_Limit_in_VsetI) qed lemma (in tiny_semicategory) tiny_smc_in_Vset: "\ \\<^sub>\ Vset \" proof- note [smc_cs_intros] = tiny_smc_Obj_in_Vset tiny_smc_Arr_in_Vset tiny_smc_Dom_in_Vset tiny_smc_Cod_in_Vset tiny_smc_Comp_in_Vset show ?thesis by (subst smc_def) (cs_concl cs_intro: smc_cs_intros V_cs_intros) qed lemma small_tiny_semicategories[simp]: "small {\. tiny_semicategory \ \}" proof(rule down) show "{\. tiny_semicategory \ \} \ elts (set {\. semicategory \ \})" by (auto intro: smc_small_cs_intros) qed lemma tiny_semicategories_vsubset_Vset: "set {\. tiny_semicategory \ \} \\<^sub>\ Vset \" by (rule vsubsetI) (simp add: tiny_semicategory.tiny_smc_in_Vset) lemma (in semicategory) smc_tiny_semicategory_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "tiny_semicategory \ \" proof(intro tiny_semicategoryI) show "tiny_digraph \ (smc_dg \)" by (rule digraph.dg_tiny_digraph_if_ge_Limit, rule smc_digraph; intro assms) qed ( auto simp: assms(1) smc_cs_simps smc_cs_intros smc_digraph digraph.dg_tiny_digraph_if_ge_Limit smc_Comp_vdomain vfsequence_axioms ) subsubsection\Opposite tiny semicategory\ lemma (in tiny_semicategory) tiny_semicategory_op: "tiny_semicategory \ (op_smc \)" by (intro tiny_semicategoryI', unfold smc_op_simps) (auto simp: smc_op_intros smc_small_cs_intros) lemmas tiny_semicategory_op[smc_op_intros] = tiny_semicategory.tiny_semicategory_op subsection\Finite semicategory\ subsubsection\Definition and elementary properties\ text\ A finite semicategory is a generalization of the concept of a finite category, as presented in nLab \cite{noauthor_nlab_nodate} \footnote{\url{https://ncatlab.org/nlab/show/finite+category}}. \ locale finite_semicategory = \ \ + vfsequence \ + Comp: vsv \\\Comp\\ for \ \ + assumes fin_smc_length[smc_cs_simps]: "vcard \ = 5\<^sub>\" and fin_smc_finite_digraph[slicing_intros]: "finite_digraph \ (smc_dg \)" and fin_smc_Comp_vdomain: "gf \\<^sub>\ \\<^sub>\ (\\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>\\<^esub> c \ f : a \\<^bsub>\\<^esub> b)" and fin_smc_Comp_is_arr[smc_cs_intros]: "\ g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" and fin_smc_assoc[smc_cs_simps]: "\ h : c \\<^bsub>\\<^esub> d; g : b \\<^bsub>\\<^esub> c; f : a \\<^bsub>\\<^esub> b \ \ (h \\<^sub>A\<^bsub>\\<^esub> g) \\<^sub>A\<^bsub>\\<^esub> f = h \\<^sub>A\<^bsub>\\<^esub> (g \\<^sub>A\<^bsub>\\<^esub> f)" lemmas [smc_cs_simps] = finite_semicategory.fin_smc_length finite_semicategory.fin_smc_assoc lemmas [slicing_intros] = finite_semicategory.fin_smc_Comp_is_arr text\Rules.\ lemma (in finite_semicategory) finite_semicategory_axioms'[smc_small_cs_intros]: assumes "\' = \" shows "finite_semicategory \' \" unfolding assms by (rule finite_semicategory_axioms) mk_ide rf finite_semicategory_def[unfolded finite_semicategory_axioms_def] |intro finite_semicategoryI| |dest finite_semicategoryD[dest]| |elim finite_semicategoryE[elim]| lemma finite_semicategoryI': assumes "semicategory \ \" and "vfinite (\\Obj\)" and "vfinite (\\Arr\)" shows "finite_semicategory \ \" proof- interpret semicategory \ \ by (rule assms(1)) show ?thesis proof(intro finite_semicategoryI) show "vfsequence \" by (simp add: vfsequence_axioms) from assms show "finite_digraph \ (smc_dg \)" by (intro finite_digraphI) (auto simp: slicing_simps) qed (auto simp: smc_cs_simps intro: smc_cs_intros) qed lemma finite_semicategoryI'': assumes "tiny_semicategory \ \" and "vfinite (\\Obj\)" and "vfinite (\\Arr\)" shows "finite_semicategory \ \" using assms by (intro finite_semicategoryI') (auto intro: smc_cs_intros smc_small_cs_intros) text\Slicing.\ context finite_semicategory begin interpretation dg: finite_digraph \ \smc_dg \\ by (rule fin_smc_finite_digraph) lemmas_with [unfolded slicing_simps]: fin_smc_Obj_vfinite[smc_small_cs_intros] = dg.fin_dg_Obj_vfinite and fin_smc_Arr_vfinite[smc_small_cs_intros] = dg.fin_dg_Arr_vfinite end text\Elementary properties.\ sublocale finite_semicategory \ tiny_semicategory by (rule tiny_semicategoryI) ( auto simp: vfsequence_axioms fin_smc_Comp_vdomain fin_smc_finite_digraph finite_digraph.fin_dg_tiny_digraph intro: smc_cs_intros smc_cs_simps ) lemmas (in finite_semicategory) fin_smc_tiny_semicategory = tiny_semicategory_axioms lemmas [smc_small_cs_intros] = finite_semicategory.fin_smc_tiny_semicategory lemma (in finite_semicategory) fin_smc_in_Vset: "\ \\<^sub>\ Vset \" by (rule tiny_smc_in_Vset) text\Size.\ lemma small_finite_semicategories[simp]: "small {\. finite_semicategory \ \}" proof(rule down) show "{\. finite_semicategory \ \} \ elts (set {\. semicategory \ \})" by (auto intro: smc_small_cs_intros) qed lemma finite_semicategories_vsubset_Vset: "set {\. finite_semicategory \ \} \\<^sub>\ Vset \" by (rule vsubsetI) (simp add: finite_semicategory.fin_smc_in_Vset) subsubsection\Opposite finite semicategory\ lemma (in finite_semicategory) finite_semicategory_op: "finite_semicategory \ (op_smc \)" by (intro finite_semicategoryI', unfold smc_op_simps) (auto simp: smc_op_intros smc_small_cs_intros) lemmas finite_semicategory_op[smc_op_intros] = finite_semicategory.finite_semicategory_op text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/document/root.bib b/thys/CZH_Foundations/document/root.bib --- a/thys/CZH_Foundations/document/root.bib +++ b/thys/CZH_Foundations/document/root.bib @@ -1,510 +1,515 @@ @book{bloch_real_2010, address = {Heidelberg}, title = {The {Real} {Numbers} and {Real} {Analysis}}, isbn = {978-0-387-72176-7}, publisher = {Springer Science + Business Media}, author = {Bloch, Ethan D.}, year = {2010}, } @book{takeuti_introduction_1971, address = {Heidelberg}, title = {Introduction to {Axiomatic} {Set} {Theory}}, isbn = {0-387-05302-6}, publisher = {Springer-Verlag}, author = {Takeuti, Gaisi and Zaring, Wilson M.}, year = {1971}, } @book{bourbaki_elements_1970, title = {Elements of {Mathematics}, {Theory} of {Sets}}, isbn = {978-3-540-22525-6}, publisher = {Originally published as {\'E}l{\'e}ments de Math{\'e}matique Th{\'e}orie des Ensembles; Paris: N. Bourbaki. Reprint, Heidelberg: Springer-Verlag, 2004.}, author = {Bourbaki, Nicolas}, year = {1970}, } @book{hungerford_algebra_2003, address = {New York}, title = {Algebra}, isbn = {978-0-387-90518-1}, publisher = {Springer}, author = {Hungerford, Thomas W.}, year = {2003}, } @book{kelley_general_1955, title = {General {Topology}}, isbn = {978-0-486-81544-2}, publisher = {Originally published as General Topology; New York: Van Nostrand Reinhold Company. Reprint, New York: Dover Publications, 2017.}, author = {Kelley, John L.}, year = {1955}, } @book{adamek_abstract_2006, title = {Abstract and {Concrete} {Categories} - {The} {Joy} of {Cats}}, author = {Adamek, Jiri and Herrlich, Horst and Strecker, George}, year = {2006}, } @misc{noauthor_wikipedia_2001, title = {Wikipedia}, url = {https://www.wikipedia.org/}, year = {2001}, } @misc{noauthor_encyclopedia_nodate, title = {Encyclopedia of {Mathematics}}, url = {https://www.encyclopediaofmath.org/index.php/Main_Page}, } @misc{noauthor_proofwiki_nodate, title = {{ProofWiki}}, url = {https://proofwiki.org/wiki/Main_Page}, } @article{paulson_hereditarily_2013, title = {The {Hereditarily} {Finite} {Sets}}, journal = {Archive of Formal Proofs}, author = {Paulson, Lawrence C.}, year = {2013}, } @book{mac_lane_categories_2010, address = {New York}, edition = {2}, series = {Graduate {Texts} in {Mathematics}}, title = {Categories for the {Working} {Mathematician}}, isbn = {978-1-4419-3123-8}, number = {5}, publisher = {Springer}, author = {Mac Lane, Saunders}, year = {2010}, } @article{stark_category_2016, title = {Category {Theory} with {Adjunctions} and {Limits}}, journal = {Archive of Formal Proofs}, author = {Stark, Eugene W.}, year = {2016}, } @article{paulson_zermelo_2019, title = {Zermelo {Fraenkel} {Set} {Theory} in {Higher}-{Order} {Logic}}, journal = {Archive of Formal Proofs}, author = {Paulson, Lawrence C.}, year = {2019}, } @article{katovsky_category_2010, title = {Category {Theory}}, journal = {Archive of Formal Proofs}, author = {Katovsky, Alexander}, year = {2010}, } @article{okeefe_category_2005, title = {Category {Theory} to {Yoneda}'s {Lemma}}, journal = {Archive of Formal Proofs}, author = {O'Keefe, Greg}, year = {2005}, } @misc{noauthor_nlab_nodate, title = {{nLab}}, url = {https://ncatlab.org/nlab/show/HomePage}, - journal = {nLab}, } @book{bodo_categories_1970, address = {New York}, title = {Categories and {Functors}}, publisher = {Academic Press}, author = {Bodo, Pareigis}, year = {1970}, } @article{mitchell_dominion_1972, title = {The {Dominion} of {Isbell}}, volume = {167}, journal = {Transactions of the American Mathematical Society}, author = {Mitchell, Barry}, year = {1972}, } @inproceedings{brown_higher-order_2019, address = {Portland, USA}, title = {Higher-{Order} {Tarski} {Grothendieck} as a {Foundation} for {Formal} {Proof}}, booktitle = {10th {International} {Conference} on {Interactive} {Theorem} {Proving} ({ITP} 2019)}, author = {Brown, Chad E. and Kaliszyk, Cezary and Pak, Karol}, editor = {Harrison, John and O'Leary, John and Tolmach, Andrew}, year = {2019}, keywords = {000 Computer science, knowledge, general works, Computer Science}, pages = {9:1--9:16}, } @inproceedings{feferman_set-theoretical_1969, address = {Heidelberg}, series = {Lecture {Notes} in {Mathematics}}, title = {Set-{Theoretical} {Foundations} of {Category} {Theory}}, booktitle = {Reports of the {Midwest} {Category} {Seminar} {III}}, publisher = {Springer}, author = {Feferman, Solomon and Kreisel, Georg}, editor = {Barr, M. and Berthiaume, P. and Day, B. J. and Duskin, J. and Feferman, S. and Kelly, G. M. and Mac Lane, S. and Tierney, M. and Walters, R. F. C.}, year = {1969}, keywords = {Abelian Group, Category Theory, Closure Condition, Free Variable, Mathematical Practice}, pages = {201--247}, } @article{shulman_set_2008, title = {Set {Theory} for {Category} {Theory}}, url = {http://arxiv.org/abs/0810.1279}, journal = {arXiv:0810.1279 [math]}, author = {Shulman, Michael A.}, year = {2008}, keywords = {Mathematics - Category Theory, Mathematics - Logic}, } @incollection{barkaoui_partizan_2006, address = {Berlin}, title = {Partizan {Games} in {Isabelle}/{HOLZF}}, volume = {4281}, isbn = {978-3-540-48815-6}, booktitle = {{ICTAC} 2006}, publisher = {Springer}, author = {Obua, Steven}, editor = {Barkaoui, Kamel and Cavalcanti, Ana and Cerone, Antonio}, year = {2006}, pages = {272--286}, } @book{jech_set_2006, address = {Heidelberg}, edition = {3}, series = {Pure and applied mathematics}, title = {Set {Theory}}, isbn = {3-540-44085-2}, number = {79}, publisher = {Springer}, author = {Jech, Thomas}, year = {2006}, keywords = {Set theory}, } @incollection{sica_doing_2006, title = {On {Doing} {Category} {Theory} within {Set} {Theoretic} {Foundations}}, isbn = {978-88-7699-031-1}, booktitle = {What is {Category} {Theory}?}, publisher = {Polimetrica s.a.s.}, author = {Rao, Vidhyanath K.}, editor = {Sica, Giandomenico}, year = {2006}, keywords = {Mathematics / General}, } @unpublished{ballarin_tutorial_2020, title = {Tutorial to {Locales} and {Locale} {Interpretation}}, url = {https://isabelle.in.tum.de/website-Isabelle2020/dist/Isabelle2020/doc/locales.pdf}, author = {Ballarin, Clemens}, year = {2020}, } @incollection{berardi_locales_2004, address = {Heidelberg}, title = {Locales and {Locale} {Expressions} in {Isabelle}/{Isar}}, volume = {3085}, isbn = {978-3-540-22164-7}, booktitle = {Types for {Proofs} and {Programs}}, publisher = {Springer}, author = {Ballarin, Clemens}, editor = {Berardi, Stefano and Coppo, Mario and Damiani, Ferruccio}, year = {2004}, pages = {34--50}, } @misc{noauthor_isabellehol_2020, title = {Isabelle/{HOL} {Standard} {Library}}, url = {https://isabelle.in.tum.de/website-Isabelle2020/dist/library/HOL/HOL/index.html}, journal = {Isabelle/HOL}, year = {2020}, } @misc{eberl_syntax_2021, title = {Syntax proposal: multiway if}, url = {https://lists.cam.ac.uk/pipermail/cl-isabelle-users/2021-February/msg00034.html}, journal = {The Cl-isabelle-users Archives}, author = {Eberl, Manuel}, year = {2021}, } @misc{noauthor_association_nodate, title = {Association of {Mizar} {Users}. {Mizar} home page.}, url = {http://mizar.org/}, } @article{bylinski_introduction_1990, title = {Introduction to {Categories} and {Functors}}, volume = {1}, number = {2}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1990}, pages = {409--420}, } @article{bylinski_subcategories_1990, title = {Subcategories and {Products} of {Categories}}, volume = {1}, number = {4}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1990}, pages = {725--732}, } @article{bylinski_opposite_1991, title = {Opposite {Categories} and {Contravariant} {Functors}}, volume = {2}, number = {3}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1991}, pages = {419--424}, } @article{trybulec_natural_1991, title = {Natural {Transformations}. {Discrete} {Categories}}, volume = {2}, number = {4}, journal = {Formalized Mathematics}, author = {Trybulec, Andrzej}, year = {1991}, pages = {467--474}, } @article{bylinski_category_1991, title = {Category {Ens}}, volume = {2}, number = {4}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1991}, pages = {527--533}, } @article{muzalewski_categories_1991, title = {Categories of {Groups}}, volume = {2}, number = {4}, journal = {Formalized Mathematics}, author = {Muzalewski, Micha{\l }}, year = {1991}, pages = {563--571}, } @article{trybulec_isomorphisms_1991, title = {Isomorphisms of {Categories}}, volume = {2}, number = {5}, journal = {Formalized Mathematics}, author = {Trybulec, Andrzej}, year = {1991}, pages = {629--634}, } @article{muzalewski_category_1991, title = {Category of {Rings}}, volume = {2}, number = {5}, journal = {Formalized Mathematics}, author = {Muzalewski, Micha{\l }}, year = {1991}, pages = {643--648}, } @article{muzalewski_category_1991-1, title = {Category of {Left} {Modules}}, volume = {2}, number = {5}, journal = {Formalized Mathematics}, author = {Muzalewski, Micha{\l }}, year = {1991}, pages = {649--652}, } @article{bancerek_comma_1991, title = {Comma {Category}}, volume = {2}, number = {5}, journal = {Formalized Mathematics}, author = {Bancerek, Grzegorz and Darmochwa{\l }, Agata}, year = {1991}, pages = {679--681}, } @article{bylinski_products_1991, title = {Products and {Coproducts} in {Categories}}, volume = {2}, number = {5}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1991}, pages = {701--709}, } @article{trybulec_isomorphisms_1992, title = {Some {Isomorphisms} {Between} {Functor} {Categories}}, volume = {3}, number = {1}, journal = {Formalized Mathematics}, author = {Trybulec, Andrzej}, year = {1992}, pages = {33--40}, } @article{bylinski_cartesian_1992, title = {Cartesian {Categories}}, volume = {3}, number = {2}, journal = {Formalized Mathematics}, author = {Byli{\'n}ski, Czes{\l }aw}, year = {1992}, pages = {161--169}, } @article{bancerek_categorial_1996, title = {Categorial {Categories} and {Slice} {Categories}}, volume = {5}, number = {2}, journal = {Formalized Mathematics}, author = {Bancerek, Grzegorz}, year = {1996}, pages = {157--165}, } @article{trybulec_categories_1996, title = {Categories without {Uniqueness} of cod and dom}, volume = {5}, number = {2}, journal = {Formalized Mathematics}, author = {Trybulec, Andrzej}, year = {1996}, pages = {259--267}, } @article{bancerek_indexed_1996, title = {Indexed {Category}}, volume = {5}, number = {3}, journal = {Formalized Mathematics}, author = {Bancerek, Grzegorz}, year = {1996}, pages = {329--337}, } @article{trybulec_functors_1996, title = {Functors for {Alternative} {Categories}}, volume = {5}, number = {4}, journal = {Formalized Mathematics}, author = {Trybulec, Andrzej}, year = {1996}, pages = {595--608}, } @article{nieszczerzewski_category_1997, title = {Category of {Functors} {Between} {Alternative} {Categories}}, volume = {6}, number = {3}, journal = {Formalized Mathematics}, author = {Nieszczerzewski, Robert}, year = {1997}, pages = {371--375}, } @article{kornilowicz_categories_1997, title = {On the {Categories} {Without} {Uniqueness} of cod and dom. {Some} {Properties} of the {Morphisms} and the {Functors}}, volume = {6}, number = {4}, journal = {Formalized Mathematics}, author = {Korni{\l }owicz, Artur}, year = {1997}, pages = {475--481}, } @article{kornilowicz_composition_1998, title = {The {Composition} of {Functors} and {Transformations} in {Alternative} {Categories}}, volume = {7}, number = {1}, journal = {Formalized Mathematics}, author = {Korni{\l }owicz, Artur}, year = {1998}, pages = {1--7}, } @article{bancerek_concrete_2001, title = {Concrete {Categories}}, volume = {9}, number = {3}, journal = {Formalized Mathematics}, author = {Bancerek, Grzegorz}, year = {2001}, pages = {605--621}, } @article{kornilowicz_products_2012, title = {Products in {Categories} without {Uniqueness} of cod and dom}, volume = {20}, number = {4}, journal = {Formalized Mathematics}, author = {Korni{\l }owicz, Artur}, year = {2012}, pages = {303--307}, } @article{riccardi_object-free_2013, title = {Object-{Free} {Definition} of {Categories}}, volume = {21}, number = {3}, journal = {Formalized Mathematics}, author = {Riccardi, Marco}, year = {2013}, keywords = {correspondence between different approaches to category, object-free category}, pages = {193--205}, } @article{golinski_coproducts_2013, title = {Coproducts in {Categories} without {Uniqueness} of cod and dom}, volume = {21}, number = {4}, journal = {Formalized Mathematics}, author = {Goli{\'n}ski, Maciej and Korni{\l }owicz, Artur}, year = {2013}, keywords = {coproducts, disjoined union}, pages = {235--239}, } @book{grabowski_preface_2014, title = {Preface}, volume = {22}, number = {2}, author = {Grabowski, Adam and Shidama, Yasunari}, year = {2014}, keywords = {Mizar, Mizar Mathematical Library}, } @article{riccardi_categorical_2015, title = {Categorical {Pullbacks}}, volume = {23}, number = {1}, journal = {Formalized Mathematics}, author = {Riccardi, Marco}, year = {2015}, keywords = {category pullback, pullback lemma}, pages = {1--14}, } @article{riccardi_exponential_2015, title = {Exponential {Objects}}, volume = {23}, number = {4}, journal = {Formalized Mathematics}, author = {Riccardi, Marco}, year = {2015}, keywords = {exponential objects, functor category, natural transformation}, pages = {351--369}, } @article{paulson_natural_1986, title = {Natural {Deduction} as {Higher}-{Order} {Resolution}}, volume = {3}, number = {3}, journal = {The Journal of Logic Programming}, author = {Paulson, Lawrence C.}, year = {1986}, pages = {237--258}, } @book{riehl_category_2016, title = {Category {Theory} in {Context}}, publisher = {Emily Riehl}, author = {Riehl, Emily}, year = {2016}, keywords = {Mathematics / Logic}, } @misc{chen_hotg_2021, title = {{HOTG}}, url = {https://bitbucket.org/cezaryka/tyset/src}, journal = {HOTG}, author = {Chen, Joshua and Kappelmann, Kevin and Krauss, Alexander}, year = {2021}, } @misc{haftmann_sketch-and-explore_2021, title = {Sketch-and-{Explore}}, url = {https://isabelle.in.tum.de/library/HOL/HOL-ex/Sketch_and_Explore.html}, author = {Haftmann, Florian}, year = {2021}, } @book{megill_metamath_2019, address = {Morrisville, North Carolina}, title = {Metamath: {A} {Computer} {Language} for {Mathematical} {Proofs}}, isbn = {978-0-359-70223-7}, publisher = {Lulu Press}, author = {Megill, Norman and Wheeler, David A.}, year = {2019}, } @inproceedings{kammuller_locales_1999, address = {Heidelberg}, series = {Lecture {Notes} in {Computer} {Science}}, title = {Locales {A} {Sectioning} {Concept} for {Isabelle}}, isbn = {978-3-540-48256-7}, booktitle = {Theorem {Proving} in {Higher} {Order} {Logics}}, publisher = {Springer}, author = {Kamm{\"u}ller, Florian and Wenzel, Markus and Paulson, Lawrence C.}, editor = {Bertot, Yves and Dowek, Gilles and Th{\'e}ry, Laurent and Hirschowitz, Andr{\'e} and Paulin-Mohring, Christine}, year = {1999}, pages = {149--165}, } +@article{milehins_category_2021, + title = {Category {Theory} for {ZFC} in {HOL} {II}: {Elementary} {Theory} of 1-{Categories}}, + journal = {Archive of Formal Proofs}, + author = {Milehins, Mihails}, + year = {2021}, +} diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Complete.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Complete.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Complete.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Complete.thy @@ -1,910 +1,910 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Completeness for categories\ theory CZH_UCAT_Complete imports CZH_UCAT_Limit begin subsection\Small-complete category\ subsubsection\Definition and elementary properties\ locale cat_small_complete = category \ \ for \ \ + assumes cat_small_complete: "\\ \. \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \ \ \u r. u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" locale cat_small_cocomplete = category \ \ for \ \ + assumes cat_small_cocomplete: "\\ \. \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \ \ \u r. u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" text\Rules.\ mk_ide rf cat_small_complete_def[unfolded cat_small_complete_axioms_def] |intro cat_small_completeI| |dest cat_small_completeD[dest]| |elim cat_small_completeE[elim]| lemma cat_small_completeE'[elim]: assumes "cat_small_complete \ \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" obtains u r where "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms by auto mk_ide rf cat_small_cocomplete_def[unfolded cat_small_cocomplete_axioms_def] |intro cat_small_cocompleteI| |dest cat_small_cocompleteD[dest]| |elim cat_small_cocompleteE[elim]| lemma cat_small_cocompleteE'[elim]: assumes "cat_small_cocomplete \ \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" obtains u r where "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms by auto subsubsection\Duality\ lemma (in cat_small_complete) cat_small_cocomplete_op[cat_op_intros]: "cat_small_cocomplete \ (op_cat \)" proof(intro cat_small_cocompleteI) fix \ \ assume "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> op_cat \" then interpret \: is_tm_functor \ \ \op_cat \\ \ . from cat_small_complete[OF \.is_tm_functor_op[unfolded cat_op_simps]] obtain u r where u: "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto then interpret u: is_cat_limit \ \op_cat \\ \ \op_cf \\ r u . from u.is_cat_colimit_op[unfolded cat_op_simps] show "\u r. u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by auto qed (auto intro: cat_cs_intros) lemmas [cat_op_intros] = cat_small_complete.cat_small_cocomplete_op lemma (in cat_small_cocomplete) cat_small_complete_op[cat_op_intros]: "cat_small_complete \ (op_cat \)" proof(intro cat_small_completeI) fix \ \ assume prems: "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> op_cat \" then interpret \: is_tm_functor \ \ \op_cat \\ \ . from cat_small_cocomplete[OF \.is_tm_functor_op[unfolded cat_op_simps]] obtain u r where u: "u : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto interpret u: is_cat_colimit \ \op_cat \\ \ \op_cf \\ r u by (rule u) from u.is_cat_limit_op[unfolded cat_op_simps] show "\u r. u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by auto qed (auto intro: cat_cs_intros) lemmas [cat_op_intros] = cat_small_cocomplete.cat_small_complete_op subsubsection\A category with equalizers and small products is small-complete\ lemma (in category) cat_small_complete_if_eq_and_obj_prod: \\See Corollary 2 in Chapter V-2 in \cite{mac_lane_categories_2010}\ assumes "\\ \ \ \. \ \ : \ \\<^bsub>\\<^esub> \; \ : \ \\<^bsub>\\<^esub> \ \ \ \E \. \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\A I. tm_cf_discrete \ I A \ \ \P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "cat_small_complete \ \" proof(intro cat_small_completeI) fix \ \ assume prems: "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" then interpret \: is_tm_functor \ \ \ \ . show "\u r. u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule cat_limit_of_cat_prod_obj_and_cat_equalizer[OF prems assms(1)]) (auto intro: assms(2)) qed (auto simp: cat_cs_intros) lemma (in category) cat_small_cocomplete_if_eq_and_obj_prod: assumes "\\ \ \ \. \ \ : \ \\<^bsub>\\<^esub> \; \ : \ \\<^bsub>\\<^esub> \ \ \ \E \. \ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\A I. tm_cf_discrete \ I A \ \ \P \. \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "cat_small_cocomplete \ \" proof- have "\E \. \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" if "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" for \ \ \ \ proof- from assms(1)[OF that] obtain \ E where \: "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \: is_cat_coequalizer \ \ \ \ \ \ E \ by (rule \) from \.is_cat_equalizer_op show ?thesis by auto qed moreover have "\P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" if "tm_cf_discrete \ I A (op_cat \)" for A I proof- interpret tm_cf_discrete \ I A \op_cat \\ by (rule that) from assms(2)[OF tm_cf_discrete_op[unfolded cat_op_simps]] obtain P \ where \: "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : I \\\<^sub>C\<^bsub>\\<^esub> \" by auto interpret \: is_cat_obj_coprod \ I A \ P \ by (rule \) from \.is_cat_obj_prod_op show ?thesis by auto qed ultimately interpret cat_small_complete \ \op_cat \\ by ( rule category.cat_small_complete_if_eq_and_obj_prod[ OF category_op, unfolded cat_op_simps ] ) show ?thesis by (rule cat_small_cocomplete_op[unfolded cat_op_simps]) qed subsection\Finite-complete category\ locale cat_finite_complete = category \ \ for \ \ + assumes cat_finite_complete: "\\ \. \ finite_category \ \; \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ \u r. u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" locale cat_finite_cocomplete = category \ \ for \ \ + assumes cat_finite_cocomplete: "\\ \. \ finite_category \ \; \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ \u r. u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" text\Rules.\ mk_ide rf cat_finite_complete_def[unfolded cat_finite_complete_axioms_def] |intro cat_finite_completeI| |dest cat_finite_completeD[dest]| |elim cat_finite_completeE[elim]| lemma cat_finite_completeE'[elim]: assumes "cat_finite_complete \ \" and "finite_category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains u r where "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms by auto mk_ide rf cat_finite_cocomplete_def[unfolded cat_finite_cocomplete_axioms_def] |intro cat_finite_cocompleteI| |dest cat_finite_cocompleteD[dest]| |elim cat_finite_cocompleteE[elim]| lemma cat_finite_cocompleteE'[elim]: assumes "cat_finite_cocomplete \ \" and "finite_category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains u r where "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms by auto text\Elementary properties.\ sublocale cat_small_complete \ cat_finite_complete proof(intro cat_finite_completeI) fix \ \ assume prems: "finite_category \ \" "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" interpret \: is_functor \ \ \ \ by (rule prems(2)) from cat_small_complete_axioms show "\u r. u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: \.cf_is_tm_functor_if_HomDom_finite_category[OF prems(1)]) qed (auto intro: cat_cs_intros) sublocale cat_small_cocomplete \ cat_finite_cocomplete proof(intro cat_finite_cocompleteI) fix \ \ assume prems: "finite_category \ \" "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" interpret \: is_functor \ \ \ \ by (rule prems(2)) from cat_small_cocomplete_axioms show "\u r. u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: \.cf_is_tm_functor_if_HomDom_finite_category[OF prems(1)]) qed (auto intro: cat_cs_intros) subsection\Discrete functor with tiny maps to the category \Set\\ lemma (in \) tm_cf_discrete_cat_Set_if_VLambda_in_Vset: assumes "VLambda I F \\<^sub>\ Vset \" shows "tm_cf_discrete \ I F (cat_Set \)" proof(intro tm_cf_discreteI) from assms have vrange_F_in_Vset: "\\<^sub>\ (VLambda I F) \\<^sub>\ Vset \" by (auto intro: vrange_in_VsetI) show "(\i\\<^sub>\I. cat_Set \\CId\\F i\) \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) from assms show "\\<^sub>\ (\i\\<^sub>\I. cat_Set \\CId\\F i\) \\<^sub>\ Vset \" by (metis vdomain_VLambda vdomain_in_VsetI) define Q where "Q i = ( if i = 0 then VPow ((\\<^sub>\i\\<^sub>\I. F i) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)) else set (F ` elts I) )" for i :: V have "\\<^sub>\ (\i\\<^sub>\I. cat_Set \\CId\\F i\) \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro vsubsetI, unfold cat_Set_components) fix y assume "y \\<^sub>\ \\<^sub>\ (\i\\<^sub>\I. VLambda (Vset \) id_Set\F i\)" then obtain i where i: "i \\<^sub>\ I" and y_def: "y = VLambda (Vset \) id_Set\F i\" by auto from i have "F i \\<^sub>\ \\<^sub>\ (VLambda I F)" by auto with vrange_F_in_Vset have "F i \\<^sub>\ Vset \" by auto then have y_def: "y = id_Set (F i)" unfolding y_def by auto show "y \\<^sub>\ (\\<^sub>\i\\<^sub>\set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" unfolding y_def proof(intro vproductI, unfold Ball_def; (intro allI impI)?) show "\\<^sub>\ (id_Rel (F i)) = set {0, 1\<^sub>\, 2\<^sub>\}" by (simp add: id_Rel_def incl_Rel_def three nat_omega_simps) fix j assume "j \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" then consider \j = 0\ | \j = 1\<^sub>\\ | \j = 2\<^sub>\\ by auto then show "id_Rel (F i)\j\ \\<^sub>\ Q j" proof cases case 1 from i show ?thesis unfolding 1 by ( subst arr_field_simps(1)[symmetric], unfold id_Rel_components Q_def ) force next case 2 from i show ?thesis unfolding 2 by ( subst arr_field_simps(2)[symmetric], unfold id_Rel_components Q_def ) auto next case 3 from i show ?thesis unfolding 3 by ( subst arr_field_simps(3)[symmetric], unfold id_Rel_components Q_def ) auto qed qed (auto simp: id_Rel_def cat_Set_cs_intros) qed moreover have "(\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i) \\<^sub>\ Vset \" proof(rule Limit_vproduct_in_VsetI) show "set {0, 1\<^sub>\, 2\<^sub>\} \\<^sub>\ Vset \" unfolding three[symmetric] by simp from assms have "VPow ((\\<^sub>\i\\<^sub>\I. F i) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)) \\<^sub>\ Vset \" by ( intro Limit_VPow_in_VsetI Limit_vtimes_in_VsetI Limit_vifunion_in_Vset_if_VLambda_in_VsetI ) auto then show "Q i \\<^sub>\ Vset \" if "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" for i using that vrange_VLambda by (auto intro!: vrange_F_in_Vset simp: Q_def nat_omega_simps) qed auto ultimately show "\\<^sub>\ (\i\\<^sub>\I. cat_Set \\CId\\F i\) \\<^sub>\ Vset \" by (meson vsubset_in_VsetI) qed auto fix i assume prems: "i \\<^sub>\ I" from assms have "\\<^sub>\ (VLambda I F) \\<^sub>\ Vset \" by (auto simp: vrange_in_VsetI) moreover from prems have "F i \\<^sub>\ \\<^sub>\ (VLambda I F)" by auto ultimately show "F i \\<^sub>\ cat_Set \\Obj\" unfolding cat_Set_components by auto qed (cs_concl cs_intro: cat_cs_intros assms)+ subsection\Product cone for the category \Set\\ subsubsection\Definition and elementary properties\ definition ntcf_Set_obj_prod :: "V \ V \ (V \ V) \ V" where "ntcf_Set_obj_prod \ I F = ntcf_obj_prod_base (cat_Set \) I F (\\<^sub>\i\\<^sub>\I. F i) (\i. vprojection_arrow I F i)" text\Components.\ lemma ntcf_Set_obj_prod_components: shows "ntcf_Set_obj_prod \ I F\NTMap\ = (\i\\<^sub>\:\<^sub>C I\Obj\. vprojection_arrow I F i)" and "ntcf_Set_obj_prod \ I F\NTDom\ = cf_const (:\<^sub>C I) (cat_Set \) (\\<^sub>\i\\<^sub>\I. F i)" and "ntcf_Set_obj_prod \ I F\NTCod\ = :\: I F (cat_Set \)" and "ntcf_Set_obj_prod \ I F\NTDGDom\ = :\<^sub>C I" and "ntcf_Set_obj_prod \ I F\NTDGCod\ = cat_Set \" unfolding ntcf_Set_obj_prod_def ntcf_obj_prod_base_components by simp_all subsubsection\Natural transformation map\ mk_VLambda ntcf_Set_obj_prod_components(1) |vsv ntcf_Set_obj_prod_NTMap_vsv[cat_cs_intros]| |vdomain ntcf_Set_obj_prod_NTMap_vdomain[cat_cs_simps]| |app ntcf_Set_obj_prod_NTMap_app[cat_cs_simps]| subsubsection\Product cone for the category \Set\ is a universal cone\ lemma (in \) tm_cf_discrete_ntcf_obj_prod_base_is_cat_obj_prod: \\See Theorem 5.2 in Chapter Introduction in \cite{hungerford_algebra_2003}.\ assumes "VLambda I F \\<^sub>\ Vset \" shows "ntcf_Set_obj_prod \ I F : (\\<^sub>\i\\<^sub>\I. F i) <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ F : I \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" -proof(intro is_cat_obj_prodI is_cat_limitI') +proof(intro is_cat_obj_prodI is_cat_limitI) interpret Set: tm_cf_discrete \ I F \cat_Set \\ by (rule tm_cf_discrete_cat_Set_if_VLambda_in_Vset[OF assms]) let ?F = \ntcf_Set_obj_prod \ I F\ show "cf_discrete \ I F (cat_Set \)" by (auto simp: cat_small_discrete_cs_intros) show F_is_cat_cone: "?F : (\\<^sub>\i\\<^sub>\I. F i) <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: I F (cat_Set \) : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" unfolding ntcf_Set_obj_prod_def proof(rule Set.tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone) show "(\\<^sub>\i\\<^sub>\I. F i) \\<^sub>\ cat_Set \\Obj\" unfolding cat_Set_components by ( intro Limit_vproduct_in_Vset_if_VLambda_in_VsetI Set.tm_cf_discrete_ObjMap_in_Vset ) auto qed (intro vprojection_arrow_is_arr Set.tm_cf_discrete_ObjMap_in_Vset) interpret F: is_cat_cone \ \\\<^sub>\i\\<^sub>\I. F i\ \:\<^sub>C I\ \cat_Set \\ \:\: I F (cat_Set \)\ \?F\ by (rule F_is_cat_cone) fix \' P' assume prems: "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: I F (cat_Set \) : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" let ?\'i = \\i. \'\NTMap\\i\\ let ?up' = \cat_Set_obj_prod_up I F P' ?\'i\ interpret \': is_cat_cone \ P' \:\<^sub>C I\ \cat_Set \\ \:\: I F (cat_Set \)\ \' by (rule prems(1)) show "\!f'. f' : P' \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i) \ \' = ?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) f'" proof(intro ex1I conjI; (elim conjE)?) show up': "?up' : P' \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" proof(rule cat_Set_obj_prod_up_cat_Set_is_arr) show "P' \\<^sub>\ Vset \" by (auto intro: cat_cs_intros cat_lim_cs_intros) fix i assume "i \\<^sub>\ I" then show "\'\NTMap\\i\ : P' \\<^bsub>cat_Set \\<^esub> F i" by ( cs_concl cs_simp: the_cat_discrete_components(1) cat_cs_simps cat_discrete_cs_simps cs_intro: cat_cs_intros ) qed (rule assms) then have P': "P' \\<^sub>\ Vset \" by (auto intro: cat_cs_intros cat_lim_cs_intros) have \'i_i: "?\'i i : P' \\<^bsub>cat_Set \\<^esub> F i" if "i \\<^sub>\ I" for i using \'.ntcf_NTMap_is_arr[unfolded the_cat_discrete_components(1), OF that] that by ( cs_prems cs_simp: cat_cs_simps cat_discrete_cs_simps the_cat_discrete_components(1) ) from cat_Set_obj_prod_up_cat_Set_is_arr[OF P' assms(1) \'i_i] have \'i: "cat_Set_obj_prod_up I F P' ?\'i : P' \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)". show "\' = ?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) ?up'" proof(rule ntcf_eqI, rule \'.is_ntcf_axioms) from F_is_cat_cone \'i show "?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) ?up' : cf_const (:\<^sub>C I) (cat_Set \) P' \\<^sub>C\<^sub>F :\: I F (cat_Set \) : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ (\'\NTMap\) = :\<^sub>C I\Obj\" by (cs_concl cs_simp: cat_cs_simps) from F_is_cat_cone \'i have dom_rhs: "\\<^sub>\ ((?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) ?up')\NTMap\) = :\<^sub>C I\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\'\NTMap\ = (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) ?up')\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix i assume prems': "i \\<^sub>\ :\<^sub>C I\Obj\" then have i: "i \\<^sub>\ I" unfolding the_cat_discrete_components by simp have [cat_cs_simps]: "vprojection_arrow I F i \\<^sub>A\<^bsub>cat_Set \\<^esub> ?up' = \'\NTMap\\i\" by ( - rule pdg_dghm_comp_dghm_proj_dghm_up[ + rule cat_Set_cf_comp_proj_obj_prod_up[ OF P' assms \'i_i i, symmetric ] ) auto from \'i prems' show "\'\NTMap\\i\ = (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) ?up')\NTMap\\i\" by ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_cs_intros ) qed (auto simp: cat_cs_intros) qed simp_all fix f' assume prems: "f' : P' \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" "\' = ?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) f'" from prems(2) have \'_eq_F_f': "\'\NTMap\\i\\ArrVal\\a\ = (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) (cat_Set \) f')\NTMap\\i\\ArrVal\\a\" if "i \\<^sub>\ I" and "a \\<^sub>\ P'" for i a by simp have [cat_Set_cs_simps]: "\'\NTMap\\i\\ArrVal\\a\ = f'\ArrVal\\a\\i\" if "i \\<^sub>\ I" and "a \\<^sub>\ P'" for i a using \'_eq_F_f'[OF that] assms prems that vprojection_arrow_is_arr[OF that(1) assms] by ( cs_prems cs_simp: cat_Set_cs_simps cat_cs_simps vprojection_arrow_app the_cat_discrete_components(1) cs_intro: cat_Set_cs_intros cat_cs_intros ) note f' = cat_Set_is_arrD[OF prems(1)] note up' = cat_Set_is_arrD[OF up'] interpret f': arr_Set \ f' by (rule f'(1)) interpret u': arr_Set \ \(cat_Set_obj_prod_up I F P' (app (\'\NTMap\)))\ by (rule up'(1)) show "f' = ?up'" proof(rule arr_Set_eqI[of \]) have dom_lhs: "\\<^sub>\ (f'\ArrVal\) = P'" by (simp add: cat_Set_cs_simps cat_cs_simps f') have dom_rhs: "\\<^sub>\ (cat_Set_obj_prod_up I F P' (app (\'\NTMap\))\ArrVal\) = P'" by (simp add: cat_Set_cs_simps cat_cs_simps up') show "f'\ArrVal\ = cat_Set_obj_prod_up I F P' (app (\'\NTMap\))\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems': "a \\<^sub>\ P'" from prems(1) prems' have "f'\ArrVal\\a\ \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" by (cs_concl cs_intro: cat_Set_cs_intros) note f'a = vproductD[OF this] from prems' have dom_rhs: "\\<^sub>\ (cat_Set_obj_prod_up I F P' (app (\'\NTMap\))\ArrVal\\a\) = I" by (cs_concl cs_simp: cat_Set_cs_simps) show "f'\ArrVal\\a\ = cat_Set_obj_prod_up I F P' (app (\'\NTMap\))\ArrVal\\a\" proof(rule vsv_eqI, unfold f'a dom_rhs) fix i assume "i \\<^sub>\ I" with prems' show "f'\ArrVal\\a\\i\ = cat_Set_obj_prod_up I F P' (app (\'\NTMap\))\ArrVal\\a\\i\" by (cs_concl cs_simp: cat_Set_cs_simps) qed (simp_all add: prems' f'a(1) cat_Set_obj_prod_up_ArrVal_app) qed auto qed (simp_all add: cat_Set_obj_prod_up_components f' up'(1)) qed qed subsection\Equalizer for the category \Set\\ subsubsection\Definition and elementary properties\ abbreviation ntcf_Set_equalizer_map :: "V \ V \ V \ V \ V \ V" where "ntcf_Set_equalizer_map \ a g f i \ ( i = \\<^sub>P\<^sub>L ? incl_Set (vequalizer a g f) a : g \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer a g f) a )" definition ntcf_Set_equalizer :: "V \ V \ V \ V \ V \ V" where "ntcf_Set_equalizer \ a b g f = ntcf_equalizer_base (cat_Set \) a b g f (vequalizer a g f) (ntcf_Set_equalizer_map \ a g f)" text\Components.\ context fixes a g f \ :: V begin lemmas ntcf_Set_equalizer_components = ntcf_equalizer_base_components[ where \=\cat_Set \\ and e=\ntcf_Set_equalizer_map \ a g f\ and E=\vequalizer a g f\ and \=a and \=g and \=f, folded ntcf_Set_equalizer_def ] end subsubsection\Natural transformation map\ mk_VLambda ntcf_Set_equalizer_components(1) |vsv ntcf_Set_equalizer_NTMap_vsv[cat_Set_cs_intros]| |vdomain ntcf_Set_equalizer_NTMap_vdomain[cat_Set_cs_simps]| |app ntcf_Set_equalizer_NTMap_app| lemma ntcf_Set_equalizer_2_NTMap_app_\[cat_Set_cs_simps]: assumes "x = \\<^sub>P\<^sub>L" shows "ntcf_Set_equalizer \ a b g f\NTMap\\x\ = incl_Set (vequalizer a g f) a" unfolding assms the_cat_parallel_components(1) ntcf_Set_equalizer_components by simp lemma ntcf_Set_equalizer_2_NTMap_app_\[cat_Set_cs_simps]: assumes "x = \\<^sub>P\<^sub>L" shows "ntcf_Set_equalizer \ a b g f\NTMap\\x\ = g \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer a g f) a" unfolding assms the_cat_parallel_components(1) ntcf_Set_equalizer_components using cat_PL_ineq by auto subsubsection\Equalizer for the category \Set\ is an equalizer\ lemma (in \) ntcf_Set_equalizer_2_is_cat_equalizer_2: assumes "\ : \ \\<^bsub>cat_Set \\<^esub> \" and "\ : \ \\<^bsub>cat_Set \\<^esub> \" shows "ntcf_Set_equalizer \ \ \ \ \ : vequalizer \ \ \ <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" -proof(intro is_cat_equalizerI is_cat_equalizerI is_cat_limitI') +proof(intro is_cat_equalizerI is_cat_equalizerI is_cat_limitI) let ?II_II = \\\\\\ (cat_Set \) \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ and ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ note \ = cat_Set_is_arrD[OF assms(1)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note \ = cat_Set_is_arrD[OF assms(2)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note [cat_Set_cs_intros] = \.arr_Set_ArrDom_in_Vset \.arr_Set_ArrCod_in_Vset let ?incl = \incl_Set (vequalizer \ \ \) \\ show \\\\_is_cat_cone: "ntcf_Set_equalizer \ \ \ \ \ : vequalizer \ \ \ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" unfolding ntcf_Set_equalizer_def proof ( intro category.cat_ntcf_equalizer_base_is_cat_cone category.cat_cf_parallel_cat_equalizer ) from assms show "(\\<^sub>P\<^sub>L = \\<^sub>P\<^sub>L ? ?incl : \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl) : vequalizer \ \ \ \\<^bsub>cat_Set \\<^esub> \" by ( cs_concl cs_simp: V_cs_simps cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros cat_PL_ineq[symmetric] ) show "(\\<^sub>P\<^sub>L = \\<^sub>P\<^sub>L ? ?incl : \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl) = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> (\\<^sub>P\<^sub>L = \\<^sub>P\<^sub>L ? ?incl : \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by ( cs_concl cs_simp: V_cs_simps cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros cat_PL_ineq[symmetric] ) from assms show "(\\<^sub>P\<^sub>L = \\<^sub>P\<^sub>L ? ?incl : \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl) = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> (\\<^sub>P\<^sub>L = \\<^sub>P\<^sub>L ? ?incl : \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by ( cs_concl cs_simp: V_cs_simps cat_Set_incl_Set_commute cs_intro: V_cs_intros cat_PL_ineq[symmetric] ) qed ( cs_concl cs_intro: cat_cs_intros V_cs_intros cat_Set_cs_intros assms cs_simp: V_cs_simps cat_cs_simps )+ interpret \\\\: is_cat_cone \ \vequalizer \ \ \\ ?II \cat_Set \\ ?II_II \ntcf_Set_equalizer \ \ \ \ \\ by (rule \\\\_is_cat_cone) show "\!f'. f' : r' \\<^bsub>cat_Set \\<^esub> vequalizer \ \ \ \ u' = ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) f'" if "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" for u' r' proof- interpret u': is_cat_cone \ r' ?II \cat_Set \\ ?II_II u' by (rule that(1)) have "\\<^sub>P\<^sub>L \\<^sub>\ \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\" unfolding the_cat_parallel_components(1) by simp from u'.ntcf_NTMap_is_arr[OF this] \\\\.NTDom.HomCod.cat_cf_parallel_cat_equalizer[OF assms] have u'_\\<^sub>P\<^sub>L_is_arr: "u'\NTMap\\\\<^sub>P\<^sub>L\ : r' \\<^bsub>cat_Set \\<^esub> \" by (cs_prems_atom_step cat_cs_simps) ( cs_prems cs_simp: cat_parallel_cs_simps cs_intro: cat_parallel_cs_intros cat_cs_intros category.cat_cf_parallel_cat_equalizer ) note u'_\\<^sub>P\<^sub>L = cat_Set_is_arrD[OF u'_\\<^sub>P\<^sub>L_is_arr] interpret u'_\\<^sub>P\<^sub>L: arr_Set \ \u'\NTMap\\\\<^sub>P\<^sub>L\\ by (rule u'_\\<^sub>P\<^sub>L(1)) have "\\<^sub>P\<^sub>L \\<^sub>\ ?II\Obj\" by (cs_concl cs_intro: cat_parallel_cs_intros) from u'.ntcf_NTMap_is_arr[OF this] \\\\.NTDom.HomCod.cat_cf_parallel_cat_equalizer[OF assms] have "u'\NTMap\\\\<^sub>P\<^sub>L\ : r' \\<^bsub>cat_Set \\<^esub> \" by ( cs_prems cs_simp: cat_cs_simps cat_parallel_cs_simps cs_intro: cat_parallel_cs_intros ) note u'_\u' = cat_cone_cf_par_eps_NTMap_app(1)[OF that(1) assms] define q where "q = [u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\, r', vequalizer \ \ \]\<^sub>\" have q_components[cat_Set_cs_simps]: "q\ArrVal\ = u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\" "q\ArrDom\ = r'" "q\ArrCod\ = vequalizer \ \ \" unfolding q_def arr_field_simps by (simp_all add: nat_omega_simps) from cat_cone_cf_par_eps_NTMap_app[OF that(1) assms] have \u'_eq_\u': "(\ \\<^sub>A\<^bsub>cat_Set \\<^esub> u'\NTMap\\\\<^sub>P\<^sub>L\)\ArrVal\\x\ = (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> u'\NTMap\\\\<^sub>P\<^sub>L\)\ArrVal\\x\" for x by simp show ?thesis proof(intro ex1I conjI; (elim conjE)?) have u'_NTMap_vrange: "\\<^sub>\ (u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\) \\<^sub>\ vequalizer \ \ \" proof(rule vsubsetI) fix y assume prems: "y \\<^sub>\ \\<^sub>\ (u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\)" then obtain x where x: "x \\<^sub>\ \\<^sub>\ (u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\)" and y_def: "y = u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\\x\" by (blast dest: u'_\\<^sub>P\<^sub>L.ArrVal.vrange_atD) have x: "x \\<^sub>\ r'" by (use x u'_\\<^sub>P\<^sub>L_is_arr in \cs_prems cs_simp: cat_cs_simps\) from \u'_eq_\u'[of x] assms x u'_\\<^sub>P\<^sub>L_is_arr have [simp]: "\\ArrVal\\u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\\x\\ = \\ArrVal\\u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\\x\\" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from prems u'_\\<^sub>P\<^sub>L.arr_Set_ArrVal_vrange[unfolded u'_\\<^sub>P\<^sub>L] show "y \\<^sub>\ vequalizer \ \ \" by (intro vequalizerI, unfold y_def) auto qed show q_is_arr: "q : r' \\<^bsub>cat_Set \\<^esub> vequalizer \ \ \" proof(intro cat_Set_is_arrI arr_SetI) show "q\ArrCod\ \\<^sub>\ Vset \" by (auto simp: q_components intro: cat_cs_intros cat_lim_cs_intros) qed ( auto simp: cat_Set_cs_simps nat_omega_simps u'_\\<^sub>P\<^sub>L q_def u'_NTMap_vrange \\\\.NTDom.HomCod.cat_in_Obj_in_Vset intro: cat_cs_intros cat_lim_cs_intros ) from q_is_arr have \_q: "incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q : r' \\<^bsub>cat_Set \\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_Set_cs_intros ) interpret arr_Set \ \incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q\ using \_q by (auto dest: cat_Set_is_arrD) show "u' = ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q" proof(rule ntcf_eqI) from q_is_arr show "ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q : cf_const ?II (cat_Set \) r' \\<^sub>C\<^sub>F ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ (u'\NTMap\) = ?II\Obj\" by (cs_concl cs_simp: cat_cs_simps) from q_is_arr have dom_rhs: "\\<^sub>\ ( (ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\) = ?II\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\ = ( ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) show "vsv (( ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\)" by (cs_concl cs_intro: cat_cs_intros) fix a assume prems: "a \\<^sub>\ ?II\Obj\" have [symmetric, cat_Set_cs_simps]: "u'\NTMap\\\\<^sub>P\<^sub>L\ = incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q" proof(rule arr_Set_eqI[of \]) from u'_\\<^sub>P\<^sub>L_is_arr have dom_lhs: "\\<^sub>\ (u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\) = r'" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \_q have dom_rhs: "\\<^sub>\ ((incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q)\ArrVal\) = r'" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\ = (incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ r'" with u'_NTMap_vrange dom_lhs u'_\\<^sub>P\<^sub>L.ArrVal.vsv_vimageI2 have "u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\\a\ \\<^sub>\ vequalizer \ \ \" by blast with prems q_is_arr u'_\\<^sub>P\<^sub>L_is_arr show "u'\NTMap\\\\<^sub>P\<^sub>L\\ArrVal\\a\ = (incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> q)\ArrVal\\a\" by ( cs_concl cs_simp: cat_Set_cs_simps cat_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_Set_cs_intros ) qed auto qed ( use u'_\\<^sub>P\<^sub>L \_q in \ cs_concl cs_intro: cat_Set_is_arrD(1) cs_simp: cat_cs_simps \ )+ from q_is_arr have u'_NTMap_app_I: "u'\NTMap\\\\<^sub>P\<^sub>L\ = ( ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\\\\<^sub>P\<^sub>L\" by ( cs_concl cs_intro: cat_cs_intros cat_parallel_cs_intros cs_simp: cat_Set_cs_simps cat_cs_simps V_cs_simps ) from q_is_arr assms have u'_NTMap_app_sI: "u'\NTMap\\\\<^sub>P\<^sub>L\ = ( ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\\\\<^sub>P\<^sub>L\" by ( cs_concl cs_simp: cat_Set_cs_simps cat_cs_simps u'_\u' cs_intro: V_cs_intros cat_cs_intros cat_Set_cs_intros cat_parallel_cs_intros ) from prems consider \a = \\<^sub>P\<^sub>L\ | \a = \\<^sub>P\<^sub>L\ by (elim the_cat_parallel_ObjE) then show "u'\NTMap\\a\ = ( ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) q )\NTMap\\a\" by cases (simp_all add: u'_NTMap_app_I u'_NTMap_app_sI) qed auto qed (simp_all add: u'.is_ntcf_axioms) fix f' assume prems: "f' : r' \\<^bsub>cat_Set \\<^esub> vequalizer \ \ \" "u' = ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) f'" from prems(2) have u'_NTMap_app: "u'\NTMap\\x\ = (ntcf_Set_equalizer \ \ \ \ \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II (cat_Set \) f')\NTMap\\x\" for x by simp have u'_f': "u'\NTMap\\\\<^sub>P\<^sub>L\ = incl_Set (vequalizer \ \ \) \ \\<^sub>A\<^bsub>cat_Set \\<^esub> f'" using u'_NTMap_app[of \\<^sub>P\<^sub>L] prems(1) by ( cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) (cs_prems cs_simp: cat_Set_cs_simps cs_intro: cat_parallel_cs_intros) note f' = cat_Set_is_arrD[OF prems(1)] note q = cat_Set_is_arrD[OF q_is_arr] interpret f': arr_Set \ f' using prems(1) by (auto dest: cat_Set_is_arrD) interpret q: arr_Set \ q using q by (auto dest: cat_Set_is_arrD) show "f' = q" proof(rule arr_Set_eqI[of \]) have dom_lhs: "\\<^sub>\ (f'\ArrVal\) = r'" by (simp add: cat_Set_cs_simps f') from q_is_arr have dom_rhs: "\\<^sub>\ (q\ArrVal\) = r'" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_Set_cs_intros) show "f'\ArrVal\ = q\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix i assume "i \\<^sub>\ r'" with prems(1) show "f'\ArrVal\\i\ = q\ArrVal\\i\" by ( cs_concl cs_simp: cat_Set_cs_simps cat_cs_simps q_components u'_f' cs_intro: V_cs_intros cat_cs_intros cat_Set_cs_intros ) qed auto qed ( use prems(1) q_is_arr in \ cs_concl cs_simp: cat_cs_simps cs_intro: q cat_Set_is_arrD \ )+ qed qed qed (auto intro: assms) subsection\The category \Set\ is small-complete\ lemma (in \) cat_small_complete_cat_Set: "cat_small_complete \ (cat_Set \)" \\This lemma appears as a remark on page 113 in \cite{mac_lane_categories_2010}.\ proof(rule category.cat_small_complete_if_eq_and_obj_prod) show "\E \. \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" if "\ : \ \\<^bsub>cat_Set \\<^esub> \" and "\ : \ \\<^bsub>cat_Set \\<^esub> \" for \ \ \ \ using ntcf_Set_equalizer_2_is_cat_equalizer_2[OF that(2,1)] by auto show "\P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" if "tm_cf_discrete \ I A (cat_Set \)" for A I proof(intro exI, rule tm_cf_discrete_ntcf_obj_prod_base_is_cat_obj_prod) interpret tm_cf_discrete \ I A \cat_Set \\ by (rule that) show "VLambda I A \\<^sub>\ Vset \" by (rule tm_cf_discrete_ObjMap_in_Vset) qed qed (rule category_cat_Set) text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Kan.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Kan.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Kan.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Kan.thy @@ -1,3088 +1,3067 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Simple Kan extensions\ theory CZH_UCAT_Kan imports CZH_Elementary_Categories.CZH_ECAT_Comma CZH_UCAT_Limit CZH_UCAT_Adjoints begin subsection\Background\ named_theorems ua_field_simps definition UObj :: V where [ua_field_simps]: "UObj = 0" definition UArr :: V where [ua_field_simps]: "UArr = 1\<^sub>\" named_theorems cat_Kan_cs_simps named_theorems cat_Kan_cs_intros subsection\Kan extension\ subsubsection\Definition and elementary properties\ text\See Chapter X-3 in \cite{mac_lane_categories_2010}.\ locale is_cat_rKe = AG: is_functor \ \ \ \ + Ran: is_functor \ \ \ \ + ntcf_rKe: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \ \ for \ \ \ \ \ \ \ \ + assumes cat_rKe_ua_fo: "universal_arrow_fo (exp_cat_cf \ \ \) (cf_map \) (cf_map \) (ntcf_arrow \)" syntax "_is_cat_rKe" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ \\<^sub>C\<^sub>F _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\ _ :/ _ \\<^sub>C _ \\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51] 51) translations "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" \ "CONST is_cat_rKe \ \ \ \ \ \ \ \" locale is_cat_lKe = AG: is_functor \ \ \ \ + Lan: is_functor \ \ \ \ + ntcf_lKe: is_ntcf \ \ \ \ \\ \\<^sub>C\<^sub>F \\ \ for \ \ \ \ \ \ \ \ + assumes cat_lKe_ua_fo: "universal_arrow_fo (exp_cat_cf \ (op_cat \) (op_cf \)) (cf_map \) (cf_map \) (ntcf_arrow (op_ntcf \))" syntax "_is_cat_lKe" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\ _ \\<^sub>C\<^sub>F _ :/ _ \\<^sub>C _ \\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51] 51) translations "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \" \ "CONST is_cat_lKe \ \ \ \ \ \ \ \" text\Rules.\ lemma (in is_cat_rKe) is_cat_rKe_axioms'[cat_Kan_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\'\<^esub> \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_rKe_axioms) mk_ide rf is_cat_rKe_def[unfolded is_cat_rKe_axioms_def] |intro is_cat_rKeI| |dest is_cat_rKeD[dest]| |elim is_cat_rKeE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_rKeD(1-3) lemma (in is_cat_lKe) is_cat_lKe_axioms'[cat_Kan_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_lKe_axioms) mk_ide rf is_cat_lKe_def[unfolded is_cat_lKe_axioms_def] |intro is_cat_lKeI| |dest is_cat_lKeD[dest]| |elim is_cat_lKeE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_lKeD(1-3) text\Duality.\ lemma (in is_cat_rKe) is_cat_lKe_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \" by (intro is_cat_lKeI, unfold cat_op_simps; (intro cat_rKe_ua_fo)?) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_rKe) is_cat_lKe_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_lKe_op) lemmas [cat_op_intros] = is_cat_rKe.is_cat_lKe_op' lemma (in is_cat_lKe) is_cat_rKe_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \" by (intro is_cat_rKeI, unfold cat_op_simps; (intro cat_lKe_ua_fo)?) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_lKe) is_cat_lKe_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_rKe_op) lemmas [cat_op_intros] = is_cat_lKe.is_cat_lKe_op' text\Elementary properties.\ lemma (in is_cat_rKe) cat_rKe_exp_cat_cf_cat_FUNCT_is_arr: assumes "\ \" and "\ \\<^sub>\ \" shows "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( rule exp_cat_cf_is_tiny_functor[ OF assms Ran.HomCod.category_axioms AG.is_functor_axioms ] ) lemma (in is_cat_lKe) cat_lKe_exp_cat_cf_cat_FUNCT_is_arr: assumes "\ \" and "\ \\<^sub>\ \" shows "exp_cat_cf \ \ \ : cat_FUNCT \ \ \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( rule exp_cat_cf_is_tiny_functor[ OF assms Lan.HomCod.category_axioms AG.is_functor_axioms ] ) subsubsection\Universal property\ text\ See Chapter X-3 in \cite{mac_lane_categories_2010} and \cite{noauthor_wikipedia_2001}\footnote{ \url{https://en.wikipedia.org/wiki/Kan_extension} }. \ lemma is_cat_rKeI': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\\' \'. \ \' : \ \\\<^sub>C\<^bsub>\\<^esub> \; \' : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ \!\. \ : \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" shows "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \ \ by (rule assms(3)) let ?\\ = \exp_cat_cf \ \ \\ and ?\ = \cf_map \\ and ?\ = \cf_map \\ show ?thesis proof(intro is_cat_rKeI is_functor.universal_arrow_foI assms) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \.\_Limit_\\ \.\_\_\\ \_def \.\_\_\\) then interpret \: \ \ by simp show "?\\ : cat_FUNCT \ \ \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" by ( cs_concl cs_intro: cat_small_cs_intros exp_cat_cf_is_tiny_functor[ OF \.\_axioms \\ \.HomCod.category_axioms assms(1) ] ) from \\ assms(2) show "cf_map \ \\<^sub>\ cat_FUNCT \ \ \\Obj\" unfolding cat_FUNCT_components by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_FUNCT_cs_intros) from assms(1-3) show "ntcf_arrow \ : ?\\\ObjMap\\?\\ \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_FUNCT_cs_simps cat_FUNCT_components(1) cs_intro: cat_FUNCT_cs_intros ) fix \' \' assume prems: "\' \\<^sub>\ cat_FUNCT \ \ \\Obj\" "\' : ?\\\ObjMap\\\'\ \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\" from prems(1) have "\' \\<^sub>\ cf_maps \ \ \" unfolding cat_FUNCT_components(1) by simp then obtain \ where \'_def: "\' = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp note \' = cat_FUNCT_is_arrD[OF prems(2)] from \'(1) \ have \'_is_ntcf: "ntcf_of_ntcf_arrow \ \ \' : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_prems cs_simp: \'_def cat_Kan_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from assms(4)[OF \ \'_is_ntcf] obtain \ where \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \'_def': "ntcf_of_ntcf_arrow \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" and unique_\: "\\'. \ \' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \; ntcf_of_ntcf_arrow \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \ \ \' = \" by metis show "\!f'. f' : \' \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\ \ \' = umap_fo ?\\ ?\ ?\ (ntcf_arrow \) \'\ArrVal\\f'\" proof(intro ex1I conjI; (elim conjE)?, unfold \'_def) from \ show "ntcf_arrow \ : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\" by (cs_concl cs_intro: cat_FUNCT_cs_intros) from \\ assms(1-3) \ \'(1) show "\' = umap_fo ?\\ ?\ ?\ (ntcf_arrow \) (cf_map \)\ArrVal\\ntcf_arrow \\" by (subst \') ( cs_concl cs_simp: - \'_def'[symmetric] cat_cs_simps cat_FUNCT_cs_simps cat_Kan_cs_simps + \'_def'[symmetric] + cat_cs_simps + cat_FUNCT_cs_simps + cat_Kan_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_Kan_cs_intros cat_FUNCT_cs_intros ) fix \' assume prems: "\' : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\" "\' = umap_fo ?\\ ?\ ?\ (ntcf_arrow \) (cf_map \)\ArrVal\\\'\" note \' = cat_FUNCT_is_arrD[OF prems(1)] from \'(1) \ have "ntcf_of_ntcf_arrow \ \ \' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) moreover from prems(2) prems(1) \\ assms(1-3) this \'(1) have "ntcf_of_ntcf_arrow \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_of_ntcf_arrow \ \ \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" by (subst (asm) \'(2)) ( cs_prems cs_simp: cat_Kan_cs_simps cat_FUNCT_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) ultimately have \_def: "\ = ntcf_of_ntcf_arrow \ \ \'" by (rule unique_\[symmetric]) show "\' = ntcf_arrow \" by (subst \'(2), use nothing in \subst \_def\) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed qed qed lemma is_cat_lKeI': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\\' \'. \ \' : \ \\\<^sub>C\<^bsub>\\<^esub> \; \' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ \!\. \ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \\ \\<^sub>C\<^sub>F \\ \ by (rule assms(3)) have "\!\. \ : \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \ \ \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" if "\' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" and "\' : \' \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" for \' \' proof- interpret \': is_functor \ \op_cat \\ \op_cat \\ \' by (rule that(1)) interpret \': is_ntcf \ \op_cat \\ \op_cat \\ \\' \\<^sub>C\<^sub>F op_cf \\ \op_cf \\ \' by (rule that(2)) from assms(4)[ OF is_functor.is_functor_op[OF that(1), unfolded cat_op_simps], OF is_ntcf.is_ntcf_op[OF that(2), unfolded cat_op_simps] ] obtain \ where \: "\ : \ \\<^sub>C\<^sub>F op_cf \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and op_\'_def: "op_ntcf \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" and unique_\': "\ \' : \ \\<^sub>C\<^sub>F op_cf \' : \ \\\<^sub>C\<^bsub>\\<^esub> \; op_ntcf \' = \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \ \ \' = \" for \' by metis interpret \: is_ntcf \ \ \ \ \op_cf \'\ \ by (rule \) show ?thesis proof(intro ex1I conjI; (elim conjE)?) show "op_ntcf \ : \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (rule \.is_ntcf_op[unfolded cat_op_simps]) from op_\'_def have "op_ntcf (op_ntcf \') = op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by simp from this \ assms(1-3) show \'_def: "\' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) fix \' assume prems: "\' : \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" "\' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" interpret \': is_ntcf \ \op_cat \\ \op_cat \\ \' \op_cf \\ \' by (rule prems(1)) from prems(2) have "op_ntcf \' = op_ntcf (op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \))" by simp also have "\ = op_ntcf \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) finally have "op_ntcf \' = op_ntcf \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by simp from unique_\'[OF \'.is_ntcf_op[unfolded cat_op_simps] this] show "\' = op_ntcf \" by (auto simp: cat_op_simps) qed qed from is_cat_rKeI' [ OF \.is_functor_op \.is_functor_op \.is_ntcf_op[unfolded cat_op_simps], unfolded cat_op_simps, OF this ] interpret \: is_cat_rKe \ \op_cat \\ \op_cat \\ \op_cat \\ \op_cf \\ \op_cf \\ \op_cf \\ \op_ntcf \\ by simp show "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \" by (rule \.is_cat_lKe_op[unfolded cat_op_simps]) qed lemma (in is_cat_rKe) cat_rKe_unique: assumes "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!\. \ : \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" proof- interpret \': is_functor \ \ \ \' by (rule assms(1)) interpret \': is_ntcf \ \ \ \\' \\<^sub>C\<^sub>F \\ \ \' by (rule assms(2)) let ?\ = \cf_map \\ and ?\ = \cf_map \\ and ?\' = \cf_map \'\ and ?\ = \ntcf_arrow \\ and ?\' = \ntcf_arrow \'\ define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def AG.\_Limit_\\ AG.\_\_\\ \_def AG.\_\_\\) then interpret \: \ \ by simp interpret \\: is_tiny_functor \ \cat_FUNCT \ \ \\ \cat_FUNCT \ \ \\ \exp_cat_cf \ \ \\ by (rule cat_rKe_exp_cat_cf_cat_FUNCT_is_arr[OF \.\_axioms \\]) from assms(1) have \': "?\' \\<^sub>\ cat_FUNCT \ \ \\Obj\" by (cs_concl cs_simp: cat_FUNCT_components(1) cs_intro: cat_FUNCT_cs_intros) with assms(2) have "?\' : exp_cat_cf \ \ \\ObjMap\\?\'\ \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from is_functor.universal_arrow_foD(3)[ OF \\.is_functor_axioms cat_rKe_ua_fo \' this ] obtain f' where f': "f' : cf_map \' \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" and \'_def: "?\' = umap_fo (exp_cat_cf \ \ \) ?\ ?\ ?\ ?\'\ArrVal\\f'\" and f'_unique: "\ f'' : ?\' \\<^bsub>cat_FUNCT \ \ \\<^esub> ?\; ntcf_arrow \' = umap_fo (exp_cat_cf \ \ \) ?\ ?\ ?\ ?\'\ArrVal\\f''\ \ \ f'' = f'" for f'' by metis show ?thesis proof(intro ex1I conjI; (elim conjE)?) from \'_def cat_FUNCT_is_arrD(1)[OF f'] show "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_of_ntcf_arrow \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" by (subst (asm) cat_FUNCT_is_arrD(2)[OF f']) (*slow*) ( cs_prems cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from cat_FUNCT_is_arrD(1)[OF f'] show f'_is_arr: "ntcf_of_ntcf_arrow \ \ f' : \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_prems cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) fix \ assume prems: "\ : \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" interpret \: is_ntcf \ \ \ \' \ \ by (rule prems(1)) from prems(1) have \: "ntcf_arrow \ : cf_map \' \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" by (cs_concl cs_intro: cat_FUNCT_cs_intros) from prems have \'_def: "ntcf_arrow \' = umap_fo (exp_cat_cf \ \ \) ?\ ?\ ?\ ?\'\ArrVal\\ntcf_arrow \\" by ( cs_concl cs_simp: prems(2) cat_Kan_cs_simps cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) show "\ = ntcf_of_ntcf_arrow \ \ f'" unfolding f'_unique[OF \ \'_def, symmetric] by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) qed qed lemma (in is_cat_lKe) cat_lKe_unique: assumes "\' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!\. \ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof- interpret \': is_functor \ \ \ \' by (rule assms(1)) interpret \': is_ntcf \ \ \ \ \\' \\<^sub>C\<^sub>F \\ \' by (rule assms(2)) interpret \: is_cat_rKe \ \op_cat \\ \op_cat \\ \op_cat \\ \op_cf \\ \op_cf \\ \op_cf \\ \op_ntcf \\ by (rule is_cat_rKe_op) from \.cat_rKe_unique[OF \'.is_functor_op \'.is_ntcf_op[unfolded cat_op_simps]] obtain \ where \: "\ : op_cf \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" and \'_def: "op_ntcf \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" and unique_\': "\\'. \ \' : op_cf \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \; op_ntcf \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \) \ \ \' = \" by metis interpret \: is_ntcf \ \op_cat \\ \op_cat \\ \op_cf \'\ \op_cf \\ \ by (rule \) show ?thesis proof(intro ex1I conjI; (elim conjE)?) show "op_ntcf \ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule \.is_ntcf_op[unfolded cat_op_simps]) have "\' = op_ntcf (op_ntcf \')" by (cs_concl cs_simp: cat_op_simps) also from \'_def have "\ = op_ntcf (op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \))" by simp also have "\ = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros) finally show "\' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by simp fix \' assume prems: "\' : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" "\' = \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" interpret \': is_ntcf \ \ \ \ \' \' by (rule prems(1)) from prems(2) have "op_ntcf \' = op_ntcf (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by simp also have "\ = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_ntcf \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros) finally have "op_ntcf \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_ntcf \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \)" by simp from unique_\'[OF \'.is_ntcf_op this] show "\' = op_ntcf \" by (auto simp: cat_op_simps) qed qed subsubsection\Further properties\ lemma (in is_cat_rKe) cat_rKe_ntcf_ua_fo_is_iso_ntcf_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "ntcf_ua_fo \ (exp_cat_cf \ \ \) (cf_map \) (cf_map \) (ntcf_arrow \) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ \(-,cf_map \) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ \(-,cf_map \) \\<^sub>C\<^sub>F op_cf (exp_cat_cf \ \ \) : op_cat (cat_FUNCT \ \ \) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \_\: is_tiny_functor \ \cat_FUNCT \ \ \\ \cat_FUNCT \ \ \\ \exp_cat_cf \ \ \\ by ( rule exp_cat_cf_is_tiny_functor[ OF assms Ran.HomCod.category_axioms AG.is_functor_axioms ] ) show ?thesis by ( rule is_functor.cf_ntcf_ua_fo_is_iso_ntcf[ OF \_\.is_functor_axioms cat_rKe_ua_fo ] ) qed lemma (in is_cat_lKe) cat_lKe_ntcf_ua_fo_is_iso_ntcf_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" defines "\\ \ exp_cat_cf \ (op_cat \) (op_cf \)" and "\\ \ cat_FUNCT \ (op_cat \) (op_cat \)" and "\\ \ cat_FUNCT \ (op_cat \) (op_cat \)" shows "ntcf_ua_fo \ \\ (cf_map \) (cf_map \) (ntcf_arrow (op_ntcf \)) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\(-,cf_map \) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\(-,cf_map \) \\<^sub>C\<^sub>F op_cf \\ : op_cat \\ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- note simps = \\_def \\_def \\_def interpret \_\: is_tiny_functor \ \\ \\ \\ unfolding simps by ( rule exp_cat_cf_is_tiny_functor[ OF assms(1,2) Lan.HomCod.category_op AG.is_functor_op ] ) show ?thesis unfolding simps by ( rule is_functor.cf_ntcf_ua_fo_is_iso_ntcf[ OF \_\.is_functor_axioms[unfolded simps] cat_lKe_ua_fo ] ) qed subsection\The Kan extension\ text\ The following subsection is based on the statement and proof of Theorem 1 in Chapter X-3 in \cite{mac_lane_categories_2010}. In what follows, only the right Kan extension is considered for simplicity. \ subsubsection\Definition and elementary properties\ definition the_cf_rKe :: "V \ V \ V \ (V \ V) \ V" where "the_cf_rKe \ \ \ lim_Obj = [ (\c\\<^sub>\\\HomCod\\Obj\. lim_Obj c\UObj\), ( \g\\<^sub>\\\HomCod\\Arr\. THE f. f : lim_Obj (\\HomCod\\Dom\\g\)\UObj\ \\<^bsub>\\HomCod\\<^esub> lim_Obj (\\HomCod\\Cod\\g\)\UObj\ \ lim_Obj (\\HomCod\\Dom\\g\)\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj (\\HomCod\\Cod\\g\)\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ((\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \) (\\HomCod\) f ), \\HomCod\, \\HomCod\ ]\<^sub>\" definition the_ntcf_rKe :: "V \ V \ V \ (V \ V) \ V" where "the_ntcf_rKe \ \ \ lim_Obj = [ ( \c\\<^sub>\\\HomDom\\Obj\. lim_Obj (\\ObjMap\\c\)\UArr\\NTMap\\0, c, \\HomCod\\CId\\\\ObjMap\\c\\\\<^sub>\ ), the_cf_rKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \, \, \\HomDom\, \\HomCod\ ]\<^sub>\" text\Components.\ lemma the_cf_rKe_components: shows "the_cf_rKe \ \ \ lim_Obj\ObjMap\ = (\c\\<^sub>\\\HomCod\\Obj\. lim_Obj c\UObj\)" and "the_cf_rKe \ \ \ lim_Obj\ArrMap\ = ( \g\\<^sub>\\\HomCod\\Arr\. THE f. f : lim_Obj (\\HomCod\\Dom\\g\)\UObj\ \\<^bsub>\\HomCod\\<^esub> lim_Obj (\\HomCod\\Cod\\g\)\UObj\ \ lim_Obj (\\HomCod\\Dom\\g\)\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj (\\HomCod\\Cod\\g\)\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ((\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \) (\\HomCod\) f )" and "the_cf_rKe \ \ \ lim_Obj\HomDom\ = \\HomCod\" and "the_cf_rKe \ \ \ lim_Obj\HomCod\ = \\HomCod\" unfolding the_cf_rKe_def dghm_field_simps by (simp_all add: nat_omega_simps) lemma the_ntcf_rKe_components: shows "the_ntcf_rKe \ \ \ lim_Obj\NTMap\ = ( \c\\<^sub>\\\HomDom\\Obj\. lim_Obj (\\ObjMap\\c\)\UArr\\NTMap\\0, c, \\HomCod\\CId\\\\ObjMap\\c\\\\<^sub>\ )" and "the_ntcf_rKe \ \ \ lim_Obj\NTDom\ = the_cf_rKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \" and "the_ntcf_rKe \ \ \ lim_Obj\NTCod\ = \" and "the_ntcf_rKe \ \ \ lim_Obj\NTDGDom\ = \\HomDom\" and "the_ntcf_rKe \ \ \ lim_Obj\NTDGCod\ = \\HomCod\" unfolding the_ntcf_rKe_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) lemmas the_cf_rKe_components' = the_cf_rKe_components[ where \=\ and \=\ and \=\, unfolded \.cf_HomCod \.cf_HomCod ] lemmas [cat_Kan_cs_simps] = the_cf_rKe_components'(3,4) lemmas the_ntcf_rKe_components' = the_ntcf_rKe_components[ where \=\ and \=\ and \=\, unfolded \.cf_HomCod \.cf_HomCod \.cf_HomDom ] lemmas [cat_Kan_cs_simps] = the_ntcf_rKe_components'(2-5) end subsubsection\Functor: object map\ mk_VLambda the_cf_rKe_components(1) |vsv the_cf_rKe_ObjMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda the_cf_rKe_components'(1)[OF \ \] |vdomain the_cf_rKe_ObjMap_vdomain[cat_Kan_cs_simps]| |app the_cf_rKe_ObjMap_impl_app[cat_Kan_cs_simps]| lemma the_cf_rKe_ObjMap_vrange: assumes "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UObj\ \\<^sub>\ \\Obj\" shows "\\<^sub>\ (the_cf_rKe \ \ \ lim_Obj\ObjMap\) \\<^sub>\ \\Obj\" unfolding the_cf_rKe_components'[OF \ \] by (intro vrange_VLambda_vsubset assms) end subsubsection\Functor: arrow map\ mk_VLambda the_cf_rKe_components(2) |vsv the_cf_rKe_ArrMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda the_cf_rKe_components(2)[where \=\ and \=\, unfolded \.cf_HomCod] |vdomain the_cf_rKe_ArrMap_vdomain[cat_Kan_cs_simps]| context fixes \ \ c c' g assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and g: "g : c \\<^bsub>\\<^esub> c'" begin interpretation \: is_functor \ \ \ \ by (rule \) lemma g': "g \\<^sub>\ \\Arr\" using g by auto mk_VLambda the_cf_rKe_components(2)[ where \=\ and \=\ and \=\, unfolded \.cf_HomCod \.cf_HomCod ] |app the_cf_rKe_ArrMap_app_impl'| lemmas the_cf_rKe_ArrMap_app' = the_cf_rKe_ArrMap_app_impl'[ OF g', unfolded \.HomCod.cat_is_arrD[OF g] ] end end lemma the_cf_rKe_ArrMap_app_impl: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g : c \\<^bsub>\\<^esub> c'" and "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f. f : r \\<^bsub>\\<^esub> r' \ u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ f" proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret u: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ r u by (rule assms(4)) interpret u': is_cat_limit \ \c' \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \\ r' u' by (rule assms(5)) have const_r_def: "cf_const (c' \\<^sub>C\<^sub>F \) \ r = cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \" proof(rule cf_eqI) show const_r: "cf_const (c' \\<^sub>C\<^sub>F \) \ r : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_lim_cs_intros) from assms(3) show const_r_g\: "cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) have ObjMap_dom_lhs: "\\<^sub>\ (cf_const (c' \\<^sub>C\<^sub>F \) \ r\ObjMap\) = c' \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) have ObjMap_dom_rhs: "\\<^sub>\ ((cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = c' \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) have ArrMap_dom_lhs: "\\<^sub>\ (cf_const (c' \\<^sub>C\<^sub>F \) \ r\ArrMap\) = c' \\<^sub>C\<^sub>F \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) have ArrMap_dom_rhs: "\\<^sub>\ ((cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\) = c' \\<^sub>C\<^sub>F \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) show "cf_const (c' \\<^sub>C\<^sub>F \) \ r\ObjMap\ = (cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix A assume prems: "A \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" from prems assms obtain b f where A_def: "A = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from assms(1,3) prems f b show "cf_const (c' \\<^sub>C\<^sub>F \) \ r\ObjMap\\A\ = (cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>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(3) in \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\)+ show "cf_const (c' \\<^sub>C\<^sub>F \) \ r\ArrMap\ = (cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) show "vsv (cf_const (c' \\<^sub>C\<^sub>F \) \ r\ArrMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) show "vsv ((cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\)" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) fix F assume prems: "F \\<^sub>\ c' \\<^sub>C\<^sub>F \\Arr\" with prems obtain A B where F: "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" by (auto intro: is_arrI) with assms obtain b f b' f' h' where F_def: "F = [[0, b, f]\<^sub>\, [0, b', f']\<^sub>\, [0, h']\<^sub>\]\<^sub>\" and A_def: "A = [0, b, f]\<^sub>\" and B_def: "B = [0, b', f']\<^sub>\" and h': "h' : b \\<^bsub>\\<^esub> b'" and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : c' \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'_def: "\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f = f'" by auto from prems assms(3) F g' h' f f' show "cf_const (c' \\<^sub>C\<^sub>F \) \ r\ArrMap\\F\ = (cf_const (c \\<^sub>C\<^sub>F \) \ r \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\\F\" unfolding F_def A_def B_def by (*slow*) ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps f'_def[symmetric] cs_intro: cat_cs_intros cat_comma_cs_intros ) qed simp qed simp_all have \c'\: "\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \" proof(rule cf_eqI) show "\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show " \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) have ObjMap_dom_lhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\) = c' \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have ObjMap_dom_rhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = c' \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) show "(\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) from assms show "vsv ((\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\)" by ( cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from assms show "vsv ((\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\)" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) fix A assume prems: "A \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" from assms(3) prems obtain b f where A_def: "A = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from prems assms b f show "(\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\A\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>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 simp have ArrMap_dom_lhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\) = c' \\<^sub>C\<^sub>F \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have ArrMap_dom_rhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\) = c' \\<^sub>C\<^sub>F \\Arr\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) show "(\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) from assms show "vsv ((\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\)" by ( cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from assms show "vsv ((\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\)" by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_comma_cs_intros) fix F assume prems: "F \\<^sub>\ c' \\<^sub>C\<^sub>F \\Arr\" with prems obtain A B where F: "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" unfolding cat_comma_cs_simps by (auto intro: is_arrI) with assms(3) obtain b f b' f' h' where F_def: "F = [[0, b, f]\<^sub>\, [0, b', f']\<^sub>\, [0, h']\<^sub>\]\<^sub>\" and A_def: "A = [0, b, f]\<^sub>\" and B_def: "B = [0, b', f']\<^sub>\" and h': "h' : b \\<^bsub>\\<^esub> b'" and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : c' \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'_def: "\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f = f'" by auto from prems assms(3) F g' h' f f' show "(\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\\F\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\\F\" unfolding F_def A_def B_def by (*slow*) ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps f'_def[symmetric] cs_intro: cat_cs_intros cat_comma_cs_intros ) qed simp qed simp_all from assms(1-3) have "u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (intro is_cat_coneI is_tm_ntcfI') + by (intro is_cat_coneI) ( cs_concl - cs_intro: - cat_cs_intros - cat_comma_cs_intros - cat_lim_cs_intros - cat_small_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros cat_lim_cs_intros cs_simp: const_r_def \c'\ )+ - with u'.cat_lim_unique_cone show + with u'.cat_lim_ua_fo show "\!G. G : r \\<^bsub>\\<^esub> r' \ u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ G" by simp qed lemma the_cf_rKe_ArrMap_app: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g : c \\<^bsub>\\<^esub> c'" and "lim_Obj c\UArr\ : lim_Obj c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "lim_Obj c'\UArr\ : lim_Obj c'\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\ : lim_Obj c\UObj\ \\<^bsub>\\<^esub> lim_Obj c'\UObj\" and "lim_Obj c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj c'\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ (the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\)" and "\ f : lim_Obj c\UObj\ \\<^bsub>\\<^esub> lim_Obj c'\UObj\; lim_Obj c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj c'\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ f \ \ f = the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\" proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret u: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \lim_Obj c\UObj\\ \lim_Obj c\UArr\\ by (rule assms(4)) interpret u': is_cat_limit \ \c' \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \\ \lim_Obj c'\UObj\\ \lim_Obj c'\UArr\\ by (rule assms(5)) from assms(3) have c: "c \\<^sub>\ \\Obj\" and c': "c' \\<^sub>\ \\Obj\" by auto note the_cf_rKe_ArrMap_app_impl' = the_cf_rKe_ArrMap_app_impl[OF assms] note the_f = theI'[OF the_cf_rKe_ArrMap_app_impl[OF assms]] note the_f_is_arr = the_f[THEN conjunct1] and the_f_commutes = the_f[THEN conjunct2] from assms(3) the_f_is_arr show "the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\ : lim_Obj c\UObj\ \\<^bsub>\\<^esub> lim_Obj c'\UObj\" by (cs_concl cs_simp: the_cf_rKe_ArrMap_app' cs_intro: cat_cs_intros) moreover from assms(3) the_f_commutes show "lim_Obj c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj c'\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ (the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\)" by (cs_concl cs_simp: the_cf_rKe_ArrMap_app' cs_intro: cat_cs_intros) ultimately show "f = the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\" if "f : lim_Obj c\UObj\ \\<^bsub>\\<^esub> lim_Obj c'\UObj\" and "lim_Obj c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F g \<^sub>A\\<^sub>C\<^sub>F \ = lim_Obj c'\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c' \\<^sub>C\<^sub>F \) \ f" by (metis that the_cf_rKe_ArrMap_app_impl') qed lemma the_cf_rKe_ArrMap_is_arr'[cat_Kan_cs_intros]: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g : c \\<^bsub>\\<^esub> c'" and "lim_Obj c\UArr\ : lim_Obj c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "lim_Obj c'\UArr\ : lim_Obj c'\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c' \<^sub>O\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a = lim_Obj c\UObj\" and "b = lim_Obj c'\UObj\" shows "the_cf_rKe \ \ \ lim_Obj\ArrMap\\g\ : a \\<^bsub>\\<^esub> b" unfolding assms(6,7) by (rule the_cf_rKe_ArrMap_app[OF assms(1-5)]) lemma lim_Obj_the_cf_rKe_commute: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "lim_Obj a\UArr\ : lim_Obj a\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \ : a \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "lim_Obj b\UArr\ : lim_Obj b\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f : a \\<^bsub>\\<^esub> b" and "[a', b', f']\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" shows "lim_Obj a\UArr\\NTMap\\a', b', f' \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\ = lim_Obj b\UArr\\NTMap\\a', b', f'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> the_cf_rKe \ \ \ lim_Obj\ArrMap\\f\" proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) note f = \.HomCod.cat_is_arrD[OF assms(5)] interpret lim_a: is_cat_limit \ \a \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \\ \lim_Obj a\UObj\\ \lim_Obj a\UArr\\ by (rule assms(3)) interpret lim_b: is_cat_limit \ \b \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \\ \lim_Obj b\UObj\\ \lim_Obj b\UArr\\ by (rule assms(4)) note f_app = the_cf_rKe_ArrMap_app[ where lim_Obj=lim_Obj, OF assms(1,2,5,3,4) ] from f_app(2) have lim_a_f\_NTMap_app: "(lim_Obj a\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\\A\ = ( lim_Obj b\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (b \\<^sub>C\<^sub>F \) \ (the_cf_rKe \ \ \ lim_Obj\ArrMap\\f\) )\NTMap\\A\" if \A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\\ for A by simp show "lim_Obj a\UArr\\NTMap\\a', b', f' \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\ = lim_Obj b\UArr\\NTMap\\a', b', f'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> the_cf_rKe \ \ \ lim_Obj\ArrMap\\f\" proof- from assms(5,6) have a'_def: "a' = 0" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto show "lim_Obj a\UArr\\NTMap\\a', b', f' \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\ = lim_Obj b\UArr\\NTMap\\a', b', f'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> the_cf_rKe \ \ \ lim_Obj\ArrMap\\f\" using lim_a_f\_NTMap_app[OF assms(6)] f' assms(3,4,5,6) unfolding a'_def by ( cs_prems cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps - cs_intro: - cat_small_cs_intros - cat_cs_intros - cat_comma_cs_intros - cat_Kan_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros ) qed qed subsubsection\Natural transformation: natural transformation map\ mk_VLambda the_ntcf_rKe_components(1) |vsv the_ntcf_rKe_NTMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda the_ntcf_rKe_components'(1)[OF \ \] |vdomain the_ntcf_rKe_ObjMap_vdomain[cat_Kan_cs_simps]| |app the_ntcf_rKe_ObjMap_impl_app[cat_Kan_cs_simps]| end subsubsection\The Kan extension is a Kan extension\ lemma the_cf_rKe_is_functor: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : lim_Obj c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "the_cf_rKe \ \ \ lim_Obj : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?UObj = \\a. lim_Obj a\UObj\\ let ?UArr = \\a. lim_Obj a\UArr\\ let ?const_comma = \\a b. cf_const (a \\<^sub>C\<^sub>F \) \ (?UObj b)\ let ?the_cf_rKe = \the_cf_rKe \ \ \ lim_Obj\ - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) note [cat_lim_cs_intros] = is_cat_cone.cat_cone_obj show ?thesis proof(intro is_functorI') show "vfsequence ?the_cf_rKe" unfolding the_cf_rKe_def by simp show "vcard ?the_cf_rKe = 4\<^sub>\" unfolding the_cf_rKe_def by (simp add: nat_omega_simps) show "vsv (?the_cf_rKe\ObjMap\)" by (cs_concl cs_intro: cat_Kan_cs_intros) moreover show "\\<^sub>\ (?the_cf_rKe\ObjMap\) = \\Obj\" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) moreover show "\\<^sub>\ (?the_cf_rKe\ObjMap\) \\<^sub>\ \\Obj\" proof ( intro the_cf_rKe_ObjMap_vrange; (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)? ) fix c assume "c \\<^sub>\ \\Obj\" with assms(3)[OF this] show "?UObj c \\<^sub>\ \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_lim_cs_intros) qed ultimately have [cat_Kan_cs_intros]: "?the_cf_rKe\ObjMap\\c\ \\<^sub>\ \\Obj\" if \c \\<^sub>\ \\Obj\\ for c by (metis that vsubsetE vsv.vsv_value) show "?the_cf_rKe\ArrMap\\f\ : ?the_cf_rKe\ObjMap\\a\ \\<^bsub>\\<^esub> ?the_cf_rKe\ObjMap\\b\" if "f : a \\<^bsub>\\<^esub> b" for a b f using assms(2) that by ( cs_concl cs_simp: cat_Kan_cs_simps - cs_intro: - assms(3) cat_small_cs_intros cat_cs_intros cat_Kan_cs_intros + cs_intro: assms(3) cat_cs_intros cat_Kan_cs_intros ) then have [cat_Kan_cs_intros]: "?the_cf_rKe\ArrMap\\f\ : A \\<^bsub>\\<^esub> B" if "A = ?the_cf_rKe\ObjMap\\a\" and "B = ?the_cf_rKe\ObjMap\\b\" and "f : a \\<^bsub>\\<^esub> b" for A B a b f by (simp add: that) show "?the_cf_rKe\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = ?the_cf_rKe\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> ?the_cf_rKe\ArrMap\\f\" (is \?the_cf_rKe\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = ?the_rKe_g \\<^sub>A\<^bsub>\\<^esub> ?the_rKe_f\) if g_is_arr: "g : b \\<^bsub>\\<^esub> c" and f_is_arr: "f : a \\<^bsub>\\<^esub> b" for b c g a f proof- let ?ntcf_const_c = \\f. ntcf_const (c \\<^sub>C\<^sub>F \) \ f\ note g = \.HomCod.cat_is_arrD[OF that(1)] and f = \.HomCod.cat_is_arrD[OF that(2)] note lim_a = assms(3)[OF f(2)] and lim_b = assms(3)[OF g(2)] and lim_c = assms(3)[OF g(3)] from that have gf: "g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" by (cs_concl cs_intro: cat_cs_intros) interpret lim_a: is_cat_limit \ \a \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj a\ \?UArr a\ by (rule lim_a) interpret lim_c: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj c\ \?UArr c\ by (rule lim_c) show ?thesis proof ( rule sym, rule the_cf_rKe_ArrMap_app(3)[OF assms(1,2) gf lim_a lim_c] ) from assms(1,2) that lim_a lim_b lim_c show "?the_rKe_g \\<^sub>A\<^bsub>\\<^esub> ?the_rKe_f : ?UObj a \\<^bsub>\\<^esub> ?UObj c" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) show "?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c (?the_rKe_g \\<^sub>A\<^bsub>\\<^esub> ?the_rKe_f)" ( is \ ?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c ?the_rKe_gf \ ) proof(rule ntcf_eqI) from that show "?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ : cf_const (a \\<^sub>C\<^sub>F \) \ (?UObj a) \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ((g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \) : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) have [cat_comma_cs_simps]: "?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ = ?const_comma c a" proof(rule cf_eqI) from g_is_arr f_is_arr show "?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros cat_comma_cs_intros ) from g_is_arr f_is_arr show "?const_comma c a : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros cat_comma_cs_intros ) from g_is_arr f_is_arr have ObjMap_dom_lhs: "\\<^sub>\ ((?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_comma_cs_intros cat_lim_cs_intros cat_cs_intros ) from g_is_arr f_is_arr have ObjMap_dom_rhs: "\\<^sub>\ (?const_comma c a\ObjMap\) = c \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps) show "(?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\ = ?const_comma c a\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) from f_is_arr g_is_arr show "vsv ((?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\)" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros cat_comma_cs_intros ) fix A assume prems: "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" with g_is_arr obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto from prems b' f' g_is_arr f_is_arr show "(?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\\A\ = ?const_comma c a\ObjMap\\A\" unfolding A_def by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros cat_comma_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) from g_is_arr f_is_arr have ArrMap_dom_lhs: "\\<^sub>\ ((?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\) = c \\<^sub>C\<^sub>F \\Arr\" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_comma_cs_intros cat_lim_cs_intros cat_cs_intros ) from g_is_arr f_is_arr have ArrMap_dom_rhs: "\\<^sub>\ (?const_comma c a\ArrMap\) = c \\<^sub>C\<^sub>F \\Arr\" by (cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps) show "(?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\ = ?const_comma c a\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) from f_is_arr g_is_arr show "vsv ((?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\)" by ( cs_concl cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros cat_comma_cs_intros ) fix F assume "F \\<^sub>\ c \\<^sub>C\<^sub>F \\Arr\" then obtain A B where F: "F : A \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> B" unfolding cat_comma_cs_simps by (auto intro: is_arrI) with g_is_arr obtain b' f' b'' f'' h' where F_def: "F = [[0, b', f']\<^sub>\, [0, b'', f'']\<^sub>\, [0, h']\<^sub>\]\<^sub>\" and A_def: "A = [0, b', f']\<^sub>\" and B_def: "B = [0, b'', f'']\<^sub>\" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'': "f'' : c \\<^bsub>\\<^esub> \\ObjMap\\b''\" and f''_def: "\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" by auto from F f_is_arr g_is_arr g' h' f' f'' show "(?const_comma a a \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\\F\ = ?const_comma c a\ArrMap\\F\" unfolding F_def A_def B_def by ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cat_comma_cs_intros cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric] ) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all from that show "?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c ?the_rKe_gf : cf_const (a \\<^sub>C\<^sub>F \) \ (?UObj a) \\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ((g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \) : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_comma_cs_simps cat_cs_simps cs_intro: cat_comma_cs_intros cat_Kan_cs_intros cat_cs_intros ) from that have dom_lhs: "\\<^sub>\ ((?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cs_simp: cat_cs_simps cat_comma_cs_simps ) from that have dom_rhs: "\\<^sub>\ ((?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c ?the_rKe_gf)\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_intro: cat_cs_intros cat_Kan_cs_intros cat_comma_cs_intros cs_simp: cat_Kan_cs_simps cat_cs_simps cat_comma_cs_simps ) show "(?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\ = (?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c ?the_rKe_gf)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix A assume prems: "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" with g_is_arr obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto note \.HomCod.cat_Comp_assoc[cat_cs_simps del] and \.HomCod.cat_Comp_assoc[cat_cs_simps del] and category.cat_Comp_assoc[cat_cs_simps del] note [symmetric, cat_cs_simps] = lim_Obj_the_cf_rKe_commute[where lim_Obj=lim_Obj] \.HomCod.cat_Comp_assoc \.HomCod.cat_Comp_assoc from assms(1,2) that prems lim_a lim_b lim_c b' f' show "(?UArr a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (g \\<^sub>A\<^bsub>\\<^esub> f) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\\A\ = (?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c ?the_rKe_gf)\NTMap\\A\" unfolding A_def by (*very slow*) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros cat_comma_cs_intros )+ qed (cs_concl cs_simp: cs_intro: cat_cs_intros)+ qed simp_all qed qed show "?the_cf_rKe\ArrMap\\\\CId\\c\\ = \\CId\\?the_cf_rKe\ObjMap\\c\\" if "c \\<^sub>\ \\Obj\" for c proof- let ?ntcf_const_c = \ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\CId\\?UObj c\)\ note lim_c = assms(3)[OF that] from that have CId_c: "\\CId\\c\ : c \\<^bsub>\\<^esub> c" by (cs_concl cs_intro: cat_cs_intros) interpret lim_c: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj c\ \?UArr c\ by (rule lim_c) show ?thesis proof ( rule sym, rule the_cf_rKe_ArrMap_app(3)[ where lim_Obj=lim_Obj, OF assms(1,2) CId_c lim_c lim_c ] ) from that lim_c show "\\CId\\?the_cf_rKe\ObjMap\\c\\ : ?UObj c \\<^bsub>\\<^esub> ?UObj c" by ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros ) have "?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c" proof(rule ntcf_eqI) from lim_c that show "?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ : cf_const (c \\<^sub>C\<^sub>F \) \ (?UObj c) \\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_comma_cs_intros) from lim_c that show "?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c : cf_const (c \\<^sub>C\<^sub>F \) \ (?UObj c) \\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - by + by (*very slow*) ( cs_concl - cs_intro: cat_cs_intros cat_lim_cs_intros - cs_simp: \.cf_cf_arr_comma_CId cat_cs_simps + cs_intro: cat_cs_intros + cs_simp: \.cf_cf_arr_comma_CId cat_cs_simps + cs_intro: cat_lim_cs_intros ) from that have dom_lhs: "\\<^sub>\ ((?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from that have dom_rhs: "\\<^sub>\ ((?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c)\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cs_simp: cat_cs_simps ) show "(?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\ = (?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix A assume prems: "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" with that obtain b f where A_def: "A = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from that prems f have "?UArr c\NTMap\\0, b, f\\<^sub>\ : ?UObj c \\<^bsub>\\<^esub> \\ObjMap\\b\" unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) from that prems f show "(?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \)\NTMap\\A\ = (?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_const_c)\NTMap\\A\" unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_lim_cs_intros cat_comma_cs_intros cat_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all with that show "?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\\CId\\c\) \<^sub>A\\<^sub>C\<^sub>F \ = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\CId\\?the_cf_rKe\ObjMap\\c\\)" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) qed qed qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma the_ntcf_rKe_is_ntcf: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : lim_Obj c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "the_ntcf_rKe \ \ \ lim_Obj : the_cf_rKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?UObj = \\a. lim_Obj a\UObj\\ let ?UArr = \\a. lim_Obj a\UArr\\ let ?const_comma = \\a b. cf_const (a \\<^sub>C\<^sub>F \) \ (?UObj b)\ let ?the_cf_rKe = \the_cf_rKe \ \ \ lim_Obj\ let ?the_ntcf_rKe = \the_ntcf_rKe \ \ \ lim_Obj\ - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret cf_rKe: is_functor \ \ \ \?the_cf_rKe\ by (rule the_cf_rKe_is_functor[OF assms, simplified]) show ?thesis proof(rule is_ntcfI') show "vfsequence ?the_ntcf_rKe" unfolding the_ntcf_rKe_def by simp show "vcard ?the_ntcf_rKe = 5\<^sub>\" unfolding the_ntcf_rKe_def by (simp add: nat_omega_simps) show "?the_ntcf_rKe\NTMap\\b\ : (?the_cf_rKe \\<^sub>C\<^sub>F \)\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" for b proof- let ?\b = \\\ObjMap\\b\\ from that have \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" by (cs_concl cs_intro: cat_cs_intros) note lim_\b = assms(3)[OF \b] interpret lim_\b: is_cat_limit \ \?\b \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F ?\b \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj ?\b\ \?UArr ?\b\ by (rule lim_\b) from that lim_\b show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros )+ qed show "?the_ntcf_rKe\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> (?the_cf_rKe \\<^sub>C\<^sub>F \)\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> ?the_ntcf_rKe\NTMap\\a\" if "f : a \\<^bsub>\\<^esub> b" for a b f proof- let ?\a = \\\ObjMap\\a\\ and ?\b = \\\ObjMap\\b\\ and ?\f = \\\ArrMap\\f\\ from that have \a: "?\a \\<^sub>\ \\Obj\" and \b: "?\b \\<^sub>\ \\Obj\" and \f: "?\f : ?\a \\<^bsub>\\<^esub> ?\b" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ note lim_\a = assms(3)[OF \a] and lim_\b = assms(3)[OF \b] from that have z_b_\b: "[0, b, \\CId\\?\b\]\<^sub>\ \\<^sub>\ ?\b \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) from lim_Obj_the_cf_rKe_commute[ OF assms(1,2) lim_\a lim_\b \f z_b_\b, symmetric ] that have [cat_Kan_cs_simps]: "?UArr ?\b\NTMap\\0, b, \\CId\\?\b\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> ?the_cf_rKe\ArrMap\\?\f\ = ?UArr ?\a\NTMap\\0, b, ?\f\\<^sub>\" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) interpret lim_\a: is_cat_limit \ \?\a \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F ?\a \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj ?\a\ \?UArr ?\a\ by (rule lim_\a) interpret lim_\b: is_cat_limit \ \?\b \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F ?\b \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj ?\b\ \?UArr ?\b\ by (rule lim_\b) from that have "[[0, a, \\CId\\?\a\]\<^sub>\, [0, b, ?\f]\<^sub>\, [0, f]\<^sub>\]\<^sub>\ : [0, a, \\CId\\?\a\]\<^sub>\ \\<^bsub>(?\a) \\<^sub>C\<^sub>F \\<^esub> [0, b, ?\f]\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from lim_\a.ntcf_Comp_commute[OF this, symmetric] that have [cat_Kan_cs_simps]: "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> ?UArr (?\a)\NTMap\ \0, a, \\CId\\?\a\\\<^sub>\ = ?UArr ?\a\NTMap\\0, b, ?\f\\<^sub>\" by ( cs_prems cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros \.cat_1_is_arrI ) from that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros ) qed qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma the_ntcf_rKe_is_cat_rKe: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : lim_Obj c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "the_ntcf_rKe \ \ \ lim_Obj : the_cf_rKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" proof- let ?UObj = \\a. lim_Obj a\UObj\\ let ?UArr = \\a. lim_Obj a\UArr\\ let ?the_cf_rKe = \the_cf_rKe \ \ \ lim_Obj\ let ?the_ntcf_rKe = \the_ntcf_rKe \ \ \ lim_Obj\ - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret cf_rKe: is_functor \ \ \ ?the_cf_rKe by (rule the_cf_rKe_is_functor[OF assms, simplified]) interpret ntcf_rKe: is_ntcf \ \ \ \?the_cf_rKe \\<^sub>C\<^sub>F \\ \ ?the_ntcf_rKe by (intro the_ntcf_rKe_is_ntcf assms(3)) - (cs_concl cs_intro: cat_cs_intros cat_small_cs_intros)+ + (cs_concl cs_intro: cat_cs_intros)+ show ?thesis proof(rule is_cat_rKeI') fix \ \ assume prems: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" interpret \: is_functor \ \ \ \ by (rule prems(1)) interpret \: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \ \ by (rule prems(2)) define \' where "\' c = [ (\A\\<^sub>\c \\<^sub>C\<^sub>F \\Obj\. \\NTMap\\A\1\<^sub>\\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\A\2\<^sub>\\\), cf_const (c \\<^sub>C\<^sub>F \) \ (\\ObjMap\\c\), \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \, c \\<^sub>C\<^sub>F \, \ ]\<^sub>\" for c have \'_components: "\' c\NTMap\ = (\A\\<^sub>\c \\<^sub>C\<^sub>F \\Obj\. \\NTMap\\A\1\<^sub>\\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\A\2\<^sub>\\\)" "\' c\NTDom\ = cf_const (c \\<^sub>C\<^sub>F \) \ (\\ObjMap\\c\)" "\' c\NTCod\ = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \" "\' c\NTDGDom\ = c \\<^sub>C\<^sub>F \" "\' c\NTDGCod\ = \" for c unfolding \'_def nt_field_simps by (simp_all add: nat_omega_simps) note [cat_Kan_cs_simps] = \'_components(2-5) have [cat_Kan_cs_simps]: "\' c\NTMap\\A\ = \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" if "A = [a, b, f]\<^sub>\" and "[a, b, f]\<^sub>\ \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" for A a b c f using that unfolding \'_components by (auto simp: nat_omega_simps) have \': "\' c : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and \'_unique: "\!f'. f' : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f'" if c: "c \\<^sub>\ \\Obj\" for c proof- from that have "?the_cf_rKe\ObjMap\\c\ = ?UObj c" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) interpret lim_c: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj c\ \?UArr c\ by (rule assms(3)[OF that]) show "\' c : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') + proof(intro is_cat_coneI is_ntcfI') show "vfsequence (\' c)" unfolding \'_def by simp show "vcard (\' c) = 5\<^sub>\" unfolding \'_def by (simp add: nat_omega_simps) show "vsv (\' c\NTMap\)" unfolding \'_components by simp show "\\<^sub>\ (\' c\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" unfolding \'_components by simp show "\' c\NTMap\\A\ : cf_const (c \\<^sub>C\<^sub>F \) \ (\\ObjMap\\c\)\ObjMap\\A\ \\<^bsub>\\<^esub> (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\A\" if "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" for A proof- from that prems c obtain b f where A_def: "A = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from that prems f c that b f show ?thesis unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) qed show "\' c\NTMap\\B\ \\<^sub>A\<^bsub>\\<^esub> cf_const (c \\<^sub>C\<^sub>F \) \ (\\ObjMap\\c\)\ArrMap\\F\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\\F\ \\<^sub>A\<^bsub>\\<^esub> \' c\NTMap\\A\" if "F : A \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> B" for A B F proof- from that c obtain b f b' f' k where F_def: "F = [[0, b, f]\<^sub>\, [0, b', f']\<^sub>\, [0, k]\<^sub>\]\<^sub>\" and A_def: "A = [0, b, f]\<^sub>\" and B_def: "B = [0, b', f']\<^sub>\" and k: "k : b \\<^bsub>\\<^esub> b'" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'_def: "\\ArrMap\\k\ \\<^sub>A\<^bsub>\\<^esub> f = f'" by auto from c that k f f' show ?thesis unfolding F_def A_def B_def by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps \.ntcf_Comp_commute'' f'_def[symmetric] cs_intro: cat_cs_intros cat_comma_cs_intros ) qed qed ( use c that in - \ - cs_concl - cs_simp: cat_Kan_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros - \ + \cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros\ )+ - from is_cat_limit.cat_lim_unique_cone[OF assms(3)[OF that] this] show + from is_cat_limit.cat_lim_ua_fo[OF assms(3)[OF that] this] show "\!f'. f' : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f'" by simp qed define \ :: V where "\ = [ ( \c\\<^sub>\\\Obj\. THE f. f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f ), \, ?the_cf_rKe, \, \ ]\<^sub>\" have \_components: "\\NTMap\ = ( \c\\<^sub>\\\Obj\. THE f. f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f )" "\\NTDom\ = \" "\\NTCod\ = ?the_cf_rKe" "\\NTDGDom\ = \" "\\NTDGCod\ = \" unfolding \_def nt_field_simps by (simp_all add: nat_omega_simps) note [cat_Kan_cs_simps] = \_components(2-5) have \_NTMap_app_def: "\\NTMap\\c\ = ( THE f. f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f )" if "c \\<^sub>\ \\Obj\" for c using that unfolding \_components by simp have \_NTMap_app_is_arr: "\\NTMap\\c\ : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c" and \'_\_commute: "\' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\NTMap\\c\)" and \_NTMap_app_unique: "\ f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c; \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f \ \ f = \\NTMap\\c\" if c: "c \\<^sub>\ \\Obj\" for c f proof- have "\\NTMap\\c\ : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c \ \' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\NTMap\\c\)" by ( cs_concl cs_simp: cat_Kan_cs_simps \_NTMap_app_def cs_intro: theI' \'_unique that ) then show "\\NTMap\\c\ : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c" and "\' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\NTMap\\c\)" by simp_all with c \'_unique[OF c] show "f = \\NTMap\\c\" if "f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c" and "\' c = ?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ f" using that by metis qed have \_NTMap_app_is_arr'[cat_Kan_cs_intros]: "\\NTMap\\c\ : a \\<^bsub>\'\<^esub> b" if "c \\<^sub>\ \\Obj\" and "a = \\ObjMap\\c\" and "b = ?UObj c" and "\' = \" for \' a b c by (simp add: that \_NTMap_app_is_arr) have \'_NTMap_app_def: "\' c\NTMap\\A\ = (?UArr c \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (c \\<^sub>C\<^sub>F \) \ (\\NTMap\\c\))\NTMap\\A\" if "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" and "c \\<^sub>\ \\Obj\" for A c using \'_\_commute[OF that(2)] by simp have \b_\f: "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = ?UArr c\NTMap\\a, b, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\c\" if "A = [a, b, f]\<^sub>\" and "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" and "c \\<^sub>\ \\Obj\" for A a b c f proof- interpret lim_c: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj c\ \?UArr c\ by (rule assms(3)[OF that(3)]) from that have b: "b \\<^sub>\ \\Obj\" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" by blast+ show "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = ?UArr c\NTMap\\a, b, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\c\" using \'_NTMap_app_def[OF that(2,3)] that(2,3) unfolding that(1) by ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) qed show "\!\. \ : \ \\<^sub>C\<^sub>F ?the_cf_rKe : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ = ?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" proof(intro ex1I[where a=\] conjI; (elim conjE)?) define \ where "\ a b f = [ ( \F\\<^sub>\b \\<^sub>C\<^sub>F \\Obj\. ?UArr b\NTMap\\F\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ ), cf_const (b \\<^sub>C\<^sub>F \) \ (\\ObjMap\\a\), \ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \, b \\<^sub>C\<^sub>F \, \ ]\<^sub>\" for a b f have \_components: "\ a b f\NTMap\ = ( \F\\<^sub>\b \\<^sub>C\<^sub>F \\Obj\. ?UArr b\NTMap\\F\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ )" "\ a b f\NTDom\ = cf_const (b \\<^sub>C\<^sub>F \) \ (\\ObjMap\\a\)" "\ a b f\NTCod\ = \ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \" "\ a b f\NTDGDom\ = b \\<^sub>C\<^sub>F \" "\ a b f\NTDGCod\ = \" for a b f unfolding \_def nt_field_simps by (simp_all add: nat_omega_simps) note [cat_Kan_cs_simps] = \_components(2-5) have \_NTMap_app[cat_Kan_cs_simps]: "\ a b f\NTMap\\F\ = ?UArr b\NTMap\\F\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" if "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" for a b f F using that unfolding \_components by auto have \: "\ a b f : \\ObjMap\\a\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" if f_is_arr: "f : a \\<^bsub>\\<^esub> b" for a b f proof- note f = \.HomCod.cat_is_arrD[OF that] note lim_a = assms(3)[OF f(2)] and lim_b = assms(3)[OF f(3)] interpret lim_b: is_cat_limit \ \b \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj b\ \?UArr b\ by (rule lim_b) from f have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by auto show ?thesis - proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') + proof(intro is_cat_coneI is_ntcfI') show "vfsequence (\ a b f)" unfolding \_def by simp show "vcard (\ a b f) = 5\<^sub>\" unfolding \_def by (simp add: nat_omega_simps) show "vsv (\ a b f\NTMap\)" unfolding \_components by auto show "\\<^sub>\ (\ a b f\NTMap\) = b \\<^sub>C\<^sub>F \\Obj\" by (auto simp: \_components) show "\ a b f\NTMap\\A\ : cf_const (b \\<^sub>C\<^sub>F \) \ (\\ObjMap\\a\)\ObjMap\\A\ \\<^bsub>\\<^esub> (\ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\A\" if "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" for A proof- from that f_is_arr obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto from f_is_arr that b' f' a b show ?thesis unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros ) qed show "\ a b f\NTMap\\B\ \\<^sub>A\<^bsub>\\<^esub> cf_const (b \\<^sub>C\<^sub>F \) \ (\\ObjMap\\a\)\ArrMap\\F\ = (\ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\\F\ \\<^sub>A\<^bsub>\\<^esub> \ a b f\NTMap\\A\" if "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" for A B F proof- from that have F: "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" by (auto intro: is_arrI) with f_is_arr obtain b' f' b'' f'' h' where F_def: "F = [[0, b', f']\<^sub>\, [0, b'', f'']\<^sub>\, [0, h']\<^sub>\]\<^sub>\" and A_def: "A = [0, b', f']\<^sub>\" and B_def: "B = [0, b'', f'']\<^sub>\" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'': "f'' : b \\<^bsub>\\<^esub> \\ObjMap\\b''\" and f''_def: "\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" by auto from lim_b.ntcf_Comp_commute[OF that] that f_is_arr g' h' f' f'' have [cat_Kan_cs_simps]: "?UArr b\NTMap\\0, b'', \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\\<^sub>\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> ?UArr b\NTMap\\0, b', f'\\<^sub>\" unfolding F_def A_def B_def by ( cs_prems cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric] cs_intro: cat_cs_intros cat_comma_cs_intros ) from f_is_arr that g' h' f' f'' show ?thesis unfolding F_def A_def B_def (*very slow*) by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps f''_def[symmetric] cs_intro: cat_cs_intros cat_Kan_cs_intros cat_comma_cs_intros )+ qed qed ( use that f_is_arr in \ cs_concl - cs_simp: cat_cs_simps cat_Kan_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros + cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros \ )+ qed show \: "\ : \ \\<^sub>C\<^sub>F ?the_cf_rKe : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_ntcfI') show "vfsequence \" unfolding \_def by simp show "vcard \ = 5\<^sub>\" unfolding \_def by (simp add: nat_omega_simps) show "vsv (\\NTMap\)" unfolding \_components by auto show "\\<^sub>\ (\\NTMap\) = \\Obj\" unfolding \_components by simp show "\\NTMap\\a\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> ?the_cf_rKe\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using that by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) then have [cat_Kan_cs_intros]: "\\NTMap\\a\ : b \\<^bsub>\\<^esub> c" if "a \\<^sub>\ \\Obj\" and "b = \\ObjMap\\a\" and "c = ?the_cf_rKe\ObjMap\\a\" for a b c using that(1) unfolding that(2,3) by simp show "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = ?the_cf_rKe\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" if f_is_arr: "f : a \\<^bsub>\\<^esub> b" for a b f proof- note f = \.HomCod.cat_is_arrD[OF that] note lim_a = assms(3)[OF f(2)] and lim_b = assms(3)[OF f(3)] interpret lim_a: is_cat_limit \ \a \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F a \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj a\ \?UArr a\ by (rule lim_a) interpret lim_b: is_cat_limit \ \b \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F b \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj b\ \?UArr b\ by (rule lim_b) from f have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by auto from lim_b.cat_lim_unique_cone'[OF \[OF that]] obtain g' where g': "g' : \\ObjMap\\a\ \\<^bsub>\\<^esub> ?UObj b" and \_NTMap_app: "\A. A \\<^sub>\ (b \\<^sub>C\<^sub>F \\Obj\) \ \ a b f\NTMap\\A\ = ?UArr b\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> g'" and g'_unique: "\g''. \ g'' : \\ObjMap\\a\ \\<^bsub>\\<^esub> ?UObj b; \A. A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\ \ \ a b f\NTMap\\A\ = ?UArr b\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> g'' \ \ g'' = g'" by metis have lim_Obj_a_f\[symmetric, cat_Kan_cs_simps]: "?UArr a\NTMap\\a', b', f' \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\ = ?UArr b\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> ?the_cf_rKe\ArrMap\\f\" if "A = [a', b', f']\<^sub>\" and "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" for A a' b' f' proof- from that(2) f_is_arr have a'_def: "a' = 0" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" unfolding that(1) by auto show ?thesis unfolding that(1) by ( rule lim_Obj_the_cf_rKe_commute [ where lim_Obj=lim_Obj, OF assms(1,2) lim_a lim_b f_is_arr that(2)[unfolded that(1)] ] ) qed { fix a' b' f' A note \.HomCod.cat_assoc_helper[ where h=\?UArr b\NTMap\\a',b',f'\\<^sub>\\ and g=\?the_cf_rKe\ArrMap\\f\\ and q=\?UArr a\NTMap\\a', b', f' \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\\ ] } note [cat_Kan_cs_simps] = this show ?thesis proof(rule trans_sym[where s=g']) show "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = g'" proof(rule g'_unique) from that show "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> ?UObj b" by (cs_concl cs_intro: cat_cs_intros cat_Kan_cs_intros) fix A assume prems': "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" with f_is_arr obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto from f_is_arr prems' show "\ a b f\NTMap\\A\ = ?UArr b\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\)" unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) qed show "?the_cf_rKe\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = g'" proof(rule g'_unique) fix A assume prems': "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" with f_is_arr obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto { fix a' b' f' A note \.HomCod.cat_assoc_helper [ where h=\?UArr b\NTMap\\a', b', f'\\<^sub>\\ and g=\\\NTMap\\b\\ and q=\\\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\\ ] } note [cat_Kan_cs_simps] = this \b_\f[OF A_def prems' b, symmetric] \b_\f[symmetric] from f_is_arr prems' b' f' show "\ a b f\NTMap\\A\ = ?UArr b\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> (?the_cf_rKe\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\)" unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_Kan_cs_intros cat_comma_cs_intros cat_op_intros ) qed ( use that in \ cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros \ ) qed qed qed ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros )+ then interpret \: is_ntcf \ \ \ \ \?the_cf_rKe\ \ by simp show "\ = ?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" proof(rule ntcf_eqI) have dom_lhs: "\\<^sub>\ (\\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) have dom_rhs: "\\<^sub>\ ((?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\\NTMap\ = (?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix b assume prems': "b \\<^sub>\ \\Obj\" note [cat_Kan_cs_simps] = \b_\f[ where f=\\\CId\\\\ObjMap\\b\\\ and c=\\\ObjMap\\b\\, symmetric ] from prems' \ show "\\NTMap\\b\ = (?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\\b\" by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros V_cs_intros) qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ fix \' assume prems': "\' : \ \\<^sub>C\<^sub>F ?the_cf_rKe : \ \\\<^sub>C\<^bsub>\\<^esub> \" "\ = ?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" interpret \': is_ntcf \ \ \ \ \?the_cf_rKe\ \' by (rule prems'(1)) have \_NTMap_app[symmetric, cat_Kan_cs_simps]: "\\NTMap\\b'\ = ?UArr (\\ObjMap\\b'\)\NTMap\\a', b', \\CId\\\\ObjMap\\b'\\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \'\NTMap\\\\ObjMap\\b'\\" if "b' \\<^sub>\ \\Obj\" and "a' = 0" for a' b' proof- from prems'(2) have \_NTMap_app: "\\NTMap\\b'\ = (?the_ntcf_rKe \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\NTMap\\b'\" for b' by simp show ?thesis using \_NTMap_app[of b'] that(1) unfolding that(2) by ( cs_prems cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) qed { fix a' b' f' A note \.HomCod.cat_assoc_helper [ where h= \?UArr (\\ObjMap\\b'\)\NTMap\\a', b', \\CId\\\\ObjMap\\b'\\\\<^sub>\\ and g=\\'\NTMap\\\\ObjMap\\b'\\\ and q=\\\NTMap\\b'\\ ] } note [cat_Kan_cs_simps] = this \b_\f[symmetric] { fix a' b' f' A note \.HomCod.cat_assoc_helper [ where h=\ ?UArr (\\ObjMap\\b'\)\NTMap\\ a', b', \\CId\\\\ObjMap\\b'\\ \\<^sub>\\ and g=\\\NTMap\\\\ObjMap\\b'\\\ and q=\\\NTMap\\b'\\ ] } note [cat_Kan_cs_simps] = this show "\' = \" proof(rule ntcf_eqI) show "\' : \ \\<^sub>C\<^sub>F ?the_cf_rKe : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule prems'(1)) show "\ : \ \\<^sub>C\<^sub>F ?the_cf_rKe : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule \) have dom_lhs: "\\<^sub>\ (\\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) have dom_rhs: "\\<^sub>\ (\'\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "\'\NTMap\ = \\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix c assume prems': "c \\<^sub>\ \\Obj\" note lim_c = assms(3)[OF prems'] interpret lim_c: is_cat_limit \ \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj c\ \?UArr c\ by (rule lim_c) from prems' have CId_c: "\\CId\\c\ : c \\<^bsub>\\<^esub> c" by (cs_concl cs_intro: cat_cs_intros) from lim_c.cat_lim_unique_cone'[OF \[OF CId_c]] obtain f where f: "f : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c" and "\A. A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\ \ \ c c (\\CId\\c\)\NTMap\\A\ = ?UArr c\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> f" and f_unique: "\f'. \ f' : \\ObjMap\\c\ \\<^bsub>\\<^esub> ?UObj c; \A. A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\ \ \ c c (\\CId\\c\)\NTMap\\A\ = ?UArr c\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> f' \ \ f' = f" by metis note [symmetric, cat_cs_simps] = \.ntcf_Comp_commute \'.ntcf_Comp_commute show "\'\NTMap\\c\ = \\NTMap\\c\" proof(rule trans_sym[where s=f]) show "\'\NTMap\\c\ = f" proof(rule f_unique) fix A assume prems'': "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" with prems' obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto let ?\b' = \\\ObjMap\\b'\\ from b' have \b': "?\b' \\<^sub>\ \\Obj\" by (cs_concl cs_intro: cat_cs_intros) interpret lim_\b': is_cat_limit \ \?\b' \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F ?\b' \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj ?\b'\ \?UArr ?\b'\ by (rule assms(3)[OF \b']) from \b' have CId_\b': "\\CId\\?\b'\ : ?\b' \\<^bsub>\\<^esub> ?\b'" by (cs_concl cs_intro: cat_cs_intros) from CId_\b' b' have a'_b'_CId_\b': "[0, b', \\CId\\?\b'\]\<^sub>\ \\<^sub>\ ?\b' \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from lim_Obj_the_cf_rKe_commute[ where lim_Obj=lim_Obj, OF assms(1,2) lim_c assms(3)[OF \b'] f' a'_b'_CId_\b' ] f' have [cat_Kan_cs_simps]: "?UArr c\NTMap\\0, b', f'\\<^sub>\ = ?UArr ?\b'\NTMap\\0, b', \\CId\\?\b'\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> ?the_cf_rKe\ArrMap\\f'\" by (cs_prems cs_simp: cat_cs_simps) from prems' prems'' b' f' show "\ c c (\\CId\\c\)\NTMap\\A\ = ?UArr c\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> \'\NTMap\\c\" unfolding A_def (*very slow*) by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros ) qed ( use prems' in \cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros\ ) show "\\NTMap\\c\ = f" proof(rule f_unique) fix A assume prems'': "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" from this prems' obtain b' f' where A_def: "A = [0, b', f']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" by auto let ?\b' = \\\ObjMap\\b'\\ from b' have \b': "?\b' \\<^sub>\ \\Obj\" by (cs_concl cs_intro: cat_cs_intros) interpret lim_\b': is_cat_limit \ \?\b' \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F ?\b' \<^sub>O\\<^sub>C\<^sub>F \\ \?UObj ?\b'\ \?UArr ?\b'\ by (rule assms(3)[OF \b']) from \b' have CId_\b': "\\CId\\?\b'\ : ?\b' \\<^bsub>\\<^esub> ?\b'" by (cs_concl cs_intro: cat_cs_intros) from CId_\b' b' have a'_b'_CId_\b': "[0, b', \\CId\\?\b'\]\<^sub>\ \\<^sub>\ ?\b' \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) from lim_Obj_the_cf_rKe_commute [ where lim_Obj=lim_Obj, OF assms(1,2) lim_c assms(3)[OF \b'] f' a'_b'_CId_\b' ] f' have [cat_Kan_cs_simps]: "?UArr c\NTMap\\0, b', f'\\<^sub>\ = ?UArr (?\b')\NTMap\\0, b', \\CId\\?\b'\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> ?the_cf_rKe\ArrMap\\f'\" by (cs_prems cs_simp: cat_cs_simps) from prems' prems'' b' f' show "\ c c (\\CId\\c\)\NTMap\\A\ = ?UArr c\NTMap\\A\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\c\" unfolding A_def (*very slow*) by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros cat_Kan_cs_intros ) qed ( use prems' in \cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros\ ) qed qed auto qed simp_all qed qed (cs_concl cs_intro: cat_cs_intros)+ qed subsection\Preservation of Kan extension\ text\ The following definitions are similar to the definitions that can be found in \cite{riehl_category_2016} or \cite{lehner_all_2014}. \ locale is_cat_rKe_preserves = is_cat_rKe \ \ \ \ \ \ \ \ + is_functor \ \ \
\ for \ \ \ \ \
\ \ \ \ \ + assumes cat_rKe_preserves: "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : (\ \\<^sub>C\<^sub>F \) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \
" syntax "_is_cat_rKe_preserves" :: "V \ V \ V \ V \ V \ V \ V \ V \ V \ V \ bool" ( \(_ :/ _ \\<^sub>C\<^sub>F _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\ _ :/ _ \\<^sub>C _ \\<^sub>C _ : _ \\\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51, 51, 51] 51 ) translations "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C (\ : \ \\\<^sub>C \
)" \ "CONST is_cat_rKe_preserves \ \ \ \ \
\ \ \ \ \" locale is_cat_lKe_preserves = is_cat_lKe \ \ \ \ \ \ \ \ + is_functor \ \ \
\ for \ \ \ \ \
\ \ \ \ \ + assumes cat_lKe_preserves: "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> (\ \\<^sub>C\<^sub>F \) \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \
" syntax "_is_cat_lKe_preserves" :: "V \ V \ V \ V \ V \ V \ V \ V \ V \ V \ bool" ( \(_ :/ _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\ _ \\<^sub>C\<^sub>F _ :/ _ \\<^sub>C _ \\<^sub>C _ : _ \\\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51, 51, 51] 51 ) translations "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C (\ : \ \\\<^sub>C \
)" \ "CONST is_cat_lKe_preserves \ \ \ \ \
\ \ \ \ \" text\Rules.\ lemma (in is_cat_rKe_preserves) is_cat_rKe_preserves_axioms': assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\
' = \
" shows "\ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\'\<^esub> \' : \' \\<^sub>C \' \\<^sub>C (\' : \' \\\<^sub>C \
')" unfolding assms by (rule is_cat_rKe_preserves_axioms) mk_ide rf is_cat_rKe_preserves_def[unfolded is_cat_rKe_preserves_axioms_def] |intro is_cat_rKe_preservesI| |dest is_cat_rKe_preservesD[dest]| |elim is_cat_rKe_preservesE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_rKeD(1-3) lemma (in is_cat_lKe_preserves) is_cat_lKe_preserves_axioms': assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\
' = \
" shows "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C (\' : \' \\\<^sub>C \
')" unfolding assms by (rule is_cat_lKe_preserves_axioms) mk_ide rf is_cat_lKe_preserves_def[unfolded is_cat_lKe_preserves_axioms_def] |intro is_cat_lKe_preservesI| |dest is_cat_lKe_preservesD[dest]| |elim is_cat_lKe_preservesE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_lKe_preservesD(1-3) text\Duality.\ lemma (in is_cat_rKe_preserves) is_cat_rKe_preserves_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C (op_cf \ : op_cat \ \\\<^sub>C op_cat \
)" proof(intro is_cat_lKe_preservesI) from cat_rKe_preserves show "op_cf \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> (op_cf \ \\<^sub>C\<^sub>F op_cf \) \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \
" by (cs_concl_step op_ntcf_cf_ntcf_comp[symmetric]) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros) qed (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_rKe_preserves) is_cat_lKe_preserves_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" and "\
' = op_cat \
" and "\' = op_cf \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C (\' : \' \\\<^sub>C \
')" unfolding assms by (rule is_cat_rKe_preserves_op) lemmas [cat_op_intros] = is_cat_rKe_preserves.is_cat_lKe_preserves_op' lemma (in is_cat_lKe_preserves) is_cat_rKe_preserves_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C (op_cf \ : op_cat \ \\\<^sub>C op_cat \
)" proof(intro is_cat_rKe_preservesI) from cat_lKe_preserves show "op_cf \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \ : (op_cf \ \\<^sub>C\<^sub>F op_cf \) \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \
" by (cs_concl_step op_ntcf_cf_ntcf_comp[symmetric]) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros) qed (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_lKe_preserves) is_cat_rKe_preserves_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" and "\
' = op_cat \
" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' : \' \\<^sub>C \' \\<^sub>C (\' : \' \\\<^sub>C \
')" unfolding assms by (rule is_cat_rKe_preserves_op) subsection\All concepts are Kan extensions\ text\ Background information for this subsection is provided in Chapter X-7 in \cite{mac_lane_categories_2010} and section 6.5 in \cite{riehl_category_2016}. It should be noted that only the connections between the Kan extensions, limits and adjunctions are exposed (an alternative proof of the Yoneda lemma using Kan extensions is not provided in the context of this work). \ subsubsection\Limits\ lemma cat_rKe_is_cat_limit: \\The statement of the theorem is similar to the statement of a part of Theorem 1 in Chapter X-7 in \cite{mac_lane_categories_2010} or Proposition 6.5.1 in \cite{riehl_category_2016}.\ assumes "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C cat_1 \ \ \\<^sub>C \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \\ObjMap\\\\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_cat_rKe \ \ \cat_1 \ \\ \ \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) from cat_1_components(1) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.AG.HomCod.cat_in_Obj_in_Vset) from cat_1_components(2) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.AG.HomCod.cat_in_Arr_in_Vset) have \_def: "\ = cf_const \ (cat_1 \ \) \" by (rule cf_const_if_HomCod_is_cat_1) (cs_concl cs_intro: cat_cs_intros) have \\_def: "\ \\<^sub>C\<^sub>F \ = cf_const \ \ (\\ObjMap\\\\)" by ( cs_concl cs_simp: cat_1_components(1) \_def cat_cs_simps cs_intro: V_cs_intros cat_cs_intros ) - interpret \: is_tm_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \ \ - by - ( - intro is_tm_ntcfI' assms(2) \.ntcf_rKe.is_ntcf_axioms, - unfold \\_def - ) - ( - cs_concl - cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros - ) + interpret \: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \ \ + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\ : \\ObjMap\\\\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - proof(intro is_cat_limitI' is_cat_coneI) + proof(intro is_cat_limitI is_cat_coneI) - show "\ : cf_const \ \ (\\ObjMap\\\\) \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - proof(intro is_tm_ntcfI' \.ntcf_rKe.is_ntcf_axioms[unfolded \\_def]) - from assms(2) show "cf_const \ \ (\\ObjMap\\\\) : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros) - qed (rule assms(2)) + show "\ : cf_const \ \ (\\ObjMap\\\\) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (rule \.ntcf_rKe.is_ntcf_axioms[unfolded \\_def]) fix u' r' assume prems: "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" interpret u': is_cat_cone \ r' \ \ \ u' by (rule prems) have \_def: "\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)" by (rule cf_const_if_HomDom_is_cat_1[OF \.Ran.is_functor_axioms]) from prems have const_r': "cf_const (cat_1 \ \) \ r' : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_lim_cs_intros cat_cs_intros ) have cf_comp_cf_const_r_\_def: "cf_const (cat_1 \ \) \ r' \\<^sub>C\<^sub>F \ = cf_const \ \ r'" by ( cs_concl cs_simp: cat_cs_simps \_def cs_intro: cat_cs_intros cat_lim_cs_intros ) from \.cat_rKe_unique[ OF const_r', unfolded cf_comp_cf_const_r_\_def, OF u'.is_ntcf_axioms ] obtain \ where \: "\ : cf_const (cat_1 \ \) \ r' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" and u'_def: "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" and unique_\: "\\'. \ \' : cf_const (cat_1 \ \) \ r' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \; u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \ \ \' = \" by auto interpret \: is_ntcf \ \cat_1 \ \\ \ \cf_const (cat_1 \ \) \ r'\ \ \ by (rule \) show "\!f'. f' : r' \\<^bsub>\\<^esub> \\ObjMap\\\\ \ u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" proof(intro ex1I conjI; (elim conjE)?) fix f' assume prems': "f' : r' \\<^bsub>\\<^esub> \\ObjMap\\\\" "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" from prems'(1) have "ntcf_const (cat_1 \ \) \ f' : cf_const (cat_1 \ \) \ r' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" by (subst \_def) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) moreover then have "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_const (cat_1 \ \) \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" by ( cs_concl cs_simp: cat_cs_simps prems'(2) \_def cs_intro: cat_cs_intros ) ultimately have \_def: "\ = ntcf_const (cat_1 \ \) \ f'" by (auto simp: unique_\[symmetric]) show "f' = \\NTMap\\\\" by (cs_concl cs_simp: cat_cs_simps \_def cs_intro: cat_cs_intros) qed (cs_concl cs_simp: cat_cs_simps u'_def \_def cs_intro: cat_cs_intros)+ qed (cs_concl cs_simp: \_def cs_intro: cat_cs_intros) qed lemma cat_lKe_is_cat_colimit: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C cat_1 \ \ \\<^sub>C \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m \\ObjMap\\\\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- - interpret \: is_cat_lKe \ \ \cat_1 \ \\ \ \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) + interpret \: is_cat_lKe \ \ \cat_1 \ \\ \ \ \ \ \ by (rule assms(1)) from cat_1_components(1) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.AG.HomCod.cat_in_Obj_in_Vset) from cat_1_components(2) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.AG.HomCod.cat_in_Arr_in_Vset) show ?thesis by ( rule is_cat_limit.is_cat_colimit_op [ OF cat_rKe_is_cat_limit[ - OF \.is_cat_rKe_op[unfolded \.AG.cat_1_op[OF \ \]] - \.is_tm_functor_op + OF + \.is_cat_rKe_op[unfolded \.AG.cat_1_op[OF \ \]] + \.ntcf_lKe.NTDom.is_functor_op ], unfolded cat_op_simps ] ) qed lemma cat_limit_is_rKe: \\The statement of the theorem is similar to the statement of a part of Theorem 1 in Chapter X-7 in \cite{mac_lane_categories_2010} or Proposition 6.5.1 in \cite{riehl_category_2016}.\ assumes "\ : \\ObjMap\\\\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_1 \ \" and "\ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C cat_1 \ \ \\<^sub>C \" proof- interpret \: is_cat_limit \ \ \ \ \\\ObjMap\\\\\ \ by (rule assms) interpret \: is_functor \ \ \cat_1 \ \\ \ by (rule assms(2)) interpret \: is_functor \ \cat_1 \ \\ \ \ by (rule assms(3)) show ?thesis proof(rule is_cat_rKeI') note \_def = cf_const_if_HomCod_is_cat_1[OF assms(2)] note \_def = cf_const_if_HomDom_is_cat_1[OF assms(3)] have \\_def: "\ \\<^sub>C\<^sub>F \ = cf_const \ \ (\\ObjMap\\\\)" by (subst \_def, use nothing in \subst \_def\) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps \\_def cs_intro: cat_cs_intros) + fix \' \' assume prems: "\' : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" "\' : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" interpret is_functor \ \cat_1 \ \\ \ \' by (rule prems(1)) note \'_def = cf_const_if_HomDom_is_cat_1[OF prems(1)] from prems(2) have \': "\' : cf_const \ \ (\'\ObjMap\\\\) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding \_def by (subst (asm) \'_def) (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from prems(2) have "\' : \'\ObjMap\\\\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (intro is_cat_coneI is_tm_ntcfI' \') - (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros)+ + by (intro is_cat_coneI \') (cs_concl cs_intro: cat_cs_intros)+ - from \.cat_lim_unique_cone[OF this] obtain f' + from \.cat_lim_ua_fo[OF this] obtain f' where f': "f' : \'\ObjMap\\\\ \\<^bsub>\\<^esub> \\ObjMap\\\\" and \_def: "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" and unique_f': "\ f'' : \'\ObjMap\\\\ \\<^bsub>\\<^esub> \\ObjMap\\\\; \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'' \ \ f'' = f'" for f'' by metis show "\!\. \ : \' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" proof(intro ex1I conjI; (elim conjE)?) from f' show "ntcf_const (cat_1 \ \) \ f' : \' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" by (subst \'_def, use nothing in \subst \_def\) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) then show "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_const (cat_1 \ \) \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" by (cs_concl cs_simp: cat_cs_simps \_def \_def cs_intro: cat_cs_intros) fix \ assume prems: "\ : \' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" interpret \: is_ntcf \ \cat_1 \ \\ \ \' \ \ by (rule prems(1)) have "\\NTMap\\\\ : \'\ObjMap\\\\ \\<^bsub>\\<^esub> \\ObjMap\\\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) moreover have "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ (\\NTMap\\\\)" by ( cs_concl cs_simp: cat_cs_simps prems(2) \_def cs_intro: cat_cs_intros ) ultimately have \\: "\\NTMap\\\\ = f'" by (rule unique_f') show "\ = ntcf_const (cat_1 \ \) \ f'" proof(rule ntcf_eqI) from f' show "ntcf_const (cat_1 \ \) \ f' : \' \\<^sub>C\<^sub>F \ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" by (subst \'_def, use nothing in \subst \_def\) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ (\\NTMap\) = cat_1 \ \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro:cat_cs_intros) have dom_rhs: "\\<^sub>\ (ntcf_const (cat_1 \ \) \ f'\NTMap\) = cat_1 \ \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro:cat_cs_intros) show "\\NTMap\ = ntcf_const (cat_1 \ \) \ f'\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ cat_1 \ \\Obj\" then have a_def: "a = \" unfolding cat_1_components by simp from f' show "\\NTMap\\a\ = ntcf_const (cat_1 \ \) \ f'\NTMap\\a\" unfolding a_def \\ by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: cat_cs_intros) qed (simp_all add: prems) qed + qed (auto simp: assms) qed lemma cat_colimit_is_lKe: assumes "\ : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m \\ObjMap\\\\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_1 \ \" and "\ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C cat_1 \ \ \\<^sub>C \" proof- interpret \: is_cat_colimit \ \ \ \ \\\ObjMap\\\\\ \ by (rule assms(1)) interpret \: is_functor \ \ \cat_1 \ \\ \ by (rule assms(2)) interpret \: is_functor \ \cat_1 \ \\ \ \ by (rule assms(3)) from cat_1_components(1) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.HomCod.cat_in_Obj_in_Vset) from cat_1_components(2) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.HomCod.cat_in_Arr_in_Vset) have \\: "\\ObjMap\\\\ = op_cf \\ObjMap\\\\" unfolding cat_op_simps by simp note cat_1_op = \.cat_1_op[OF \ \] show ?thesis by ( rule is_cat_rKe.is_cat_lKe_op [ OF cat_limit_is_rKe [ OF \.is_cat_limit_op[unfolded \\] \.is_functor_op[unfolded cat_1_op] \.is_functor_op[unfolded cat_1_op] ], unfolded cat_op_simps cat_1_op ] ) qed subsubsection\Adjoints\ lemma (in is_cf_adjunction) cf_adjunction_counit_is_rKe: \\The statement of the theorem is similar to the statement of a part of Theorem 2 in Chapter X-7 in \cite{mac_lane_categories_2010} or Proposition 6.5.2 in \cite{riehl_category_2016}. The proof follows (approximately) the proof in \cite{riehl_category_2016}.\ shows "\\<^sub>C \ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> cf_id \
: \
\\<^sub>C \ \\<^sub>C \
" proof- define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) then interpret \: \ \ by simp note exp_adj = cf_adj_exp_cf_cat_exp_cf_cat[OF \ \\ R.category_axioms] let ?\ = \\\<^sub>C \\ let ?\ = \\\<^sub>C \\ let ?\
\ = \exp_cat_ntcf \ \
?\\ let ?\
\ = \exp_cat_cf \ \
\\ let ?\
\ = \exp_cat_cf \ \
\\ let ?\
\
= \cat_FUNCT \ \
\
\ let ?\\
= \cat_FUNCT \ \ \
\ let ?adj_\
\ = \cf_adjunction_of_unit \ ?\
\ ?\
\ ?\
\\ interpret \
\: is_cf_adjunction \ ?\\
?\
\
?\
\ ?\
\ ?adj_\
\ by (rule exp_adj) show ?thesis proof(intro is_cat_rKeI) have id_\
: "cf_map (cf_id \
) \\<^sub>\ cat_FUNCT \ \
\
\Obj\" by ( cs_concl cs_simp: cat_FUNCT_components(1) cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) then have exp_id_\
: "exp_cat_cf \ \
\\ObjMap\\cf_map (cf_id \
)\ = cf_map \" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros ) have \: "cf_map \ \\<^sub>\ cat_FUNCT \ \ \
\Obj\" by ( cs_concl cs_simp: cat_FUNCT_components(1) cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) have \: "ntcf_arrow (\\<^sub>C \) \\<^sub>\ ntcf_arrows \ \
\
" by (cs_concl cs_intro: cat_FUNCT_cs_intros adj_cs_intros) have \
\
: "category \ (cat_FUNCT \ \
\
)" by (cs_concl cs_intro: cat_cs_intros) have \\
: "category \ (cat_FUNCT \ \ \
)" by (cs_concl cs_intro: cat_cs_intros) from \ \ \\ id_\
\
\
\\
LR.is_functor_axioms RL.is_functor_axioms R.cat_cf_id_is_functor NT.is_iso_ntcf_axioms have \_id_\
: "\\<^sub>C ?adj_\
\\NTMap\\cf_map (cf_id \
)\ = ntcf_arrow ?\" by ( cs_concl cs_simp: cat_Set_the_inverse[symmetric] cat_op_simps cat_cs_simps cat_FUNCT_cs_simps adj_cs_simps cs_intro: \
\.NT.iso_ntcf_is_arr_isomorphism'' cat_op_intros adj_cs_intros - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros ) show "universal_arrow_fo ?\
\ (cf_map (cf_id \
)) (cf_map \) (ntcf_arrow ?\)" by ( rule is_cf_adjunction.cf_adjunction_counit_component_is_ua_fo[ OF exp_adj id_\
, unfolded exp_id_\
\_id_\
] ) qed (cs_concl cs_intro: cat_cs_intros adj_cs_intros)+ qed lemma (in is_cf_adjunction) cf_adjunction_unit_is_lKe: shows "\\<^sub>C \ : cf_id \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \
\\<^sub>C \" by ( rule is_cat_rKe.is_cat_lKe_op [ OF is_cf_adjunction.cf_adjunction_counit_is_rKe [ OF is_cf_adjunction_op, folded op_ntcf_cf_adjunction_unit op_cf_cf_id ], unfolded cat_op_simps ntcf_op_ntcf_op_ntcf[OF cf_adjunction_unit_is_ntcf] ] ) lemma cf_adjunction_if_lKe_preserves: \\The statement of the theorem is similar to the statement of a part of Theorem 2 in Chapter X-7 in \cite{mac_lane_categories_2010} or Proposition 6.5.2 in \cite{riehl_category_2016}.\ assumes "\ : cf_id \
\\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \
\\<^sub>C \ \\<^sub>C (\ : \
\\\<^sub>C \)" shows "cf_adjunction_of_unit \ \ \ \ : \ \\<^sub>C\<^sub>F \ : \
\\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_cat_lKe_preserves \ \
\ \
\ \ \cf_id \
\ \ \ \ by (rule assms) from \.cat_lKe_preserves interpret \\: is_cat_lKe \ \
\ \ \ \ \\ \\<^sub>C\<^sub>F \\ \\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \\ by (cs_prems cs_simp: cat_cs_simps) from \\.cat_lKe_unique [ OF \.AG.HomCod.cat_cf_id_is_functor, unfolded \.AG.cf_cf_comp_cf_id_left, OF \.AG.cf_ntcf_id_is_ntcf ] obtain \ where \: "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F cf_id \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and ntcf_id_\_def: "ntcf_id \ = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by metis interpret \: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \cf_id \\ \ by (rule \) show ?thesis proof(rule counit_unit_is_cf_adjunction) show [cat_cs_simps]: "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = ntcf_id \" by (rule ntcf_id_\_def[symmetric]) have \_def: "\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by ( cs_concl cs_simp: cat_cs_simps ntcf_id_cf_comp[symmetric] cs_intro: cat_cs_intros ) note [cat_cs_simps] = this[symmetric] let ?\\\ = \\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \\ let ?\\\ = \\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \\ let ?\\\ = \\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \\ have "(?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = (?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof(rule ntcf_eqI) have dom_lhs: "\\<^sub>\ (((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \
\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_rhs: "\\<^sub>\ (((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \
\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) note is_ntcf.ntcf_Comp_commute[cat_cs_simps del] note category.cat_Comp_assoc[cat_cs_simps del] show "((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = ((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \
\Obj\" then show "((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = ((?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\" by ( cs_concl cs_simp: cat_cs_simps \.ntcf_lKe.ntcf_Comp_commute[symmetric] cs_intro: cat_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros)+ qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ also have "\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by ( cs_concl cs_simp: cat_cs_simps cf_comp_cf_ntcf_comp_assoc cf_ntcf_comp_ntcf_cf_comp_assoc cf_ntcf_comp_ntcf_vcomp[symmetric] cs_intro: cat_cs_intros ) also have "\ = \" by (cs_concl cs_simp: cat_cs_simps) finally have "(?\\\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\\\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = \" by simp then have \_def': "\ = (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by ( cs_concl cs_simp: cat_cs_simps ntcf_vcomp_ntcf_cf_comp[symmetric] cs_intro: cat_cs_intros )+ have \\\\: "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \.cat_lKe_unique[OF \.Lan.is_functor_axioms \.ntcf_lKe.is_ntcf_axioms] obtain \ where "\ \' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
; \ = \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \ \ \' = \" for \' by metis from this[OF \.Lan.cf_ntcf_id_is_ntcf \_def] this[OF \\\\ \_def'] show "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) = ntcf_id \" by simp qed (cs_concl cs_intro: cat_cs_intros)+ qed lemma cf_adjunction_if_rKe_preserves: assumes "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> cf_id \
: \
\\<^sub>C \ \\<^sub>C (\ : \
\\\<^sub>C \)" shows "cf_adjunction_of_counit \ \ \ \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: is_cat_rKe_preserves \ \
\ \
\ \ \cf_id \
\ \ \ \ by (rule assms) have "op_cf (cf_id \
) = cf_id (op_cat \
)" unfolding cat_op_simps by simp show ?thesis by ( rule is_cf_adjunction.is_cf_adjunction_op [ OF cf_adjunction_if_lKe_preserves[ OF \.is_cat_rKe_preserves_op[unfolded op_cf_cf_id] ], folded cf_adjunction_of_counit_def, unfolded cat_op_simps ] ) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Limit.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Limit.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Limit.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Limit.thy @@ -1,3473 +1,3706 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Limits\ theory CZH_UCAT_Limit imports CZH_UCAT_Universal CZH_Elementary_Categories.CZH_ECAT_Discrete CZH_Elementary_Categories.CZH_ECAT_SS CZH_Elementary_Categories.CZH_ECAT_Parallel begin subsection\Background\ named_theorems cat_lim_cs_simps named_theorems cat_lim_cs_intros subsection\Cone and cocone\ text\ In the context of this work, the concept of a cone corresponds to that of a cone to the base of a functor from a vertex, as defined in Chapter III-4 in \cite{mac_lane_categories_2010}; the concept of a cocone corresponds to that of a cone from the base of a functor to a vertex, as defined in Chapter III-3 in \cite{mac_lane_categories_2010}. - -In this body of work, only limits and colimits of functors with tiny maps -are considered. The definitions of a cone and a cocone also reflect this. -However, this restriction may be removed in the future. \ -(*TODO: remove the size limitation; see TODO in the next subsection*) -locale is_cat_cone = is_tm_ntcf \ \ \ \cf_const \ \ c\ \ \ for \ c \ \ \ \ + +locale is_cat_cone = is_ntcf \ \ \ \cf_const \ \ c\ \ \ for \ c \ \ \ \ + assumes cat_cone_obj[cat_lim_cs_intros]: "c \\<^sub>\ \\Obj\" syntax "_is_cat_cone" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : c <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_cone \ c \ \ \ \" -locale is_cat_cocone = is_tm_ntcf \ \ \ \ \cf_const \ \ c\ \ for \ c \ \ \ \ + +locale is_cat_cocone = is_ntcf \ \ \ \ \cf_const \ \ c\ \ for \ c \ \ \ \ + assumes cat_cocone_obj[cat_lim_cs_intros]: "c \\<^sub>\ \\Obj\" syntax "_is_cat_cocone" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_cocone \ c \ \ \ \" text\Rules.\ lemma (in is_cat_cone) is_cat_cone_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "c' = c" and "\' = \" and "\' = \" and "\' = \" shows "\ : c' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_cone_axioms) mk_ide rf is_cat_cone_def[unfolded is_cat_cone_axioms_def] |intro is_cat_coneI| |dest is_cat_coneD[dest!]| |elim is_cat_coneE[elim!]| lemma (in is_cat_cone) is_cat_coneD'[cat_lim_cs_intros]: assumes "c' = cf_const \ \ c" - shows "\ : c' \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - unfolding assms by (cs_concl cs_intro: cat_small_cs_intros) + shows "\ : c' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + unfolding assms by (cs_concl cs_intro: cat_cs_intros) lemmas [cat_lim_cs_intros] = is_cat_cone.is_cat_coneD' lemma (in is_cat_cocone) is_cat_cocone_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "c' = c" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_cocone_axioms) mk_ide rf is_cat_cocone_def[unfolded is_cat_cocone_axioms_def] |intro is_cat_coconeI| |dest is_cat_coconeD[dest!]| |elim is_cat_coconeE[elim!]| lemma (in is_cat_cocone) is_cat_coconeD'[cat_lim_cs_intros]: assumes "c' = cf_const \ \ c" - shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m c' : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - unfolding assms by (cs_concl cs_intro: cat_small_cs_intros) + shows "\ : \ \\<^sub>C\<^sub>F c' : \ \\\<^sub>C\<^bsub>\\<^esub> \" + unfolding assms by (cs_concl cs_intro: cat_cs_intros) lemmas [cat_lim_cs_intros] = is_cat_cocone.is_cat_coconeD' text\Duality.\ lemma (in is_cat_cone) is_cat_cocone_op: "op_ntcf \ : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_coconeI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_lim_cs_intros cat_op_intros)+ lemma (in is_cat_cone) is_cat_cocone_op'[cat_op_intros]: assumes "\' = \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cf \" shows "op_ntcf \ : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_cocone_op) lemmas [cat_op_intros] = is_cat_cone.is_cat_cocone_op' lemma (in is_cat_cocone) is_cat_cone_op: "op_ntcf \ : c <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_coneI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_lim_cs_intros cat_op_intros) lemma (in is_cat_cocone) is_cat_cone_op'[cat_op_intros]: assumes "\' = \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cf \" shows "op_ntcf \ : c <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_cone_op) lemmas [cat_op_intros] = is_cat_cocone.is_cat_cone_op' text\Elementary properties.\ lemma (in is_cat_cone) cat_cone_LArr_app_is_arr: assumes "j \\<^sub>\ \\Obj\" shows "\\NTMap\\j\ : c \\<^bsub>\\<^esub> \\ObjMap\\j\" proof- from assms have [simp]: "cf_const \ \ c\ObjMap\\j\ = c" by (cs_concl cs_simp: cat_cs_simps) from ntcf_NTMap_is_arr[OF assms] show ?thesis by simp qed lemma (in is_cat_cone) cat_cone_LArr_app_is_arr'[cat_lim_cs_intros]: assumes "j \\<^sub>\ \\Obj\" and "\j = \\ObjMap\\j\" shows "\\NTMap\\j\ : c \\<^bsub>\\<^esub> \j" using assms(1) unfolding assms(2) by (rule cat_cone_LArr_app_is_arr) lemmas [cat_lim_cs_intros] = is_cat_cone.cat_cone_LArr_app_is_arr' lemma (in is_cat_cocone) cat_cocone_LArr_app_is_arr: assumes "j \\<^sub>\ \\Obj\" shows "\\NTMap\\j\ : \\ObjMap\\j\ \\<^bsub>\\<^esub> c" proof- from assms have [simp]: "cf_const \ \ c\ObjMap\\j\ = c" by (cs_concl cs_simp: cat_cs_simps) from ntcf_NTMap_is_arr[OF assms] show ?thesis by simp qed lemma (in is_cat_cocone) cat_cocone_LArr_app_is_arr'[cat_lim_cs_intros]: assumes "j \\<^sub>\ \\Obj\" and "\j = \\ObjMap\\j\" shows "\\NTMap\\j\ : \j \\<^bsub>\\<^esub> c" using assms(1) unfolding assms(2) by (rule cat_cocone_LArr_app_is_arr) lemmas [cat_lim_cs_intros] = is_cat_cocone.cat_cocone_LArr_app_is_arr' lemma (in is_cat_cone) cat_cone_Comp_commute[cat_lim_cs_simps]: assumes "f : a \\<^bsub>\\<^esub> b" shows "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = \\NTMap\\b\" using ntcf_Comp_commute[symmetric, OF assms] assms by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemmas [cat_lim_cs_simps] = is_cat_cone.cat_cone_Comp_commute lemma (in is_cat_cocone) cat_cocone_Comp_commute[cat_lim_cs_simps]: assumes "f : a \\<^bsub>\\<^esub> b" shows "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\NTMap\\a\" using ntcf_Comp_commute[OF assms] assms by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemmas [cat_lim_cs_simps] = is_cat_cocone.cat_cocone_Comp_commute text\Utilities/helper lemmas.\ lemma (in is_cat_cone) helper_cat_cone_ntcf_vcomp_Comp: assumes "\' : c' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f' : c' \\<^bsub>\\<^esub> c" and "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" and "j \\<^sub>\ \\Obj\" shows "\'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'" proof- from assms(3) have "\'\NTMap\\j\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\\j\" by simp from this assms(1,2,4) show "\'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma (in is_cat_cone) helper_cat_cone_Comp_ntcf_vcomp: assumes "\' : c' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f' : c' \\<^bsub>\\<^esub> c" and "\j. j \\<^sub>\ \\Obj\ \ \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'" shows "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" proof- interpret \': is_cat_cone \ c' \ \ \ \' by (rule assms(1)) show ?thesis proof(rule ntcf_eqI[OF \'.is_ntcf_axioms]) from assms(2) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' : cf_const \ \ c' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\'\NTMap\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\" proof(rule vsv_eqI, unfold cat_cs_simps) show "vsv ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\)" by (cs_concl cs_intro: cat_cs_intros) from assms show "\\Obj\ = \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) fix j assume prems': "j \\<^sub>\ \\Obj\" with assms(1,2) show "\'\NTMap\\j\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\\j\" by (cs_concl cs_simp: cat_cs_simps assms(3) cs_intro: cat_cs_intros) qed auto qed simp_all qed lemma (in is_cat_cone) helper_cat_cone_Comp_ntcf_vcomp_iff: assumes "\' : c' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "f' : c' \\<^bsub>\\<^esub> c \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' \ f' : c' \\<^bsub>\\<^esub> c \ (\j\\<^sub>\\\Obj\. \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f')" using helper_cat_cone_ntcf_vcomp_Comp[OF assms] helper_cat_cone_Comp_ntcf_vcomp[OF assms] by (intro iffI; elim conjE; intro conjI) metis+ lemma (in is_cat_cocone) helper_cat_cocone_ntcf_vcomp_Comp: assumes "\' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f' : c \\<^bsub>\\<^esub> c'" and "\' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" and "j \\<^sub>\ \\Obj\" shows "\'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\" proof- interpret \': is_cat_cocone \ c' \ \ \ \' by (rule assms(1)) from assms(3) have "op_ntcf \' = op_ntcf (ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by simp from this assms(2) have op_\': "op_ntcf \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f'" by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) have "\'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>op_cat \\<^esub> f'" by ( rule is_cat_cone.helper_cat_cone_ntcf_vcomp_Comp[ OF is_cat_cone_op \'.is_cat_cone_op, unfolded cat_op_simps, OF assms(2) op_\' assms(4) ] ) from this assms(2,4) show "\'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\" by (cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) qed lemma (in is_cat_cocone) helper_cat_cocone_Comp_ntcf_vcomp: assumes "\' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f' : c \\<^bsub>\\<^esub> c'" and "\j. j \\<^sub>\ \\Obj\ \ \'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\" shows "\' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof- interpret \': is_cat_cocone \ c' \ \ \ \' by (rule assms(1)) from assms(2) have \'j: "\'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>op_cat \\<^esub> f'" if "j \\<^sub>\ \\Obj\" for j using that unfolding assms(3)[OF that] by (cs_concl cs_simp: cat_op_simps cat_cs_simps cs_intro: cat_cs_intros) have op_\': "op_ntcf \' = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f'" by ( rule is_cat_cone.helper_cat_cone_Comp_ntcf_vcomp[ OF is_cat_cone_op \'.is_cat_cone_op, unfolded cat_op_simps, OF assms(2) \'j, simplified ] ) from assms(2) show "\' = (ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by ( cs_concl cs_simp: cat_op_simps op_\' eq_op_ntcf_iff[symmetric, OF \'.is_ntcf_axioms] cs_intro: cat_cs_intros ) qed lemma (in is_cat_cocone) helper_cat_cocone_Comp_ntcf_vcomp_iff: assumes "\' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "f' : c \\<^bsub>\\<^esub> c' \ \' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \ f' : c \\<^bsub>\\<^esub> c' \ (\j\\<^sub>\\\Obj\. \'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\)" using helper_cat_cocone_ntcf_vcomp_Comp[OF assms] helper_cat_cocone_Comp_ntcf_vcomp[OF assms] by (intro iffI; elim conjE; intro conjI) metis+ subsection\Limit and colimit\ subsubsection\Definition and elementary properties\ text\ The concept of a limit is introduced in Chapter III-4 in \cite{mac_lane_categories_2010}; the concept of a colimit is introduced in Chapter III-3 in \cite{mac_lane_categories_2010}. \ -(*TODO: remove the size limitation*) -locale is_cat_limit = is_cat_cone \ r \ \ \ u for \ \ \ \ r u + - assumes cat_lim_ua_fo: - "universal_arrow_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u)" +locale is_cat_limit = is_cat_cone \ r \ \ \ u for \ \ \ \ r u + + assumes cat_lim_ua_fo: "\u' r'. u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ + \!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" syntax "_is_cat_limit" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_limit \ \ \ \ r u" locale is_cat_colimit = is_cat_cocone \ r \ \ \ u for \ \ \ \ r u + - assumes cat_colim_ua_fo: "universal_arrow_fo - (\\<^sub>C \ (op_cat \) (op_cat \)) (cf_map \) r (ntcf_arrow (op_ntcf u))" + assumes cat_colim_ua_of: "\u' r'. u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e r' : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ + \!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" syntax "_is_cat_colimit" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_colimit \ \ \ \ r u" text\Rules.\ lemma (in is_cat_limit) is_cat_limit_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "r' = r" and "\' = \" and "\' = \" and "\' = \" shows "u : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_limit_axioms) mk_ide rf is_cat_limit_def[unfolded is_cat_limit_axioms_def] |intro is_cat_limitI| |dest is_cat_limitD[dest]| |elim is_cat_limitE[elim]| lemmas [cat_lim_cs_intros] = is_cat_limitD(1) lemma (in is_cat_colimit) is_cat_colimit_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "r' = r" and "\' = \" and "\' = \" and "\' = \" shows "u : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_colimit_axioms) mk_ide rf is_cat_colimit_def[unfolded is_cat_colimit_axioms_def] |intro is_cat_colimitI| |dest is_cat_colimitD[dest]| |elim is_cat_colimitE[elim]| lemmas [cat_lim_cs_intros] = is_cat_colimitD(1) -text\Duality\ +text\Limits, colimits and universal arrows.\ + +lemma (in is_cat_limit) cat_lim_is_universal_arrow_fo: + "universal_arrow_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u)" +proof(intro is_functor.universal_arrow_foI) + + define \ where "\ = \ + \" + have \: "\ \" and \\: "\ \\<^sub>\ \" + by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) + then interpret \: \ \ by simp + + show "\\<^sub>C\<^sub>F \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" + by + ( + intro + \ \\ + cf_diagonal_is_functor + NTDom.HomDom.category_axioms + NTDom.HomCod.category_axioms + ) + + show "r \\<^sub>\ \\Obj\" by (intro cat_cone_obj) + then show "ntcf_arrow u : \\<^sub>C\<^sub>F \ \ \\ObjMap\\r\ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + fix r' u' assume prems: + "r' \\<^sub>\ \\Obj\" "u' : \\<^sub>C\<^sub>F \ \ \\ObjMap\\r'\ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" + from prems(1) have [cat_cs_simps]: + "cf_of_cf_map \ \ (cf_map \) = \" + "cf_of_cf_map \ \ (cf_map (cf_const \ \ r')) = cf_const \ \ r'" + by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros)+ + from prems(2,1) have + "u' : cf_map (cf_const \ \ r') \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" + by (cs_prems cs_simp: cat_cs_simps) + note u'[unfolded cat_cs_simps] = cat_FUNCT_is_arrD[OF this] + + from cat_lim_ua_fo[OF is_cat_coneI[OF u'(1) prems(1)]] obtain f + where f: "f : r' \\<^bsub>\\<^esub> r" + and [symmetric, cat_cs_simps]: + "ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" + and f_unique: + "\ + f' : r' \\<^bsub>\\<^esub> r; + ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' + \ \ f' = f" + for f' + by metis + + show "\!f'. + f' : r' \\<^bsub>\\<^esub> r \ + u' = umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" + proof(intro ex1I conjI; (elim conjE)?) + show "f : r' \\<^bsub>\\<^esub> r" by (rule f) + with \\ cat_cone_obj show u'_def: + "u' = umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f\" + by + ( + cs_concl + cs_simp: u'(2)[symmetric] cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + fix f' assume prems': + "f' : r' \\<^bsub>\\<^esub> r" + "u' = umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" + from prems'(2) \\ f prems' cat_cone_obj have u'_def': + "u' = ntcf_arrow (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + from prems'(1) have "ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" + by (cs_concl cs_simp: cat_FUNCT_cs_simps u'_def' cs_intro: cat_cs_intros) + from f_unique[OF prems'(1) this] show "f' = f" . + + qed + +qed + +lemma (in is_cat_cone) cat_cone_is_cat_limit: + assumes "universal_arrow_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \)" + shows "\ : c <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + + define \ where "\ = \ + \" + have \: "\ \" and \\: "\ \\<^sub>\ \" + by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) + then interpret \: \ \ by simp + + show ?thesis + proof(intro is_cat_limitI is_cat_cone_axioms) + fix u' c' assume prems: "u' : c' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + + interpret u': is_cat_cone \ c' \ \ \ u' by (rule prems) + + from u'.cat_cone_obj have u'_is_arr: + "ntcf_arrow u' : \\<^sub>C\<^sub>F \ \ \\ObjMap\\c'\ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map \" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + from is_functor.universal_arrow_foD(3) + [ + OF + cf_diagonal_is_functor[ + OF \ \\ NTDom.HomDom.category_axioms NTDom.HomCod.category_axioms + ] + assms + u'.cat_cone_obj + u'_is_arr + ] + obtain f where f: "f : c' \\<^bsub>\\<^esub> c" + and u'_def': "ntcf_arrow u' = + umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f\" + and f'_unique: + "\ + f' : c' \\<^bsub>\\<^esub> c; + ntcf_arrow u' = + umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f'\ + \ \ f' = f" + for f' + by metis + + from u'_def' \\ f cat_cone_obj have u'_def: + "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + show "\!f'. f' : c' \\<^bsub>\\<^esub> c \ u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" + proof(intro ex1I conjI; (elim conjE)?, (rule f)?, (rule u'_def)?) + fix f'' assume prems': + "f'' : c' \\<^bsub>\\<^esub> c" "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f''" + from \\ prems' have + "ntcf_arrow u' = + umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f''\" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + from f'_unique[OF prems'(1) this] show "f'' = f". + qed + + qed + +qed + +lemma (in is_cat_colimit) cat_colim_is_universal_arrow_of: + "universal_arrow_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u)" +proof(intro is_functor.universal_arrow_ofI) + + define \ where "\ = \ + \" + have \: "\ \" and \\: "\ \\<^sub>\ \" + by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) + then interpret \: \ \ by simp + + show "\\<^sub>C\<^sub>F \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" + by + ( + intro + \ \\ + cf_diagonal_is_functor + NTDom.HomDom.category_axioms + NTDom.HomCod.category_axioms + ) + + show "r \\<^sub>\ \\Obj\" by (intro cat_cocone_obj) + + then show "ntcf_arrow u : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \\<^sub>C\<^sub>F \ \ \\ObjMap\\r\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + fix r' u' assume prems: + "r' \\<^sub>\ \\Obj\" "u' : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \\<^sub>C\<^sub>F \ \ \\ObjMap\\r'\" + from prems(1) have [cat_cs_simps]: + "cf_of_cf_map \ \ (cf_map \) = \" + "cf_of_cf_map \ \ (cf_map (cf_const \ \ r')) = cf_const \ \ r'" + by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros)+ + from prems(2,1) have + "u' : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> cf_map (cf_const \ \ r')" + by (cs_prems cs_simp: cat_cs_simps) + note u'[unfolded cat_cs_simps] = cat_FUNCT_is_arrD[OF this] + + from cat_colim_ua_of[OF is_cat_coconeI[OF u'(1) prems(1)]] obtain f + where f: "f : r \\<^bsub>\\<^esub> r'" + and [symmetric, cat_cs_simps]: + "ntcf_of_ntcf_arrow \ \ u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" + and f_unique: + "\ + f' : r \\<^bsub>\\<^esub> r'; + ntcf_of_ntcf_arrow \ \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u + \ \ f' = f" + for f' + by metis + + show " \!f'. + f' : r \\<^bsub>\\<^esub> r' \ + u' = umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" + proof(intro ex1I conjI; (elim conjE)?) + + show "f : r \\<^bsub>\\<^esub> r'" by (rule f) + with \\ cat_cocone_obj show u'_def: + "u' = umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f\" + by + ( + cs_concl + cs_simp: u'(2)[symmetric] cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + fix f' assume prems': + "f' : r \\<^bsub>\\<^esub> r'" + "u' = umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" + from prems'(2) \\ f prems' cat_cocone_obj have u'_def': + "u' = ntcf_arrow (ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u)" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + from prems'(1) have "ntcf_of_ntcf_arrow \ \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" + by (cs_concl cs_simp: cat_FUNCT_cs_simps u'_def' cs_intro: cat_cs_intros) + from f_unique[OF prems'(1) this] show "f' = f" . + + qed + +qed + +lemma (in is_cat_cocone) cat_cocone_is_cat_colimit: + assumes "universal_arrow_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \)" + shows "\ : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m c : \ \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + + define \ where "\ = \ + \" + have \: "\ \" and \\: "\ \\<^sub>\ \" + by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) + then interpret \: \ \ by simp + + show ?thesis + proof(intro is_cat_colimitI is_cat_cocone_axioms) + + fix u' c' assume prems: "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e c' : \ \\\<^sub>C\<^bsub>\\<^esub> \" + + interpret u': is_cat_cocone \ c' \ \ \ u' by (rule prems) + + from u'.cat_cocone_obj have u'_is_arr: + "ntcf_arrow u' : cf_map \ \\<^bsub>cat_FUNCT \ \ \\<^esub> \\<^sub>C\<^sub>F \ \ \\ObjMap\\c'\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + from is_functor.universal_arrow_ofD(3) + [ + OF + cf_diagonal_is_functor[ + OF \ \\ NTDom.HomDom.category_axioms NTDom.HomCod.category_axioms + ] + assms + u'.cat_cocone_obj + u'_is_arr + ] + obtain f where f: "f : c \\<^bsub>\\<^esub> c'" + and u'_def': "ntcf_arrow u' = + umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f\" + and f'_unique: + "\ + f' : c \\<^bsub>\\<^esub> c'; + ntcf_arrow u' = + umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f'\ + \ \ f' = f" + for f' + by metis + + from u'_def' \\ f cat_cocone_obj have u'_def: + "u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + + show "\!f'. f' : c \\<^bsub>\\<^esub> c' \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" + proof(intro ex1I conjI; (elim conjE)?, (rule f)?, (rule u'_def)?) + fix f'' assume prems': + "f'' : c \\<^bsub>\\<^esub> c'" "u' = ntcf_const \ \ f'' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" + from \\ prems' have + "ntcf_arrow u' = + umap_of (\\<^sub>C\<^sub>F \ \ \) (cf_map \) c (ntcf_arrow \) c'\ArrVal\\f''\" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros + ) + from f'_unique[OF prems'(1) this] show "f'' = f". + qed + + qed + +qed + + +text\Duality.\ lemma (in is_cat_limit) is_cat_colimit_op: "op_ntcf u : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" - using cat_lim_ua_fo - by (intro is_cat_colimitI) - (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) +proof(intro is_cat_colimitI) + show "op_ntcf u : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e r : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by (cs_concl cs_simp: cs_intro: cat_op_intros) + fix u' r' assume prems: + "u' : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e r' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + interpret u': is_cat_cocone \ r' \op_cat \\ \op_cat \\ \op_cf \\ u' + by (rule prems) + from cat_lim_ua_fo[OF u'.is_cat_cone_op[unfolded cat_op_simps]] obtain f + where f: "f : r' \\<^bsub>\\<^esub> r" + and op_u'_def: "op_ntcf u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" + and f_unique: + "\ f' : r' \\<^bsub>\\<^esub> r; op_ntcf u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' \ \ + f' = f" + for f' + by metis + from op_u'_def have "op_ntcf (op_ntcf u') = op_ntcf (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f)" + by simp + from this f have u'_def: + "u' = ntcf_const (op_cat \) (op_cat \) f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf u" + by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) + show "\!f'. + f' : r \\<^bsub>op_cat \\<^esub> r' \ + u' = ntcf_const (op_cat \) (op_cat \) f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf u" + proof(intro ex1I conjI; (elim conjE)?, (unfold cat_op_simps)?) + fix f' assume prems': + "f' : r' \\<^bsub>\\<^esub> r" + "u' = ntcf_const (op_cat \) (op_cat \) f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf u" + from prems'(2) have + "op_ntcf u' = op_ntcf (ntcf_const (op_cat \) (op_cat \) f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf u)" + by simp + from this prems'(1) have "op_ntcf u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_op_intros + ) + from f_unique[OF prems'(1) this] show "f' = f". + qed (intro u'_def f)+ +qed lemma (in is_cat_limit) is_cat_colimit_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf u : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_colimit_op) lemmas [cat_op_intros] = is_cat_limit.is_cat_colimit_op' lemma (in is_cat_colimit) is_cat_limit_op: "op_ntcf u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" - using cat_colim_ua_fo - by (intro is_cat_limitI) - (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) +proof(intro is_cat_limitI) + show "op_ntcf u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by (cs_concl cs_simp: cs_intro: cat_op_intros) + fix u' r' assume prems: + "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + interpret u': is_cat_cone \ r' \op_cat \\ \op_cat \\ \op_cf \\ u' + by (rule prems) + from cat_colim_ua_of[OF u'.is_cat_cocone_op[unfolded cat_op_simps]] obtain f + where f: "f : r \\<^bsub>\\<^esub> r'" + and op_u'_def: "op_ntcf u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" + and f_unique: + "\ f' : r \\<^bsub>\\<^esub> r'; op_ntcf u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u \ \ + f' = f" + for f' + by metis + from op_u'_def have "op_ntcf (op_ntcf u') = op_ntcf (ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u)" + by simp + from this f have u'_def: + "u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f" + by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) + show "\!f'. + f' : r' \\<^bsub>op_cat \\<^esub> r \ + u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f'" + proof(intro ex1I conjI; (elim conjE)?, (unfold cat_op_simps)?) + fix f' assume prems': + "f' : r \\<^bsub>\\<^esub> r'" + "u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f'" + from prems'(2) have + "op_ntcf u' = op_ntcf (op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f')" + by simp + from this prems'(1) have "op_ntcf u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_op_intros + ) + from f_unique[OF prems'(1) this] show "f' = f". + qed (intro u'_def f)+ +qed lemma (in is_cat_colimit) is_cat_colimit_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_limit_op) lemmas [cat_op_intros] = is_cat_colimit.is_cat_colimit_op' -text\Elementary properties of limits and colimits.\ - -sublocale is_cat_limit \ \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_concl cs_intro: cat_cs_intros cat_small_cs_intros) - -sublocale is_cat_colimit \ \: is_functor - \ \op_cat \\ \cat_Funct \ (op_cat \) (op_cat \)\ \\\<^sub>C \ (op_cat \) (op_cat \)\ - by (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_op_intros) - - subsubsection\Universal property\ -lemma is_cat_limitI': - assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "\u' r'. \ u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ \ \ - \!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - shows "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" -proof(intro is_cat_limitI is_functor.universal_arrow_foI) - interpret u: is_cat_cone \ r \ \ \ u by (rule assms(1)) - show "r \\<^sub>\ \\Obj\" by (cs_concl cs_intro: cat_lim_cs_intros) - show "\\<^sub>C \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Funct \ \ \" - by (cs_concl cs_intro: cat_cs_intros cat_small_cs_intros) - show "ntcf_arrow u : \\<^sub>C \ \ \\ObjMap\\r\ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" - by - ( - cs_concl - cs_simp: cat_cs_simps - cs_intro: cat_lim_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros - ) - fix r' u' assume prems: - "r' \\<^sub>\ \\Obj\" "u' : \\<^sub>C \ \ \\ObjMap\\r'\ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" - note u' = cat_Funct_is_arrD[OF prems(2)] - from u'(1) prems(1) have u'_is_tm_ntcf: - "ntcf_of_ntcf_arrow \ \ u' : cf_const \ \ r' \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m \ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by - ( - cs_prems - cs_simp: cat_cs_simps cat_small_cs_simps cat_FUNCT_cs_simps - cs_intro: cat_cs_intros - ) - from this prems(1) have u'_is_cat_cone: - "ntcf_of_ntcf_arrow \ \ u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (intro is_cat_coneI) - interpret u': is_cat_cone \ r' \ \ \ \ntcf_of_ntcf_arrow \ \ u'\ - by (rule u'_is_cat_cone) - from assms(2)[OF u'_is_cat_cone] obtain f' where f': "f' : r' \\<^bsub>\\<^esub> r" - and u'_def: "ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - and unique: "\f''. - \ - f'' : r' \\<^bsub>\\<^esub> r; - ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'' - \ \ f'' = f'" - by (meson prems(1)) - from u'_def have u'_NTMap_app: - "ntcf_of_ntcf_arrow \ \ u'\NTMap\\j\ = (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\\j\" - if "j \\<^sub>\ \\Obj\" for j - by simp - have u'_NTMap_app: "u'\NTMap\\j\ = u\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'" - if "j \\<^sub>\ \\Obj\" for j - using u'_NTMap_app[OF that] that f' - by (cs_prems cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros) - show "\!f'. - f' : r' \\<^bsub>\\<^esub> r \ - u' = umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" - proof(intro ex1I conjI; (elim conjE)?) - show "f' : r' \\<^bsub>\\<^esub> r" by (rule f') - have u'_def'[symmetric, cat_cs_simps]: - "ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - proof(rule ntcf_eqI) - from u'_is_tm_ntcf show - "ntcf_of_ntcf_arrow \ \ u' : cf_const \ \ r' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_small_cs_intros) - from f' show - "u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' : cf_const \ \ r' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) - show - "ntcf_of_ntcf_arrow \ \ u'\NTMap\ = (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\" - proof(rule vsv_eqI) - from f' show "\\<^sub>\ (ntcf_of_ntcf_arrow \ \ u'\NTMap\) = - \\<^sub>\ ((u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\)" - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) - show "ntcf_of_ntcf_arrow \ \ u'\NTMap\\a\ = - (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\\a\" - if "a \\<^sub>\ \\<^sub>\ (ntcf_of_ntcf_arrow \ \ u'\NTMap\)" for a - proof- - from that have "a \\<^sub>\ \\Obj\" by (cs_prems cs_simp: cat_cs_simps) - with f' show - "ntcf_of_ntcf_arrow \ \ u'\NTMap\\a\ = - (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f')\NTMap\\a\" - by - ( - cs_concl - cs_simp: cat_cs_simps cat_FUNCT_cs_simps u'_NTMap_app - cs_intro: cat_cs_intros - ) - qed - qed (auto intro: cat_cs_intros) - qed simp_all - from f' u'(1) show - "u' = umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" - by (subst u'(2)) - ( - cs_concl - cs_simp: cat_cs_simps cat_FUNCT_cs_simps - cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros - ) - fix f'' assume prems': - "f'' : r' \\<^bsub>\\<^esub> r" - "u' = umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f''\" - from prems'(2,1) u'(1) have - "ntcf_of_ntcf_arrow \ \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f''" - by (subst (asm) u'(2)) - ( - cs_prems - cs_simp: cat_cs_simps cat_FUNCT_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros - ) - from unique[OF prems'(1) this] show "f'' = f'" . - qed -qed (intro assms)+ - -lemma (in is_cat_limit) cat_lim_unique_cone: - assumes "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - shows "\!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" -proof- - interpret u': is_cat_cone \ r' \ \ \ u' by (rule assms(1)) - have "ntcf_arrow u' : \\<^sub>C \ \ \\ObjMap\\r'\ \\<^bsub>cat_Funct \ \ \\<^esub> cf_map \" - by - ( - cs_concl - cs_intro: cat_lim_cs_intros cat_FUNCT_cs_intros cs_simp: cat_cs_simps - ) - from \.universal_arrow_foD(3)[OF cat_lim_ua_fo u'.cat_cone_obj this] obtain f' - where f': "f' : r' \\<^bsub>\\<^esub> r" - and u': "ntcf_arrow u' = - umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f'\" - and unique: - "\ - f'' : r' \\<^bsub>\\<^esub> r; - ntcf_arrow u' = - umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f''\ - \ \ f'' = f'" - for f'' - by metis - show "\!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - proof(intro ex1I conjI; (elim conjE)?) - show "f' : r' \\<^bsub>\\<^esub> r" by (rule f') - with u' show "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - by - ( - cs_prems - cs_simp: cat_cs_simps cat_FUNCT_cs_simps - cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros - ) - fix f'' assume prems: "f'' : r' \\<^bsub>\\<^esub> r" "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f''" - from prems(1) have "ntcf_arrow u' = - umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f''\" - by - ( - cs_concl - cs_simp: cat_cs_simps cat_FUNCT_cs_simps prems(2)[symmetric] - cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros - ) - from prems(1) this show "f'' = f'" by (intro unique) - qed -qed - lemma (in is_cat_limit) cat_lim_unique_cone': assumes "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r' \\<^bsub>\\<^esub> r \ (\j\\<^sub>\\\Obj\. u'\NTMap\\j\ = u\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f')" by (fold helper_cat_cone_Comp_ntcf_vcomp_iff[OF assms(1)]) - (intro cat_lim_unique_cone assms) + (intro cat_lim_ua_fo assms) lemma (in is_cat_limit) cat_lim_unique: assumes "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" - by (intro cat_lim_unique_cone[OF is_cat_limitD(1)[OF assms]]) + by (intro cat_lim_ua_fo[OF is_cat_limitD(1)[OF assms]]) lemma (in is_cat_limit) cat_lim_unique': assumes "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r' \\<^bsub>\\<^esub> r \ (\j\\<^sub>\\\Obj\. u'\NTMap\\j\ = u\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f')" by (intro cat_lim_unique_cone'[OF is_cat_limitD(1)[OF assms]]) lemma (in is_cat_colimit) cat_colim_unique_cocone: assumes "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" proof- interpret u': is_cat_cocone \ r' \ \ \ u' by (rule assms(1)) from u'.cat_cocone_obj have op_r': "r' \\<^sub>\ op_cat \\Obj\" unfolding cat_op_simps by simp from - is_cat_limit.cat_lim_unique_cone[ + is_cat_limit.cat_lim_ua_fo[ OF is_cat_limit_op u'.is_cat_cone_op, folded op_ntcf_ntcf_const ] obtain f' where f': "f' : r' \\<^bsub>op_cat \\<^esub> r" and [cat_cs_simps]: "op_ntcf u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf (ntcf_const \ \ f')" and unique: "\ f'' : r' \\<^bsub>op_cat \\<^esub> r; op_ntcf u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf (ntcf_const \ \ f'') \ \ f'' = f'" for f'' by metis show ?thesis proof(intro ex1I conjI; (elim conjE)?) from f' show f': "f' : r \\<^bsub>\\<^esub> r'" unfolding cat_op_simps by simp show "u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" by (rule eq_op_ntcf_iff[THEN iffD1], insert f') (cs_concl cs_intro: cat_cs_intros cs_simp: cat_cs_simps cat_op_simps)+ fix f'' assume prems: "f'' : r \\<^bsub>\\<^esub> r'" "u' = ntcf_const \ \ f'' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" from prems(1) have "f'' : r' \\<^bsub>op_cat \\<^esub> r" unfolding cat_op_simps by simp moreover from prems(1) have "op_ntcf u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf (ntcf_const \ \ f'')" unfolding prems(2) by (cs_concl cs_intro: cat_cs_intros cs_simp: cat_cs_simps cat_op_simps) ultimately show "f'' = f'" by (rule unique) qed qed lemma (in is_cat_colimit) cat_colim_unique_cocone': assumes "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r \\<^bsub>\\<^esub> r' \ (\j\\<^sub>\\\Obj\. u'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> u\NTMap\\j\)" by (fold helper_cat_cocone_Comp_ntcf_vcomp_iff[OF assms(1)]) (intro cat_colim_unique_cocone assms) lemma (in is_cat_colimit) cat_colim_unique: assumes "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = ntcf_const \ \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" by (intro cat_colim_unique_cocone[OF is_cat_colimitD(1)[OF assms]]) lemma (in is_cat_colimit) cat_colim_unique': assumes "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : r \\<^bsub>\\<^esub> r' \ (\j\\<^sub>\\\Obj\. u'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> u\NTMap\\j\)" proof- interpret u': is_cat_colimit \ \ \ \ r' u' by (rule assms(1)) show ?thesis by (fold helper_cat_cocone_Comp_ntcf_vcomp_iff[OF u'.is_cat_cocone_axioms]) (intro cat_colim_unique assms) qed lemma cat_lim_ex_is_arr_isomorphism: - assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" proof- interpret u: is_cat_limit \ \ \ \ r u by (rule assms(1)) interpret u': is_cat_limit \ \ \ \ r' u' by (rule assms(2)) + define \ where "\ = \ + \" + have \: "\ \" and \\: "\ \\<^sub>\ \" + by (simp_all add: \_def u.\_Limit_\\ u.\_\_\\ \_def u.\_\_\\) + then interpret \: \ \ by simp + have \: "\\<^sub>C\<^sub>F \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ \" + by + ( + intro + \ \\ + cf_diagonal_is_functor + u.NTDom.HomDom.category_axioms + u.NTDom.HomCod.category_axioms + ) + then interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ by simp + from is_functor.cf_universal_arrow_fo_ex_is_arr_isomorphism[ + OF \ u.cat_lim_is_universal_arrow_fo u'.cat_lim_is_universal_arrow_fo + ] obtain f where f: "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and u': "ntcf_arrow u' = - umap_fo (\\<^sub>C \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f\" - by - ( - elim u.\.cf_universal_arrow_fo_ex_is_arr_isomorphism[ - OF u.cat_lim_ua_fo u'.cat_lim_ua_fo - ] - ) + umap_fo (\\<^sub>C\<^sub>F \ \ \) (cf_map \) r (ntcf_arrow u) r'\ArrVal\\f\" + by auto from f have "f : r' \\<^bsub>\\<^esub> r" by auto from u' this have "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" by ( - cs_prems - cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_small_cs_simps - cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_small_cs_intros + cs_prems + cs_simp: cat_cs_simps cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) with f that show ?thesis by simp qed lemma cat_lim_ex_is_arr_isomorphism': - assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and "\j. j \\<^sub>\ \\Obj\ \ u'\NTMap\\j\ = u\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f" proof- interpret u: is_cat_limit \ \ \ \ r u by (rule assms(1)) interpret u': is_cat_limit \ \ \ \ r' u' by (rule assms(2)) from assms obtain f where iso_f: "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and u'_def: "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" by (rule cat_lim_ex_is_arr_isomorphism) then have f: "f : r' \\<^bsub>\\<^esub> r" by auto then have "u'\NTMap\\j\ = u\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f" if "j \\<^sub>\ \\Obj\" for j by ( intro u.helper_cat_cone_ntcf_vcomp_Comp[ OF u'.is_cat_cone_axioms f u'_def that ] ) with iso_f that show ?thesis by simp qed lemma cat_colim_ex_is_arr_isomorphism: assumes "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" and "u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" proof- interpret u: is_cat_colimit \ \ \ \ r u by (rule assms(1)) interpret u': is_cat_colimit \ \ \ \ r' u' by (rule assms(2)) obtain f where f: "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>op_cat \\<^esub> r" and [cat_cs_simps]: "op_ntcf u' = op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (op_cat \) (op_cat \) f" by ( elim cat_lim_ex_is_arr_isomorphism[ OF u.is_cat_limit_op u'.is_cat_limit_op ] ) from f have iso_f: "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" unfolding cat_op_simps by simp then have f: "f : r \\<^bsub>\\<^esub> r'" by auto have "u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" by (rule eq_op_ntcf_iff[THEN iffD1], insert f) (cs_concl cs_intro: cat_cs_intros cs_simp: cat_cs_simps cat_op_simps)+ from iso_f this that show ?thesis by simp qed lemma cat_colim_ex_is_arr_isomorphism': assumes "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "u' : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r' : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" and "\j. j \\<^sub>\ \\Obj\ \ u'\NTMap\\j\ = f \\<^sub>A\<^bsub>\\<^esub> u\NTMap\\j\" proof- interpret u: is_cat_colimit \ \ \ \ r u by (rule assms(1)) interpret u': is_cat_colimit \ \ \ \ r' u' by (rule assms(2)) from assms obtain f where iso_f: "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" and u'_def: "u' = ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F u" by (rule cat_colim_ex_is_arr_isomorphism) then have f: "f : r \\<^bsub>\\<^esub> r'" by auto then have "u'\NTMap\\j\ = f \\<^sub>A\<^bsub>\\<^esub> u\NTMap\\j\" if "j \\<^sub>\ \\Obj\" for j by ( intro u.helper_cat_cocone_ntcf_vcomp_Comp[ OF u'.is_cat_cocone_axioms f u'_def that ] ) with iso_f that show ?thesis by simp qed subsection\Finite limit and finite colimit\ locale is_cat_finite_limit = is_cat_limit \ \ \ \ r u + finite_category \ \ for \ \ \ \ r u syntax "_is_cat_finite_limit" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_finite_limit \ \ \ \ r u" locale is_cat_finite_colimit = is_cat_colimit \ \ \ \ r u + finite_category \ \ for \ \ \ \ r u syntax "_is_cat_finite_colimit" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n r : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_finite_colimit \ \ \ \ r u" text\Rules.\ lemma (in is_cat_finite_limit) is_cat_finite_limit_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "r' = r" and "\' = \" and "\' = \" and "\' = \" shows "u : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_finite_limit_axioms) mk_ide rf is_cat_finite_limit_def |intro is_cat_finite_limitI| |dest is_cat_finite_limitD[dest]| |elim is_cat_finite_limitE[elim]| lemmas [cat_lim_cs_intros] = is_cat_finite_limitD lemma (in is_cat_finite_colimit) is_cat_finite_colimit_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "r' = r" and "\' = \" and "\' = \" and "\' = \" shows "u : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n r' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_finite_colimit_axioms) mk_ide rf is_cat_finite_colimit_def[unfolded is_cat_colimit_axioms_def] |intro is_cat_finite_colimitI| |dest is_cat_finite_colimitD[dest]| |elim is_cat_finite_colimitE[elim]| lemmas [cat_lim_cs_intros] = is_cat_finite_colimitD -text\Duality\ +text\Duality.\ lemma (in is_cat_finite_limit) is_cat_finite_colimit_op: "op_ntcf u : op_cf \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n r : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by ( cs_concl cs_intro: is_cat_finite_colimitI cat_op_intros cat_small_cs_intros ) lemma (in is_cat_finite_limit) is_cat_finite_colimit_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf u : \' >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n r : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_finite_colimit_op) lemmas [cat_op_intros] = is_cat_finite_limit.is_cat_finite_colimit_op' lemma (in is_cat_finite_colimit) is_cat_finite_limit_op: "op_ntcf u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by ( cs_concl cs_intro: is_cat_finite_limitI cat_op_intros cat_small_cs_intros ) lemma (in is_cat_finite_colimit) is_cat_finite_colimit_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m\<^sub>.\<^sub>f\<^sub>i\<^sub>n \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_finite_limit_op) lemmas [cat_op_intros] = is_cat_finite_colimit.is_cat_finite_colimit_op' subsection\Product and coproduct\ subsubsection\Definition and elementary properties\ text\ The definition of the product object is a specialization of the definition presented in Chapter III-4 in \cite{mac_lane_categories_2010}. In the definition presented below, the discrete category that is used in the definition presented in \cite{mac_lane_categories_2010} is parameterized by an index set and the functor from the discrete category is parameterized by a function from the index set to the set of the objects of the category. \ locale is_cat_obj_prod = is_cat_limit \ \:\<^sub>C I\ \ \:\: I A \\ P \ + cf_discrete \ I A \ for \ I A \ P \ syntax "_is_cat_obj_prod" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_obj_prod \ I A \ P \" locale is_cat_obj_coprod = is_cat_colimit \ \:\<^sub>C I\ \ \:\: I A \\ U \ + cf_discrete \ I A \ for \ I A \ U \ syntax "_is_cat_obj_coprod" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U : I \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_obj_coprod \ I A \ U \" text\Rules.\ lemma (in is_cat_obj_prod) is_cat_obj_prod_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "P' = P" and "A' = A" and "I' = I" and "\' = \" shows "\ : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A' : I' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_obj_prod_axioms) mk_ide rf is_cat_obj_prod_def |intro is_cat_obj_prodI| |dest is_cat_obj_prodD[dest]| |elim is_cat_obj_prodE[elim]| lemmas [cat_lim_cs_intros] = is_cat_obj_prodD lemma (in is_cat_obj_coprod) is_cat_obj_coprod_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "U' = U" and "A' = A" and "I' = I" and "\' = \" shows "\ : A' >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : I' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_obj_coprod_axioms) mk_ide rf is_cat_obj_coprod_def |intro is_cat_obj_coprodI| |dest is_cat_obj_coprodD[dest]| |elim is_cat_obj_coprodE[elim]| lemmas [cat_lim_cs_intros] = is_cat_obj_coprodD text\Duality.\ lemma (in is_cat_obj_prod) is_cat_obj_coprod_op: "op_ntcf \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" using cf_discrete_vdomain_vsubset_Vset by (intro is_cat_obj_coprodI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) lemma (in is_cat_obj_prod) is_cat_obj_coprod_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : I \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_coprod_op) lemmas [cat_op_intros] = is_cat_obj_prod.is_cat_obj_coprod_op' lemma (in is_cat_obj_coprod) is_cat_obj_prod_op: "op_ntcf \ : U <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" using cf_discrete_vdomain_vsubset_Vset by (intro is_cat_obj_prodI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) lemma (in is_cat_obj_coprod) is_cat_obj_prod_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : U <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_prod_op) lemmas [cat_op_intros] = is_cat_obj_coprod.is_cat_obj_prod_op' subsubsection\Universal property\ (*cat_obj_prod_unique_cone already proven*) lemma (in is_cat_obj_prod) cat_obj_prod_unique_cone': assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: I A \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ (\j\\<^sub>\I. \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f')" by ( rule cat_lim_unique_cone'[ OF assms, unfolded the_cat_discrete_components(1) ] ) lemma (in is_cat_obj_prod) cat_obj_prod_unique: assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) \ f'" by (intro cat_lim_unique[OF is_cat_obj_prodD(1)[OF assms]]) lemma (in is_cat_obj_prod) cat_obj_prod_unique': assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ (\i\\<^sub>\I. \'\NTMap\\i\ = \\NTMap\\i\ \\<^sub>A\<^bsub>\\<^esub> f')" proof- interpret \': is_cat_obj_prod \ I A \ P' \' by (rule assms(1)) show ?thesis by ( rule cat_lim_unique'[ OF \'.is_cat_limit_axioms, unfolded the_cat_discrete_components(1) ] ) qed lemma (in is_cat_obj_coprod) cat_obj_coprod_unique_cocone': assumes "\' : :\: I A \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e U' : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : U \\<^bsub>\\<^esub> U' \ (\j\\<^sub>\I. \'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\)" by ( rule cat_colim_unique_cocone'[ OF assms, unfolded the_cat_discrete_components(1) ] ) lemma (in is_cat_obj_coprod) cat_obj_coprod_unique: assumes "\' : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : U \\<^bsub>\\<^esub> U' \ \' = ntcf_const (:\<^sub>C I) \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by (intro cat_colim_unique[OF is_cat_obj_coprodD(1)[OF assms]]) lemma (in is_cat_obj_coprod) cat_obj_coprod_unique': assumes "\' : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : I \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : U \\<^bsub>\\<^esub> U' \ (\j\\<^sub>\I. \'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\)" by ( rule cat_colim_unique'[ OF is_cat_obj_coprodD(1)[OF assms], unfolded the_cat_discrete_components ] ) lemma cat_obj_prod_ex_is_arr_isomorphism: assumes "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : P' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> P" and "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) \ f" proof- interpret \: is_cat_obj_prod \ I A \ P \ by (rule assms(1)) interpret \': is_cat_obj_prod \ I A \ P' \' by (rule assms(2)) from that show ?thesis by ( elim cat_lim_ex_is_arr_isomorphism[ OF \.is_cat_limit_axioms \'.is_cat_limit_axioms ] ) qed lemma cat_obj_prod_ex_is_arr_isomorphism': assumes "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : I \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : P' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> P" and "\j. j \\<^sub>\ I \ \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f" proof- interpret \: is_cat_obj_prod \ I A \ P \ by (rule assms(1)) interpret \': is_cat_obj_prod \ I A \ P' \' by (rule assms(2)) from that show ?thesis by ( elim cat_lim_ex_is_arr_isomorphism'[ OF \.is_cat_limit_axioms \'.is_cat_limit_axioms, unfolded the_cat_discrete_components(1) ] ) qed lemma cat_obj_coprod_ex_is_arr_isomorphism: assumes "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U : I \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : I \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : U \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> U'" and "\' = ntcf_const (:\<^sub>C I) \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof- interpret \: is_cat_obj_coprod \ I A \ U \ by (rule assms(1)) interpret \': is_cat_obj_coprod \ I A \ U' \' by (rule assms(2)) from that show ?thesis by ( elim cat_colim_ex_is_arr_isomorphism[ OF \.is_cat_colimit_axioms \'.is_cat_colimit_axioms ] ) qed lemma cat_obj_coprod_ex_is_arr_isomorphism': assumes "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U : I \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : I \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : U \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> U'" and "\j. j \\<^sub>\ I \ \'\NTMap\\j\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\" proof- interpret \: is_cat_obj_coprod \ I A \ U \ by (rule assms(1)) interpret \': is_cat_obj_coprod \ I A \ U' \' by (rule assms(2)) from that show ?thesis by ( elim cat_colim_ex_is_arr_isomorphism'[ OF \.is_cat_colimit_axioms \'.is_cat_colimit_axioms, unfolded the_cat_discrete_components(1) ] ) qed subsection\Finite product and finite coproduct\ locale is_cat_finite_obj_prod = is_cat_obj_prod \ I A \ P \ for \ I A \ P \ + assumes cat_fin_obj_prod_index_in_\: "I \\<^sub>\ \" syntax "_is_cat_finite_obj_prod" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n A : I \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_finite_obj_prod \ I A \ P \" locale is_cat_finite_obj_coprod = is_cat_obj_coprod \ I A \ U \ for \ I A \ U \ + assumes cat_fin_obj_coprod_index_in_\: "I \\<^sub>\ \" syntax "_is_cat_finite_obj_coprod" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ >\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n U : I \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_finite_obj_coprod \ I A \ U \" lemma (in is_cat_finite_obj_prod) cat_fin_obj_prod_index_vfinite: "vfinite I" using cat_fin_obj_prod_index_in_\ by auto sublocale is_cat_finite_obj_prod \ I: finite_category \ \:\<^sub>C I\ by (intro finite_categoryI') ( auto - simp: NTDom.HomDom.tiny_dg_category the_cat_discrete_components + simp: NTDom.HomDom.category_axioms the_cat_discrete_components intro!: cat_fin_obj_prod_index_vfinite ) lemma (in is_cat_finite_obj_coprod) cat_fin_obj_coprod_index_vfinite: "vfinite I" using cat_fin_obj_coprod_index_in_\ by auto sublocale is_cat_finite_obj_coprod \ I: finite_category \ \:\<^sub>C I\ by (intro finite_categoryI') ( auto - simp: NTDom.HomDom.tiny_dg_category the_cat_discrete_components + simp: NTDom.HomDom.category_axioms the_cat_discrete_components intro!: cat_fin_obj_coprod_index_vfinite ) text\Rules.\ lemma (in is_cat_finite_obj_prod) is_cat_finite_obj_prod_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "P' = P" and "A' = A" and "I' = I" and "\' = \" shows "\ : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n A' : I' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_finite_obj_prod_axioms) mk_ide rf is_cat_finite_obj_prod_def[unfolded is_cat_finite_obj_prod_axioms_def] |intro is_cat_finite_obj_prodI| |dest is_cat_finite_obj_prodD[dest]| |elim is_cat_finite_obj_prodE[elim]| lemmas [cat_lim_cs_intros] = is_cat_finite_obj_prodD lemma (in is_cat_finite_obj_coprod) is_cat_finite_obj_coprod_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "U' = U" and "A' = A" and "I' = I" and "\' = \" shows "\ : A' >\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n U' : I' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_finite_obj_coprod_axioms) mk_ide rf is_cat_finite_obj_coprod_def[unfolded is_cat_finite_obj_coprod_axioms_def] |intro is_cat_finite_obj_coprodI| |dest is_cat_finite_obj_coprodD[dest]| |elim is_cat_finite_obj_coprodE[elim]| lemmas [cat_lim_cs_intros] = is_cat_finite_obj_coprodD text\Duality.\ lemma (in is_cat_finite_obj_prod) is_cat_finite_obj_coprod_op: "op_ntcf \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n P : I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_finite_obj_coprodI) ( cs_concl cs_simp: cat_op_simps cs_intro: cat_fin_obj_prod_index_in_\ cat_cs_intros cat_op_intros ) lemma (in is_cat_finite_obj_prod) is_cat_finite_obj_coprod_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n P : I \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_finite_obj_coprod_op) lemmas [cat_op_intros] = is_cat_finite_obj_prod.is_cat_finite_obj_coprod_op' lemma (in is_cat_finite_obj_coprod) is_cat_finite_obj_prod_op: "op_ntcf \ : U <\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n A : I \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_finite_obj_prodI) ( cs_concl cs_simp: cat_op_simps cs_intro: cat_fin_obj_coprod_index_in_\ cat_cs_intros cat_op_intros ) lemma (in is_cat_finite_obj_coprod) is_cat_finite_obj_prod_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : U <\<^sub>C\<^sub>F\<^sub>.\<^sub>\\<^sub>.\<^sub>f\<^sub>i\<^sub>n A : I \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_finite_obj_prod_op) lemmas [cat_op_intros] = is_cat_finite_obj_coprod.is_cat_finite_obj_prod_op' subsection\Product and coproduct of two objects\ subsubsection\Definition and elementary properties\ locale is_cat_obj_prod_2 = is_cat_obj_prod \ \2\<^sub>\\ \if2 a b\ \ P \ for \ a b \ P \ syntax "_is_cat_obj_prod_2" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {_,_} :/ 2\<^sub>C \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_obj_prod_2 \ a b \ P \" locale is_cat_obj_coprod_2 = is_cat_obj_coprod \ \2\<^sub>\\ \if2 a b\ \ P \ for \ a b \ P \ syntax "_is_cat_obj_coprod_2" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ {_,_} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ _ :/ 2\<^sub>C \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_obj_coprod_2 \ a b \ U \" abbreviation proj_fst where "proj_fst \ \ vpfst (\\NTMap\)" abbreviation proj_snd where "proj_snd \ \ vpsnd (\\NTMap\)" text\Rules.\ lemma (in is_cat_obj_prod_2) is_cat_obj_prod_2_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "P' = P" and "a' = a" and "b' = b" and "\' = \" shows "\ : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a',b'} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_prod_2_axioms) mk_ide rf is_cat_obj_prod_2_def |intro is_cat_obj_prod_2I| |dest is_cat_obj_prod_2D[dest]| |elim is_cat_obj_prod_2E[elim]| lemmas [cat_lim_cs_intros] = is_cat_obj_prod_2D lemma (in is_cat_obj_coprod_2) is_cat_obj_coprod_2_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "P' = P" and "a' = a" and "b' = b" and "\' = \" shows "\ : {a',b'} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P' : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_coprod_2_axioms) mk_ide rf is_cat_obj_coprod_2_def |intro is_cat_obj_coprod_2I| |dest is_cat_obj_coprod_2D[dest]| |elim is_cat_obj_coprod_2E[elim]| lemmas [cat_lim_cs_intros] = is_cat_obj_coprod_2D text\Duality.\ lemma (in is_cat_obj_prod_2) is_cat_obj_coprod_2_op: "op_ntcf \ : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (rule is_cat_obj_coprod_2I[OF is_cat_obj_coprod_op]) lemma (in is_cat_obj_prod_2) is_cat_obj_coprod_2_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_coprod_2_op) lemmas [cat_op_intros] = is_cat_obj_prod_2.is_cat_obj_coprod_2_op' lemma (in is_cat_obj_coprod_2) is_cat_obj_prod_2_op: "op_ntcf \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (rule is_cat_obj_prod_2I[OF is_cat_obj_prod_op]) lemma (in is_cat_obj_coprod_2) is_cat_obj_prod_2_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_obj_prod_2_op) lemmas [cat_op_intros] = is_cat_obj_coprod_2.is_cat_obj_prod_2_op' text\Product/coproduct of two objects is a finite product/coproduct.\ sublocale is_cat_obj_prod_2 \ is_cat_finite_obj_prod \ \2\<^sub>\\ \if2 a b\ \ P \ proof(intro is_cat_finite_obj_prodI) show "2\<^sub>\ \\<^sub>\ \" by simp qed (cs_concl cs_simp: two[symmetric] cs_intro: cat_lim_cs_intros) sublocale is_cat_obj_coprod_2 \ is_cat_finite_obj_coprod \ \2\<^sub>\\ \if2 a b\ \ P \ proof(intro is_cat_finite_obj_coprodI) show "2\<^sub>\ \\<^sub>\ \" by simp qed (cs_concl cs_simp: two[symmetric] cs_intro: cat_lim_cs_intros) text\Elementary properties.\ lemma (in is_cat_obj_prod_2) cat_obj_prod_2_lr_in_Obj: shows cat_obj_prod_2_left_in_Obj[cat_lim_cs_intros]: "a \\<^sub>\ \\Obj\" and cat_obj_prod_2_right_in_Obj[cat_lim_cs_intros]: "b \\<^sub>\ \\Obj\" proof- have 0: "0 \\<^sub>\ 2\<^sub>\" and 1: "1\<^sub>\ \\<^sub>\ 2\<^sub>\" by simp_all show "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" by ( intro cf_discrete_selector_vrange[OF 0, simplified] cf_discrete_selector_vrange[OF 1, simplified] )+ qed lemmas [cat_lim_cs_intros] = is_cat_obj_prod_2.cat_obj_prod_2_lr_in_Obj lemma (in is_cat_obj_coprod_2) cat_obj_coprod_2_lr_in_Obj: shows cat_obj_coprod_2_left_in_Obj[cat_lim_cs_intros]: "a \\<^sub>\ \\Obj\" and cat_obj_coprod_2_right_in_Obj[cat_lim_cs_intros]: "b \\<^sub>\ \\Obj\" by ( intro is_cat_obj_prod_2.cat_obj_prod_2_lr_in_Obj[ OF is_cat_obj_prod_2_op, unfolded cat_op_simps ] )+ lemmas [cat_lim_cs_intros] = is_cat_obj_coprod_2.cat_obj_coprod_2_lr_in_Obj text\Utilities/help lemmas.\ lemma helper_I2_proj_fst_proj_snd_iff: "(\j\\<^sub>\2\<^sub>\. \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f') \ (proj_fst \' = proj_fst \ \\<^sub>A\<^bsub>\\<^esub> f' \ proj_snd \' = proj_snd \ \\<^sub>A\<^bsub>\\<^esub> f')" unfolding two by auto lemma helper_I2_proj_fst_proj_snd_iff': "(\j\\<^sub>\2\<^sub>\. \'\NTMap\\j\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\) \ (proj_fst \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_fst \ \ proj_snd \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_snd \)" unfolding two by auto subsubsection\Universal property\ lemma (in is_cat_obj_prod_2) cat_obj_prod_2_unique_cone': assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: (2\<^sub>\) (if2 a b) \ : :\<^sub>C (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ proj_fst \' = proj_fst \ \\<^sub>A\<^bsub>\\<^esub> f' \ proj_snd \' = proj_snd \ \\<^sub>A\<^bsub>\\<^esub> f'" by ( rule cat_obj_prod_unique_cone'[ OF assms, unfolded helper_I2_proj_fst_proj_snd_iff ] ) lemma (in is_cat_obj_prod_2) cat_obj_prod_2_unique: assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C (2\<^sub>\)) \ f'" by (rule cat_obj_prod_unique[OF is_cat_obj_prod_2D[OF assms]]) lemma (in is_cat_obj_prod_2) cat_obj_prod_2_unique': assumes "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P' \\<^bsub>\\<^esub> P \ proj_fst \' = proj_fst \ \\<^sub>A\<^bsub>\\<^esub> f' \ proj_snd \' = proj_snd \ \\<^sub>A\<^bsub>\\<^esub> f'" by ( rule cat_obj_prod_unique'[ OF is_cat_obj_prod_2D[OF assms], unfolded helper_I2_proj_fst_proj_snd_iff ] ) lemma (in is_cat_obj_coprod_2) cat_obj_coprod_2_unique_cocone': assumes "\' : :\: (2\<^sub>\) (if2 a b) \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e P' : :\<^sub>C (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P \\<^bsub>\\<^esub> P' \ proj_fst \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_fst \ \ proj_snd \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_snd \" by ( rule cat_obj_coprod_unique_cocone'[ OF assms, unfolded helper_I2_proj_fst_proj_snd_iff' ] ) lemma (in is_cat_obj_coprod_2) cat_obj_coprod_2_unique: assumes "\' : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P' : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P \\<^bsub>\\<^esub> P' \ \' = ntcf_const (:\<^sub>C (2\<^sub>\)) \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by (rule cat_obj_coprod_unique[OF is_cat_obj_coprod_2D[OF assms]]) lemma (in is_cat_obj_coprod_2) cat_obj_coprod_2_unique': assumes "\' : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P' : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : P \\<^bsub>\\<^esub> P' \ proj_fst \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_fst \ \ proj_snd \' = f' \\<^sub>A\<^bsub>\\<^esub> proj_snd \" by ( rule cat_obj_coprod_unique'[ OF is_cat_obj_coprod_2D[OF assms], unfolded helper_I2_proj_fst_proj_snd_iff' ] ) lemma cat_obj_prod_2_ex_is_arr_isomorphism: assumes "\ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : P' <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ {a,b} : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : P' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> P" and "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C (2\<^sub>\)) \ f" proof- interpret \: is_cat_obj_prod_2 \ a b \ P \ by (rule assms(1)) interpret \': is_cat_obj_prod_2 \ a b \ P' \' by (rule assms(2)) from that show ?thesis by ( elim cat_obj_prod_ex_is_arr_isomorphism[ OF \.is_cat_obj_prod_axioms \'.is_cat_obj_prod_axioms ] ) qed lemma cat_obj_coprod_2_ex_is_arr_isomorphism: assumes "\ : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : {a,b} >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ U' : 2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : U \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> U'" and "\' = ntcf_const (:\<^sub>C (2\<^sub>\)) \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof- interpret \: is_cat_obj_coprod_2 \ a b \ U \ by (rule assms(1)) interpret \': is_cat_obj_coprod_2 \ a b \ U' \' by (rule assms(2)) from that show ?thesis by ( elim cat_obj_coprod_ex_is_arr_isomorphism[ OF \.is_cat_obj_coprod_axioms \'.is_cat_obj_coprod_axioms ] ) qed subsection\Pullbacks and pushouts\ subsubsection\Definition and elementary properties\ text\ The definitions and the elementary properties of the pullbacks and the pushouts can be found, for example, in Chapter III-3 and Chapter III-4 in \cite{mac_lane_categories_2010}. \ locale is_cat_pullback = is_cat_limit \ \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ X x + cf_scospan \ \ \ \ \ \ \ for \ \ \ \ \ \ \ X x syntax "_is_cat_pullback" :: "V \ V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b _\_\_\_\_ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51, 51, 51, 51] 51) translations "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_pullback \ \ \ \ \ \ \ X x" locale is_cat_pushout = is_cat_colimit \ \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ X x + cf_sspan \ \ \ \ \ \ \ for \ \ \ \ \ \ \ X x syntax "_is_cat_pushout" :: "V \ V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _\_\_\_\_ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51, 51, 51, 51] 51) translations "x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_pushout \ \ \ \ \ \ \ X x" text\Rules.\ lemma (in is_cat_pullback) is_cat_pullback_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "X' = X" shows "x : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \'\\'\\'\\'\\' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_pullback_axioms) mk_ide rf is_cat_pullback_def |intro is_cat_pullbackI| |dest is_cat_pullbackD[dest]| |elim is_cat_pullbackE[elim]| lemmas [cat_lim_cs_intros] = is_cat_pullbackD lemma (in is_cat_pushout) is_cat_pushout_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "X' = X" shows "x : \'\\'\\'\\'\\' >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_pushout_axioms) mk_ide rf is_cat_pushout_def |intro is_cat_pushoutI| |dest is_cat_pushoutD[dest]| |elim is_cat_pushoutE[elim]| lemmas [cat_lim_cs_intros] = is_cat_pushoutD text\Duality.\ lemma (in is_cat_pullback) is_cat_pushout_op: "op_ntcf x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_pushoutI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_pullback) is_cat_pushout_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_pushout_op) lemmas [cat_op_intros] = is_cat_pullback.is_cat_pushout_op' lemma (in is_cat_pushout) is_cat_pullback_op: "op_ntcf x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_pullbackI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros)+ lemma (in is_cat_pushout) is_cat_pullback_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_pullback_op) lemmas [cat_op_intros] = is_cat_pushout.is_cat_pullback_op' text\Elementary properties.\ lemma cat_cone_cospan: assumes "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "cf_scospan \ \ \ \ \ \ \" shows "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "\ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" proof- interpret x: is_cat_cone \ X \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x by (rule assms(1)) interpret cospan: cf_scospan \ \ \ \ \ \ \ by (rule assms(2)) have \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S : \\<^sub>S\<^sub>S \\<^bsub>\\\\<^sub>C\<^esub> \\<^sub>S\<^sub>S" and \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S : \\<^sub>S\<^sub>S \\<^bsub>\\\\<^sub>C\<^esub> \\<^sub>S\<^sub>S" by (cs_concl cs_simp: cs_intro: cat_ss_cs_intros)+ from x.ntcf_Comp_commute[OF \\<^sub>S\<^sub>S] \\<^sub>S\<^sub>S \\<^sub>S\<^sub>S show "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" by (cs_prems cs_simp: cat_ss_cs_simps cat_cs_simps cs_intro: cat_cs_intros) moreover from x.ntcf_Comp_commute[OF \\<^sub>S\<^sub>S] \\<^sub>S\<^sub>S \\<^sub>S\<^sub>S show "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" by (cs_prems cs_simp: cat_ss_cs_simps cat_cs_simps cs_intro: cat_cs_intros) ultimately show "\ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" by simp qed lemma (in is_cat_pullback) cat_pb_cone_cospan: shows "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "\ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ = \ \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" by (all\rule cat_cone_cospan[OF is_cat_cone_axioms cf_scospan_axioms]\) lemma cat_cocone_span: assumes "x : \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e X : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "cf_sspan \ \ \ \ \ \ \" shows "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" and "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" and "x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" proof- interpret x: is_cat_cocone \ X \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x by (rule assms(1)) interpret span: cf_sspan \ \ \ \ \ \ \ by (rule assms(2)) note op = cat_cone_cospan [ OF x.is_cat_cone_op[unfolded cat_op_simps] span.cf_scospan_op, unfolded cat_op_simps ] from op(1) show "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" by ( cs_prems cs_simp: cat_ss_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_ss_cs_intros ) moreover from op(2) show "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" by ( cs_prems cs_simp: cat_ss_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_ss_cs_intros ) ultimately show "x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" by auto qed lemma (in is_cat_pushout) cat_po_cocone_span: shows "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" and "x\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" and "x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> \" by (all\rule cat_cocone_span[OF is_cat_cocone_axioms cf_sspan_axioms]\) subsubsection\Universal property\ lemma is_cat_pullbackI': assumes "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "cf_scospan \ \ \ \ \ \ \" and "\x' X'. x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \ \ \!f'. f' : X' \\<^bsub>\\<^esub> X \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" shows "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" -proof(intro is_cat_pullbackI is_cat_limitI') +proof(intro is_cat_pullbackI is_cat_limitI) show "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms(1)) interpret x: is_cat_cone \ X \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x by (rule assms(1)) show "cf_scospan \ \ \ \ \ \ \" by (rule assms(2)) interpret cospan: cf_scospan \ \ \ \ \ \ \ by (rule assms(2)) fix u' r' assume prems: "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" interpret u': is_cat_cone \ r' \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ u' by (rule prems) from assms(3)[OF prems] obtain f' where f': "f' : r' \\<^bsub>\\<^esub> X" and u'_\\<^sub>S\<^sub>S: "u'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" and u'_\\<^sub>S\<^sub>S: "u'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" and unique_f': "\f''. \ f'' : r' \\<^bsub>\\<^esub> X; u'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f''; u'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'' \ \ f'' = f'" by metis show "\!f'. f' : r' \\<^bsub>\\<^esub> X \ u' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'" proof(intro ex1I conjI; (elim conjE)?) show "u' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'" proof(rule ntcf_eqI) show "u' : cf_const \\\\<^sub>C \ r' \\<^sub>C\<^sub>F \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by (rule u'.is_ntcf_axioms) from f' show "x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f' : cf_const \\\\<^sub>C \ r' \\<^sub>C\<^sub>F \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f' have dom_rhs: "\\<^sub>\ ((x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f')\NTMap\) = \\\\<^sub>C\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f')\NTMap\" proof(rule vsv_eqI, unfold cat_cs_simps dom_rhs) fix a assume prems': "a \\<^sub>\ \\\\<^sub>C\Obj\" from this f' x.is_ntcf_axioms show "u'\NTMap\\a\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f')\NTMap\\a\" by (elim the_cat_scospan_ObjE; simp only:) ( cs_concl cs_simp: cat_cs_simps cat_ss_cs_simps u'_\\<^sub>S\<^sub>S u'_\\<^sub>S\<^sub>S cat_cone_cospan(1)[OF assms(1,2)] cat_cone_cospan(1)[OF prems assms(2)] cs_intro: cat_cs_intros cat_ss_cs_intros )+ qed (cs_concl cs_intro: cat_cs_intros | auto)+ qed simp_all fix f'' assume prems: "f'' : r' \\<^bsub>\\<^esub> X" "u' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f''" have \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S \\<^sub>\ \\\\<^sub>C\Obj\" and \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S \\<^sub>\ \\\\<^sub>C\Obj\" by (cs_concl cs_simp: cs_intro: cat_ss_cs_intros)+ have "u'\NTMap\\a\ = x\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> f''" if "a \\<^sub>\ \\\\<^sub>C\Obj\" for a proof- from prems(2) have "u'\NTMap\\a\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'')\NTMap\\a\" by simp from this that prems(1) show "u'\NTMap\\a\ = x\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> f''" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed from unique_f'[OF prems(1) this[OF \\<^sub>S\<^sub>S] this[OF \\<^sub>S\<^sub>S]] show "f'' = f'". qed (intro f') qed lemma is_cat_pushoutI': assumes "x : \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e X : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "cf_sspan \ \ \ \ \ \ \" and "\x' X'. x' : \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e X' : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \ \ \!f'. f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" shows "x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret x: is_cat_cocone \ X \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x by (rule assms(1)) interpret span: cf_sspan \ \ \ \ \ \ \ by (rule assms(2)) have assms_3': "\!f'. f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f'" if "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>op_cat \\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" for x' X' proof- from that(1) have [cat_op_simps]: "f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" for f' by (intro iffI conjI; (elim conjE)?) ( cs_concl cs_simp: category.op_cat_Comp[symmetric] cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_ss_cs_intros )+ interpret x': is_cat_cone \ X' \\\\\<^sub>C\ \op_cat \\ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>op_cat \\<^esub>\ x' by (rule that) show ?thesis unfolding cat_op_simps by ( rule assms(3)[ OF x'.is_cat_cocone_op[unfolded cat_op_simps], unfolded cat_op_simps ] ) qed interpret op_x: is_cat_pullback \ \ \ \ \ \ \op_cat \\ X \op_ntcf x\ using is_cat_pullbackI' [ OF x.is_cat_cone_op[unfolded cat_op_simps] span.cf_scospan_op, unfolded cat_op_simps, OF assms_3' ] by simp show "x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \" by (rule op_x.is_cat_pushout_op[unfolded cat_op_simps]) qed lemma (in is_cat_pullback) cat_pb_unique_cone: assumes "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X' \\<^bsub>\\<^esub> X \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" proof- interpret x': is_cat_cone \ X' \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x' by (rule assms) - from cat_lim_unique_cone[OF assms] obtain f' + from cat_lim_ua_fo[OF assms] obtain f' where f': "f' : X' \\<^bsub>\\<^esub> X" and x'_def: "x' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'" and unique_f': "\f''. \ f'' : X' \\<^bsub>\\<^esub> X; x' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'' \ \ f'' = f'" by auto have \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S \\<^sub>\ \\\\<^sub>C\Obj\" and \\<^sub>S\<^sub>S: "\\<^sub>S\<^sub>S \\<^sub>\ \\\\<^sub>C\Obj\" by (cs_concl cs_intro: cat_ss_cs_intros)+ show ?thesis proof(intro ex1I conjI; (elim conjE)?) show "f' : X' \\<^bsub>\\<^esub> X" by (rule f') have "x'\NTMap\\a\ = x\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> f'" if "a \\<^sub>\ \\\\<^sub>C\Obj\" for a proof- from x'_def have "x'\NTMap\\a\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f')\NTMap\\a\" by simp from this that f' show "x'\NTMap\\a\ = x\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> f'" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed from this[OF \\<^sub>S\<^sub>S] this[OF \\<^sub>S\<^sub>S] show "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto fix f'' assume prems': "f'' : X' \\<^bsub>\\<^esub> X" "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f''" "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f''" have "x' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f''" proof(rule ntcf_eqI) show "x' : cf_const \\\\<^sub>C \ X' \\<^sub>C\<^sub>F \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by (rule x'.is_ntcf_axioms) from prems'(1) show "x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'' : cf_const \\\\<^sub>C \ X' \\<^sub>C\<^sub>F \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ (x'\NTMap\) = \\\\<^sub>C\Obj\" by (cs_concl cs_simp: cat_cs_simps) from prems'(1) have dom_rhs: "\\<^sub>\ ((x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'')\NTMap\) = \\\\<^sub>C\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "x'\NTMap\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'')\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems'': "a \\<^sub>\ \\\\<^sub>C\Obj\" from this prems'(1) show "x'\NTMap\\a\ = (x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'')\NTMap\\a\" by (elim the_cat_scospan_ObjE; simp only:) ( cs_concl cs_simp: prems'(2,3) cat_cone_cospan(1,2)[OF assms cf_scospan_axioms] cat_pb_cone_cospan cat_ss_cs_simps cat_cs_simps cs_intro: cat_ss_cs_intros cat_cs_intros )+ qed (auto simp: cat_cs_intros) qed simp_all from unique_f'[OF prems'(1) this] show "f'' = f'". qed qed lemma (in is_cat_pullback) cat_pb_unique: assumes "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X' \\<^bsub>\\<^esub> X \ x' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f'" by (rule cat_lim_unique[OF is_cat_pullbackD(1)[OF assms]]) lemma (in is_cat_pullback) cat_pb_unique': assumes "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X' \\<^bsub>\\<^esub> X \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f'" proof- interpret x': is_cat_pullback \ \ \ \ \ \ \ X' x' by (rule assms(1)) show ?thesis by (rule cat_pb_unique_cone[OF x'.is_cat_cone_axioms]) qed lemma (in is_cat_pushout) cat_po_unique_cocone: assumes "x' : \\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub> >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e X' : \\\\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" proof- interpret x': is_cat_cocone \ X' \\\\\<^sub>C\ \ \\\\\\\\\\\\\<^sub>C\<^sub>F\<^bsub>\\<^esub>\ x' by (rule assms(1)) have [cat_op_simps]: "f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" for f' by (intro iffI conjI; (elim conjE)?) ( cs_concl cs_simp: category.op_cat_Comp[symmetric] cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_ss_cs_intros )+ show ?thesis by ( rule is_cat_pullback.cat_pb_unique_cone[ OF is_cat_pullback_op x'.is_cat_cone_op[unfolded cat_op_simps], unfolded cat_op_simps ] ) qed lemma (in is_cat_pushout) cat_po_unique: assumes "x' : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X \\<^bsub>\\<^esub> X' \ x' = ntcf_const \\\\<^sub>C \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F x" by (rule cat_colim_unique[OF is_cat_pushoutD(1)[OF assms]]) lemma (in is_cat_pushout) cat_po_unique': assumes "x' : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : X \\<^bsub>\\<^esub> X' \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\ \ x'\NTMap\\\\<^sub>S\<^sub>S\ = f' \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" proof- interpret x': is_cat_pushout \ \ \ \ \ \ \ X' x' by (rule assms(1)) show ?thesis by (rule cat_po_unique_cocone[OF x'.is_cat_cocone_axioms]) qed lemma cat_pullback_ex_is_arr_isomorphism: assumes "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" and "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : X' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X" and "x' = x \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \\\\<^sub>C \ f" proof- interpret x: is_cat_pullback \ \ \ \ \ \ \ X x by (rule assms(1)) interpret x': is_cat_pullback \ \ \ \ \ \ \ X' x' by (rule assms(2)) from that show ?thesis by ( elim cat_lim_ex_is_arr_isomorphism[ OF x.is_cat_limit_axioms x'.is_cat_limit_axioms ] ) qed lemma cat_pullback_ex_is_arr_isomorphism': assumes "x : X <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" and "x' : X' <\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>b \\\\\\\\\ \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : X' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X" and "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f" and "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f" proof- interpret x: is_cat_pullback \ \ \ \ \ \ \ X x by (rule assms(1)) interpret x': is_cat_pullback \ \ \ \ \ \ \ X' x' by (rule assms(2)) obtain f where f: "f : X' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X" and "j \\<^sub>\ \\\\<^sub>C\Obj\ \ x'\NTMap\\j\ = x\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f" for j by ( elim cat_lim_ex_is_arr_isomorphism'[ OF x.is_cat_limit_axioms x'.is_cat_limit_axioms ] ) then have "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f" "x'\NTMap\\\\<^sub>S\<^sub>S\ = x\NTMap\\\\<^sub>S\<^sub>S\ \\<^sub>A\<^bsub>\\<^esub> f" by (auto simp: cat_ss_cs_intros) with f show ?thesis using that by simp qed lemma cat_pushout_ex_is_arr_isomorphism: assumes "x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \" and "x' : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X' \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : X \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X'" and "x' = ntcf_const \\\\<^sub>C \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F x" proof- interpret x: is_cat_pushout \ \ \ \ \ \ \ X x by (rule assms(1)) interpret x': is_cat_pushout \ \ \ \ \ \ \ X' x' by (rule assms(2)) from that show ?thesis by ( elim cat_colim_ex_is_arr_isomorphism[ OF x.is_cat_colimit_axioms x'.is_cat_colimit_axioms ] ) qed lemma cat_pushout_ex_is_arr_isomorphism': assumes "x : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X \\\<^sub>C\<^bsub>\\<^esub> \" and "x' : \\\\\\\\\ >\<^sub>C\<^sub>F\<^sub>.\<^sub>p\<^sub>o X' \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : X \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X'" and "x'\NTMap\\\\<^sub>S\<^sub>S\ = f \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "x'\NTMap\\\\<^sub>S\<^sub>S\ = f \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" proof- interpret x: is_cat_pushout \ \ \ \ \ \ \ X x by (rule assms(1)) interpret x': is_cat_pushout \ \ \ \ \ \ \ X' x' by (rule assms(2)) obtain f where f: "f : X \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> X'" and "j \\<^sub>\ \\\\<^sub>C\Obj\ \ x'\NTMap\\j\ = f \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\j\" for j by ( elim cat_colim_ex_is_arr_isomorphism'[ OF x.is_cat_colimit_axioms x'.is_cat_colimit_axioms, unfolded the_cat_parallel_components(1) ] ) then have "x'\NTMap\\\\<^sub>S\<^sub>S\ = f \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" and "x'\NTMap\\\\<^sub>S\<^sub>S\ = f \\<^sub>A\<^bsub>\\<^esub> x\NTMap\\\\<^sub>S\<^sub>S\" by (auto simp: cat_ss_cs_intros) with f show ?thesis using that by simp qed subsection\Equalizers and coequalizers\ subsubsection\Definition and elementary properties\ text\ See \cite{noauthor_wikipedia_2001}\footnote{ \url{https://en.wikipedia.org/wiki/Equaliser_(mathematics)} }. \ locale is_cat_equalizer = is_cat_limit \ \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ \ \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ E \ for \ \ \ \ \ \ E \ + assumes cat_eq_\[cat_lim_cs_intros]: "\ : \ \\<^bsub>\\<^esub> \" and cat_eq_\[cat_lim_cs_intros]: "\ : \ \\<^bsub>\\<^esub> \" syntax "_is_cat_equalizer" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q '(_,_,_,_') :/ \\\<^sup>2\<^sub>C \\\<^sub>C\ _)\ [51, 51, 51, 51, 51, 51] 51) translations "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_equalizer \ \ \ \ \ \ E \" locale is_cat_coequalizer = is_cat_colimit \ \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ \ \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ E \ for \ \ \ \ \ \ E \ + assumes cat_coeq_\[cat_lim_cs_intros]: "\ : \ \\<^bsub>\\<^esub> \" and cat_coeq_\[cat_lim_cs_intros]: "\ : \ \\<^bsub>\\<^esub> \" syntax "_is_cat_coequalizer" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" (\(_ :/ '(_,_,_,_') >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q _ :/ \\\<^sup>2\<^sub>C \\\<^sub>C\ _)\ [51, 51, 51, 51, 51, 51] 51) translations "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_cat_coequalizer \ \ \ \ \ \ E \" text\Rules.\ lemma (in is_cat_equalizer) is_cat_equalizer_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "E' = E" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\',\',\',\') : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_equalizer_axioms) mk_ide rf is_cat_equalizer_def[unfolded is_cat_equalizer_axioms_def] |intro is_cat_equalizerI| |dest is_cat_equalizerD[dest]| |elim is_cat_equalizerE[elim]| lemmas [cat_lim_cs_intros] = is_cat_equalizerD(1) lemma (in is_cat_coequalizer) is_cat_coequalizer_axioms'[cat_lim_cs_intros]: assumes "\' = \" and "E' = E" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : (\',\',\',\') >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E' : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_cat_coequalizer_axioms) mk_ide rf is_cat_coequalizer_def[unfolded is_cat_coequalizer_axioms_def] |intro is_cat_coequalizerI| |dest is_cat_coequalizerD[dest]| |elim is_cat_coequalizerE[elim]| lemmas [cat_lim_cs_intros] = is_cat_coequalizerD(1) text\Elementary properties.\ sublocale is_cat_equalizer \ cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (intro cf_parallelI cat_parallelI) (simp_all add: cat_parallel_cs_intros cat_lim_cs_intros cat_cs_intros) sublocale is_cat_coequalizer \ cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (intro cf_parallelI cat_parallelI) ( simp_all add: cat_parallel_cs_intros cat_lim_cs_intros cat_cs_intros cat_PL_ineq[symmetric] ) text\Duality.\ lemma (in is_cat_equalizer) is_cat_coequalizer_op: "op_ntcf \ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_coequalizerI) (cs_concl cs_simp: cat_op_simps cs_intro: cat_op_intros cat_lim_cs_intros)+ lemma (in is_cat_equalizer) is_cat_coequalizer_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_coequalizer_op) lemmas [cat_op_intros] = is_cat_equalizer.is_cat_coequalizer_op' lemma (in is_cat_coequalizer) is_cat_equalizer_op: "op_ntcf \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (intro is_cat_equalizerI) ( cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_lim_cs_intros )+ lemma (in is_cat_coequalizer) is_cat_equalizer_op'[cat_op_intros]: assumes "\' = op_cat \" shows "op_ntcf \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_cat_equalizer_op) lemmas [cat_op_intros] = is_cat_coequalizer.is_cat_equalizer_op' text\Elementary properties.\ lemma cf_parallel_if_is_cat_cone: assumes "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \" proof- let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret is_cat_cone \ E ?II \ ?II_II \ by (rule assms(1)) show ?thesis by (intro cf_parallelI cat_parallelI) ( simp_all add: assms cat_parallel_cs_intros cat_lim_cs_intros cat_cs_intros ) qed lemma cf_parallel_if_is_cat_cocone: assumes "\' : \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E' : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \" proof- let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret is_cat_cocone \ E' ?II \ ?II_II \' by (rule assms(1)) show ?thesis by (intro cf_parallelI cat_parallelI) ( simp_all add: assms cat_parallel_cs_intros cat_lim_cs_intros cat_cs_intros cat_PL_ineq[symmetric] ) qed lemma (in category) cat_cf_parallel_cat_equalizer: assumes "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \" using assms by (intro cf_parallelI cat_parallelI) (auto simp: cat_parallel_cs_intros cat_cs_intros) lemma (in category) cat_cf_parallel_cat_coequalizer: assumes "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \" using assms by (intro cf_parallelI cat_parallelI) (simp_all add: cat_parallel_cs_intros cat_cs_intros cat_PL_ineq[symmetric]) lemma cat_cone_cf_par_eps_NTMap_app: assumes "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" proof- let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret \: is_cat_cone \ E ?II \ ?II_II \ by (rule assms(1)) from assms(2,3) have \: "\ \\<^sub>\ \\Obj\" and \: "\ \\<^sub>\ \\Obj\" by auto interpret par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (intro cf_parallel_if_is_cat_cone, rule assms) (auto intro: assms \ \) have \\<^sub>P\<^sub>L: "\\<^sub>P\<^sub>L : \\<^sub>P\<^sub>L \\<^bsub>\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\<^esub> \\<^sub>P\<^sub>L" and \\<^sub>P\<^sub>L: "\\<^sub>P\<^sub>L : \\<^sub>P\<^sub>L \\<^bsub>\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\<^esub> \\<^sub>P\<^sub>L" by ( simp_all add: par.the_cat_parallel_is_arr_\\\ par.the_cat_parallel_is_arr_\\\ ) from \.ntcf_Comp_commute[OF \\<^sub>P\<^sub>L] show "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" by (*slow*) ( cs_prems cs_simp: cat_parallel_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) from \.ntcf_Comp_commute[OF \\<^sub>P\<^sub>L] show "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" by (*slow*) ( cs_prems cs_simp: cat_parallel_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) qed lemma cat_cocone_cf_par_eps_NTMap_app: assumes "\ : \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" proof- let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret \: is_cat_cocone \ E ?II \ ?II_II \ by (rule assms(1)) from assms(2,3) have \: "\ \\<^sub>\ \\Obj\" and \: "\ \\<^sub>\ \\Obj\" by auto interpret par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (intro cf_parallel_if_is_cat_cocone, rule assms) (auto intro: assms \ \) note \_NTMap_app = cat_cone_cf_par_eps_NTMap_app[ OF \.is_cat_cone_op[unfolded cat_op_simps], unfolded cat_op_simps, OF assms(2,3) ] from \_NTMap_app show \_NTMap_app: "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_parallel_cs_simps category.op_cat_Comp[symmetric] cs_intro: cat_cs_intros cat_parallel_cs_intros )+ qed lemma (in is_cat_equalizer) cat_eq_2_eps_NTMap_app: "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" "\\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" by ( intro cat_cone_cf_par_eps_NTMap_app[ OF is_cat_cone_axioms cat_eq_\ cat_eq_\ ] )+ lemma (in is_cat_coequalizer) cat_coeq_2_eps_NTMap_app: "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" "\\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" by ( intro cat_cocone_cf_par_eps_NTMap_app[ OF is_cat_cocone_axioms cat_coeq_\ cat_coeq_\ ] )+ lemma (in is_cat_equalizer) cat_eq_Comp_eq: "\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" "\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\ = \ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" unfolding cat_eq_2_eps_NTMap_app[symmetric] by simp_all lemma (in is_cat_coequalizer) cat_coeq_Comp_eq: "\\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" "\\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> \" unfolding cat_coeq_2_eps_NTMap_app[symmetric] by simp_all subsubsection\Universal property\ lemma is_cat_equalizerI': assumes "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\\' E'. \' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \ \ \!f'. f' : E' \\<^bsub>\\<^esub> E \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'" shows "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret \: is_cat_cone \ E ?II \ ?II_II \ by (rule assms(1)) interpret \: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (rule \.NTDom.HomCod.cat_cf_parallel_cat_equalizer[OF assms(2,3)]) show ?thesis - proof(intro is_cat_equalizerI is_cat_limitI' assms(1-3)) + proof(intro is_cat_equalizerI is_cat_limitI assms(1-3)) fix u' r' assume prems: "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" interpret u': is_cat_cone \ r' ?II \ ?II_II u' by (rule prems) from assms(4)[OF prems] obtain f' where f': "f' : r' \\<^bsub>\\<^esub> E" and u'_NTMap_app_\: "u'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'" and unique_f': "\f''. \ f'' : r' \\<^bsub>\\<^esub> E; u'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'' \ \ f'' = f'" by metis show "\!f'. f' : r' \\<^bsub>\\<^esub> E \ u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'" proof(intro ex1I conjI; (elim conjE)?) show "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'" proof(rule ntcf_eqI) show "u' : cf_const ?II \ r' \\<^sub>C\<^sub>F ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by (rule u'.is_ntcf_axioms) from f' show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f' : cf_const ?II \ r' \\<^sub>C\<^sub>F ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cat_ss_cs_simps cs_intro: cat_cs_intros cat_ss_cs_intros ) have dom_lhs: "\\<^sub>\ (u'\NTMap\) = ?II\Obj\" unfolding cat_cs_simps by simp from f' have dom_rhs: "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f')\NTMap\) = ?II\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f')\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems': "a \\<^sub>\ ?II\Obj\" note [cat_parallel_cs_simps] = cat_cone_cf_par_eps_NTMap_app[OF u'.is_cat_cone_axioms assms(2-3)] cat_cone_cf_par_eps_NTMap_app[OF assms(1-3)] u'_NTMap_app_\ from prems' f' assms(2,3) show "u'\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f')\NTMap\\a\" by (elim the_cat_parallel_ObjE; simp only:) ( cs_concl cs_simp: cat_parallel_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) qed (cs_concl cs_intro: V_cs_intros cat_cs_intros)+ qed simp_all fix f'' assume prems'': "f'' : r' \\<^bsub>\\<^esub> E" "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f''" from prems''(2) have u'_NTMap_a: "u'\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'')\NTMap\\a\" for a by simp have "u'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f''" using u'_NTMap_a[of \\<^sub>P\<^sub>L] prems''(1) by ( cs_prems cs_simp: cat_parallel_cs_simps cat_cs_simps cs_intro: cat_parallel_cs_intros cat_cs_intros ) from unique_f'[OF prems''(1) this] show "f'' = f'". qed (rule f') qed qed lemma is_cat_coequalizerI': assumes "\ : \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" and "\\' E'. \' : \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E' : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \ \ \!f'. f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" shows "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?op_II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?op_II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ and ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ (op_cat \) \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ interpret \: is_cat_cocone \ E ?op_II \ ?op_II_II \ by (rule assms(1)) interpret par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (rule \.NTDom.HomCod.cat_cf_parallel_cat_coequalizer[OF assms(2,3)]) interpret op_par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \op_cat \\ by (rule par.cf_parallel_op) have assms_4': "\!f'. f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>op_cat \\<^esub> f'" if "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> op_cat \" for \' E' proof- have [cat_op_simps]: "f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" for f' by (intro iffI conjI; (elim conjE)?) ( cs_concl cs_simp: category.op_cat_Comp[symmetric] cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros )+ interpret \': is_cat_cone \ E' ?II \op_cat \\ ?II_II \' by (rule that) show ?thesis unfolding cat_op_simps by ( rule assms(4)[ OF \'.is_cat_cocone_op[unfolded cat_op_simps], unfolded cat_op_simps ] ) qed interpret op_\: is_cat_equalizer \ \ \ \ \ \op_cat \\ E \op_ntcf \\ by ( rule is_cat_equalizerI' [ OF \.is_cat_cone_op[unfolded cat_op_simps], unfolded cat_op_simps, OF assms(2,3) assms_4' ] ) show ?thesis by (rule op_\.is_cat_coequalizer_op[unfolded cat_op_simps]) qed lemma (in is_cat_equalizer) cat_eq_unique_cone: assumes "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" (is \\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \\) shows "\!f'. f' : E' \\<^bsub>\\<^esub> E \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'" proof- interpret \': is_cat_cone \ E' ?II \ ?II_II \' by (rule assms(1)) - from cat_lim_unique_cone[OF assms(1)] obtain f' where f': "f' : E' \\<^bsub>\\<^esub> E" + from cat_lim_ua_fo[OF assms(1)] obtain f' where f': "f' : E' \\<^bsub>\\<^esub> E" and \'_def: "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'" and unique: "\ f'' : E' \\<^bsub>\\<^esub> E; \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'' \ \ f'' = f'" for f'' by auto show ?thesis proof(intro ex1I conjI; (elim conjE)?) show f': "f' : E' \\<^bsub>\\<^esub> E" by (rule f') from \'_def have "\'\NTMap\\\\<^sub>P\<^sub>L\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f')\NTMap\\\\<^sub>P\<^sub>L\" by simp from this f' show \'_NTMap_app_I: "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'" by ( cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) fix f'' assume prems: "f'' : E' \\<^bsub>\\<^esub> E" "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f''" have "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f''" proof(rule ntcf_eqI[OF ]) show "\' : cf_const ?II \ E' \\<^sub>C\<^sub>F ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by (rule \'.is_ntcf_axioms) from f' prems(1) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'' : cf_const ?II \ E' \\<^sub>C\<^sub>F ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\'\NTMap\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'')\NTMap\" proof(rule vsv_eqI, unfold cat_cs_simps) show "vsv ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'')\NTMap\)" by (cs_concl cs_intro: cat_cs_intros) from prems(1) show "?II\Obj\ = \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'')\NTMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) fix a assume prems': "a \\<^sub>\ ?II\Obj\" note [cat_cs_simps] = cat_eq_2_eps_NTMap_app cat_cone_cf_par_eps_NTMap_app[ OF \'.is_cat_cone_axioms cf_parallel_\' cf_parallel_\' ] from prems' prems(1) have [cat_cs_simps]: "\'\NTMap\\a\ = \\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> f''" by (elim the_cat_parallel_ObjE; simp only:) ( cs_concl cs_simp: cat_cs_simps cat_parallel_cs_simps prems(2) cs_intro: cat_cs_intros cat_parallel_cs_intros )+ from prems' prems show "\'\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const ?II \ f'')\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed auto qed simp_all from unique[OF prems(1) this] show "f'' = f'" . qed qed lemma (in is_cat_equalizer) cat_eq_unique: assumes "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : E' \\<^bsub>\\<^esub> E \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ f'" by (rule cat_lim_unique[OF is_cat_equalizerD(1)[OF assms]]) lemma (in is_cat_equalizer) cat_eq_unique': assumes "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : E' \\<^bsub>\\<^esub> E \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f'" proof- interpret \': is_cat_equalizer \ \ \ \ \ \ E' \' by (rule assms(1)) show ?thesis by (rule cat_eq_unique_cone[OF \'.is_cat_cone_axioms]) qed lemma (in is_cat_coequalizer) cat_coeq_unique_cocone: assumes "\' : \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E' : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" (is \\' : ?II_II >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>c\<^sub>o\<^sub>n\<^sub>e E' : ?II \\\<^sub>C\<^bsub>\\<^esub> \\) shows "\!f'. f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" proof- interpret \': is_cat_cocone \ E' ?II \ ?II_II \' by (rule assms(1)) have [cat_op_simps]: "f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>op_cat \\<^esub> f' \ f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" for f' by (intro iffI conjI; (elim conjE)?) ( cs_concl cs_simp: category.op_cat_Comp[symmetric] cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros )+ show ?thesis by ( rule is_cat_equalizer.cat_eq_unique_cone[ OF is_cat_equalizer_op \'.is_cat_cone_op[unfolded cat_op_simps], unfolded cat_op_simps ] ) qed lemma (in is_cat_coequalizer) cat_coeq_unique: assumes "\' : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E' : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : E \\<^bsub>\\<^esub> E' \ \' = ntcf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ f' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by (rule cat_colim_unique[OF is_cat_coequalizerD(1)[OF assms]]) lemma (in is_cat_coequalizer) cat_coeq_unique': assumes "\' : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E' : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" shows "\!f'. f' : E \\<^bsub>\\<^esub> E' \ \'\NTMap\\\\<^sub>P\<^sub>L\ = f' \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" proof- interpret \': is_cat_coequalizer \ \ \ \ \ \ E' \' by (rule assms(1)) show ?thesis by (rule cat_coeq_unique_cocone[OF \'.is_cat_cocone_axioms]) qed lemma cat_equalizer_2_ex_is_arr_isomorphism: assumes "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : E' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E" and "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ f" proof- interpret \: is_cat_equalizer \ \ \ \ \ \ E \ by (rule assms(1)) interpret \': is_cat_equalizer \ \ \ \ \ \ E' \' by (rule assms(2)) from that show ?thesis by ( elim cat_lim_ex_is_arr_isomorphism[ OF \.is_cat_limit_axioms \'.is_cat_limit_axioms ] ) qed lemma cat_equalizer_2_ex_is_arr_isomorphism': assumes "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : E' <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : E' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E" and "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f" and "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f" proof- interpret \: is_cat_equalizer \ \ \ \ \ \ E \ by (rule assms(1)) interpret \': is_cat_equalizer \ \ \ \ \ \ E' \' by (rule assms(2)) obtain f where f: "f : E' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E" and "j \\<^sub>\ \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\ \ \'\NTMap\\j\ = \\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f" for j by ( elim cat_lim_ex_is_arr_isomorphism'[ OF \.is_cat_limit_axioms \'.is_cat_limit_axioms ] ) then have "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f" "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> f" unfolding the_cat_parallel_components by auto with f show ?thesis using that by simp qed lemma cat_coequalizer_2_ex_is_arr_isomorphism: assumes "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E' : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : E \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E'" and "\' = ntcf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" proof- interpret \: is_cat_coequalizer \ \ \ \ \ \ E \ by (rule assms(1)) interpret \': is_cat_coequalizer \ \ \ \ \ \ E' \' by (rule assms(2)) from that show ?thesis by ( elim cat_colim_ex_is_arr_isomorphism[ OF \.is_cat_colimit_axioms \'.is_cat_colimit_axioms ] ) qed lemma cat_coequalizer_2_ex_is_arr_isomorphism': assumes "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E' : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" obtains f where "f : E \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E'" and "\'\NTMap\\\\<^sub>P\<^sub>L\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" and "\'\NTMap\\\\<^sub>P\<^sub>L\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" proof- interpret \: is_cat_coequalizer \ \ \ \ \ \ E \ by (rule assms(1)) interpret \': is_cat_coequalizer \ \ \ \ \ \ E' \' by (rule assms(2)) obtain f where f: "f : E \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> E'" and "j \\<^sub>\ \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\ \ \'\NTMap\\j\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\j\" for j by ( elim cat_colim_ex_is_arr_isomorphism'[ OF \.is_cat_colimit_axioms \'.is_cat_colimit_axioms ] ) then have "\'\NTMap\\\\<^sub>P\<^sub>L\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" "\'\NTMap\\\\<^sub>P\<^sub>L\ = f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" unfolding the_cat_parallel_components by auto with f show ?thesis using that by simp qed subsection\Projection cone\ subsubsection\Definition and elementary properties\ definition ntcf_obj_prod_base :: "V \ V \ (V \ V) \ V \ (V \ V) \ V" where "ntcf_obj_prod_base \ I F P f = [(\j\\<^sub>\:\<^sub>C I\Obj\. f j), cf_const (:\<^sub>C I) \ P, :\: I F \, :\<^sub>C I, \]\<^sub>\" text\Components.\ lemma ntcf_obj_prod_base_components: shows "ntcf_obj_prod_base \ I F P f\NTMap\ = (\j\\<^sub>\:\<^sub>C I\Obj\. f j)" and "ntcf_obj_prod_base \ I F P f\NTDom\ = cf_const (:\<^sub>C I) \ P" and "ntcf_obj_prod_base \ I F P f\NTCod\ = :\: I F \" and "ntcf_obj_prod_base \ I F P f\NTDGDom\ = :\<^sub>C I" and "ntcf_obj_prod_base \ I F P f\NTDGCod\ = \" unfolding ntcf_obj_prod_base_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda ntcf_obj_prod_base_components(1) |vsv ntcf_obj_prod_base_NTMap_vsv[cat_cs_intros]| |vdomain ntcf_obj_prod_base_NTMap_vdomain[cat_cs_simps]| |app ntcf_obj_prod_base_NTMap_app[cat_cs_simps]| subsubsection\Projection natural transformation is a cone\ lemma (in tm_cf_discrete) tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone: assumes "P \\<^sub>\ \\Obj\" and "\a. a \\<^sub>\ I \ f a : P \\<^bsub>\\<^esub> F a" shows "ntcf_obj_prod_base \ I F P f : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: I F \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') from assms(2) have [cat_cs_intros]: "\ a \\<^sub>\ I; P' = P; Fa = F a \ \ f a : P' \\<^bsub>\\<^esub> Fa" for a P' Fa by simp show "vfsequence (ntcf_obj_prod_base \ I F P f)" unfolding ntcf_obj_prod_base_def by auto show "vcard (ntcf_obj_prod_base \ I F P f) = 5\<^sub>\" unfolding ntcf_obj_prod_base_def by (auto simp: nat_omega_simps) from assms show "cf_const (:\<^sub>C I) \ P : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_intro: cf_discrete_vdomain_vsubset_Vset cat_discrete_cs_intros cat_cs_intros ) show ":\: I F \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_discrete_cs_intros) show "ntcf_obj_prod_base \ I F P f\NTMap\\a\ : cf_const (:\<^sub>C I) \ P\ObjMap\\a\ \\<^bsub>\\<^esub> :\: I F \\ObjMap\\a\" if "a \\<^sub>\ :\<^sub>C I\Obj\" for a proof- from that have "a \\<^sub>\ I" unfolding the_cat_discrete_components by simp from that this show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps cs_intro: cat_cs_intros ) qed show "ntcf_obj_prod_base \ I F P f\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> cf_const (:\<^sub>C I) \ P\ArrMap\\g\ = :\: I F \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> ntcf_obj_prod_base \ I F P f\NTMap\\a\" if "g : a \\<^bsub>:\<^sub>C I\<^esub> b" for a b g proof- note g = the_cat_discrete_is_arrD[OF that] from that g(4)[unfolded g(7-9)] g(1)[unfolded g(7-9)] show ?thesis unfolding g(7-9) by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps cs_intro: cf_discrete_vdomain_vsubset_Vset cat_cs_intros cat_discrete_cs_intros ) qed - from assms(1) show "cf_const (:\<^sub>C I) \ P : :\<^sub>C I \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by - ( - cs_concl cs_intro: - cat_cs_intros cat_small_cs_intros cat_small_discrete_cs_intros - ) qed ( auto simp: assms ntcf_obj_prod_base_components tm_cf_discrete_the_cf_discrete_is_tm_functor ) lemma (in tm_cf_discrete) tm_cf_discrete_ntcf_obj_prod_base_is_cat_obj_prod: assumes "P \\<^sub>\ \\Obj\" and "\a. a \\<^sub>\ I \ f a : P \\<^bsub>\\<^esub> F a" and "\u' r'. \ u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: I F \ : :\<^sub>C I \\\<^sub>C\<^bsub>\\<^esub> \ \ \ \!f'. f' : r' \\<^bsub>\\<^esub> P \ u' = ntcf_obj_prod_base \ I F P f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const (:\<^sub>C I) \ f'" shows "ntcf_obj_prod_base \ I F P f : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ F : I \\\<^sub>C\<^bsub>\\<^esub> \" proof ( intro is_cat_obj_prodI - is_cat_limitI' + is_cat_limitI tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone[OF assms(1,2), simplified] assms(1,3) ) show "cf_discrete \ I F \" by (cs_concl cs_intro: cat_small_discrete_cs_intros) qed subsection\Equalizer cone\ subsubsection\Definition and elementary properties\ definition ntcf_equalizer_base :: "V \ V \ V \ V \ V \ V \ (V \ V) \ V" where "ntcf_equalizer_base \ \ \ \ \ E e = [ (\x\\<^sub>\\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\. e x), cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E, \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \, \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L, \ ]\<^sub>\" text\Components.\ lemma ntcf_equalizer_base_components: shows "ntcf_equalizer_base \ \ \ \ \ E e\NTMap\ = (\x\\<^sub>\\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\. e x)" and [cat_lim_cs_simps]: "ntcf_equalizer_base \ \ \ \ \ E e\NTDom\ = cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E" and [cat_lim_cs_simps]: "ntcf_equalizer_base \ \ \ \ \ E e\NTCod\ = \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \" and [cat_lim_cs_simps]: "ntcf_equalizer_base \ \ \ \ \ E e\NTDGDom\ = \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L" and [cat_lim_cs_simps]: "ntcf_equalizer_base \ \ \ \ \ E e\NTDGCod\ = \" unfolding ntcf_equalizer_base_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda ntcf_equalizer_base_components(1) |vsv ntcf_equalizer_base_NTMap_vsv[cat_lim_cs_intros]| |vdomain ntcf_equalizer_base_NTMap_vdomain[cat_lim_cs_simps]| |app ntcf_equalizer_base_NTMap_app[cat_lim_cs_simps]| subsubsection\Equalizer cone is a cone\ lemma (in category) cat_ntcf_equalizer_base_is_cat_cone: assumes "e \\<^sub>P\<^sub>L : E \\<^bsub>\\<^esub> \" and "e \\<^sub>P\<^sub>L : E \\<^bsub>\\<^esub> \" and "e \\<^sub>P\<^sub>L = \ \\<^sub>A\<^bsub>\\<^esub> e \\<^sub>P\<^sub>L" and "e \\<^sub>P\<^sub>L = \ \\<^sub>A\<^bsub>\\<^esub> e \\<^sub>P\<^sub>L" and "\ : \ \\<^bsub>\\<^esub> \" and "\ : \ \\<^bsub>\\<^esub> \" shows "ntcf_equalizer_base \ \ \ \ \ E e : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ \ by (intro cf_parallelI cat_parallelI assms(5,6)) (simp_all add: cat_parallel_cs_intros cat_cs_intros) show ?thesis proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') show "vfsequence (ntcf_equalizer_base \ \ \ \ \ E e)" unfolding ntcf_equalizer_base_def by auto show "vcard (ntcf_equalizer_base \ \ \ \ \ E e) = 5\<^sub>\" unfolding ntcf_equalizer_base_def by (simp add: nat_omega_simps) from assms(2) show - "cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + "cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_parallel_cs_intros cat_cs_intros ) - then show "cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_small_cs_intros) from assms show - "\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_parallel_cs_intros) - then show "\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_small_cs_intros) + "\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \ : \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cs_intro: cat_parallel_cs_intros cat_small_cs_intros) show "ntcf_equalizer_base \ \ \ \ \ E e\NTMap\\i\ : cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E\ObjMap\\i\ \\<^bsub>\\<^esub> \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ObjMap\\i\" if "i \\<^sub>\ \\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\Obj\" for i proof- from that assms(1,2,5,6) show ?thesis by (elim the_cat_parallel_ObjE; simp only:) ( cs_concl cs_simp: cat_lim_cs_simps cat_cs_simps cat_parallel_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) qed show "ntcf_equalizer_base \ \ \ \ \ E e\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> cf_const (\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L) \ E\ArrMap\\f'\ = \\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \ \ \ \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> ntcf_equalizer_base \ \ \ \ \ E e\NTMap\\a'\" if "f' : a' \\<^bsub>\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\<^esub> b'" for a' b' f' using that assms(1,2,5,6) by (elim par.the_cat_parallel_is_arrE; simp only:) ( cs_concl cs_simp: cat_cs_simps cat_lim_cs_simps cat_parallel_cs_simps assms(3,4)[symmetric] cs_intro: cat_parallel_cs_intros )+ qed ( use assms(2) in \ cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cs_simp: cat_lim_cs_simps \ )+ qed subsection\Limits by products and equalizers\ lemma cat_limit_of_cat_prod_obj_and_cat_equalizer: \\See Theorem 1 in Chapter V-2 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\\ \ \ \. \ \ : \ \\<^bsub>\\<^esub> \; \ : \ \\<^bsub>\\<^esub> \ \ \ \E \. \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\A. tm_cf_discrete \ (\\Obj\) A \ \ \P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : \\Obj\ \\\<^sub>C\<^bsub>\\<^esub> \" and "\A. tm_cf_discrete \ (\\Arr\) A \ \ \P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : \\Arr\ \\\<^sub>C\<^bsub>\\<^esub> \" obtains r u where "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?L =\\u. \\ObjMap\\\\Cod\\u\\\ and ?R =\\i. \\ObjMap\\i\\ interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) have "?R j \\<^sub>\ \\Obj\" if "j \\<^sub>\ \\Obj\" for j by (cs_concl cs_intro: cat_cs_intros that) have "tm_cf_discrete \ (\\Obj\) ?R \" proof(intro tm_cf_discreteI) show "\\ObjMap\\i\ \\<^sub>\ \\Obj\" if "i \\<^sub>\ \\Obj\" for i by (cs_concl cs_intro: cat_cs_intros that) show "VLambda (\\Obj\) ?R \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (VLambda (\\Obj\) ?R) \\<^sub>\ Vset \" proof- have "\\<^sub>\ (VLambda (\\Obj\) ?R) \\<^sub>\ \\<^sub>\ (\\ObjMap\)" by (auto simp: \.cf_ObjMap_vdomain) moreover have "\\<^sub>\ (\\ObjMap\) \\<^sub>\ Vset \" by (force intro: vrange_in_VsetI \.tm_cf_ObjMap_in_Vset) ultimately show ?thesis by auto qed qed (auto simp: cat_small_cs_intros) show "(\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ Vset \" proof- have "\\<^sub>\ (\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ \\<^sub>\ (\\ArrMap\)" proof(rule vrange_VLambda_vsubset) fix x assume x: "x \\<^sub>\ \\Obj\" then have "\\CId\\x\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" by (auto intro: cat_cs_intros simp: cat_cs_simps) moreover from x have "\\CId\\\\ObjMap\\x\\ = \\ArrMap\\\\CId\\x\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show "\\CId\\\\ObjMap\\x\\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" by (simp add: \.ArrMap.vsv_vimageI2) qed moreover have "\\<^sub>\ (\\ArrMap\) \\<^sub>\ Vset \" by (force intro: vrange_in_VsetI \.tm_cf_ArrMap_in_Vset) ultimately show ?thesis by auto qed qed (auto simp: cat_small_cs_intros) qed (auto intro: cat_cs_intros) from assms(3)[where A=\?R\, OF this] obtain P\<^sub>O \\<^sub>O where \\<^sub>O: "\\<^sub>O : P\<^sub>O <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ ?R : \\Obj\ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \\<^sub>O: is_cat_obj_prod \ \\\Obj\\ ?R \ P\<^sub>O \\<^sub>O by (rule \\<^sub>O) have P\<^sub>O: "P\<^sub>O \\<^sub>\ \\Obj\" by (intro \\<^sub>O.cat_cone_obj) have "?L u \\<^sub>\ \\Obj\" if "u \\<^sub>\ \\Arr\" for u proof- from that obtain a b where "u : a \\<^bsub>\\<^esub> b" by auto then show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed have tm_cf_discrete: "tm_cf_discrete \ (\\Arr\) ?L \" proof(intro tm_cf_discreteI) show "\\ObjMap\\\\Cod\\f\\ \\<^sub>\ \\Obj\" if "f \\<^sub>\ \\Arr\" for f proof- from that obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed show "(\u\\<^sub>\\\Arr\. \\ObjMap\\\\Cod\\u\\) \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\u\\<^sub>\\\Arr\. ?L u) \\<^sub>\ Vset \" proof- have "\\<^sub>\ (\u\\<^sub>\\\Arr\. ?L u) \\<^sub>\ \\<^sub>\ (\\ObjMap\)" proof(rule vrange_VLambda_vsubset) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then show "\\ObjMap\\\\Cod\\f\\ \\<^sub>\ \\<^sub>\ (\\ObjMap\)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed moreover have "\\<^sub>\ (\\ObjMap\) \\<^sub>\ Vset \" by (auto intro: vrange_in_VsetI \.tm_cf_ObjMap_in_Vset) ultimately show ?thesis by auto qed qed (auto simp: cat_small_cs_intros) show "(\i\\<^sub>\\\Arr\. \\CId\\\\ObjMap\\\\Cod\\i\\\) \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\i\\<^sub>\\\Arr\. \\CId\\\\ObjMap\\\\Cod\\i\\\) \\<^sub>\ Vset \" proof- have "\\<^sub>\ (\i\\<^sub>\\\Arr\. \\CId\\\\ObjMap\\\\Cod\\i\\\) \\<^sub>\ \\<^sub>\ (\\ArrMap\)" proof(rule vrange_VLambda_vsubset) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto then have "\\CId\\b\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" by (auto intro: cat_cs_intros simp: cat_cs_simps) moreover from f have "\\CId\\\\ObjMap\\\\Cod\\f\\\ = \\ArrMap\\\\CId\\b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show "\\CId\\\\ObjMap\\\\Cod\\f\\\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" by (simp add: \.ArrMap.vsv_vimageI2) qed moreover have "\\<^sub>\ (\\ArrMap\) \\<^sub>\ Vset \" by (force intro: vrange_in_VsetI \.tm_cf_ArrMap_in_Vset) ultimately show ?thesis by auto qed qed (auto simp: cat_small_cs_intros) qed (auto intro: cat_cs_intros) from assms(4)[where A=\?L\, OF this, simplified] obtain P\<^sub>A \\<^sub>A where \\<^sub>A: "\\<^sub>A : P\<^sub>A <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ ?L : \\Arr\ \\\<^sub>C\<^bsub>\\<^esub> \" by auto interpret \\<^sub>A: is_cat_obj_prod \ \\\Arr\\ ?L \ P\<^sub>A \\<^sub>A by (rule \\<^sub>A) let ?F = \\u. \\ObjMap\\\\Cod\\u\\\ and ?f = \\u. \\<^sub>O\NTMap\\\\Cod\\u\\\ let ?\\<^sub>O' = \ntcf_obj_prod_base \ (:\<^sub>C (\\Arr\)\Obj\) ?F P\<^sub>O ?f\ have \\<^sub>O': "?\\<^sub>O' : P\<^sub>O <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: (\\Arr\) (\u. \\ObjMap\\\\Cod\\u\\) \ : :\<^sub>C (\\Arr\) \\\<^sub>C\<^bsub>\\<^esub> \" unfolding the_cat_discrete_components(1) proof ( intro tm_cf_discrete.tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone tm_cf_discrete ) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then show "\\<^sub>O\NTMap\\\\Cod\\f\\ : P\<^sub>O \\<^bsub>\\<^esub> \\ObjMap\\\\Cod\\f\\" by ( cs_concl cs_simp: the_cat_discrete_components(1) cat_discrete_cs_simps cat_cs_simps cs_intro: cat_cs_intros ) qed (intro P\<^sub>O) from \\<^sub>A.cat_obj_prod_unique_cone'[OF \\<^sub>O'] obtain f' where f': "f' : P\<^sub>O \\<^bsub>\\<^esub> P\<^sub>A" and \\<^sub>O'_NTMap_app: "\j. j \\<^sub>\ \\Arr\ \ ?\\<^sub>O'\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'" and unique_f': "\ f'' : P\<^sub>O \\<^bsub>\\<^esub> P\<^sub>A; \j. j \\<^sub>\ \\Arr\ \ ?\\<^sub>O'\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> f'' \ \ f'' = f'" for f'' by metis have \\<^sub>O_NTMap_app_Cod: "\\<^sub>O\NTMap\\b\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> f'" if "f : a \\<^bsub>\\<^esub> b" for f a b proof- from that have "f \\<^sub>\ \\Arr\" by auto from \\<^sub>O'_NTMap_app[OF this] that show ?thesis by ( cs_prems cs_simp: cat_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros ) qed from this[symmetric] have \\<^sub>A_NTMap_Comp_app: "\\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (f' \\<^sub>A\<^bsub>\\<^esub> q) = \\<^sub>O\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> q" if "f : a \\<^bsub>\\<^esub> b" and "q : c \\<^bsub>\\<^esub> P\<^sub>O" for q f a b c using that f' by (intro \.HomCod.cat_assoc_helper) ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros )+ let ?g = \\u. \\ArrMap\\u\ \\<^sub>A\<^bsub>\\<^esub> \\<^sub>O\NTMap\\\\Dom\\u\\\ let ?\\<^sub>O'' = \ntcf_obj_prod_base \ (:\<^sub>C (\\Arr\)\Obj\) ?F P\<^sub>O ?g\ have \\<^sub>O'': "?\\<^sub>O'' : P\<^sub>O <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: (\\Arr\) ?L \ : :\<^sub>C (\\Arr\) \\\<^sub>C\<^bsub>\\<^esub> \" unfolding the_cat_discrete_components(1) proof ( intro tm_cf_discrete.tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone tm_cf_discrete ) show "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\<^sub>O\NTMap\\\\Dom\\f\\ : P\<^sub>O \\<^bsub>\\<^esub> \\ObjMap\\\\Cod\\f\\" if "f \\<^sub>\ \\Arr\" for f proof- from that obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros ) qed qed (intro P\<^sub>O) from \\<^sub>A.cat_obj_prod_unique_cone'[OF \\<^sub>O''] obtain g' where g': "g' : P\<^sub>O \\<^bsub>\\<^esub> P\<^sub>A" and \\<^sub>O''_NTMap_app: "\j. j \\<^sub>\ \\Arr\ \ ?\\<^sub>O''\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> g'" and unique_g': "\ g'' : P\<^sub>O \\<^bsub>\\<^esub> P\<^sub>A; \j. j \\<^sub>\ \\Arr\ \ ?\\<^sub>O''\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> g'' \ \ g'' = g'" for g'' by (metis (lifting)) have "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\<^sub>O\NTMap\\a\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> g'" if "f : a \\<^bsub>\\<^esub> b" for f a b proof- from that have "f \\<^sub>\ \\Arr\" by auto from \\<^sub>O''_NTMap_app[OF this] that show ?thesis by ( cs_prems cs_simp: cat_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros ) qed then have \\<^sub>O_NTMap_app_Dom: "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\<^sub>O\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> q) = (\\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> g') \\<^sub>A\<^bsub>\\<^esub> q" if "f : a \\<^bsub>\\<^esub> b" and "q : c \\<^bsub>\\<^esub> P\<^sub>O" for q f a b c using that g' by (intro \.HomCod.cat_assoc_helper) ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros ) from assms(2)[OF f' g'] obtain E \ where \: "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (P\<^sub>O,P\<^sub>A,g',f') : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \: is_cat_equalizer \ P\<^sub>O P\<^sub>A g' f' \ E \ by (rule \) define \ where "\ = [(\i\\<^sub>\\\Obj\. \\<^sub>O\NTMap\\i\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\), cf_const \ \ E, \, \, \]\<^sub>\" have \_components: "\\NTMap\ = (\i\\<^sub>\\\Obj\. \\<^sub>O\NTMap\\i\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\)" "\\NTDom\ = cf_const \ \ E" "\\NTCod\ = \" "\\NTDGDom\ = \" "\\NTDGCod\ = \" unfolding \_def nt_field_simps by (simp_all add: nat_omega_simps) have [cat_cs_simps]: "\\NTMap\\i\ = \\<^sub>O\NTMap\\i\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\<^sub>P\<^sub>L\" if "i \\<^sub>\ \\Obj\" for i using that unfolding \_components by simp have "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - proof(intro is_cat_limitI') + proof(intro is_cat_limitI) show \: "\ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') show "vfsequence \" unfolding \_def by simp show "vcard \ = 5\<^sub>\" unfolding \_def by (simp add: nat_omega_simps) show "cf_const \ \ E : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_lim_cs_intros) show "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "\\NTMap\\a\ : cf_const \ \ E\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using that by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps cat_parallel_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros cat_lim_cs_intros cat_parallel_cs_intros ) show "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> cf_const \ \ E\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" if "f : a \\<^bsub>\\<^esub> b" for a b f using that \ g' f' by ( cs_concl cs_simp: cat_parallel_cs_simps cat_cs_simps the_cat_discrete_components(1) \\<^sub>O_NTMap_app_Cod \\<^sub>O_NTMap_app_Dom \.cat_eq_Comp_eq(1) cs_intro: cat_lim_cs_intros cat_cs_intros cat_parallel_cs_intros ) - show "cf_const \ \ E : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by - ( - cs_concl cs_simp: cs_intro: - cat_lim_cs_intros cat_cs_intros cat_small_cs_intros - ) - show "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by (cs_concl cs_simp: cs_intro: cat_small_cs_intros) qed (auto simp: \_components cat_lim_cs_intros) interpret \: is_cat_cone \ E \ \ \ \ by (rule \) show "\!f'. f' : r' \\<^bsub>\\<^esub> E \ u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" if "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" for u' r' proof- interpret u': is_cat_cone \ r' \ \ \ u' by (rule that) let ?u' = \\j. u'\NTMap\\j\\ let ?\' = \ntcf_obj_prod_base \ (\\Obj\) ?R r' ?u'\ have \'_NTMap_app: "?\'\NTMap\\j\ = u'\NTMap\\j\" if "j \\<^sub>\ \\Obj\" for j using that unfolding ntcf_obj_prod_base_components the_cat_discrete_components by auto have \': "?\' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: (\\Obj\) ?R \ : :\<^sub>C (\\Obj\) \\\<^sub>C\<^bsub>\\<^esub> \" unfolding the_cat_discrete_components(1) proof(intro tm_cf_discrete.tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone) show "tm_cf_discrete \ (\\Obj\) ?R \" proof(intro tm_cf_discreteI) show "\\ObjMap\\i\ \\<^sub>\ \\Obj\" if "i \\<^sub>\ \\Obj\" for i - using that - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + by (cs_concl cs_simp: cat_cs_simps cs_intro: that cat_cs_intros) + show "category \ \" by (auto intro: cat_cs_intros) + from \.tm_cf_ObjMap_in_Vset show "(\x\\<^sub>\\\Obj\. \\ObjMap\\x\) \\<^sub>\ Vset \" + by (auto simp: \.cf_ObjMap_vdomain) + show "(\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ Vset \" + proof(rule vbrelation.vbrelation_Limit_in_VsetI) + have "\\<^sub>\ (\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ \\<^sub>\ (\\ArrMap\)" + proof(intro vsubsetI) + fix x assume "x \\<^sub>\ \\<^sub>\ (\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\)" + then obtain i where i: "i \\<^sub>\ \\Obj\" + and x_def: "x = \\CId\\\\ObjMap\\i\\" + by auto + from i have "x = \\ArrMap\\\\CId\\i\\" + by (simp add: x_def \.cf_ObjMap_CId) + moreover from i have "\\CId\\i\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + ultimately show "x \\<^sub>\ \\<^sub>\ (\\ArrMap\)" + by (auto intro: \.ArrMap.vsv_vimageI2) + qed + then show "\\<^sub>\ (\i\\<^sub>\\\Obj\. \\CId\\\\ObjMap\\i\\) \\<^sub>\ Vset \" + by + ( + auto simp: + \.tm_cf_ArrMap_in_Vset vrange_in_VsetI vsubset_in_VsetI + ) + qed (auto intro: \.HomDom.tiny_cat_Obj_in_Vset) qed - ( - auto intro: - cat_cs_intros - P\<^sub>O - \\<^sub>O.NTCod.tm_cf_ArrMap_in_Vset[unfolded the_cf_discrete_components] - \\<^sub>O.NTCod.tm_cf_ObjMap_in_Vset[unfolded the_cf_discrete_components] - ) show "u'\NTMap\\j\ : r' \\<^bsub>\\<^esub> \\ObjMap\\j\" if "j \\<^sub>\ \\Obj\" for j using that by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: cat_lim_cs_intros) from \\<^sub>O.cat_obj_prod_unique_cone'[OF this] obtain h' where h': "h' : r' \\<^bsub>\\<^esub> P\<^sub>O" and \'_NTMap_app': "\j. j \\<^sub>\ (\\Obj\) \ ?\'\NTMap\\j\ = \\<^sub>O\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> h'" and unique_h': "\h''. \ h'' : r' \\<^bsub>\\<^esub> P\<^sub>O; \j. j \\<^sub>\ (\\Obj\) \ ?\'\NTMap\\j\ = \\<^sub>O\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> h'' \ \ h'' = h'" by metis interpret \': is_cat_cone \ r' \:\<^sub>C (\\Obj\)\ \ \:\: (\\Obj\) (app (\\ObjMap\)) \\ ?\' by (rule \') let ?u'' = \\u. u'\NTMap\\\\Cod\\u\\\ let ?\'' = \ntcf_obj_prod_base \ (\\Arr\) ?L r' ?u''\ have \''_NTMap_app: "?\''\NTMap\\f\ = u'\NTMap\\b\" if "f : a \\<^bsub>\\<^esub> b" for f a b using that unfolding ntcf_obj_prod_base_components the_cat_discrete_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: cat_cs_intros) have \'': "?\'' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e :\: (\\Arr\) ?L \ : :\<^sub>C (\\Arr\) \\\<^sub>C\<^bsub>\\<^esub> \" unfolding the_cat_discrete_components(1) proof ( intro tm_cf_discrete.tm_cf_discrete_ntcf_obj_prod_base_is_cat_cone tm_cf_discrete ) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then show "u'\NTMap\\\\Cod\\f\\ : r' \\<^bsub>\\<^esub> \\ObjMap\\\\Cod\\f\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (simp add: cat_lim_cs_intros) from \\<^sub>A.cat_obj_prod_unique_cone'[OF this] obtain h'' where h'': "h'' : r' \\<^bsub>\\<^esub> P\<^sub>A" and \''_NTMap_app': "\j. j \\<^sub>\ \\Arr\ \ ?\''\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> h''" and unique_h'': "\h'''. \ h''' : r' \\<^bsub>\\<^esub> P\<^sub>A; \j. j \\<^sub>\ \\Arr\ \ ?\''\NTMap\\j\ = \\<^sub>A\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> h''' \ \ h''' = h''" by metis interpret \'': is_cat_cone \ r' \:\<^sub>C (\\Arr\)\ \ \:\: (\\Arr\) ?L \\ ?\'' by (rule \'') have g'h'_f'h': "g' \\<^sub>A\<^bsub>\\<^esub> h' = f' \\<^sub>A\<^bsub>\\<^esub> h'" proof- from g' h' have g'h': "g' \\<^sub>A\<^bsub>\\<^esub> h' : r' \\<^bsub>\\<^esub> P\<^sub>A" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f' h' have f'h': "f' \\<^sub>A\<^bsub>\\<^esub> h' : r' \\<^bsub>\\<^esub> P\<^sub>A" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have "?\''\NTMap\\f\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (g' \\<^sub>A\<^bsub>\\<^esub> h')" if "f \\<^sub>\ \\Arr\" for f proof- from that obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto then have "?\''\NTMap\\f\ = u'\NTMap\\b\" by (cs_concl cs_simp: \''_NTMap_app cat_cs_simps) also from f have "\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> ?\'\NTMap\\a\" by ( cs_concl cs_simp: \'_NTMap_app cat_lim_cs_simps cs_intro: cat_cs_intros ) also from f g' h' have "\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (g' \\<^sub>A\<^bsub>\\<^esub> h')" by ( cs_concl cs_simp: cat_cs_simps cat_discrete_cs_simps the_cat_discrete_components(1) \'_NTMap_app' \\<^sub>O_NTMap_app_Dom cs_intro: cat_cs_intros ) finally show ?thesis by simp qed from unique_h''[OF g'h' this, simplified] have g'h'_h'': "g' \\<^sub>A\<^bsub>\\<^esub> h' = h''". have "?\''\NTMap\\f\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (f' \\<^sub>A\<^bsub>\\<^esub> h')" if "f \\<^sub>\ \\Arr\" for f proof- from that obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto then have "?\''\NTMap\\f\ = u'\NTMap\\b\" by (cs_concl cs_simp: \''_NTMap_app cat_cs_simps) also from f have "\ = ?\'\NTMap\\b\" by (cs_concl cs_simp: \'_NTMap_app cs_intro: cat_cs_intros) also from f have "\ = \\<^sub>O\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> h'" by (cs_concl cs_simp: \'_NTMap_app' cs_intro: cat_cs_intros) also from f g' h' have "\ = (\\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> f') \\<^sub>A\<^bsub>\\<^esub> h'" by (cs_concl cs_simp: \\<^sub>O_NTMap_app_Cod cs_intro: cat_cs_intros) also from that f' h' have "\ = \\<^sub>A\NTMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (f' \\<^sub>A\<^bsub>\\<^esub> h')" by ( cs_concl cs_simp: cat_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros ) finally show ?thesis by simp qed from unique_h''[OF f'h' this, simplified] have f'h'_h'': "f' \\<^sub>A\<^bsub>\\<^esub> h' = h''". from g'h'_h'' f'h'_h'' show ?thesis by simp qed let ?II = \\\\<^sub>C \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L\ and ?II_II = \\\\\\ \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L P\<^sub>O P\<^sub>A g' f'\ define \' where "\' = [ (\f\\<^sub>\?II\Obj\. (f = \\<^sub>P\<^sub>L ? h' : (f' \\<^sub>A\<^bsub>\\<^esub> h'))), cf_const ?II \ r', ?II_II, ?II, \ ]\<^sub>\" have \'_components: "\'\NTMap\ = (\f\\<^sub>\?II\Obj\. (f = \\<^sub>P\<^sub>L ? h' : (f' \\<^sub>A\<^bsub>\\<^esub> h')))" "\'\NTDom\ = cf_const ?II \ r'" "\'\NTCod\ = ?II_II" "\'\NTDGDom\ = ?II" "\'\NTDGCod\ = \" unfolding \'_def nt_field_simps by (simp_all add: nat_omega_simps) have \'_NTMap_app_I2: "\'\NTMap\\x\ = h'" if "x = \\<^sub>P\<^sub>L" for x proof- have "x \\<^sub>\ ?II\Obj\" unfolding that by (cs_concl cs_intro: cat_parallel_cs_intros) then show ?thesis unfolding \'_components that by simp qed have \'_NTMap_app_sI2: "\'\NTMap\\x\ = f' \\<^sub>A\<^bsub>\\<^esub> h'" if "x = \\<^sub>P\<^sub>L" for x proof- have "x \\<^sub>\ ?II\Obj\" unfolding that by (cs_concl cs_intro: cat_parallel_cs_intros) with \.cat_parallel_\\ show ?thesis unfolding \'_components by (cs_concl cs_simp: V_cs_simps that) qed interpret par: cf_parallel \ \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L \\<^sub>P\<^sub>L P\<^sub>O P\<^sub>A g' f' \ by (intro cf_parallelI cat_parallelI) ( simp_all add: cat_cs_intros cat_parallel_cs_intros cat_PL_ineq[symmetric] ) have "\' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') show "vfsequence \'" unfolding \'_def by auto show "vcard \' = 5\<^sub>\" unfolding \'_def by (simp add: nat_omega_simps) from h' show "cf_const (?II) \ r' : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "?II_II : ?II \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_parallel_cs_simps cs_intro: cat_cs_intros) from h' show "\'\NTMap\\a\ : cf_const ?II \ r'\ObjMap\\a\ \\<^bsub>\\<^esub> ?II_II\ObjMap\\a\" if "a \\<^sub>\ ?II\Obj\" for a using that by (elim the_cat_parallel_ObjE; simp only:) ( cs_concl cs_simp: \'_NTMap_app_I2 \'_NTMap_app_sI2 cat_cs_simps cat_parallel_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) from h' f' g'h'_f'h' show "\'\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> cf_const ?II \ r'\ArrMap\\f\ = ?II_II\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \'\NTMap\\a\" if "f : a \\<^bsub>?II\<^esub> b" for a b f using that by (elim \.the_cat_parallel_is_arrE; simp only:) ( cs_concl cs_intro: cat_cs_intros cat_parallel_cs_intros cs_simp: cat_cs_simps cat_parallel_cs_simps \'_NTMap_app_I2 \'_NTMap_app_sI2 )+ qed ( simp add: \'_components | cs_concl cs_simp: cat_cs_simps cs_intro: cat_lim_cs_intros cat_cs_intros cat_small_cs_intros )+ from \.cat_eq_unique_cone[OF this] obtain t' where t': "t' : r' \\<^bsub>\\<^esub> E" and \'_NTMap_app: "\'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> t'" and unique_t': "\ t'' : r' \\<^bsub>\\<^esub> E; \'\NTMap\\\\<^sub>P\<^sub>L\ = \\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> t''\ \ t'' = t'" for t'' by metis show "\!f'. f' : r' \\<^bsub>\\<^esub> E \ u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" proof(intro ex1I conjI; (elim conjE)?, (rule t')?) show [symmetric, cat_cs_simps]: "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t'" proof(rule ntcf_eqI[OF u'.is_ntcf_axioms]) from t' show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t' : cf_const \ \ r' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t')\NTMap\" proof(rule vsv_eqI) show "vsv ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t')\NTMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from t' show "\\<^sub>\ (u'\NTMap\) = \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t')\NTMap\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "u'\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t')\NTMap\\a\" if "a \\<^sub>\ \\<^sub>\ (u'\NTMap\)" for a proof- from that have "a \\<^sub>\ \\Obj\" by (cs_prems cs_simp: cat_cs_simps) with t' show "u'\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t')\NTMap\\a\" by ( cs_concl cs_simp: cat_cs_simps \'_NTMap_app cat_parallel_cs_simps the_cat_discrete_components(1) \'_NTMap_app[symmetric] \'_NTMap_app_I2 \'_NTMap_app'[symmetric] cs_intro: cat_cs_intros cat_parallel_cs_intros ) qed qed auto qed simp_all fix t'' assume prems': "t'' : r' \\<^bsub>\\<^esub> E" "u' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t''" then have u'_NTMap_app_x: "u'\NTMap\\x\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ t'')\NTMap\\x\" for x by simp have "?\'\NTMap\\j\ = \\<^sub>O\NTMap\\j\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> t'')" if "j \\<^sub>\ \\Obj\" for j using u'_NTMap_app_x[of j] prems'(1) that by ( cs_prems cs_simp: cat_cs_simps cat_discrete_cs_simps cat_parallel_cs_simps the_cat_discrete_components(1) cs_intro: cat_cs_intros cat_parallel_cs_intros ) (simp add: \'_NTMap_app[OF that, symmetric]) moreover from prems'(1) have "\\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> t'' : r' \\<^bsub>\\<^esub> P\<^sub>O" by ( cs_concl cs_simp: cat_cs_simps cat_parallel_cs_simps cs_intro: cat_cs_intros cat_parallel_cs_intros ) ultimately have [cat_cs_simps]: "\\NTMap\\\\<^sub>P\<^sub>L\ \\<^sub>A\<^bsub>\\<^esub> t'' = h'" by (intro unique_h') simp show "t'' = t'" by (rule unique_t', intro prems'(1)) (cs_concl cs_simp: \'_NTMap_app_I2 cat_cs_simps) qed qed qed then show ?thesis using that by clarsimp qed lemma cat_colimit_of_cat_prod_obj_and_cat_coequalizer: \\See Theorem 1 in Chapter V-2 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\\ \ \ \. \ \ : \ \\<^bsub>\\<^esub> \; \ : \ \\<^bsub>\\<^esub> \ \ \ \E \. \ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" and "\A. tm_cf_discrete \ (\\Obj\) A \ \ \P \. \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : \\Obj\ \\\<^sub>C\<^bsub>\\<^esub> \" and "\A. tm_cf_discrete \ (\\Arr\) A \ \ \P \. \ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : \\Arr\ \\\<^sub>C\<^bsub>\\<^esub> \" obtains r u where "u : \ >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) have "\E \. \ : E <\<^sub>C\<^sub>F\<^sub>.\<^sub>e\<^sub>q (\,\,\,\) : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> op_cat \" if "\ : \ \\<^bsub>\\<^esub> \" "\ : \ \\<^bsub>\\<^esub> \" for \ \ \ \ proof- from assms(2)[OF that(1,2)] obtain E \ where \: "\ : (\,\,\,\) >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>e\<^sub>q E : \\\<^sup>2\<^sub>C \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \: is_cat_coequalizer \ \ \ \ \ \ E \ by (rule \) from \.is_cat_equalizer_op[unfolded cat_op_simps] show ?thesis by auto qed moreover have "\P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : \\Obj\ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" if "tm_cf_discrete \ (\\Obj\) A (op_cat \)" for A proof- interpret tm_cf_discrete \ \\\Obj\\ A \op_cat \\ by (rule that) from assms(3)[OF tm_cf_discrete_op[unfolded cat_op_simps]] obtain P \ where \: "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : \\Obj\ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \: is_cat_obj_coprod \ \\\Obj\\ A \ P \ by (rule \) from \.is_cat_obj_prod_op show ?thesis by auto qed moreover have "\P \. \ : P <\<^sub>C\<^sub>F\<^sub>.\<^sub>\ A : \\Arr\ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" if "tm_cf_discrete \ (\\Arr\) A (op_cat \)" for A proof- interpret tm_cf_discrete \ \\\Arr\\ A \op_cat \\ by (rule that) from assms(4)[OF tm_cf_discrete_op[unfolded cat_op_simps]] obtain P \ where \: "\ : A >\<^sub>C\<^sub>F\<^sub>.\<^sub>\ P : \\Arr\ \\\<^sub>C\<^bsub>\\<^esub> \" by clarsimp interpret \: is_cat_obj_coprod \ \\\Arr\\ A \ P \ by (rule \) from \.is_cat_obj_prod_op show ?thesis by auto qed ultimately obtain u r where u: "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by ( rule cat_limit_of_cat_prod_obj_and_cat_equalizer[ OF \.is_tm_functor_op, unfolded cat_op_simps ] ) interpret u: is_cat_limit \ \op_cat \\ \op_cat \\ \op_cf \\ r u by (rule u) from u.is_cat_colimit_op[unfolded cat_op_simps] that show ?thesis by simp qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan.thy @@ -1,3574 +1,3452 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Pointwise Kan extensions\ theory CZH_UCAT_PWKan imports CZH_UCAT_Kan begin subsection\Pointwise Kan extensions\ text\ The following subsection is based on elements of the content of section 6.3 in \cite{riehl_category_2016} and Chapter X-5 in \cite{mac_lane_categories_2010}. \ locale is_cat_pw_rKe = is_cat_rKe \ \ \ \ \ \ \ \ for \ \ \ \ \ \ \ \ + assumes cat_pw_rKe_preserved: "a \\<^sub>\ \\Obj\ \ \ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C cat_Set \)" syntax "_is_cat_pw_rKe" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" ( \(_ :/ _ \\<^sub>C\<^sub>F _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\ _ :/ _ \\<^sub>C _ \\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51] 51 ) translations "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" \ "CONST is_cat_pw_rKe \ \ \ \ \ \ \ \" locale is_cat_pw_lKe = is_cat_lKe \ \ \ \ \ \ \ \ for \ \ \ \ \ \ \ \ + assumes cat_pw_lKe_preserved: "a \\<^sub>\ op_cat \\Obj\ \ op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) : op_cat \ \\\<^sub>C cat_Set \)" syntax "_is_cat_pw_lKe" :: "V \ V \ V \ V \ V \ V \ V \ V \ bool" ( \(_ :/ _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\ _ \\<^sub>C\<^sub>F _ :/ _ \\<^sub>C _ \\<^sub>C _)\ [51, 51, 51, 51, 51, 51, 51] 51 ) translations "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \" \ "CONST is_cat_pw_lKe \ \ \ \ \ \ \ \" lemma (in is_cat_pw_rKe) cat_pw_rKe_preserved'[cat_Kan_cs_intros]: assumes "a \\<^sub>\ \\Obj\" and "\' = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" and "\' = cat_Set \" shows "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C (\' : \' \\\<^sub>C \')" using assms(1) unfolding assms(2-4) by (rule cat_pw_rKe_preserved) lemmas [cat_Kan_cs_intros] = is_cat_pw_rKe.cat_pw_rKe_preserved' lemma (in is_cat_pw_lKe) cat_pw_lKe_preserved'[cat_Kan_cs_intros]: assumes "a \\<^sub>\ op_cat \\Obj\" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a)" and "\' = cat_Set \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \' : \' \\<^sub>C \' \\<^sub>C (\' : \' \\\<^sub>C \')" using assms(1) unfolding assms by (rule cat_pw_lKe_preserved) lemmas [cat_Kan_cs_intros] = is_cat_pw_lKe.cat_pw_lKe_preserved' text\Rules.\ lemma (in is_cat_pw_rKe) is_cat_pw_rKe_axioms'[cat_Kan_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\'\<^esub> \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_pw_rKe_axioms) mk_ide rf is_cat_pw_rKe_def[unfolded is_cat_pw_rKe_axioms_def] |intro is_cat_pw_rKeI| |dest is_cat_pw_rKeD[dest]| |elim is_cat_pw_rKeE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_pw_rKeD(1) lemma (in is_cat_pw_lKe) is_cat_pw_lKe_axioms'[cat_Kan_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\'\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_pw_lKe_axioms) mk_ide rf is_cat_pw_lKe_def[unfolded is_cat_pw_lKe_axioms_def] |intro is_cat_pw_lKeI| |dest is_cat_pw_lKeD[dest]| |elim is_cat_pw_lKeE[elim]| lemmas [cat_Kan_cs_intros] = is_cat_pw_lKeD(1) text\Duality.\ lemma (in is_cat_pw_rKe) is_cat_pw_lKe_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \" proof(intro is_cat_pw_lKeI, unfold cat_op_simps) fix a assume prems: "a \\<^sub>\ \\Obj\" from cat_pw_rKe_preserved[OF prems] prems show "\ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat \(-,a) : \ \\\<^sub>C cat_Set \)" by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros) qed (cs_concl cs_intro: cat_op_intros) lemma (in is_cat_pw_rKe) is_cat_pw_lKe_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> \' \\<^sub>C\<^sub>F \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_pw_lKe_op) lemmas [cat_op_intros] = is_cat_pw_rKe.is_cat_pw_lKe_op' lemma (in is_cat_pw_lKe) is_cat_pw_rKe_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C op_cat \" proof(intro is_cat_pw_rKeI, unfold cat_op_simps) fix a assume prems: "a \\<^sub>\ \\Obj\" from cat_pw_lKe_preserved[unfolded cat_op_simps, OF prems] prems show "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat \ \\<^sub>C op_cat \ \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat \(a,-) : op_cat \ \\\<^sub>C cat_Set \)" by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros) qed (cs_concl cs_intro: cat_op_intros) lemma (in is_cat_pw_lKe) is_cat_pw_lKe_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> \' : \' \\<^sub>C \' \\<^sub>C \'" unfolding assms by (rule is_cat_pw_rKe_op) lemmas [cat_op_intros] = is_cat_pw_lKe.is_cat_pw_lKe_op' (*FIXME: any reason not to generalize and include in CZH_UCAT_Hom?*) subsection\Cone functor\ subsubsection\Definition and elementary properties\ definition cf_Cone :: "V \ V \ V \ V" where "cf_Cone \ \ \ = - Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_Funct \ (\\HomDom\) (\\HomCod\)(-,cf_map \) \\<^sub>C\<^sub>F - op_cf (\\<^sub>C \ (\\HomDom\) (\\HomCod\))" + Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (\\HomCod\)(-,cf_map \) \\<^sub>C\<^sub>F + op_cf (\\<^sub>C\<^sub>F \ (\\HomDom\) (\\HomCod\))" text\An alternative form of the definition.\ context is_functor begin lemma cf_Cone_def': - "cf_Cone \ \ \ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_Funct \ \ \(-,cf_map \) \\<^sub>C\<^sub>F op_cf (\\<^sub>C \ \ \)" + "cf_Cone \ \ \ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ \(-,cf_map \) \\<^sub>C\<^sub>F op_cf (\\<^sub>C\<^sub>F \ \ \)" unfolding cf_Cone_def cat_cs_simps by simp end subsubsection\Object map\ -lemma (in is_tm_functor) cf_Cone_ObjMap_vsv[cat_Kan_cs_intros]: +lemma (in is_functor) cf_Cone_ObjMap_vsv[cat_Kan_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" shows "vsv (cf_Cone \ \ \\ObjMap\)" proof- from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ - from assms(2) show ?thesis - unfolding cf_Cone_def - by - ( - cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros - ) -qed - -lemmas [cat_Kan_cs_intros] = is_tm_functor.cf_Cone_ObjMap_vsv - -lemma (in is_tm_functor) cf_Cone_ObjMap_vdomain[cat_Kan_cs_simps]: - assumes "\ \" and "\ \\<^sub>\ \" and "b \\<^sub>\ \\Obj\" - shows "\\<^sub>\ (cf_Cone \ \ \\ObjMap\) = \\Obj\" -proof- - from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ - from assms(2) show ?thesis - unfolding cf_Cone_def' - by - ( - cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros - ) -qed - -lemmas [cat_Kan_cs_simps] = is_tm_functor.cf_Cone_ObjMap_vdomain - -lemma (in is_tm_functor) cf_Cone_ObjMap_app[cat_Kan_cs_simps]: - assumes "\ \" - and "\ \\<^sub>\ \" - and "b \\<^sub>\ \\Obj\" - shows "cf_Cone \ \ \\ObjMap\\b\ = - Hom (cat_Funct \ \ \) (cf_map (cf_const \ \ b)) (cf_map \)" -proof- - from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ - from assms(2,3) show ?thesis - unfolding cf_Cone_def - by - ( - cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros - ) -qed - -lemmas [cat_Kan_cs_simps] = is_tm_functor.cf_Cone_ObjMap_app - - -subsubsection\Arrow map\ - -lemma (in is_tm_functor) cf_Cone_ArrMap_vsv[cat_Kan_cs_intros]: - assumes "\ \" and "\ \\<^sub>\ \" - shows "vsv (cf_Cone \ \ \\ArrMap\)" -proof- - from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from assms(2) show ?thesis unfolding cf_Cone_def by ( cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed -lemmas [cat_Kan_cs_intros] = is_tm_functor.cf_Cone_ArrMap_vsv - -lemma (in is_tm_functor) cf_Cone_ArrMap_vdomain[cat_Kan_cs_simps]: +lemmas [cat_Kan_cs_intros] = is_functor.cf_Cone_ObjMap_vsv + +lemma (in is_functor) cf_Cone_ObjMap_vdomain[cat_Kan_cs_simps]: + assumes "\ \" and "\ \\<^sub>\ \" and "b \\<^sub>\ \\Obj\" + shows "\\<^sub>\ (cf_Cone \ \ \\ObjMap\) = \\Obj\" +proof- + from assms interpret \: \ \ by simp + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ + from assms show ?thesis + unfolding cf_Cone_def' + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + ) +qed + +lemmas [cat_Kan_cs_simps] = is_functor.cf_Cone_ObjMap_vdomain + +lemma (in is_functor) cf_Cone_ObjMap_app[cat_Kan_cs_simps]: + assumes "\ \" and "\ \\<^sub>\ \" and "b \\<^sub>\ \\Obj\" + shows "cf_Cone \ \ \\ObjMap\\b\ = + Hom (cat_FUNCT \ \ \) (cf_map (cf_const \ \ b)) (cf_map \)" +proof- + from assms interpret \: \ \ by simp + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ + from assms(2,3) show ?thesis + unfolding cf_Cone_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + ) +qed + +lemmas [cat_Kan_cs_simps] = is_functor.cf_Cone_ObjMap_app + + +subsubsection\Arrow map\ + +lemma (in is_functor) cf_Cone_ArrMap_vsv[cat_Kan_cs_intros]: + assumes "\ \" and "\ \\<^sub>\ \" + shows "vsv (cf_Cone \ \ \\ArrMap\)" +proof- + from assms interpret \: \ \ by simp + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ + from assms(2) show ?thesis + unfolding cf_Cone_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + ) +qed + +lemmas [cat_Kan_cs_intros] = is_functor.cf_Cone_ArrMap_vsv + +lemma (in is_functor) cf_Cone_ArrMap_vdomain[cat_Kan_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (cf_Cone \ \ \\ArrMap\) = \\Arr\" proof- from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from assms(2) show ?thesis unfolding cf_Cone_def' by ( cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed -lemmas [cat_Kan_cs_simps] = is_tm_functor.cf_Cone_ArrMap_vdomain - -lemma (in is_tm_functor) cf_Cone_ArrMap_app[cat_Kan_cs_simps]: +lemmas [cat_Kan_cs_simps] = is_functor.cf_Cone_ArrMap_vdomain + +lemma (in is_functor) cf_Cone_ArrMap_app[cat_Kan_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "f : a \\<^bsub>\\<^esub> b" shows "cf_Cone \ \ \\ArrMap\\f\ = cf_hom - (cat_Funct \ \ \) - [ntcf_arrow (ntcf_const \ \ f), cat_Funct \ \ \\CId\\cf_map \\]\<^sub>\" + (cat_FUNCT \ \ \) + [ntcf_arrow (ntcf_const \ \ f), cat_FUNCT \ \ \\CId\\cf_map \\]\<^sub>\" proof- from assms interpret \: \ \ by simp - from assms interpret \: is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ \cat_Funct \ \ \\ \\\<^sub>C \ \ \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ + from assms interpret \: is_functor \ \ \cat_FUNCT \ \ \\ \\\<^sub>C\<^sub>F \ \ \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from assms(2,3) show ?thesis unfolding cf_Cone_def by ( cs_concl - cs_simp: cat_cs_simps cat_Funct_components(1) cat_op_simps - cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros + cs_simp: cat_cs_simps cat_FUNCT_components(1) cat_op_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed -lemmas [cat_Kan_cs_simps] = is_tm_functor.cf_Cone_ArrMap_app +lemmas [cat_Kan_cs_simps] = is_functor.cf_Cone_ArrMap_app subsubsection\The cone functor is a functor\ -lemma (in is_tm_functor) tm_cf_cf_Cone_is_functor: - "cf_Cone \ \ \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" - unfolding cf_Cone_def' - by - ( - cs_concl - cs_simp: cat_op_simps cat_Funct_components(1) - cs_intro: - cat_small_cs_intros - cat_cs_intros - cat_FUNCT_cs_intros - cat_op_intros - ) - -lemma (in is_tm_functor) tm_cf_cf_Cone_is_functor_if_ge_Limit: +lemma (in is_functor) tm_cf_cf_Cone_is_functor_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "cf_Cone \ \ \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- - from assms interpret \\: category \ \cat_Funct \ \ \\ + from assms interpret \\: category \ \cat_FUNCT \ \ \\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) - interpret \_\\: category \ \cat_Funct \ \ \\ - by (rule \\.cat_category_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros assms)+ from assms interpret op_\: - is_tiny_functor \ \op_cat \\ \op_cat (cat_Funct \ \ \)\ \op_cf (\\<^sub>C \ \ \)\ - by (intro is_functor.cf_is_tiny_functor_if_ge_Limit) - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - have "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_Funct \ \ \(-,cf_map \) : - op_cat (cat_Funct \ \ \) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" + is_functor \ \op_cat \\ \op_cat (cat_FUNCT \ \ \)\ \op_cf (\\<^sub>C\<^sub>F \ \ \)\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ + have "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ \(-,cf_map \) : + op_cat (cat_FUNCT \ \ \) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl - cs_simp: cat_Funct_components(1) - cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros + cs_simp: cat_FUNCT_cs_simps + cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) then show "cf_Cone \ \ \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" unfolding cf_Cone_def' by (cs_concl cs_intro: cat_cs_intros) qed subsection\Lemma X.5: \L_10_5_N\\label{sec:lem_X_5_start}\ text\ This subsection and several further subsections (\ref{sec:lem_X_5_start}-\ref{sec:lem_X_5_end}) expose definitions that are used in the proof of the technical lemma that was used in the proof of Theorem 3 from Chapter X-5 in \cite{mac_lane_categories_2010}. \ definition L_10_5_N :: "V \ V \ V \ V \ V \ V" where "L_10_5_N \ \ \ \ c = [ ( \a\\<^sub>\\\HomCod\\Obj\. cf_nt \ \ \\ObjMap\\cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(a,-) \\<^sub>C\<^sub>F \), c\\<^sub>\ ), ( \f\\<^sub>\\\HomCod\\Arr\. cf_nt \ \ \\ArrMap\\ ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \), \\HomCod\\CId\\c\ \\<^sub>\ ), op_cat (\\HomCod\), cat_Set \ ]\<^sub>\" text\Components.\ lemma L_10_5_N_components: shows "L_10_5_N \ \ \ \ c\ObjMap\ = ( \a\\<^sub>\\\HomCod\\Obj\. cf_nt \ \ \\ObjMap\\cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(a,-) \\<^sub>C\<^sub>F \), c\\<^sub>\ )" and "L_10_5_N \ \ \ \ c\ArrMap\ = ( \f\\<^sub>\\\HomCod\\Arr\. cf_nt \ \ \\ArrMap\\ ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \), \\HomCod\\CId\\c\ \\<^sub>\ )" and "L_10_5_N \ \ \ \ c\HomDom\ = op_cat (\\HomCod\)" and "L_10_5_N \ \ \ \ c\HomCod\ = cat_Set \" unfolding L_10_5_N_def dghm_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) lemmas L_10_5_N_components' = L_10_5_N_components[ where \=\ and \=\, unfolded cat_cs_simps ] lemmas [cat_Kan_cs_simps] = L_10_5_N_components'(3,4) end subsubsection\Object map\ mk_VLambda L_10_5_N_components(1) |vsv L_10_5_N_ObjMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ c assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin mk_VLambda L_10_5_N_components'(1)[OF \ \] |vdomain L_10_5_N_ObjMap_vdomain[cat_Kan_cs_simps]| |app L_10_5_N_ObjMap_app[cat_Kan_cs_simps]| end subsubsection\Arrow map\ mk_VLambda L_10_5_N_components(2) |vsv L_10_5_N_ArrMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ c assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin mk_VLambda L_10_5_N_components'(2)[OF \ \] |vdomain L_10_5_N_ArrMap_vdomain[cat_Kan_cs_simps]| |app L_10_5_N_ArrMap_app[cat_Kan_cs_simps]| end subsubsection\\L_10_5_N\ is a functor\ lemma L_10_5_N_is_functor: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" shows "L_10_5_N \ \ \ \ c : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- let ?FUNCT = \\\. cat_FUNCT \ \ (cat_Set \)\ interpret \: \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(3)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(4)) + interpret \: is_functor \ \ \ \ by (rule assms(3)) + interpret \: is_functor \ \ \ \ by (rule assms(4)) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) interpret cf_nt: is_functor \ \?FUNCT \ \\<^sub>C \\ \cat_Set \\ \cf_nt \ \ \\ by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show ?thesis proof(intro is_functorI') show "vfsequence (L_10_5_N \ \ \ \ c)" unfolding L_10_5_N_def by simp show "vcard (L_10_5_N \ \ \ \ c) = 4\<^sub>\" unfolding L_10_5_N_def by (simp add: nat_omega_simps) show "vsv (L_10_5_N \ \ \ \ c\ObjMap\)" by (cs_concl cs_intro: cat_Kan_cs_intros) from assms(3,4) show "vsv (L_10_5_N \ \ \ \ c\ArrMap\)" by (cs_concl cs_intro: cat_Kan_cs_intros) from assms show "\\<^sub>\ (L_10_5_N \ \ \ \ c\ObjMap\) = op_cat \\Obj\" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_op_simps cs_intro: cat_cs_intros ) show "\\<^sub>\ (L_10_5_N \ \ \ \ c\ObjMap\) \\<^sub>\ cat_Set \\Obj\" unfolding L_10_5_N_components'[OF \.is_functor_axioms \.is_functor_axioms] proof(rule vrange_VLambda_vsubset) fix a assume prems: "a \\<^sub>\ \\Obj\" from prems assms show "cf_nt \ \ \\ObjMap\\cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \), c\\<^sub>\ \\<^sub>\ cat_Set \\Obj\" by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros FUNCT_\.cat_Hom_in_Vset cat_FUNCT_cs_intros ) qed from assms show "\\<^sub>\ (L_10_5_N \ \ \ \ c\ArrMap\) = op_cat \\Arr\" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_op_simps cs_intro: cat_cs_intros ) show "L_10_5_N \ \ \ \ c\ArrMap\\f\ : L_10_5_N \ \ \ \ c\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ObjMap\\b\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f using that assms unfolding cat_op_simps by ( cs_concl cs_simp: L_10_5_N_ArrMap_app L_10_5_N_ObjMap_app cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) show "L_10_5_N \ \ \ \ c\ArrMap\\g \\<^sub>A\<^bsub>op_cat \\<^esub> f\ = L_10_5_N \ \ \ \ c\ArrMap\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ArrMap\\f\" if "g : b' \\<^bsub>op_cat \\<^esub> c'" and "f : a' \\<^bsub>op_cat \\<^esub> b'" for b' c' g a' f proof- from that assms(5) show ?thesis unfolding cat_op_simps by (*slow*) ( cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_cs_simps cat_prod_cs_simps cat_op_simps cf_nt.cf_ArrMap_Comp[symmetric] ) qed show "L_10_5_N \ \ \ \ c\ArrMap\\op_cat \\CId\\a\\ = cat_Set \\CId\\L_10_5_N \ \ \ \ c\ObjMap\\a\\" if "a \\<^sub>\ op_cat \\Obj\" for a proof- note [cat_cs_simps] = ntcf_id_cf_comp[symmetric] ntcf_arrow_id_ntcf_id[symmetric] cat_FUNCT_CId_app[symmetric] from that[unfolded cat_op_simps] assms show ?thesis by (*slow*) ( cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros cs_simp: cat_FUNCT_cs_simps cat_cs_simps cat_Kan_cs_simps cat_op_simps ) qed qed (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros)+ qed lemma L_10_5_N_is_functor'[cat_Kan_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = op_cat \" and "\' = cat_Set \" and "\' = \" shows "L_10_5_N \ \ \ \ c : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1-5) unfolding assms(6-8) by (rule L_10_5_N_is_functor) subsection\Lemma X.5: \L_10_5_\_arrow\\ subsubsection\Definition and elementary properties\ definition L_10_5_\_arrow :: "V \ V \ V \ V \ V \ V \ V" where "L_10_5_\_arrow \ \ c \ a b = [ (\f\\<^sub>\Hom (\\HomCod\) c (\\ObjMap\\b\). \\NTMap\\0, b, f\\<^sub>\), Hom (\\HomCod\) c (\\ObjMap\\b\), Hom (\\HomCod\) a (\\ObjMap\\b\) ]\<^sub>\" text\Components.\ lemma L_10_5_\_arrow_components: shows "L_10_5_\_arrow \ \ c \ a b\ArrVal\ = (\f\\<^sub>\Hom (\\HomCod\) c (\\ObjMap\\b\). \\NTMap\\0, b, f\\<^sub>\)" and "L_10_5_\_arrow \ \ c \ a b\ArrDom\ = Hom (\\HomCod\) c (\\ObjMap\\b\)" and "L_10_5_\_arrow \ \ c \ a b\ArrCod\ = Hom (\\HomCod\) a (\\ObjMap\\b\)" unfolding L_10_5_\_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) lemmas L_10_5_\_arrow_components' = L_10_5_\_arrow_components[ where \=\ and \=\, unfolded cat_cs_simps ] lemmas [cat_Kan_cs_simps] = L_10_5_\_arrow_components'(2,3) end subsubsection\Arrow value\ mk_VLambda L_10_5_\_arrow_components(1) |vsv L_10_5_\_arrow_ArrVal_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin mk_VLambda L_10_5_\_arrow_components'(1)[OF \ \] |vdomain L_10_5_\_arrow_ArrVal_vdomain[cat_Kan_cs_simps]| |app L_10_5_\_arrow_ArrVal_app[unfolded in_Hom_iff]| end lemma L_10_5_\_arrow_ArrVal_app': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" shows "L_10_5_\_arrow \ \ c \ a b\ArrVal\\f\ = \\NTMap\\0, b, f\\<^sub>\" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms(3) have c: "c \\<^sub>\ \\Obj\" by auto show ?thesis by (rule L_10_5_\_arrow_ArrVal_app[OF assms(1,2,3)]) qed subsubsection\\L_10_5_\_arrow\ is an arrow\ lemma L_10_5_\_arrow_ArrVal_is_arr: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = ntcf_arrow \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" and "b \\<^sub>\ \\Obj\" shows "L_10_5_\_arrow \ \ c \' a b\ArrVal\\f\ : a \\<^bsub>\\<^esub> \\ObjMap\\b\" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(4)) from assms(5,6) show ?thesis unfolding assms(3) by ( cs_concl cs_simp: cat_cs_simps L_10_5_\_arrow_ArrVal_app cat_comma_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) qed lemma L_10_5_\_arrow_ArrVal_is_arr'[cat_Kan_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = ntcf_arrow \" and "a' = a" and "b' = \\ObjMap\\b\" and "\' = \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" and "b \\<^sub>\ \\Obj\" shows "L_10_5_\_arrow \ \ c \' a b\ArrVal\\f\ : a' \\<^bsub>\\<^esub> b'" using assms(1-3, 7-9) unfolding assms(3-6) by (rule L_10_5_\_arrow_ArrVal_is_arr) subsubsection\Further elementary properties\ lemma L_10_5_\_arrow_is_arr: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = ntcf_arrow \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "L_10_5_\_arrow \ \ c \' a b : Hom \ c (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b\)" proof- note L_10_5_\_arrow_components = L_10_5_\_arrow_components'[OF assms(1,2)] interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(5)) show ?thesis proof(intro cat_Set_is_arrI) show "arr_Set \ (L_10_5_\_arrow \ \ c \' a b)" proof(intro arr_SetI) show "vfsequence (L_10_5_\_arrow \ \ c \' a b)" unfolding L_10_5_\_arrow_def by simp show "vcard (L_10_5_\_arrow \ \ c \' a b) = 3\<^sub>\" unfolding L_10_5_\_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (L_10_5_\_arrow \ \ c \' a b\ArrVal\) \\<^sub>\ L_10_5_\_arrow \ \ c \' a b\ArrCod\" unfolding L_10_5_\_arrow_components proof(intro vrange_VLambda_vsubset, unfold in_Hom_iff) fix f assume "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" from L_10_5_\_arrow_ArrVal_is_arr[OF assms(1,2,4,5) this assms(6)] this show "\'\NTMap\\0, b, f\\<^sub>\ : a \\<^bsub>\\<^esub> \\ObjMap\\b\" by ( cs_prems cs_simp: L_10_5_\_arrow_ArrVal_app' cat_cs_simps cs_intro: cat_cs_intros ) qed from assms(3,6) show "L_10_5_\_arrow \ \ c \' a b\ArrDom\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) from assms(1-3,6) \.cat_cone_obj show "L_10_5_\_arrow \ \ c \' a b\ArrCod\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) qed (auto simp: L_10_5_\_arrow_components) qed (simp_all add: L_10_5_\_arrow_components) qed lemma L_10_5_\_arrow_is_arr'[cat_Kan_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = ntcf_arrow \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" and "A = Hom \ c (\\ObjMap\\b\)" and "B = Hom \ a (\\ObjMap\\b\)" and "\' = cat_Set \" shows "L_10_5_\_arrow \ \ c \' a b : A \\<^bsub>\'\<^esub> B" using assms(1-6) unfolding assms(7-9) by (rule L_10_5_\_arrow_is_arr) lemma L_10_5_\_cf_hom[cat_Kan_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = ntcf_arrow \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "f : a' \\<^bsub>\\<^esub> b'" shows "L_10_5_\_arrow \ \ c \' a b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\c\, \\ArrMap\\f\]\<^sub>\ = cf_hom \ [\\CId\\a\, \\ArrMap\\f\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> L_10_5_\_arrow \ \ c \' a a'" (is "?lhs = ?rhs") proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(5)) have [cat_Kan_cs_simps]: "\\NTMap\\a'', b'', \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\\<^sub>\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a', b', f'\\<^sub>\" if F_def: "F = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" and A_def: "A = [a', b', f']\<^sub>\" and B_def: "B = [a'', b'', f'']\<^sub>\" and F: "F : A \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> B" for F A B a' b' f' a'' b'' f'' g' h' proof- from F[unfolded F_def A_def B_def] assms(3) have a'_def: "a' = 0" and a''_def: "a'' = 0" and g'_def: "g' = 0" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'': "f'' : c \\<^bsub>\\<^esub> \\ObjMap\\b''\" and f''_def: "\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" by auto from \.ntcf_Comp_commute[OF F] that(2) F g' h' f' f'' \.is_functor_axioms \.is_functor_axioms show "\\NTMap\\a'', b'', \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\\<^sub>\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a', b', f'\\<^sub>\" unfolding F_def A_def B_def a'_def a''_def g'_def by (*slow*) ( - cs_prems 1 + cs_prems cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric] cs_intro: cat_cs_intros cat_comma_cs_intros ) qed from assms(3) assms(6,7) \.HomCod.category_axioms have lhs_is_arr: "?lhs : Hom \ c (\\ObjMap\\a'\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b'\)" unfolding assms(4) by ( cs_concl cs_simp: cs_intro: cat_lim_cs_intros cat_cs_intros cat_Kan_cs_intros cat_prod_cs_intros cat_op_intros ) then have dom_lhs: "\\<^sub>\ ((?lhs)\ArrVal\) = Hom \ c (\\ObjMap\\a'\)" by (cs_concl cs_simp: cat_cs_simps) from assms(3) assms(6,7) \.HomCod.category_axioms \.HomCod.category_axioms have rhs_is_arr: "?rhs : Hom \ c (\\ObjMap\\a'\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b'\)" unfolding assms(4) by ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cat_Kan_cs_intros cat_prod_cs_intros cat_op_intros ) then have dom_rhs: "\\<^sub>\ ((?rhs)\ArrVal\) = Hom \ c (\\ObjMap\\a'\)" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI) from lhs_is_arr show arr_Set_lhs: "arr_Set \ ?lhs" by (auto dest: cat_Set_is_arrD(1)) from rhs_is_arr show arr_Set_rhs: "arr_Set \ ?rhs" by (auto dest: cat_Set_is_arrD(1)) show "?lhs\ArrVal\ = ?rhs\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix g assume prems: "g : c \\<^bsub>\\<^esub> \\ObjMap\\a'\" from prems assms(7) have \f: "\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> g : c \\<^bsub>\\<^esub> \\ObjMap\\b'\" by (cs_concl cs_intro: cat_cs_intros) with assms(6,7) prems \.HomCod.category_axioms \.HomCod.category_axioms show "?lhs\ArrVal\\g\ = ?rhs\ArrVal\\g\" by (*slow*) ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cat_Kan_cs_intros cat_comma_cs_intros cat_prod_cs_intros cat_op_intros cat_1_is_arrI cs_simp: L_10_5_\_arrow_ArrVal_app' cat_cs_simps cat_Kan_cs_simps cat_op_simps cat_FUNCT_cs_simps cat_comma_cs_simps assms(4) )+ qed (use arr_Set_lhs arr_Set_rhs in auto) qed ( use lhs_is_arr rhs_is_arr in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\ )+ qed subsection\Lemma X.5: \L_10_5_\\\ subsubsection\Definition and elementary properties\ definition L_10_5_\ where "L_10_5_\ \ \ c \ a = [ (\bf\\<^sub>\c \\<^sub>C\<^sub>F \\Obj\. \\NTMap\\bf\1\<^sub>\\\\ArrVal\\bf\2\<^sub>\\\), cf_const (c \\<^sub>C\<^sub>F \) (\\HomCod\) a, \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \, c \\<^sub>C\<^sub>F \, (\\HomCod\) ]\<^sub>\" text\Components.\ lemma L_10_5_\_components: shows "L_10_5_\ \ \ c \ a\NTMap\ = (\bf\\<^sub>\c \\<^sub>C\<^sub>F \\Obj\. \\NTMap\\bf\1\<^sub>\\\\ArrVal\\bf\2\<^sub>\\\)" and "L_10_5_\ \ \ c \ a\NTDom\ = cf_const (c \\<^sub>C\<^sub>F \) (\\HomCod\) a" and "L_10_5_\ \ \ c \ a\NTCod\ = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \" and "L_10_5_\ \ \ c \ a\NTDGDom\ = c \\<^sub>C\<^sub>F \" and "L_10_5_\ \ \ c \ a\NTDGCod\ = (\\HomCod\)" unfolding L_10_5_\_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) lemmas L_10_5_\_components' = L_10_5_\_components[ where \=\ and \=\, unfolded cat_cs_simps ] lemmas [cat_Kan_cs_simps] = L_10_5_\_components'(2-5) end subsubsection\Natural transformation map\ mk_VLambda L_10_5_\_components(1) |vsv L_10_5_\_NTMap_vsv[cat_Kan_cs_intros]| |vdomain L_10_5_\_NTMap_vdomain[cat_Kan_cs_simps]| lemma L_10_5_\_NTMap_app[cat_Kan_cs_simps]: assumes "bf = [0, b, f]\<^sub>\" and "bf \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" shows "L_10_5_\ \ \ c \ a\NTMap\\bf\ = \\NTMap\\b\\ArrVal\\f\" using assms unfolding L_10_5_\_components by (simp add: nat_omega_simps) subsubsection\\L_10_5_\\ is a cone\ lemma L_10_5_\_is_cat_cone[cat_cs_intros]: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and \'_def: "\' = ntcf_arrow \" and \: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and a: "a \\<^sub>\ \\Obj\" shows "L_10_5_\ \ \ c \' a : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?H_\ = \\c. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-)\ let ?H_\ = \\a. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)\ - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) - - from assms(3) interpret c\: tiny_category \ \c \\<^sub>C\<^sub>F \\ + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + + from assms(3) interpret c\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) - from assms(3) interpret \c: is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ + from assms(3) interpret \c: is_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by ( cs_concl cs_simp: cat_comma_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_comma_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros ) interpret \: is_ntcf \ \ \cat_Set \\ \?H_\ c \\<^sub>C\<^sub>F \\ \?H_\ a \\<^sub>C\<^sub>F \\ \ by (rule \) show ?thesis - proof(intro is_cat_coneI is_tm_ntcfI' is_ntcfI') + proof(intro is_cat_coneI is_ntcfI') show "vfsequence (L_10_5_\ \ \ c \' a)" unfolding L_10_5_\_def by simp show "vcard (L_10_5_\ \ \ c \' a) = 5\<^sub>\" unfolding L_10_5_\_def by (simp add: nat_omega_simps) from a interpret cf_const: - is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \cf_const (c \\<^sub>C\<^sub>F \) \ a\ - by (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros) + is_functor \ \c \\<^sub>C\<^sub>F \\ \ \cf_const (c \\<^sub>C\<^sub>F \) \ a\ + by (cs_concl cs_intro: cat_cs_intros) show "L_10_5_\ \ \ c \' a\NTMap\\bf\ : cf_const (c \\<^sub>C\<^sub>F \) \ a\ObjMap\\bf\ \\<^bsub>\\<^esub> (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\bf\" if "bf \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" for bf proof- from that assms(3) obtain b f where bf_def: "bf = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from \.ntcf_NTMap_is_arr[OF b] a b assms(3) f have "\\NTMap\\b\ : Hom \ c (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b\)" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) with that b f show "L_10_5_\ \ \ c \' a\NTMap\\bf\ : cf_const (c \\<^sub>C\<^sub>F \) \ a\ObjMap\\bf\ \\<^bsub>\\<^esub> (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\bf\" unfolding bf_def \'_def by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) qed show "L_10_5_\ \ \ c \' a\NTMap\\B\ \\<^sub>A\<^bsub>\\<^esub> cf_const (c \\<^sub>C\<^sub>F \) \ a\ArrMap\\F\ = (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ArrMap\\F\ \\<^sub>A\<^bsub>\\<^esub> L_10_5_\ \ \ c \' a\NTMap\\A\" if "F : A \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> B" for A B F proof- from \.is_functor_axioms that assms(3) obtain a' f a'' f' g where F_def: "F = [[0, a', f]\<^sub>\, [0, a'', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" and A_def: "A = [0, a', f]\<^sub>\" and B_def: "B = [0, a'', f']\<^sub>\" and g: "g : a' \\<^bsub>\\<^esub> a''" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\a'\" and f': "f' : c \\<^bsub>\\<^esub> \\ObjMap\\a''\" and f'_def: "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" by auto from \.ntcf_Comp_commute[OF g] have "(\\NTMap\\a''\ \\<^sub>A\<^bsub>cat_Set \\<^esub> (?H_\ c \\<^sub>C\<^sub>F \)\ArrMap\\g\)\ArrVal\\f\ = ((?H_\ a \\<^sub>C\<^sub>F \)\ArrMap\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\a'\)\ArrVal\\f\" by simp from this a g f f' \.HomCod.category_axioms \.HomCod.category_axioms have [cat_cs_simps]: "\\NTMap\\a''\\ArrVal\\\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a'\\ArrVal\\f\" by (*slow*) ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros ) from that a g f f' \.HomCod.category_axioms \.HomCod.category_axioms show ?thesis unfolding F_def A_def B_def \'_def (*slow*) by ( cs_concl cs_simp: f'_def[symmetric] cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) qed qed ( use assms in \ cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_Kan_cs_intros a + cs_intro: cat_cs_intros cat_Kan_cs_intros a \ )+ qed lemma L_10_5_\_is_cat_cone'[cat_Kan_cs_intros]: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = ntcf_arrow \" and "\' = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \" and "c\ = c \\<^sub>C\<^sub>F \" and "\' = \" and "\' = \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "a \\<^sub>\ \\Obj\" shows "L_10_5_\ \ \ c \' a : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \' : c\ \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1-4,9,10) unfolding assms(5-8) by (rule L_10_5_\_is_cat_cone) subsection\Lemma X.5: \L_10_5_\\\ subsubsection\Definition and elementary properties\ definition L_10_5_\ :: "V \ V \ V \ V \ V \ V \ V" where "L_10_5_\ \ \ \ c \ a = [ (\b\\<^sub>\\\HomDom\\Obj\. L_10_5_\_arrow \ \ c \ a b), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(c,-) \\<^sub>C\<^sub>F \, Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(a,-) \\<^sub>C\<^sub>F \, \\HomDom\, cat_Set \ ]\<^sub>\" text\Components.\ lemma L_10_5_\_components: shows "L_10_5_\ \ \ \ c \ a\NTMap\ = (\b\\<^sub>\\\HomDom\\Obj\. L_10_5_\_arrow \ \ c \ a b)" and "L_10_5_\ \ \ \ c \ a\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(c,-) \\<^sub>C\<^sub>F \" and "L_10_5_\ \ \ \ c \ a\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(a,-) \\<^sub>C\<^sub>F \" and "L_10_5_\ \ \ \ c \ a\NTDGDom\ = \\HomDom\" and "L_10_5_\ \ \ \ c \ a\NTDGCod\ = cat_Set \" unfolding L_10_5_\_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \ - assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) lemmas L_10_5_\_components' = L_10_5_\_components[ where \=\ and \=\, unfolded cat_cs_simps ] lemmas [cat_Kan_cs_simps] = L_10_5_\_components'(2-5) end subsubsection\Natural transformation map\ mk_VLambda L_10_5_\_components(1) |vsv L_10_5_\_NTMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda L_10_5_\_components'(1)[OF \ \] |vdomain L_10_5_\_NTMap_vdomain[cat_Kan_cs_simps]| |app L_10_5_\_NTMap_app[cat_Kan_cs_simps]| end subsubsection\\L_10_5_\\ is a natural transformation\ lemma L_10_5_\_is_ntcf: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and \'_def: "\' = ntcf_arrow \" and \: "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and a: "a \\<^sub>\ \\Obj\" shows "L_10_5_\ \ \ \ c \' a : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" (is \?L_10_5_\ : ?H_\ c \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?H_\ a \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \\) proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(5)) - from assms(3) interpret c\: tiny_category \ \c \\<^sub>C\<^sub>F \\ + from assms(3) interpret c\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) - from assms(3) interpret \c: is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ + from assms(3) interpret \c: is_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by ( cs_concl cs_simp: cat_comma_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_comma_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros ) show "?L_10_5_\ : ?H_\ c \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?H_\ a \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof(intro is_ntcfI') show "vfsequence ?L_10_5_\" unfolding L_10_5_\_def by auto show "vcard ?L_10_5_\ = 5\<^sub>\" unfolding L_10_5_\_def by (simp add: nat_omega_simps) show "?L_10_5_\\NTMap\\b\ : (?H_\ c \\<^sub>C\<^sub>F \)\ObjMap\\b\ \\<^bsub>cat_Set \\<^esub> (?H_\ a \\<^sub>C\<^sub>F \)\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" for b proof- from a that assms(3) show ?thesis unfolding \'_def by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_Kan_cs_intros cat_lim_cs_intros cat_cs_intros cat_op_intros ) qed show "?L_10_5_\\NTMap\\b'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> (?H_\ c \\<^sub>C\<^sub>F \)\ArrMap\\f\ = (?H_\ a \\<^sub>C\<^sub>F \)\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\\NTMap\\a'\" if "f : a' \\<^bsub>\\<^esub> b'" for a' b' f proof- from that a assms(3) show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_op_simps \'_def cs_intro: cat_lim_cs_intros cat_cs_intros ) qed qed ( use assms(3,6) in \ cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros \ )+ qed lemma L_10_5_\_is_ntcf'[cat_Kan_cs_intros]: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "\' = ntcf_arrow \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \" and "\' = \" and "\' = cat_Set \" and "\' = \" and "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "L_10_5_\ \ \ \ c \' a : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1-4,10,11) unfolding assms(5-9) by (rule L_10_5_\_is_ntcf) subsection\Lemma X.5: \L_10_5_\_arrow\\ subsubsection\Definition and elementary properties\ definition L_10_5_\_arrow where "L_10_5_\_arrow \ \ \ \ c a = [ (\\\\<^sub>\L_10_5_N \ \ \ \ c\ObjMap\\a\. ntcf_arrow (L_10_5_\ \ \ c \ a)), L_10_5_N \ \ \ \ c\ObjMap\\a\, cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\ ]\<^sub>\" text\Components.\ lemma L_10_5_\_arrow_components: shows "L_10_5_\_arrow \ \ \ \ c a\ArrVal\ = (\\\\<^sub>\L_10_5_N \ \ \ \ c\ObjMap\\a\. ntcf_arrow (L_10_5_\ \ \ c \ a))" and "L_10_5_\_arrow \ \ \ \ c a\ArrDom\ = L_10_5_N \ \ \ \ c\ObjMap\\a\" and "L_10_5_\_arrow \ \ \ \ c a\ArrCod\ = cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" unfolding L_10_5_\_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) lemmas [cat_Kan_cs_simps] = L_10_5_\_arrow_components(2,3) subsubsection\Arrow value\ mk_VLambda L_10_5_\_arrow_components(1) |vsv L_10_5_\_arrow_vsv[cat_Kan_cs_intros]| |vdomain L_10_5_\_arrow_vdomain| |app L_10_5_\_arrow_app| lemma L_10_5_\_arrow_vdomain'[cat_Kan_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" shows "\\<^sub>\ (L_10_5_\_arrow \ \ \ \ c a\ArrVal\) = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \))" using assms by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps L_10_5_\_arrow_vdomain cs_intro: cat_cs_intros ) lemma L_10_5_\_arrow_app'[cat_Kan_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and \'_def: "\' = ntcf_arrow \" and \: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and a: "a \\<^sub>\ \\Obj\" shows "L_10_5_\_arrow \ \ \ \ c a\ArrVal\\\'\ = ntcf_arrow (L_10_5_\ \ \ c \' a)" using assms by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps L_10_5_\_arrow_app cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) lemma \\a_def: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and \\a'_def: "\\a' = ntcf_arrow \\a" and \\a: "\\a : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and a: "a \\<^sub>\ \\Obj\" shows "\\a = L_10_5_\ \ \ \ c (ntcf_arrow (L_10_5_\ \ \ c \\a' a)) a" (is \\\a = ?L_10_5_\ (ntcf_arrow ?L_10_5_\) a\) proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \\a: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \\ \\a by (rule \\a) show ?thesis proof(rule ntcf_eqI) show "\\a : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (rule \\a) from assms(1-3) a show "?L_10_5_\ (ntcf_arrow ?L_10_5_\) a : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_Kan_cs_simps \\a'_def cs_intro: cat_cs_intros cat_Kan_cs_intros ) have dom_lhs: "\\<^sub>\ (\\a\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) have dom_rhs: "\\<^sub>\ (?L_10_5_\ (ntcf_arrow (?L_10_5_\)) a\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros) show "\\a\NTMap\ = ?L_10_5_\ (ntcf_arrow ?L_10_5_\) a\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix b assume prems: "b \\<^sub>\ \\Obj\" from prems assms(3) a have lhs: "\\a\NTMap\\b\ : Hom \ c (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b\)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) then have dom_lhs: "\\<^sub>\ (\\a\NTMap\\b\\ArrVal\) = Hom \ c (\\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) from prems assms(3) a have rhs: "L_10_5_\_arrow \ \ c (ntcf_arrow ?L_10_5_\) a b : Hom \ c (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b\)" unfolding \\a'_def by ( cs_concl cs_simp: cat_Kan_cs_simps - cs_intro: cat_small_cs_intros cat_Kan_cs_intros cat_cs_intros + cs_intro: cat_Kan_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ (L_10_5_\_arrow \ \ c (ntcf_arrow ?L_10_5_\) a b\ArrVal\) = Hom \ c (\\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) have [cat_cs_simps]: "\\a\NTMap\\b\ = L_10_5_\_arrow \ \ c (ntcf_arrow ?L_10_5_\) a b" proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (\\a\NTMap\\b\)" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ (L_10_5_\_arrow \ \ c (ntcf_arrow (?L_10_5_\)) a b)" by (auto dest: cat_Set_is_arrD(1)) show "\\a\NTMap\\b\\ArrVal\ = L_10_5_\_arrow \ \ c (ntcf_arrow ?L_10_5_\) a b\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix f assume "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" with assms prems show "\\a\NTMap\\b\\ArrVal\\f\ = L_10_5_\_arrow \ \ c (ntcf_arrow ?L_10_5_\) a b\ArrVal\\f\" unfolding \\a'_def by ( cs_concl cs_simp: cat_Kan_cs_simps cat_FUNCT_cs_simps L_10_5_\_arrow_ArrVal_app cs_intro: cat_cs_intros cat_comma_cs_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ from prems show "\\a\NTMap\\b\ = L_10_5_\ \ \ \ c (ntcf_arrow ?L_10_5_\) a\NTMap\\b\" by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros cat_Kan_cs_intros V_cs_intros)+ qed simp_all qed subsection\Lemma X.5: \L_10_5_\'_arrow\\ subsubsection\Definition and elementary properties\ definition L_10_5_\'_arrow :: "V \ V \ V \ V \ V \ V \ V" where "L_10_5_\'_arrow \ \ \ \ c a = [ ( \\\\<^sub>\cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\. ntcf_arrow (L_10_5_\ \ \ \ c \ a) ), cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\, L_10_5_N \ \ \ \ c\ObjMap\\a\ ]\<^sub>\" text\Components.\ lemma L_10_5_\'_arrow_components: shows "L_10_5_\'_arrow \ \ \ \ c a\ArrVal\ = ( \\\\<^sub>\cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\. ntcf_arrow (L_10_5_\ \ \ \ c \ a) )" and [cat_Kan_cs_simps]: "L_10_5_\'_arrow \ \ \ \ c a\ArrDom\ = cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" and [cat_Kan_cs_simps]: "L_10_5_\'_arrow \ \ \ \ c a\ArrCod\ = L_10_5_N \ \ \ \ c\ObjMap\\a\" unfolding L_10_5_\'_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow value\ mk_VLambda L_10_5_\'_arrow_components(1) |vsv L_10_5_\'_arrow_ArrVal_vsv[cat_Kan_cs_intros]| |vdomain L_10_5_\'_arrow_ArrVal_vdomain| |app L_10_5_\'_arrow_ArrVal_app| lemma L_10_5_\'_arrow_ArrVal_vdomain'[cat_Kan_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and \: "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and a: "a \\<^sub>\ \\Obj\" shows "\\<^sub>\ (L_10_5_\'_arrow \ \ \ \ c a\ArrVal\) = Hom - (cat_Funct \ (c \\<^sub>C\<^sub>F \) \) + (cat_FUNCT \ (c \\<^sub>C\<^sub>F \) \) (cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a)) (cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \))" proof- interpret \: \ \ by (rule assms(1)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(3)) - from assms(2,4) show ?thesis + from assms(1,2,4) show ?thesis by ( - cs_concl + cs_concl cs_simp: cat_Kan_cs_simps L_10_5_\'_arrow_ArrVal_vdomain cs_intro: cat_cs_intros ) qed lemma L_10_5_\'_arrow_ArrVal_app'[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and \'_def: "\' = ntcf_arrow \" and \: "\ : a <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" and a: "a \\<^sub>\ \\Obj\" shows "L_10_5_\'_arrow \ \ \ \ c a\ArrVal\\\'\ = ntcf_arrow (L_10_5_\ \ \ \ c \' a)" proof- interpret \: \ \ by (rule assms(1)) interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ by (rule assms(4)) from assms(2,5) have "\' \\<^sub>\ cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" unfolding \'_def by ( cs_concl cs_simp: cat_Kan_cs_simps cat_Funct_components(1) - cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros cat_cs_intros + cs_intro: cat_FUNCT_cs_intros cat_cs_intros ) then show "L_10_5_\'_arrow \ \ \ \ c a\ArrVal\\\'\ = ntcf_arrow (L_10_5_\ \ \ \ c \' a)" unfolding L_10_5_\'_arrow_components by auto qed subsubsection\\L_10_5_\'_arrow\ is an isomorphism in the category \Set\\ lemma L_10_5_\'_arrow_is_arr_isomorphism: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" shows "L_10_5_\'_arrow \ \ \ \ c a : cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ObjMap\\a\" (*FIXME: any reason not to evaluate ObjMap*) ( is \ ?L_10_5_\'_arrow : cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> ?L_10_5_N\ObjMap\\a\ \ ) proof- let ?FUNCT = \\\. cat_FUNCT \ \ (cat_Set \)\ - let ?c\_\ = \cat_Funct \ (c \\<^sub>C\<^sub>F \) \\ + let ?c\_\ = \cat_FUNCT \ (c \\<^sub>C\<^sub>F \) \\ let ?H_\ = \\c. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-)\ let ?H_\ = \\c. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)\ from assms(1,2) interpret \: \ \ by simp - interpret \: is_tm_functor \ \ \ \ by (rule assms(3)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(4)) - - from \.vempty_is_zet assms interpret c\: tiny_category \ \c \\<^sub>C\<^sub>F \\ + interpret \: is_functor \ \ \ \ by (rule assms(3)) + interpret \: is_functor \ \ \ \ by (rule assms(4)) + + from \.vempty_is_zet assms interpret c\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) - from assms(2,6) interpret c\_\: category \ ?c\_\ + from assms(2,6) interpret c\_\: category \ ?c\_\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \.vempty_is_zet assms interpret \c: - is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ + is_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) - have \\: "\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros) + have \\: "\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros) from assms(5,6) have [cat_cs_simps]: "cf_of_cf_map (c \\<^sub>C\<^sub>F \) \ (cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a)) = cf_const (c \\<^sub>C\<^sub>F \) \ a" "cf_of_cf_map (c \\<^sub>C\<^sub>F \) \ (cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)) = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \" "cf_of_cf_map \ (cat_Set \) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)) = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \" "cf_of_cf_map \ (cat_Set \) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)) = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros)+ - note cf_Cone_ObjMap_app = is_tm_functor.cf_Cone_ObjMap_app[OF \\ assms(1,2,6)] + note cf_Cone_ObjMap_app = is_functor.cf_Cone_ObjMap_app[OF \\ assms(1,2,6)] show ?thesis proof ( intro cat_Set_is_arr_isomorphismI cat_Set_is_arrI arr_SetI, unfold L_10_5_\'_arrow_components(3) cf_Cone_ObjMap_app ) show "vfsequence ?L_10_5_\'_arrow" unfolding L_10_5_\'_arrow_def by auto show \'_arrow_ArrVal_vsv: "vsv (?L_10_5_\'_arrow\ArrVal\)" unfolding L_10_5_\'_arrow_components by auto show "vcard ?L_10_5_\'_arrow = 3\<^sub>\" unfolding L_10_5_\'_arrow_def by (simp add: nat_omega_simps) show [cat_cs_simps]: "\\<^sub>\ (?L_10_5_\'_arrow\ArrVal\) = ?L_10_5_\'_arrow\ArrDom\" unfolding L_10_5_\'_arrow_components by simp show vrange_\'_arrow_vsubset_N'': "\\<^sub>\ (?L_10_5_\'_arrow\ArrVal\) \\<^sub>\ ?L_10_5_N\ObjMap\\a\" unfolding L_10_5_\'_arrow_components proof(rule vrange_VLambda_vsubset) fix \ assume prems: "\ \\<^sub>\ cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" from this assms c\_\.category_axioms have \_is_arr: "\ : cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a) \\<^bsub>?c\_\\<^esub> cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)" by ( cs_prems - cs_simp: cat_cs_simps cat_Kan_cs_simps cat_Funct_components(1) - cs_intro: cat_small_cs_intros + cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_components(1) + cs_intro: cat_cs_intros ) - note \ = cat_Funct_is_arrD(1,2)[OF \_is_arr, unfolded cat_cs_simps] + note \ = cat_FUNCT_is_arrD(1,2)[OF \_is_arr, unfolded cat_cs_simps] have "cf_of_cf_map (c \\<^sub>C\<^sub>F \) \ (cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)) = \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) from prems assms \(1) show "ntcf_arrow (L_10_5_\ \ \ \ c \ a) \\<^sub>\ ?L_10_5_N\ObjMap\\a\" by (subst \(2)) (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: is_cat_coneI cat_cs_intros cat_Kan_cs_intros cat_FUNCT_cs_intros ) qed show "\\<^sub>\ (?L_10_5_\'_arrow\ArrVal\) = ?L_10_5_N\ObjMap\\a\" proof ( intro vsubset_antisym[OF vrange_\'_arrow_vsubset_N''], intro vsubsetI ) fix \\a assume "\\a \\<^sub>\ ?L_10_5_N\ObjMap\\a\" from this assms have \\a: "\\a : cf_map (?H_\ c \\<^sub>C\<^sub>F \) \\<^bsub>?FUNCT \\<^esub> cf_map (?H_\ a \\<^sub>C\<^sub>F \)" by ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros ) note \\a = cat_FUNCT_is_arrD[OF this, unfolded cat_cs_simps] interpret \: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \L_10_5_\ \ \ c \\a a\ by (rule L_10_5_\_is_cat_cone[OF assms(3,4,5) \\a(2,1) assms(6)]) show "\\a \\<^sub>\ \\<^sub>\ (?L_10_5_\'_arrow\ArrVal\)" proof(rule vsv.vsv_vimageI2') show "vsv (?L_10_5_\'_arrow\ArrVal\)" by (rule \'_arrow_ArrVal_vsv) from \.is_cat_cone_axioms assms show "ntcf_arrow (L_10_5_\ \ \ c \\a a) \\<^sub>\ \\<^sub>\ (?L_10_5_\'_arrow\ArrVal\)" by ( cs_concl cs_simp: cat_Kan_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros + cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from assms \\a(1,2) show "\\a = ?L_10_5_\'_arrow\ArrVal\\ntcf_arrow (L_10_5_\ \ \ c \\a a)\" by ( subst \\a(2), cs_concl_step \\a_def[OF assms(3,4,5) \\a(2,1) assms(6)] ) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed qed from assms show "?L_10_5_\'_arrow\ArrDom\ \\<^sub>\ Vset \" - by (intro Vset_trans[OF _ Vset_in_mono[OF assms(2)]]) + by ( cs_concl - cs_simp: cat_Kan_cs_simps cat_Funct_components(1) cf_Cone_ObjMap_app - cs_intro: - cat_small_cs_intros - cat_cs_intros - cat_FUNCT_cs_intros - c\_\.cat_Hom_in_Vset + cs_simp: cat_Kan_cs_simps cat_FUNCT_components(1) cf_Cone_ObjMap_app + cs_intro: cat_cs_intros cat_FUNCT_cs_intros c\_\.cat_Hom_in_Vset ) with assms(2) have "?L_10_5_\'_arrow\ArrDom\ \\<^sub>\ Vset \" by (meson Vset_in_mono Vset_trans) from assms show "?L_10_5_N\ObjMap\\a\ \\<^sub>\ Vset \" by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros FUNCT_\.cat_Hom_in_Vset cat_FUNCT_cs_intros ) show dom_\'_arrow: "\\<^sub>\ (?L_10_5_\'_arrow\ArrVal\) = Hom ?c\_\ (cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a)) (cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \))" unfolding L_10_5_\'_arrow_components cf_Cone_ObjMap_app by simp show "?L_10_5_\'_arrow\ArrDom\ = Hom ?c\_\ (cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a)) (cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \))" unfolding L_10_5_\'_arrow_components cf_Cone_ObjMap_app by simp show "v11 (?L_10_5_\'_arrow\ArrVal\)" proof(rule vsv.vsv_valeq_v11I, unfold dom_\'_arrow in_Hom_iff) fix \' \'' assume prems: "\' : cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a) \\<^bsub>?c\_\\<^esub> cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)" "\'' : cf_map (cf_const (c \\<^sub>C\<^sub>F \) \ a) \\<^bsub>?c\_\\<^esub> cf_map (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)" "?L_10_5_\'_arrow\ArrVal\\\'\ = ?L_10_5_\'_arrow\ArrVal\\\''\" - note \' = cat_Funct_is_arrD[OF prems(1), unfolded cat_cs_simps] - and \'' = cat_Funct_is_arrD[OF prems(2), unfolded cat_cs_simps] + note \' = cat_FUNCT_is_arrD[OF prems(1), unfolded cat_cs_simps] + and \'' = cat_FUNCT_is_arrD[OF prems(2), unfolded cat_cs_simps] interpret \': is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'\ by (rule is_cat_coneI[OF \'(1) assms(6)]) interpret \'': is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ \ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \''\ by (rule is_cat_coneI[OF \''(1) assms(6)]) have \'\': "ntcf_arrow (ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \') = \'" by (subst (2) \'(2)) (cs_concl cs_simp: cat_FUNCT_cs_simps) have \''\'': "ntcf_arrow (ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'') = \''" by (subst (2) \''(2)) (cs_concl cs_simp: cat_FUNCT_cs_simps) from prems(3) \'(1) \''(1) assms have "L_10_5_\ \ \ \ c \' a = L_10_5_\ \ \ \ c \'' a" by (subst (asm) \'(2), use nothing in \subst (asm) \''(2)\) (*slow*) ( cs_prems cs_simp: \'\' \''\'' cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_lim_cs_intros cat_Kan_cs_intros cat_cs_intros ) from this have \\'a_\\''a: "L_10_5_\ \ \ \ c \' a\NTMap\\b\\ArrVal\\f\ = L_10_5_\ \ \ \ c \'' a\NTMap\\b\\ArrVal\\f\" if "b \\<^sub>\ \\Obj\" and "f : c \\<^bsub>\\<^esub> (\\ObjMap\\b\)" for b f by simp have [cat_cs_simps]: "\'\NTMap\\0, b, f\\<^sub>\ = \''\NTMap\\0, b, f\\<^sub>\" if "b \\<^sub>\ \\Obj\" and "f : c \\<^bsub>\\<^esub> (\\ObjMap\\b\)" for b f using \\'a_\\''a[OF that] that by ( cs_prems cs_simp: cat_Kan_cs_simps L_10_5_\_arrow_ArrVal_app cs_intro: cat_cs_intros ) have "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \' = ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \''" proof(rule ntcf_eqI) show "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \' : cf_const (c \\<^sub>C\<^sub>F \) \ a \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule \'.is_ntcf_axioms) then have dom_lhs: "\\<^sub>\ (ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'' : cf_const (c \\<^sub>C\<^sub>F \) \ a \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule \''.is_ntcf_axioms) then have dom_rhs: "\\<^sub>\ (ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \''\NTMap\) = c \\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'\NTMap\ = ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \''\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix A assume "A \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" with assms(5) obtain b f where A_def: "A = [0, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and f: "f : c \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from b f show "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \'\NTMap\\A\ = ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ \''\NTMap\\A\" unfolding A_def by (cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps) qed (cs_concl cs_intro: V_cs_intros)+ qed simp_all then show "\' = \''" proof(rule inj_onD[OF bij_betw_imp_inj_on[OF bij_betw_ntcf_of_ntcf_arrow]]) show "\' \\<^sub>\ ntcf_arrows \ (c \\<^sub>C\<^sub>F \) \" by (subst \'(2)) ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) show "\'' \\<^sub>\ ntcf_arrows \ (c \\<^sub>C\<^sub>F \) \" by (subst \''(2)) ( cs_concl cs_intro: cat_lim_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed qed (cs_concl cs_intro: cat_Kan_cs_intros) qed auto qed lemma L_10_5_\'_arrow_is_arr_isomorphism'[cat_Kan_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" and "A = cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" and "B = L_10_5_N \ \ \ \ c\ObjMap\\a\" and "\' = cat_Set \" shows "L_10_5_\'_arrow \ \ \ \ c a : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-6) unfolding assms(7-9) by (rule L_10_5_\'_arrow_is_arr_isomorphism) lemma L_10_5_\'_arrow_is_arr: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" shows "L_10_5_\'_arrow \ \ \ \ c a : cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ObjMap\\a\" by ( rule cat_Set_is_arr_isomorphismD(1)[ OF L_10_5_\'_arrow_is_arr_isomorphism[OF assms(1-6)] ] ) lemma L_10_5_\'_arrow_is_arr'[cat_Kan_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" and "A = cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)\ObjMap\\a\" and "B = L_10_5_N \ \ \ \ c\ObjMap\\a\" and "\' = cat_Set \" shows "L_10_5_\'_arrow \ \ \ \ c a : A \\<^bsub>\'\<^esub> B" using assms(1-6) unfolding assms(7-9) by (rule L_10_5_\'_arrow_is_arr) subsection\Lemma X.5: \L_10_5_\\\label{sec:lem_X_5_end}\ subsubsection\Definition and elementary properties\ definition L_10_5_\ :: "V \ V \ V \ V \ V \ V" where "L_10_5_\ \ \ \ \ c = [ (\a\\<^sub>\\\HomCod\\Obj\. L_10_5_\'_arrow \ \ \ \ c a), cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \), L_10_5_N \ \ \ \ c, op_cat (\\HomCod\), cat_Set \ ]\<^sub>\" text\Components.\ lemma L_10_5_\_components: shows "L_10_5_\ \ \ \ \ c\NTMap\ = (\a\\<^sub>\\\HomCod\\Obj\. L_10_5_\'_arrow \ \ \ \ c a)" and [cat_Kan_cs_simps]: "L_10_5_\ \ \ \ \ c\NTDom\ = cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \)" and [cat_Kan_cs_simps]: "L_10_5_\ \ \ \ \ c\NTCod\ = L_10_5_N \ \ \ \ c" and "L_10_5_\ \ \ \ \ c\NTDGDom\ = op_cat (\\HomCod\)" and [cat_Kan_cs_simps]: "L_10_5_\ \ \ \ \ c\NTDGCod\ = cat_Set \" unfolding L_10_5_\_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation is_functor \ \ \ \ by (rule \) lemmas L_10_5_\_components' = L_10_5_\_components[where \=\, unfolded cat_cs_simps] lemmas [cat_Kan_cs_simps] = L_10_5_\_components'(4) end subsubsection\Natural transformation map\ mk_VLambda L_10_5_\_components(1) |vsv L_10_5_\_NTMap_vsv[cat_Kan_cs_intros]| context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation is_functor \ \ \ \ by (rule \) mk_VLambda L_10_5_\_components(1)[where \=\, unfolded cat_cs_simps] |vdomain L_10_5_\_NTMap_vdomain[cat_Kan_cs_simps]| |app L_10_5_\_NTMap_app[cat_Kan_cs_simps]| end subsubsection\\L_10_5_\\ is a natural isomorphism\ lemma L_10_5_\_is_iso_ntcf: \\See lemma on page 245 in \cite{mac_lane_categories_2010}.\ assumes "\ \" and "\ \\<^sub>\ \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" shows "L_10_5_\ \ \ \ \ c : cf_Cone \ \ (\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o L_10_5_N \ \ \ \ c : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" (is \?L_10_5_\ : ?cf_Cone \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o ?L_10_5_N : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \\) proof- let ?FUNCT = \\\. cat_FUNCT \ \ (cat_Set \)\ - let ?c\_\ = \cat_Funct \ (c \\<^sub>C\<^sub>F \) \\ + let ?c\_\ = \cat_FUNCT \ (c \\<^sub>C\<^sub>F \) \\ let ?ntcf_c\_\ = \ntcf_const (c \\<^sub>C\<^sub>F \) \\ let ?\_c\ = \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ let ?H_\ = \\c. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-)\ let ?H_\ = \\a. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)\ let ?L_10_5_\'_arrow = \L_10_5_\'_arrow \ \ \ \ c\ let ?cf_c\_\ = \cf_const (c \\<^sub>C\<^sub>F \) \\ let ?L_10_5_\ = \L_10_5_\ \ \ \ c\ let ?L_10_5_\_arrow = \L_10_5_\_arrow \ \ c\ interpret \: \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(3)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(4)) - - from \.vempty_is_zet assms(5) interpret c\: tiny_category \ \c \\<^sub>C\<^sub>F \\ + interpret \: is_functor \ \ \ \ by (rule assms(3)) + interpret \: is_functor \ \ \ \ by (rule assms(4)) + + from \.vempty_is_zet assms(5) interpret c\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) - from assms(2,5) interpret c\_\: category \ ?c\_\ + from assms(1,2,5) interpret c\_\: category \ ?c\_\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) interpret \_c\_\: category \ ?c\_\ - by (rule c\_\.cat_category_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros assms(2))+ - from assms(2,5) interpret \: is_functor \ \ ?c\_\ \\\<^sub>C \ (c \\<^sub>C\<^sub>F \) \\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms assms(2) interpret \\: - is_functor \ \ ?c\_\ \\\<^sub>C \ (c \\<^sub>C\<^sub>F \) \\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ + by (cs_concl cs_intro: cat_cs_intros assms(2))+ + from assms(2,5) interpret \: is_functor \ \ ?c\_\ \\\<^sub>C\<^sub>F \ (c \\<^sub>C\<^sub>F \) \\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from \.vempty_is_zet assms(5) interpret \c: - is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ + is_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by ( cs_concl cs_simp: cat_comma_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_comma_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros ) interpret \\c: is_tiny_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by (rule \c.cf_is_tiny_functor_if_ge_Limit[OF assms(1,2)]) interpret E: is_functor \ \?FUNCT \ \\<^sub>C \\ \cat_Set \\ \cf_eval \ \ \\ by (rule \.HomCod.cat_cf_eval_is_functor[OF assms(1,2)]) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from assms(2) interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ interpret cat_Set_\\: subcategory \ \cat_Set \\ \cat_Set \\ by (rule \.subcategory_cat_Set_cat_Set[OF assms(1,2)]) show ?thesis proof(intro is_iso_ntcfI is_ntcfI', unfold cat_op_simps) show "vfsequence (?L_10_5_\)" unfolding L_10_5_\_def by auto show "vcard (?L_10_5_\) = 5\<^sub>\" unfolding L_10_5_\_def by (simp add: nat_omega_simps) from assms(2) show "?cf_Cone : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" - by (intro is_tm_functor.tm_cf_cf_Cone_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros)+ + by (intro is_functor.tm_cf_cf_Cone_is_functor_if_ge_Limit) + (cs_concl cs_intro: cat_cs_intros)+ from assms show "?L_10_5_N : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_Kan_cs_intros) show "?L_10_5_\\NTMap\\a\ : ?cf_Cone\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> ?L_10_5_N\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using assms(2,3,4,5) that by ( cs_concl cs_simp: L_10_5_\_NTMap_app cs_intro: cat_cs_intros L_10_5_\'_arrow_is_arr_isomorphism ) from cat_Set_is_arr_isomorphismD[OF this] show "?L_10_5_\\NTMap\\a\ : ?cf_Cone\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> ?L_10_5_N\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using that by auto have [cat_cs_simps]: "?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom ?c\_\ [ntcf_arrow (?ntcf_c\_\ f), ntcf_arrow (ntcf_id ?\_c\)]\<^sub>\ = cf_hom (?FUNCT \) [ ntcf_arrow (ntcf_id (?H_\ c \\<^sub>C\<^sub>F \)), ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) ]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a" ( is "?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs = ?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a" ) if "f : b \\<^bsub>\\<^esub> a" for a b f proof- let ?H_f = \Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\ - from that assms \_c\_\.category_axioms c\_\.category_axioms have lhs: + from that assms c\_\.category_axioms c\_\.category_axioms have lhs: "?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs : Hom ?c\_\ (cf_map (?cf_c\_\ a)) (cf_map ?\_c\) \\<^bsub>cat_Set \\<^esub> ?L_10_5_N\ObjMap\\b\" by (*slow*) ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps cat_FUNCT_cs_simps - cat_Funct_components(1) + cat_FUNCT_components(1) cat_op_simps cs_intro: cat_Kan_cs_intros - cat_small_cs_intros cat_FUNCT_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros ) then have dom_lhs: "\\<^sub>\ ((?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\) = Hom ?c\_\ (cf_map (?cf_c\_\ a)) (cf_map ?\_c\)" by (cs_concl cs_simp: cat_cs_simps) - from that assms \_c\_\.category_axioms c\_\.category_axioms have rhs: + from that assms c\_\.category_axioms c\_\.category_axioms have rhs: "?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a : Hom ?c\_\ (cf_map (?cf_c\_\ a)) (cf_map ?\_c\) \\<^bsub>cat_Set \\<^esub> ?L_10_5_N\ObjMap\\b\" by (*slow*) ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps - cat_Funct_components(1) + cat_FUNCT_components(1) cat_op_simps cs_intro: cat_Kan_cs_intros - cat_small_cs_intros cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros ) then have dom_rhs: "\\<^sub>\ ((?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a)\ArrVal\) = Hom ?c\_\ (cf_map (?cf_c\_\ a)) (cf_map ?\_c\)" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a)" by (auto dest: cat_Set_is_arrD(1)) show "(?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\ = (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix F assume prems: "F : cf_map (?cf_c\_\ a) \\<^bsub>?c\_\\<^esub> cf_map ?\_c\" let ?F = \ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ F\ from that have [cat_cs_simps]: "cf_of_cf_map (c \\<^sub>C\<^sub>F \) \ (cf_map (?cf_c\_\ a)) = ?cf_c\_\ a" "cf_of_cf_map (c \\<^sub>C\<^sub>F \) \ (cf_map (?\_c\)) = ?\_c\" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) - note F = cat_Funct_is_arrD[OF prems, unfolded cat_cs_simps] + note F = cat_FUNCT_is_arrD[OF prems, unfolded cat_cs_simps] from that F(1) have F_const_is_cat_cone: "?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f : b <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?\_c\ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl - cs_simp: cat_cs_simps - cs_intro: cat_small_cs_intros is_cat_coneI cat_cs_intros + cs_simp: cat_cs_simps cs_intro: is_cat_coneI cat_cs_intros ) have [cat_cs_simps]: "?L_10_5_\ (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b = ?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?L_10_5_\ (ntcf_arrow ?F) a" proof(rule ntcf_eqI) from assms that F(1) show "?L_10_5_\ (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b : ?H_\ c \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?H_\ b \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_intro: - cat_small_cs_intros - cat_Kan_cs_intros - cat_cs_intros - is_cat_coneI + cat_Kan_cs_intros cat_cs_intros is_cat_coneI ) then have dom_\: "\\<^sub>\ (?L_10_5_\ (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms that F(1) show "?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?L_10_5_\ (ntcf_arrow ?F) a : ?H_\ c \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F ?H_\ b \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_intro: cat_Kan_cs_intros cat_cs_intros is_cat_coneI ) then have dom_f\\: "\\<^sub>\ ((?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?L_10_5_\ (ntcf_arrow ?F) a)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "?L_10_5_\ (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b\NTMap\ = (?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?L_10_5_\ (ntcf_arrow ?F) a)\NTMap\" proof(rule vsv_eqI, unfold dom_\ dom_f\\) fix b' assume prems': "b' \\<^sub>\ \\Obj\" let ?Y = \Yoneda_component (?H_\ b) a f (\\ObjMap\\b'\)\ let ?\b' = \\\ObjMap\\b'\\ let ?\b' = \\\ObjMap\\b'\\ have [cat_cs_simps]: "?L_10_5_\_arrow (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b b' = ?Y \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\_arrow (ntcf_arrow ?F) a b'" (is \?\_Ffbb' = ?Y\\) proof- from assms prems' F_const_is_cat_cone have \_Ffbb': "?\_Ffbb' : Hom \ c ?\b' \\<^bsub>cat_Set \\<^esub> Hom \ b ?\b'" by ( cs_concl cs_intro: cat_cs_intros L_10_5_\_arrow_is_arr ) then have dom_\_Ffbb': "\\<^sub>\ (?\_Ffbb'\ArrVal\) = Hom \ c (?\b')" by (cs_concl cs_simp: cat_cs_simps) from assms that \.HomCod.category_axioms prems' F(1) have Y\: "?Y\ : Hom \ c ?\b' \\<^bsub>cat_Set \\<^esub> Hom \ b ?\b'" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps cat_op_simps cs_intro: is_cat_coneI cat_Kan_cs_intros cat_cs_intros ) then have dom_Y\: "\\<^sub>\ (?Y\\ArrVal\) = Hom \ c (?\b')" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI) from \_Ffbb' show arr_Set_\_Ffbb': "arr_Set \ ?\_Ffbb'" by (auto dest: cat_Set_is_arrD(1)) from Y\ show arr_Set_Y\: "arr_Set \ ?Y\" by (auto dest: cat_Set_is_arrD(1)) show "?\_Ffbb'\ArrVal\ = ?Y\\ArrVal\" proof(rule vsv_eqI, unfold dom_\_Ffbb' dom_Y\ in_Hom_iff) fix g assume "g : c \\<^bsub>\\<^esub> ?\b'" with assms(2-) \.is_functor_axioms \.is_functor_axioms \.HomCod.category_axioms \.HomCod.category_axioms that prems' F(1) show "?\_Ffbb'\ArrVal\\g\ = ?Y\\ArrVal\\g\" by (*slow*) ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps L_10_5_\_arrow_ArrVal_app cat_comma_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros is_cat_coneI cat_cs_intros cat_comma_cs_intros cat_op_intros cs_simp: cat_FUNCT_cs_simps - cs_intro: cat_small_cs_intros ) qed (use arr_Set_\_Ffbb' arr_Set_Y\ in auto) qed (use \_Ffbb' Y\ in \cs_concl cs_simp: cat_cs_simps\)+ qed from assms prems' that F(1) show "?L_10_5_\ (ntcf_arrow (?F \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?ntcf_c\_\ f)) b\NTMap\\b'\ = (?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?L_10_5_\ (ntcf_arrow ?F) a)\NTMap\\b'\" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps cs_intro: is_cat_coneI cat_Kan_cs_intros cat_cs_intros ) qed (cs_concl cs_intro: cat_Kan_cs_intros cat_cs_intros)+ qed simp_all from that F(1) interpret F: is_cat_cone \ a \c \\<^sub>C\<^sub>F \\ \ \?\_c\\ ?F by (cs_concl cs_intro: is_cat_coneI cat_cs_intros) from assms(2-) prems F(1) that \.HomCod.cat_ntcf_Hom_snd_is_ntcf[OF that] (*speedup*) - \_c\_\.category_axioms (*speedup*) + c\_\.category_axioms (*speedup*) show "(?L_10_5_\'_arrow b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\\F\ = (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\'_arrow a)\ArrVal\\F\" by (subst (1 2) F(2)) (*exceptionally slow*) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_cs_simps - cat_Funct_components(1) + cat_FUNCT_components(1) cat_op_simps cs_intro: - cat_small_cs_intros is_cat_coneI cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed show "?L_10_5_\\NTMap\\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_Cone\ArrMap\\f\ = ?L_10_5_N\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?L_10_5_\\NTMap\\a\" if "f : b \\<^bsub>\\<^esub> a" for a b f using that assms by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps - cat_Funct_components(1) + cat_FUNCT_components(1) cat_FUNCT_cs_simps cat_op_simps cs_intro: - cat_small_cs_intros cat_Kan_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed subsection\ The limit of \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ exists for every pointwise right Kan extension of \\\ along \\\ \ lemma (in is_cat_pw_rKe) cat_pw_rKe_ex_cat_limit: \\Based on the elements of Chapter X-5 in \cite{mac_lane_categories_2010}. The size conditions for the functors \\\ and \\\ are related to the choice of the definition of the limit in this work (by definition, all limits are small).\ - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" obtains UA where "UA : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- define \ where "\ = \ + \" have \: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def AG.\_Limit_\\ AG.\_\_\\ \_def AG.\_\_\\) then interpret \: \ \ by simp let ?FUNCT = \\\. cat_FUNCT \ \ (cat_Set \)\ let ?H_A = \\f. Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\ let ?H_A\ = \\f. ?H_A f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \\ let ?H_\ = \\a. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)\ let ?H_\\ = \\a. ?H_\ a \\<^sub>C\<^sub>F \\ let ?H_\\ = \\a. ?H_\ a \\<^sub>C\<^sub>F \\ let ?H_\ = \\c. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-)\ let ?H_\\ = \\c. ?H_\ c \\<^sub>C\<^sub>F \\ let ?H_\\ = \\b. ?H_\ b \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \\ let ?SET_\ = \exp_cat_cf \ (cat_Set \) \\ let ?H_FUNCT = \\\ \. Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>?FUNCT \(-,cf_map \)\ let ?ua_NTDGDom = \op_cat (?FUNCT \)\ let ?ua_NTDom = \\a. ?H_FUNCT \ (?H_\\ a)\ let ?ua_NTCod = \\a. ?H_FUNCT \ (?H_\\ a) \\<^sub>C\<^sub>F op_cf ?SET_\\ - let ?c\_\ = \cat_Funct \ (c \\<^sub>C\<^sub>F \) \\ + let ?c\_\ = \cat_FUNCT \ (c \\<^sub>C\<^sub>F \) \\ let ?ua = \ \a. ntcf_ua_fo \ ?SET_\ (cf_map (?H_\\ a)) (cf_map (?H_\\ a)) (ntcf_arrow (?H_\\ a)) \ let ?cf_nt = \cf_nt \ \ (cf_id \)\ let ?cf_eval = \cf_eval \ \ \\ let ?\_c\ = \\ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \\ let ?cf_c\_\ = \cf_const (c \\<^sub>C\<^sub>F \) \\ let ?\c = \\\ObjMap\\c\\ - let ?\ = \\\<^sub>C \ (c \\<^sub>C\<^sub>F \) \\ + let ?\ = \\\<^sub>C\<^sub>F \ (c \\<^sub>C\<^sub>F \) \\ let ?ntcf_ua_fo = \ \a. ntcf_ua_fo \ ?SET_\ (cf_map (?H_\\ a)) (cf_map (?H_\\ a)) (ntcf_arrow (?H_\\ a)) \ let ?umap_fo = \ \b. umap_fo ?SET_\ (cf_map (?H_\\ b)) (cf_map (?H_\\ b)) (ntcf_arrow (?H_\\ b)) (cf_map (?H_\ c)) \ - interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) - - from AG.vempty_is_zet assms(3) interpret c\: tiny_category \ \c \\<^sub>C\<^sub>F \\ + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + + from AG.vempty_is_zet assms(3) interpret c\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) - from \\ assms(3) interpret c\_\: category \ ?c\_\ + from \\ assms(3) interpret c\_\: category \ ?c\_\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) - interpret \_c\_\: category \ ?c\_\ - by (rule c\_\.cat_category_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros \\)+ - from \\ assms(3) interpret \: is_functor \ \ ?c\_\ ?\ - by - ( - cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_op_intros - )+ - from \.is_functor_axioms \\ interpret \\: - is_functor \ \ \?c\_\\ \?\\ - by (cs_intro_step is_functor.cf_is_functor_if_ge_Limit) - (cs_concl cs_intro: cat_cs_intros)+ + from \\ assms(3) interpret \: is_functor \ \ ?c\_\ ?\ + by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from AG.vempty_is_zet assms(3) interpret \c: - is_tm_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ + is_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by ( cs_concl cs_simp: cat_comma_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_comma_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros ) + interpret \\c: is_tiny_functor \ \c \\<^sub>C\<^sub>F \\ \ \c \<^sub>O\\<^sub>C\<^sub>F \\ by (rule \c.cf_is_tiny_functor_if_ge_Limit[OF \ \\]) interpret E: is_functor \ \?FUNCT \ \\<^sub>C \\ \cat_Set \\ ?cf_eval by (rule AG.HomCod.cat_cf_eval_is_functor[OF \ \\]) from \\ interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from \\ interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) from \\ interpret FUNCT_\: tiny_category \ \?FUNCT \\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: tiny_category \ \ by (rule category.cat_tiny_category_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: is_tiny_functor \ \ \ \ by (rule is_functor.cf_is_tiny_functor_if_ge_Limit) (use \\ in \cs_concl cs_intro: cat_cs_intros\)+ interpret cat_Set_\\: subcategory \ \cat_Set \\ \cat_Set \\ by (rule AG.subcategory_cat_Set_cat_Set[OF \ \\]) from assms(3) \\ interpret Hom_c: is_functor \ \ \cat_Set \\ \?H_\ c\ by (cs_concl cs_intro: cat_cs_intros) (** E' **) define E' :: V where "E' = [ (\a\\<^sub>\\\Obj\. ?cf_eval\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\), (\f\\<^sub>\\\Arr\. ?cf_eval\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\), op_cat \, cat_Set \ ]\<^sub>\ " have E'_components: "E'\ObjMap\ = (\a\\<^sub>\\\Obj\. ?cf_eval\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\)" "E'\ArrMap\ = (\f\\<^sub>\\\Arr\. ?cf_eval\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\)" "E'\HomDom\ = op_cat \" "E'\HomCod\ = cat_Set \" unfolding E'_def dghm_field_simps by (simp_all add: nat_omega_simps) note [cat_cs_simps] = E'_components(3,4) have E'_ObjMap_app[cat_cs_simps]: "E'\ObjMap\\a\ = ?cf_eval\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\" if "a \\<^sub>\ \\Obj\" for a using that unfolding E'_components by simp have E'_ArrMap_app[cat_cs_simps]: "E'\ArrMap\\f\ = ?cf_eval\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\" if "f \\<^sub>\ \\Arr\" for f using that unfolding E'_components by simp have E': "E' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof(intro is_functorI') show "vfsequence E'" unfolding E'_def by auto show "vcard E' = 4\<^sub>\" unfolding E'_def by (simp add: nat_omega_simps) show "vsv (E'\ObjMap\)" unfolding E'_components by simp show "vsv (E'\ArrMap\)" unfolding E'_components by simp show "\\<^sub>\ (E'\ObjMap\) = op_cat \\Obj\" unfolding E'_components by (simp add: cat_op_simps) show "\\<^sub>\ (E'\ObjMap\) \\<^sub>\ cat_Set \\Obj\" unfolding E'_components proof(rule vrange_VLambda_vsubset) fix a assume prems: "a \\<^sub>\ \\Obj\" then have "?H_\\ a : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) with assms(3) prems show "?cf_eval\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\ \\<^sub>\ cat_Set \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cat_Set_components(1) cs_intro: cat_cs_intros cat_op_intros Ran.HomCod.cat_Hom_in_Vset ) qed show "\\<^sub>\ (E'\ArrMap\) = op_cat \\Arr\" unfolding E'_components by (simp add: cat_op_simps) show "E'\ArrMap\\f\ : E'\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> E'\ObjMap\\b\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f proof- from that[unfolded cat_op_simps] assms(3) show ?thesis by (intro cat_Set_\\.subcat_is_arrD) ( cs_concl cs_simp: category.cf_eval_ObjMap_app category.cf_eval_ArrMap_app E'_ObjMap_app E'_ArrMap_app cs_intro: cat_cs_intros ) qed then have [cat_cs_intros]: "E'\ArrMap\\f\ : A \\<^bsub>cat_Set \\<^esub> B" if "A = E'\ObjMap\\a\" and "B = E'\ObjMap\\b\" and "f : b \\<^bsub>\\<^esub> a" for a b f A B using that by (simp add: cat_op_simps) show "E'\ArrMap\\g \\<^sub>A\<^bsub>op_cat \\<^esub> f\ = E'\ArrMap\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> E'\ArrMap\\f\" if "g : b \\<^bsub>op_cat \\<^esub> c" and "f : a \\<^bsub>op_cat \\<^esub> b" for b c g a f proof- note g = that(1)[unfolded cat_op_simps] and f = that(2)[unfolded cat_op_simps] from g f assms(3) \\ show ?thesis by ( cs_concl cs_intro: - cat_small_cs_intros cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_prod_cs_simps cat_op_simps E.cf_ArrMap_Comp[symmetric] )+ qed show "E'\ArrMap\\op_cat \\CId\\a\\ = cat_Set \\CId\\E'\ObjMap\\a\\" if "a \\<^sub>\ op_cat \\Obj\" for a proof(cs_concl_step cat_Set_\\.subcat_CId[symmetric]) from that[unfolded cat_op_simps] assms(3) show "E'\ObjMap\\a\ \\<^sub>\ cat_Set \\Obj\" by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cat_op_simps cs_intro: cat_cs_intros ) from that[unfolded cat_op_simps] assms(3) show "E'\ArrMap\\op_cat \\CId\\a\\ = cat_Set \\CId\\E'\ObjMap\\a\\" by ( cs_concl cs_intro: cat_cs_intros cs_simp: cat_Set_components(1) cat_cs_simps cat_op_simps ntcf_id_cf_comp[symmetric] ) qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ then interpret E': is_functor \ \op_cat \\ \cat_Set \\ E' by simp (** N' **) define N' :: V where "N' = [ (\a\\<^sub>\\\Obj\. ?cf_nt\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\), (\f\\<^sub>\\\Arr\. ?cf_nt\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\), op_cat \, cat_Set \ ]\<^sub>\ " have N'_components: "N'\ObjMap\ = (\a\\<^sub>\\\Obj\. ?cf_nt\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\)" "N'\ArrMap\ = (\f\\<^sub>\\\Arr\. ?cf_nt\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\)" "N'\HomDom\ = op_cat \" "N'\HomCod\ = cat_Set \" unfolding N'_def dghm_field_simps by (simp_all add: nat_omega_simps) note [cat_cs_simps] = N'_components(3,4) have N'_ObjMap_app[cat_cs_simps]: "N'\ObjMap\\a\ = ?cf_nt\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\" if "a \\<^sub>\ \\Obj\" for a using that unfolding N'_components by simp have N'_ArrMap_app[cat_cs_simps]: "N'\ArrMap\\f\ = ?cf_nt\ArrMap\\ntcf_arrow (?H_A\ f), \\CId\\c\\\<^sub>\" if "f \\<^sub>\ \\Arr\" for f using that unfolding N'_components by simp from \\ interpret cf_nt_\: is_functor \ \?FUNCT \ \\<^sub>C \\ \cat_Set \\ \?cf_nt\ by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have N': "N' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof(intro is_functorI') show "vfsequence N'" unfolding N'_def by simp show "vcard N' = 4\<^sub>\" unfolding N'_def by (simp add: nat_omega_simps) show "vsv (N'\ObjMap\)" unfolding N'_components by simp show "vsv (N'\ArrMap\)" unfolding N'_components by simp show "\\<^sub>\ (N'\ObjMap\) = op_cat \\Obj\" unfolding N'_components by (simp add: cat_op_simps) show "\\<^sub>\ (N'\ObjMap\) \\<^sub>\ cat_Set \\Obj\" unfolding N'_components proof(rule vrange_VLambda_vsubset) fix a assume prems: "a \\<^sub>\ \\Obj\" with assms(3) \\ show "?cf_nt\ObjMap\\cf_map (?H_\\ a), c\\<^sub>\ \\<^sub>\ cat_Set \\Obj\" by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros FUNCT_\.cat_Hom_in_Vset cat_FUNCT_cs_intros ) qed show "\\<^sub>\ (N'\ArrMap\) = op_cat \\Arr\" unfolding N'_components by (simp add: cat_op_simps) show "N'\ArrMap\\f\ : N'\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> N'\ObjMap\\b\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f using that[unfolded cat_op_simps] assms(3) by ( cs_concl cs_simp: N'_ObjMap_app N'_ArrMap_app cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) show "N'\ArrMap\\g \\<^sub>A\<^bsub>op_cat \\<^esub> f\ = N'\ArrMap\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> N'\ArrMap\\f\" if "g : b \\<^bsub>op_cat \\<^esub> c" and "f : a \\<^bsub>op_cat \\<^esub> b" for b c g a f proof- from that assms(3) \\ show ?thesis unfolding cat_op_simps by ( cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_prod_cs_simps cat_op_simps cf_nt_\.cf_ArrMap_Comp[symmetric] ) qed show "N'\ArrMap\\op_cat \\CId\\a\\ = cat_Set \\CId\\N'\ObjMap\\a\\" if "a \\<^sub>\ op_cat \\Obj\" for a proof- note [cat_cs_simps] = ntcf_id_cf_comp[symmetric] ntcf_arrow_id_ntcf_id[symmetric] cat_FUNCT_CId_app[symmetric] from that[unfolded cat_op_simps] assms(3) \\ show ?thesis by (*very slow*) ( cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros cs_simp: cat_FUNCT_cs_simps cat_cs_simps cat_op_simps )+ qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ then interpret N': is_functor \ \op_cat \\ \cat_Set \\ N' by simp (** Y' **) define Y' :: V where "Y' = [ (\a\\<^sub>\\\Obj\. ntcf_Yoneda \ \ \\NTMap\\cf_map (?H_\\ a), c\\<^sub>\), N', E', op_cat \, cat_Set \ ]\<^sub>\" have Y'_components: "Y'\NTMap\ = (\a\\<^sub>\\\Obj\. ntcf_Yoneda \ \ \\NTMap\\cf_map (?H_\\ a), c\\<^sub>\)" "Y'\NTDom\ = N'" "Y'\NTCod\ = E'" "Y'\NTDGDom\ = op_cat \" "Y'\NTDGCod\ = cat_Set \" unfolding Y'_def nt_field_simps by (simp_all add: nat_omega_simps) note [cat_cs_simps] = Y'_components(2-5) have Y'_NTMap_app[cat_cs_simps]: "Y'\NTMap\\a\ = ntcf_Yoneda \ \ \\NTMap\\cf_map (?H_\\ a), c\\<^sub>\" if "a \\<^sub>\ \\Obj\" for a using that unfolding Y'_components by simp from \ \\ interpret Y: is_iso_ntcf \ \?FUNCT \ \\<^sub>C \\ \cat_Set \\ ?cf_nt ?cf_eval \ntcf_Yoneda \ \ \\ by (rule AG.HomCod.cat_ntcf_Yoneda_is_ntcf) have Y': "Y' : N' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o E' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence Y'" unfolding Y'_def by simp show "vcard Y' = 5\<^sub>\" unfolding Y'_def by (simp add: nat_omega_simps) show "vsv (Y'\NTMap\)" unfolding Y'_components by auto show "\\<^sub>\ (Y'\NTMap\) = op_cat \\Obj\" unfolding Y'_components by (simp add: cat_op_simps) show Y'_NTMap_a: "Y'\NTMap\\a\ : N'\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> E'\ObjMap\\a\" if "a \\<^sub>\ op_cat \\Obj\" for a using that[unfolded cat_op_simps] assms(3) by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) then show "Y'\NTMap\\a\ : N'\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> E'\ObjMap\\a\" if "a \\<^sub>\ op_cat \\Obj\" for a by (intro cat_Set_is_arr_isomorphismD[OF Y'_NTMap_a[OF that]]) show "Y'\NTMap\\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> N'\ArrMap\\f\ = E'\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Y'\NTMap\\a\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f proof- note f = that[unfolded cat_op_simps] from f assms(3) show ?thesis by ( cs_concl cs_simp: cat_cs_simps Y.ntcf_Comp_commute cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros )+ qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ have E'_def: "E' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)" proof(rule cf_eqI) show "E' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) from assms(3) show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ (E'\ObjMap\) = \\Obj\" unfolding E'_components by simp from assms(3) have dom_rhs: "\\<^sub>\ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ObjMap\) = \\Obj\" unfolding E'_components by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) show "E'\ObjMap\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ObjMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with assms(3) show "E'\ObjMap\\a\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ObjMap\\a\" by ( cs_concl cs_simp: cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) qed (auto simp: E'_components cat_cs_intros assms(3)) have dom_lhs: "\\<^sub>\ (E'\ArrMap\) = \\Arr\" unfolding E'_components by simp from assms(3) have dom_rhs: "\\<^sub>\ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ArrMap\) = \\Arr\" unfolding E'_components by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) show "E'\ArrMap\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ArrMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix f assume prems: "f \\<^sub>\ \\Arr\" then obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto have [cat_cs_simps]: "cf_eval_arrow \ (ntcf_arrow (?H_A\ f)) (\\CId\\c\) = cf_hom \ [f, \\CId\\?\c\]\<^sub>\" (is \?cf_eval_arrow = ?cf_hom_f\c\) proof- have cf_eval_arrow_f_CId_\c: "?cf_eval_arrow : Hom \ b ?\c \\<^bsub>cat_Set \\<^esub> Hom \ a ?\c" proof(rule cf_eval_arrow_is_arr') from f show "?H_A\ f : ?H_\\ b \\<^sub>C\<^sub>F ?H_\\ a : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) qed ( use f assms(3) in \ cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros \ )+ from f assms(3) have dom_lhs: "\\<^sub>\ (?cf_eval_arrow\ArrVal\) = Hom \ b ?\c" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) from assms(3) f Ran.HomCod.category_axioms have cf_hom_f\c: "?cf_hom_f\c : Hom \ b ?\c \\<^bsub>cat_Set \\<^esub> Hom \ a ?\c" by ( cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros ) from f assms(3) have dom_rhs: "\\<^sub>\ (?cf_hom_f\c\ArrVal\) = Hom \ b ?\c" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) show ?thesis proof(rule arr_Set_eqI) from cf_eval_arrow_f_CId_\c show "arr_Set \ ?cf_eval_arrow" by (auto dest: cat_Set_is_arrD(1)) from cf_hom_f\c show "arr_Set \ ?cf_hom_f\c" by (auto dest: cat_Set_is_arrD(1)) show "?cf_eval_arrow\ArrVal\ = ?cf_hom_f\c\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs, unfold in_Hom_iff) from f assms(3) show "vsv (?cf_eval_arrow\ArrVal\)" by (cs_concl cs_intro: cat_cs_intros) from f assms(3) show "vsv (?cf_hom_f\c\ArrVal\)" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) fix g assume "g : b \\<^bsub>\\<^esub> ?\c" with f assms(3) show "?cf_eval_arrow\ArrVal\\g\ = ?cf_hom_f\c\ArrVal\\g\" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) qed simp qed ( use cf_eval_arrow_f_CId_\c cf_hom_f\c in \cs_concl cs_simp: cat_cs_simps\ )+ qed from f prems assms(3) show "E'\ArrMap\\f\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ArrMap\\f\" by ( cs_concl cs_simp: cat_op_simps cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) qed (auto simp: E'_components cat_cs_intros assms(3)) qed simp_all from Y' have inv_Y': "inv_ntcf Y' : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o N' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" unfolding E'_def by (auto intro: iso_ntcf_is_arr_isomorphism) interpret N'': is_functor \ \op_cat \\ \cat_Set \\ \L_10_5_N \ \ \ \ c\ by (rule L_10_5_N_is_functor[OF \ \\ assms]) (** \ **) define \ :: V where "\ = [ (\a\\<^sub>\\\Obj\. ?ntcf_ua_fo a\NTMap\\cf_map (?H_\ c)\), N', L_10_5_N \ \ \ \ c, op_cat \, cat_Set \ ]\<^sub>\" have \_components: "\\NTMap\ = (\a\\<^sub>\\\Obj\. ?ntcf_ua_fo a\NTMap\\cf_map (?H_\ c)\)" "\\NTDom\ = N'" "\\NTCod\ = L_10_5_N \ \ \ \ c" "\\NTDGDom\ = op_cat \" "\\NTDGCod\ = cat_Set \" unfolding \_def nt_field_simps by (simp_all add: nat_omega_simps) note [cat_cs_simps] = Y'_components(2-5) have \_NTMap_app[cat_cs_simps]: "\\NTMap\\a\ = ?ntcf_ua_fo a\NTMap\\cf_map (?H_\ c)\" if "a \\<^sub>\ \\Obj\" for a using that unfolding \_components by simp have \: "\ : N' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o L_10_5_N \ \ \ \ c : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- show ?thesis proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence \" unfolding \_def by auto show "vcard \ = 5\<^sub>\" unfolding \_def by (simp_all add: nat_omega_simps) show "N' : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (rule N') show "L_10_5_N \ \ \ \ c : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\\NTDom\ = N'" unfolding \_components by simp show "\\NTCod\ = L_10_5_N \ \ \ \ c" unfolding \_components by simp show "\\NTDGDom\ = op_cat \" unfolding \_components by simp show "\\NTDGCod\ = cat_Set \" unfolding \_components by simp show "vsv (\\NTMap\)" unfolding \_components by simp show "\\<^sub>\ (\\NTMap\) = op_cat \\Obj\" unfolding \_components by (simp add: cat_op_simps) show \_NTMap_is_arr_isomorphism[unfolded cat_op_simps]: "\\NTMap\\a\ : N'\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ObjMap\\a\" if "a \\<^sub>\ op_cat \\Obj\" for a proof- note a = that[unfolded cat_op_simps] interpret \: is_cat_rKe_preserves \ \ \ \ \cat_Set \\ \ \ \ \?H_\ a\ \ by (rule cat_pw_rKe_preserved[OF a]) interpret a\: is_cat_rKe \ \ \ \cat_Set \\ \ \?H_\\ a\ \?H_\\ a\ \?H_\\ a\ by (rule \.cat_rKe_preserves) interpret is_iso_ntcf \ \op_cat (?FUNCT \)\ \cat_Set \\ \?H_FUNCT \ (?H_\\ a)\ \?H_FUNCT \ (?H_\\ a) \\<^sub>C\<^sub>F op_cf ?SET_\\ \?ntcf_ua_fo a\ by (rule a\.cat_rKe_ntcf_ua_fo_is_iso_ntcf_if_ge_Limit[OF \ \\]) have "cf_map (?H_\ c) \\<^sub>\ ?FUNCT \\Obj\" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from iso_ntcf_is_arr_isomorphism[unfolded cat_op_simps, OF this] a assms \\ show ?thesis by (*very slow*) ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_small_cs_intros cat_Kan_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed show \_NTMap_is_arr[unfolded cat_op_simps]: "\\NTMap\\a\ : N'\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> L_10_5_N \ \ \ \ c\ObjMap\\a\" if "a \\<^sub>\ op_cat \\Obj\" for a by ( rule cat_Set_is_arr_isomorphismD[ OF \_NTMap_is_arr_isomorphism[OF that[unfolded cat_op_simps]] ] ) show "\\NTMap\\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> N'\ArrMap\\f\ = L_10_5_N \ \ \ \ c\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\a\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f proof- note f = that[unfolded cat_op_simps] from f have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by auto interpret p_a_\: is_cat_rKe_preserves \ \ \ \ \cat_Set \\ \ \ \ \?H_\ a\ \ by (rule cat_pw_rKe_preserved[OF a]) interpret a_\: is_cat_rKe \ \ \ \cat_Set \\ \ \?H_\\ a\ \?H_\\ a\ \?H_\\ a\ by (rule p_a_\.cat_rKe_preserves) interpret ntcf_ua_fo_a_\: is_iso_ntcf \ ?ua_NTDGDom \cat_Set \\ \?ua_NTDom a\ \?ua_NTCod a\ \?ua a\ by (rule a_\.cat_rKe_ntcf_ua_fo_is_iso_ntcf_if_ge_Limit[OF \ \\]) interpret p_b_\: is_cat_rKe_preserves \ \ \ \ \cat_Set \\ \ \ \ \?H_\ b\ \ by (rule cat_pw_rKe_preserved[OF b]) interpret b_\: is_cat_rKe \ \ \ \cat_Set \\ \ \?H_\\ b\ \?H_\\ b\ \?H_\\ b\ by (rule p_b_\.cat_rKe_preserves) interpret ntcf_ua_fo_b_\: is_iso_ntcf \ ?ua_NTDGDom \cat_Set \\ \?ua_NTDom b\ \?ua_NTCod b\ \?ua b\ by (rule b_\.cat_rKe_ntcf_ua_fo_is_iso_ntcf_if_ge_Limit[OF \ \\]) interpret \_SET: is_tiny_functor \ \?FUNCT \\ \?FUNCT \\ ?SET_\ by ( rule exp_cat_cf_is_tiny_functor[ OF \ \\ AG.category_cat_Set AG.is_functor_axioms ] ) from f interpret Hom_f: is_ntcf \ \ \cat_Set \\ \?H_\ a\ \?H_\ b\ \?H_A f\ by (cs_concl cs_intro: cat_cs_intros) let ?cf_hom_lhs = \ cf_hom (?FUNCT \) [ntcf_arrow (ntcf_id (?H_\ c)), ntcf_arrow (?H_A\ f)]\<^sub>\ \ let ?cf_hom_rhs = \ cf_hom (?FUNCT \) [ ntcf_arrow (ntcf_id (?H_\ c \\<^sub>C\<^sub>F \)), ntcf_arrow (?H_A f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) ]\<^sub>\ \ let ?dom = \Hom (?FUNCT \) (cf_map (?H_\ c)) (cf_map (?H_\\ a))\ let ?cod = \Hom (?FUNCT \) (cf_map (?H_\\ c)) (cf_map (?H_\\ b))\ let ?cf_hom_lhs_umap_fo_inter = \Hom (?FUNCT \) (cf_map (?H_\ c)) (cf_map (?H_\\ b))\ let ?umap_fo_cf_hom_rhs_inter = \Hom (?FUNCT \) (cf_map (?H_\\ c)) (cf_map (?H_\\ a))\ have [cat_cs_simps]: "?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs = ?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a" proof- from f assms(3) \\ have cf_hom_lhs: "?cf_hom_lhs : ?dom \\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs_umap_fo_inter" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) from f assms(3) \\ have umap_fo_b: "?umap_fo b : ?cf_hom_lhs_umap_fo_inter \\<^bsub>cat_Set \\<^esub> ?cod" by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) from cf_hom_lhs umap_fo_b have umap_fo_cf_hom_lhs: "?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs : ?dom \\<^bsub>cat_Set \\<^esub> ?cod" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) then have dom_umap_fo_cf_hom_lhs: "\\<^sub>\ ((?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\) = ?dom" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f assms(3) \\ have cf_hom_rhs: "?cf_hom_rhs : ?umap_fo_cf_hom_rhs_inter \\<^bsub>cat_Set \\<^esub> ?cod" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) from f assms(3) \\ have umap_fo_a: "?umap_fo a : ?dom \\<^bsub>cat_Set \\<^esub> ?umap_fo_cf_hom_rhs_inter" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: - cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) from cf_hom_rhs umap_fo_a have cf_hom_rhs_umap_fo_a: "?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a : ?dom \\<^bsub>cat_Set \\<^esub> ?cod" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros ) then have dom_cf_hom_rhs_umap_fo_a: "\\<^sub>\ ((?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a)\ArrVal\) = ?dom" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show ?thesis proof(rule arr_Set_eqI) from umap_fo_cf_hom_lhs show arr_Set_umap_fo_cf_hom_lhs: "arr_Set \ (?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)" by (auto dest: cat_Set_is_arrD(1)) from cf_hom_rhs_umap_fo_a show arr_Set_cf_hom_rhs_umap_fo_a: "arr_Set \ (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a)" by (auto dest: cat_Set_is_arrD(1)) show "(?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\ = (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a)\ArrVal\" proof ( rule vsv_eqI, unfold dom_umap_fo_cf_hom_lhs dom_cf_hom_rhs_umap_fo_a in_Hom_iff; (rule refl)? ) fix \ assume prems: "\ : cf_map (?H_\ c) \\<^bsub>?FUNCT \\<^esub> cf_map (?H_\\ a)" let ?\ = \ntcf_of_ntcf_arrow \ (cat_Set \) \\ let ?lhs = \?H_\\ b \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ((?H_A\ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\ let ?rhs = \(?H_A f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_\\ a \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (?\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \))\ let ?cf_hom_\\ = \\b b'. cf_hom \ [\\CId\\b\, \\NTMap\\b'\]\<^sub>\\ let ?Yc = \\Q. Yoneda_component (?H_\ b) a f Q\ let ?\\ = \\b'. ?\\NTMap\\\\ObjMap\\b'\\\ let ?\\ = \\b'. \\ObjMap\\\\ObjMap\\b'\\\ have [cat_cs_simps]: "cf_of_cf_map \ (cat_Set \) (cf_map (?H_\ c)) = ?H_\ c" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) have [cat_cs_simps]: "cf_of_cf_map \ (cat_Set \) (cf_map (?H_\\ a)) = ?H_\\ a" by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros) note \ = cat_FUNCT_is_arrD[OF prems, unfolded cat_cs_simps] have Hom_c: "?H_\\ c : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have [cat_cs_simps]: "?lhs = ?rhs" proof(rule ntcf_eqI) from \(1) f show lhs: "?lhs : ?H_\\ c \\<^sub>C\<^sub>F ?H_\\ b : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cs_intro: cat_cs_intros) then have dom_lhs: "\\<^sub>\ (?lhs\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps)+ from \(1) f show rhs: "?rhs : ?H_\\ c \\<^sub>C\<^sub>F ?H_\\ b : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cs_intro: cat_cs_intros) then have dom_rhs: "\\<^sub>\ (?rhs\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps)+ have [cat_cs_simps]: "?cf_hom_\\ b b' \\<^sub>A\<^bsub>cat_Set \\<^esub> (?Yc (?\\ b') \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\ b') = ?Yc (\\ObjMap\\b'\) \\<^sub>A\<^bsub>cat_Set \\<^esub> (?cf_hom_\\ a b' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\ b')" (is \?lhs_Set = ?rhs_Set\) if "b' \\<^sub>\ \\Obj\" for b' proof- let ?\b' = \\\ObjMap\\b'\\ from \(1) f that assms(3) Ran.HomCod.category_axioms have lhs_Set_is_arr: "?lhs_Set : Hom \ c (?\b') \\<^bsub>cat_Set \\<^esub> Hom \ b (\\ObjMap\\b'\)" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros ) then have dom_lhs_Set: "\\<^sub>\ (?lhs_Set\ArrVal\) = Hom \ c ?\b'" by (cs_concl cs_simp: cat_cs_simps) from \(1) f that assms(3) Ran.HomCod.category_axioms have rhs_Set_is_arr: "?rhs_Set : Hom \ c (?\b') \\<^bsub>cat_Set \\<^esub> Hom \ b (\\ObjMap\\b'\)" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros ) then have dom_rhs_Set: "\\<^sub>\ (?rhs_Set\ArrVal\) = Hom \ c ?\b'" by (cs_concl cs_simp: cat_cs_simps) - show ?thesis - proof(rule arr_Set_eqI) - from lhs_Set_is_arr show arr_Set_lhs_Set: "arr_Set \ ?lhs_Set" - by (auto dest: cat_Set_is_arrD(1)) - from rhs_Set_is_arr show arr_Set_rhs_Set: "arr_Set \ ?rhs_Set" - by (auto dest: cat_Set_is_arrD(1)) - show "?lhs_Set\ArrVal\ = ?rhs_Set\ArrVal\" - proof(rule vsv_eqI, unfold dom_lhs_Set dom_rhs_Set in_Hom_iff) - fix h assume "h : c \\<^bsub>\\<^esub> ?\b'" - with \(1) f that assms Ran.HomCod.category_axioms show - "?lhs_Set\ArrVal\\h\ = ?rhs_Set\ArrVal\\h\" - by (*exceptionally slow*) - ( - cs_concl - cs_simp: cat_cs_simps cat_op_simps - cs_intro: - cat_cs_intros cat_prod_cs_intros cat_op_intros - ) - qed (use arr_Set_lhs_Set arr_Set_rhs_Set in auto) - qed - ( - use lhs_Set_is_arr rhs_Set_is_arr in - \cs_concl cs_simp: cat_cs_simps\ - )+ + show ?thesis + proof(rule arr_Set_eqI) + from lhs_Set_is_arr show arr_Set_lhs_Set: "arr_Set \ ?lhs_Set" + by (auto dest: cat_Set_is_arrD(1)) + from rhs_Set_is_arr show arr_Set_rhs_Set: "arr_Set \ ?rhs_Set" + by (auto dest: cat_Set_is_arrD(1)) + show "?lhs_Set\ArrVal\ = ?rhs_Set\ArrVal\" + proof(rule vsv_eqI, unfold dom_lhs_Set dom_rhs_Set in_Hom_iff) + fix h assume "h : c \\<^bsub>\\<^esub> ?\b'" + with \(1) f that assms Ran.HomCod.category_axioms show + "?lhs_Set\ArrVal\\h\ = ?rhs_Set\ArrVal\\h\" + by (*exceptionally slow*) + ( + cs_concl + cs_simp: cat_cs_simps cat_op_simps + cs_intro: + cat_cs_intros cat_prod_cs_intros cat_op_intros + ) + qed (use arr_Set_lhs_Set arr_Set_rhs_Set in auto) + qed + ( + use lhs_Set_is_arr rhs_Set_is_arr in + \cs_concl cs_simp: cat_cs_simps\ + )+ qed show "?lhs\NTMap\ = ?rhs\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix b' assume "b' \\<^sub>\ \\Obj\" with \(1) f assms(3) show "?lhs\NTMap\\b'\ = ?rhs\NTMap\\b'\" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all from assms(3) f \(1) prems \\ (*speedup*) Ran.HomCod.category_axioms FUNCT_\.category_axioms FUNCT_\.category_axioms AG.is_functor_axioms Ran.is_functor_axioms Hom_f.is_ntcf_axioms show "(?umap_fo b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?cf_hom_lhs)\ArrVal\\\\ = (?cf_hom_rhs \\<^sub>A\<^bsub>cat_Set \\<^esub> ?umap_fo a)\ArrVal\\\\" by (subst (1 2) \(2)) (*exceptionally slow*) ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros cat_op_intros ) qed ( use arr_Set_umap_fo_cf_hom_lhs arr_Set_cf_hom_rhs_umap_fo_a in auto ) qed ( use umap_fo_cf_hom_lhs cf_hom_rhs_umap_fo_a in \cs_concl cs_simp: cat_cs_simps\ )+ qed from f assms \\ show ?thesis by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) qed qed auto qed (**main**) from L_10_5_\_is_iso_ntcf[OF \ \\ assms] have inv_\: "inv_ntcf (L_10_5_\ \ \ \ \ c) : L_10_5_N \ \ \ \ c \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_Cone \ \ ?\_c\ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (auto intro: iso_ntcf_is_arr_isomorphism) define \ where "\ = inv_ntcf (L_10_5_\ \ \ \ \ c) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf Y'" from inv_Y' \ inv_\ have \: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_Cone \ \ ?\_c\ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" unfolding \_def by (cs_concl cs_intro: cat_cs_intros) interpret \: is_iso_ntcf \ \op_cat \\ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,?\c)\ \cf_Cone \ \ ?\_c\\ \ by (rule \) let ?\_\c_CId = \\\NTMap\\?\c\\ArrVal\\\\CId\\?\c\\\ let ?ntcf_\_\c_CId = \ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ ?\_\c_CId\ - from AG.vempty_is_zet assms(3) have \: "?\ : \ \\\<^sub>C\<^bsub>\\<^esub> ?c\_\" + from AG.vempty_is_zet assms(3) have \: "?\ : \ \\\<^sub>C\<^bsub>\\<^esub> ?c\_\" by ( cs_concl cs_simp: cat_comma_cs_simps - cs_intro: cat_small_cs_intros cat_cs_intros cat_comma_cs_intros + cs_intro: cat_cs_intros cat_comma_cs_intros ) from assms(3) have \c: "?\c \\<^sub>\ \\Obj\" by (cs_concl cs_intro: cat_cs_intros) from AG.vempty_is_zet have \_c\: "cf_map (?\_c\) \\<^sub>\ ?c\_\\Obj\" by ( cs_concl - cs_simp: cat_Funct_components(1) - cs_intro: cat_small_cs_intros cat_FUNCT_cs_intros + cs_simp: cat_FUNCT_components(1) + cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from \.ntcf_NTMap_is_arr[unfolded cat_op_simps, OF \c] assms(3) AG.vempty_is_zet \.vempty_is_zet \\ have \_\c: "\\NTMap\\?\c\ : Hom \ ?\c?\c \\<^bsub>cat_Set \\<^esub> Hom ?c\_\ (cf_map (?cf_c\_\ ?\c)) (cf_map ?\_c\)" by (*very slow*) ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps cat_comma_cs_simps cat_op_simps - cat_Funct_components(1) + cat_FUNCT_components(1) cs_intro: - cat_small_cs_intros cat_Kan_cs_intros cat_comma_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros - category.cat_category_if_ge_Limit[where \=\ and \=\] - is_functor.cf_is_functor_if_ge_Limit[where \=\ and \=\] ) with assms(3) have \_\c_CId: "?\_\c_CId : cf_map (?cf_c\_\ ?\c) \\<^bsub>?c\_\\<^esub> cf_map ?\_c\" by (cs_concl cs_intro: cat_cs_intros) have ntcf_arrow_\_\c_CId: "ntcf_arrow ?ntcf_\_\c_CId = ?\_\c_CId" - by (rule cat_Funct_is_arrD(2)[OF \_\c_CId, symmetric]) + by (rule cat_FUNCT_is_arrD(2)[OF \_\c_CId, symmetric]) + have ua: "universal_arrow_fo ?\ (cf_map (?\_c\)) ?\c ?\_\c_CId" by ( - rule is_functor.cf_universal_arrow_fo_if_is_iso_ntcf_if_ge_Limit[ - OF \ \ \\ \c \_c\ \[unfolded cf_Cone_def cat_cs_simps] + rule is_functor.cf_universal_arrow_fo_if_is_iso_ntcf[ + OF \ \c \_c\ \[unfolded cf_Cone_def cat_cs_simps] ] ) moreover have ntcf_\_\c_CId: "?ntcf_\_\c_CId : ?\c <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e ?\_c\ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_cat_coneI) - from cat_Funct_is_arrD(1)[OF \_\c_CId] assms(3) AG.vempty_is_zet show + from cat_FUNCT_is_arrD(1)[OF \_\c_CId] assms(3) AG.vempty_is_zet show "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ ?\_\c_CId : - ?cf_c\_\ ?\c \\<^sub>C\<^sub>F\<^sub>.\<^sub>t\<^sub>m ?\_c\ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + ?cf_c\_\ ?\c \\<^sub>C\<^sub>F ?\_c\ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_prems cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (rule \c) ultimately have "?ntcf_\_\c_CId : ?\c <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m ?\_c\ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - by - ( - intro is_cat_limitI[ - where u=\?ntcf_\_\c_CId\, unfolded ntcf_arrow_\_\c_CId - ] - ) + by (intro is_cat_cone.cat_cone_is_cat_limit) + (simp_all add: ntcf_arrow_\_\c_CId) then show ?thesis using that by auto qed subsection\The limit for the pointwise Kan extension\ subsubsection\Definition and elementary properties\ text\See Theorem 3 in Chapter X-5 in \cite{mac_lane_categories_2010}.\ definition the_pw_cat_rKe_limit :: "V \ V \ V \ V \ V \ V" where "the_pw_cat_rKe_limit \ \ \ \ c = [ \\ObjMap\\c\, ( SOME UA. UA : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \\HomCod\ ) ]\<^sub>\" text\Components.\ lemma the_pw_cat_rKe_limit_components: shows "the_pw_cat_rKe_limit \ \ \ \ c\UObj\ = \\ObjMap\\c\" and "the_pw_cat_rKe_limit \ \ \ \ c\UArr\ = ( SOME UA. UA : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \\HomCod\ )" unfolding the_pw_cat_rKe_limit_def ua_field_simps by (simp_all add: nat_omega_simps) context is_functor begin lemmas the_pw_cat_rKe_limit_components' = the_pw_cat_rKe_limit_components[where \=\, unfolded cat_cs_simps] end subsubsection\The limit for the pointwise Kan extension is a limit\ lemma (in is_cat_pw_rKe) cat_pw_rKe_the_pw_cat_rKe_limit_is_limit: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "c \\<^sub>\ \\Obj\" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" shows "the_pw_cat_rKe_limit \ \ \ \ c\UArr\ : the_pw_cat_rKe_limit \ \ \ \ c\UObj\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- from cat_pw_rKe_ex_cat_limit[OF assms] obtain UA where UA: "UA : \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F \ : c \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto show ?thesis unfolding the_pw_cat_rKe_limit_components by (rule someI2, unfold cat_cs_simps, rule UA) qed lemma (in is_cat_pw_rKe) cat_pw_rKe_the_ntcf_rKe_is_cat_rKe: - assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" - and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "the_ntcf_rKe \ \ \ (the_pw_cat_rKe_limit \ \ \ \) : the_cf_rKe \ \ \ (the_pw_cat_rKe_limit \ \ \ \) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" proof- - interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) show "the_ntcf_rKe \ \ \ (the_pw_cat_rKe_limit \ \ \ \) : the_cf_rKe \ \ \ (the_pw_cat_rKe_limit \ \ \ \) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : \ \\<^sub>C \ \\<^sub>C \" by ( rule the_ntcf_rKe_is_cat_rKe [ OF assms(1) ntcf_rKe.NTCod.is_functor_axioms cat_pw_rKe_the_pw_cat_rKe_limit_is_limit[OF assms] ] ) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan_Example.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan_Example.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan_Example.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_PWKan_Example.thy @@ -1,2422 +1,2425 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Pointwise Kan extensions: application example\ theory CZH_UCAT_PWKan_Example imports CZH_Elementary_Categories.CZH_ECAT_Ordinal CZH_UCAT_PWKan begin subsection\Background\ text\ The application example presented in this section is based on -Exercise 6.1.ii in \cite{riehl_category_2016}. +Exercise 6.1.ii in \cite{riehl_category_2016}. The primary purpose +of this section is the instantiation of the locales associated +with the pointwise Kan extensions. \ (*TODO: is the explicit elimination rule necessary?*) lemma cat_ordinal_2_is_arrE: assumes "f : a \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> b" - obtains "f = [0, 0]\<^sub>\" and " a = 0" and "b = 0" + obtains "f = [0, 0]\<^sub>\" and "a = 0" and "b = 0" | "f = [0, 1\<^sub>\]\<^sub>\" and "a = 0" and "b = 1\<^sub>\" | "f = [1\<^sub>\, 1\<^sub>\]\<^sub>\" and "a = 1\<^sub>\" and "b = 1\<^sub>\" using cat_ordinal_is_arrD[OF assms] unfolding two by auto (*TODO: is the explicit elimination rule necessary?*) lemma cat_ordinal_3_is_arrE: assumes "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" obtains "f = [0, 0]\<^sub>\" and " a = 0" and "b = 0" | "f = [0, 1\<^sub>\]\<^sub>\" and "a = 0" and "b = 1\<^sub>\" | "f = [0, 2\<^sub>\]\<^sub>\" and "a = 0" and "b = 2\<^sub>\" | "f = [1\<^sub>\, 1\<^sub>\]\<^sub>\" and "a = 1\<^sub>\" and "b = 1\<^sub>\" | "f = [1\<^sub>\, 2\<^sub>\]\<^sub>\" and "a = 1\<^sub>\" and "b = 2\<^sub>\" | "f = [2\<^sub>\, 2\<^sub>\]\<^sub>\" and "a = 2\<^sub>\" and "b = 2\<^sub>\" using cat_ordinal_is_arrD[OF assms] unfolding three by auto lemma 0123: "0 \\<^sub>\ 2\<^sub>\" "1\<^sub>\ \\<^sub>\ 2\<^sub>\" "0 \\<^sub>\ 3\<^sub>\" "1\<^sub>\ \\<^sub>\ 3\<^sub>\" "2\<^sub>\ \\<^sub>\ 3\<^sub>\" by auto subsection\\\23\\ subsubsection\Definition and elementary properties\ definition \23 :: V where "\23 = [ (\a\\<^sub>\cat_ordinal (2\<^sub>\)\Obj\. if a = 0 then 0 else 2\<^sub>\), ( \f\\<^sub>\cat_ordinal (2\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ [0, 0]\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ [0, 2\<^sub>\]\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ [2\<^sub>\, 2\<^sub>\]\<^sub>\ | otherwise \ 0 ), cat_ordinal (2\<^sub>\), cat_ordinal (3\<^sub>\) ]\<^sub>\" text\Components.\ lemma \23_components: shows "\23\ObjMap\ = (\a\\<^sub>\cat_ordinal (2\<^sub>\)\Obj\. if a = 0 then 0 else 2\<^sub>\)" and "\23\ArrMap\ = ( \f\\<^sub>\cat_ordinal (2\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ [0, 0]\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ [0, 2\<^sub>\]\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ [2\<^sub>\, 2\<^sub>\]\<^sub>\ | otherwise \ 0 )" and [cat_Kan_cs_simps]: "\23\HomDom\ = cat_ordinal (2\<^sub>\)" and [cat_Kan_cs_simps]: "\23\HomCod\ = cat_ordinal (3\<^sub>\)" unfolding \23_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda \23_components(1) |vsv \23_ObjMap_vsv[cat_Kan_cs_intros]| |vdomain \23_ObjMap_vdomain[cat_Kan_cs_simps]| |app \23_ObjMap_app| lemma \23_ObjMap_app_0[cat_Kan_cs_simps]: assumes "x = 0" shows "\23\ObjMap\\x\ = 0" by ( cs_concl cs_simp: \23_ObjMap_app cat_ordinal_cs_simps V_cs_simps assms cs_intro: nat_omega_intros ) lemma \23_ObjMap_app_1[cat_Kan_cs_simps]: assumes "x = 1\<^sub>\" shows "\23\ObjMap\\x\ = 2\<^sub>\" by ( cs_concl cs_simp: cat_ordinal_cs_simps V_cs_simps omega_of_set \23_ObjMap_app assms cs_intro: nat_omega_intros V_cs_intros ) subsubsection\Arrow map\ mk_VLambda \23_components(2) |vsv \23_ArrMap_vsv[cat_Kan_cs_intros]| |vdomain \23_ArrMap_vdomain[cat_Kan_cs_simps]| |app \23_ArrMap_app| lemma \23_ArrMap_app_00[cat_Kan_cs_simps]: assumes "f = [0, 0]\<^sub>\" shows "\23\ArrMap\\f\ = [0, 0]\<^sub>\" unfolding assms by ( cs_concl cs_simp: \23_ArrMap_app cat_ordinal_cs_simps V_cs_simps cs_intro: cat_ordinal_cs_intros nat_omega_intros ) lemma \23_ArrMap_app_01[cat_Kan_cs_simps]: assumes "f = [0, 1\<^sub>\]\<^sub>\" shows "\23\ArrMap\\f\ = [0, 2\<^sub>\]\<^sub>\" proof- have "[0, 1\<^sub>\]\<^sub>\ \\<^sub>\ ordinal_arrs (2\<^sub>\)" by ( cs_concl cs_simp: omega_of_set cs_intro: cat_ordinal_cs_intros V_cs_intros nat_omega_intros ) then show ?thesis unfolding assms by (simp add: \23_components cat_ordinal_components) qed lemma \23_ArrMap_app_11[cat_Kan_cs_simps]: assumes "f = [1\<^sub>\, 1\<^sub>\]\<^sub>\" shows "\23\ArrMap\\f\ = [2\<^sub>\, 2\<^sub>\]\<^sub>\" proof- have "[1\<^sub>\, 1\<^sub>\]\<^sub>\ \\<^sub>\ ordinal_arrs (2\<^sub>\)" by ( cs_concl cs_simp: omega_of_set cs_intro: cat_ordinal_cs_intros V_cs_intros nat_omega_intros ) then show ?thesis unfolding assms by (simp add: \23_components cat_ordinal_components) qed subsubsection\\\23\ is a tiny functor\ lemma (in \) \23_is_functor: "\23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_ordinal (3\<^sub>\)" proof- from ord_of_nat_\ interpret cat_ordinal_2: finite_category \ \cat_ordinal (2\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) from ord_of_nat_\ interpret cat_ordinal_3: finite_category \ \cat_ordinal (3\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) show ?thesis proof(intro is_tiny_functorI' is_functorI') show "vfsequence \23" unfolding \23_def by auto show "vcard \23 = 4\<^sub>\" unfolding \23_def by (simp add: nat_omega_simps) show "\\<^sub>\ (\23\ObjMap\) \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" proof ( rule vsv.vsv_vrange_vsubset, unfold cat_Kan_cs_simps cat_ordinal_cs_simps, intro cat_Kan_cs_intros ) fix x assume "x \\<^sub>\ 2\<^sub>\" then consider \x = 0\ | \x = 1\<^sub>\\ unfolding two by auto then show "\23\ObjMap\\x\ \\<^sub>\ 3\<^sub>\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_Kan_cs_simps omega_of_set cs_intro: nat_omega_intros )+ qed show "\23\ArrMap\\f\ : \23\ObjMap\\a\ \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> \23\ObjMap\\b\" if "f : a \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> b" for a b f using that by (elim cat_ordinal_2_is_arrE; simp only:) ( cs_concl cs_simp: omega_of_set cat_Kan_cs_simps cs_intro: nat_omega_intros V_cs_intros cat_ordinal_cs_intros ) show "\23\ArrMap\\g \\<^sub>A\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> f\ = \23\ArrMap\\g\ \\<^sub>A\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> \23\ArrMap\\f\" if "g : b \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> c" and "f : a \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> b" for b c g a f proof- have "0 \\<^sub>\ 3\<^sub>\" "1\<^sub>\ \\<^sub>\ 3\<^sub>\" "2\<^sub>\ \\<^sub>\ 3\<^sub>\" by auto then show ?thesis using that by (elim cat_ordinal_2_is_arrE; simp only:) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cs_intro: V_cs_intros cat_ordinal_cs_intros )+ qed show "\23\ArrMap\\cat_ordinal (2\<^sub>\)\CId\\c\\ = cat_ordinal (3\<^sub>\)\CId\\\23\ObjMap\\c\\" if "c \\<^sub>\ cat_ordinal (2\<^sub>\)\Obj\" for c proof- from that consider \c = 0\ | \c = 1\<^sub>\\ unfolding cat_ordinal_components(1) two by auto then show ?thesis by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: omega_of_set cat_Kan_cs_simps cat_ordinal_cs_simps cs_intro: nat_omega_intros cat_ordinal_cs_intros ) qed qed (auto intro!: cat_cs_intros simp: \23_components) qed lemma (in \) \23_is_functor'[cat_Kan_cs_intros]: assumes "\' = cat_ordinal (2\<^sub>\)" and "\' = cat_ordinal (3\<^sub>\)" shows "\23 : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule \23_is_functor) lemmas [cat_Kan_cs_intros] = \.\23_is_functor' lemma (in \) \23_is_tiny_functor: "\23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> cat_ordinal (3\<^sub>\)" proof- from ord_of_nat_\ interpret cat_ordinal_2: finite_category \ \cat_ordinal (2\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) from ord_of_nat_\ interpret cat_ordinal_3: finite_category \ \cat_ordinal (3\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) show ?thesis by (intro is_tiny_functorI' \23_is_functor) (auto intro!: cat_small_cs_intros) qed lemma (in \) \23_is_tiny_functor'[cat_Kan_cs_intros]: assumes "\' = cat_ordinal (2\<^sub>\)" and "\' = cat_ordinal (3\<^sub>\)" shows "\23 : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>i\<^sub>n\<^sub>y\<^bsub>\\<^esub> \'" unfolding assms by (rule \23_is_tiny_functor) lemmas [cat_Kan_cs_intros] = \.\23_is_tiny_functor' subsection\ \LK23\: the functor associated with the left Kan extension along \<^const>\\23\ \ subsubsection\Definition and elementary properties\ definition LK23 :: "V \ V" where "LK23 \ = [ ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \\ObjMap\\0\ | a = 1\<^sub>\ \ \\ObjMap\\0\ | a = 2\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Obj\ ), ( \f\\<^sub>\cat_ordinal (3\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [1\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [2\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | otherwise \ \\HomCod\\Arr\ ), cat_ordinal (3\<^sub>\), \\HomCod\ ]\<^sub>\" text\Components.\ lemma LK23_components: shows "LK23 \\ObjMap\ = ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \\ObjMap\\0\ | a = 1\<^sub>\ \ \\ObjMap\\0\ | a = 2\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Obj\ )" and "LK23 \\ArrMap\ = ( \f\\<^sub>\cat_ordinal (3\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [1\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [2\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | otherwise \ \\HomCod\\Arr\ )" and "LK23 \\HomDom\ = cat_ordinal (3\<^sub>\)" and "LK23 \\HomCod\ = \\HomCod\" unfolding LK23_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_functor begin lemmas LK23_components' = LK23_components[where \=\, unfolded cat_cs_simps] lemmas [cat_Kan_cs_simps] = LK23_components'(3,4) end lemmas [cat_Kan_cs_simps] = is_functor.LK23_components'(3,4) subsubsection\Object map\ mk_VLambda LK23_components(1) |vsv LK23_ObjMap_vsv[cat_Kan_cs_intros]| |vdomain LK23_ObjMap_vdomain[cat_Kan_cs_simps]| |app LK23_ObjMap_app| lemma LK23_ObjMap_app_0[cat_Kan_cs_simps]: assumes "a = 0" shows "LK23 \\ObjMap\\a\ = \\ObjMap\\0\" unfolding LK23_components assms cat_ordinal_components by simp lemma LK23_ObjMap_app_1[cat_Kan_cs_simps]: assumes "a = 1\<^sub>\" shows "LK23 \\ObjMap\\a\ = \\ObjMap\\0\" unfolding LK23_components assms cat_ordinal_components by simp lemma LK23_ObjMap_app_2[cat_Kan_cs_simps]: assumes "a = 2\<^sub>\" shows "LK23 \\ObjMap\\a\ = \\ObjMap\\1\<^sub>\\" unfolding LK23_components assms cat_ordinal_components by simp subsubsection\Arrow map\ mk_VLambda LK23_components(2) |vsv LK23_ArrMap_vsv[cat_Kan_cs_intros]| |vdomain LK23_ArrMap_vdomain[cat_Kan_cs_simps]| |app LK23_ArrMap_app| lemma LK23_ArrMap_app_00[cat_Kan_cs_simps]: assumes "f = [0, 0]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\0, 0\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by auto qed lemma LK23_ArrMap_app_01[cat_Kan_cs_simps]: assumes "f = [0, 1\<^sub>\]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\0, 0\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by auto qed lemma LK23_ArrMap_app_02[cat_Kan_cs_simps]: assumes "f = [0, 2\<^sub>\]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\0, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by auto qed lemma LK23_ArrMap_app_11[cat_Kan_cs_simps]: assumes "f = [1\<^sub>\, 1\<^sub>\]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\0, 0\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by auto qed lemma LK23_ArrMap_app_12[cat_Kan_cs_simps]: assumes "f = [1\<^sub>\, 2\<^sub>\]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\0, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: omega_of_set cs_intro: nat_omega_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by auto qed lemma LK23_ArrMap_app_22[cat_Kan_cs_simps]: assumes "f = [2\<^sub>\, 2\<^sub>\]\<^sub>\" shows "LK23 \\ArrMap\\f\ = \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: nat_omega_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding LK23_components assms by simp qed subsubsection\\LK23\ is a functor\ lemma cat_LK23_is_functor: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "LK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) from ord_of_nat_\ interpret cat_ordinal_2: finite_category \ \cat_ordinal (2\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) from ord_of_nat_\ interpret cat_ordinal_3: finite_category \ \cat_ordinal (3\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms) show ?thesis proof(intro is_functorI') show "vfsequence (LK23 \)" unfolding LK23_def by auto show "vcard (LK23 \) = 4\<^sub>\" unfolding LK23_def by (simp add: nat_omega_simps) show "\\<^sub>\ (LK23 \\ObjMap\) \\<^sub>\ \\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_Kan_cs_simps) fix x assume prems: "x \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" then consider \x = 0\ | \x = 1\<^sub>\\ | \x = 2\<^sub>\\ unfolding cat_ordinal_cs_simps three by auto then show "LK23 \\ObjMap\\x\ \\<^sub>\ \\Obj\" by cases ( cs_concl cs_simp: cat_Kan_cs_simps cat_ordinal_cs_simps omega_of_set cs_intro: cat_cs_intros nat_omega_intros )+ qed (cs_concl cs_intro: cat_Kan_cs_intros) show "LK23 \\ArrMap\\f\ : LK23 \\ObjMap\\a\ \\<^bsub>\\<^esub> LK23 \\ObjMap\\b\" if "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for a b f proof- from 0123 that show ?thesis by (elim cat_ordinal_3_is_arrE; simp only:) ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ qed show "LK23 \\ArrMap\\g \\<^sub>A\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> f\ = LK23 \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> LK23 \\ArrMap\\f\" if "g : b \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> c" and "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for b c g a f proof- from 0123 that show ?thesis by (elim cat_ordinal_3_is_arrE; simp only:; (solves\simp\)?) (*slow*) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps \.cf_ArrMap_Comp[symmetric] cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ qed show "LK23 \\ArrMap\\cat_ordinal (3\<^sub>\)\CId\\c\\ = \\CId\\LK23 \\ObjMap\\c\\" if "c \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" for c proof- from that consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ unfolding cat_ordinal_components three by auto moreover have "0 \\<^sub>\ 2\<^sub>\" "1\<^sub>\ \\<^sub>\ 2\<^sub>\" "0 \\<^sub>\ 3\<^sub>\" "1\<^sub>\ \\<^sub>\ 3\<^sub>\" "2\<^sub>\ \\<^sub>\ 3\<^sub>\" by auto ultimately show ?thesis by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps is_functor.cf_ObjMap_CId[symmetric] cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ qed qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma cat_LK23_is_functor'[cat_Kan_cs_intros]: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = cat_ordinal (3\<^sub>\)" shows "LK23 \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cat_LK23_is_functor) subsubsection\The fundamental property of \LK23\\ lemma cf_comp_LK23_\23[cat_Kan_cs_simps]: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "LK23 \ \\<^sub>C\<^sub>F \23 = \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret LK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \LK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) show ?thesis proof(rule cf_eqI) show "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms) have ObjMap_dom_lhs: "\\<^sub>\ ((LK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\) = 2\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros ) have ObjMap_dom_rhs: "\\<^sub>\ (\\ObjMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) show "(LK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\ = \\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume prems: "a \\<^sub>\ 2\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ by force then show "(LK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\\a\ = \\ObjMap\\a\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: omega_of_set cat_cs_simps cat_ordinal_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros nat_omega_intros )+ qed (cs_concl cs_simp: cs_intro: cat_cs_intros V_cs_intros)+ have ArrMap_dom_lhs: "\\<^sub>\ ((LK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\) = cat_ordinal (2\<^sub>\)\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have ArrMap_dom_rhs: "\\<^sub>\ (\\ArrMap\) = cat_ordinal (2\<^sub>\)\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(LK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\ = \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume prems: "f \\<^sub>\ cat_ordinal (2\<^sub>\)\Arr\" then obtain a b where "f : a \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> b" by auto then show "(LK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\\f\ = \\ArrMap\\f\" by (elim cat_ordinal_2_is_arrE; simp only:) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros )+ qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros)+ qed (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) qed subsection\ \RK23\: the functor associated with the right Kan extension along \<^const>\\23\ \ subsubsection\Definition and elementary properties\ definition RK23 :: "V \ V" where "RK23 \ = [ ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \\ObjMap\\0\ | a = 1\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | a = 2\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Obj\ ), ( \f\\<^sub>\cat_ordinal (3\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [0, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | f = [2\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | otherwise \ \\HomCod\\Arr\ ), cat_ordinal (3\<^sub>\), \\HomCod\ ]\<^sub>\" text\Components.\ lemma RK23_components: shows "RK23 \\ObjMap\ = ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \\ObjMap\\0\ | a = 1\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | a = 2\<^sub>\ \ \\ObjMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Obj\ )" and "RK23 \\ArrMap\ = ( \f\\<^sub>\cat_ordinal (3\<^sub>\)\Arr\. if f = [0, 0]\<^sub>\ \ \\ArrMap\\0, 0\\<^sub>\ | f = [0, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [0, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\0, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 1\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | f = [1\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | f = [2\<^sub>\, 2\<^sub>\]\<^sub>\ \ \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ | otherwise \ \\HomCod\\Arr\ )" and "RK23 \\HomDom\ = cat_ordinal (3\<^sub>\)" and "RK23 \\HomCod\ = \\HomCod\" unfolding RK23_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_functor begin lemmas RK23_components' = RK23_components[where \=\, unfolded cat_cs_simps] lemmas [cat_Kan_cs_simps] = RK23_components'(3,4) end lemmas [cat_Kan_cs_simps] = is_functor.RK23_components'(3,4) subsubsection\Object map\ mk_VLambda RK23_components(1) |vsv RK23_ObjMap_vsv[cat_Kan_cs_intros]| |vdomain RK23_ObjMap_vdomain[cat_Kan_cs_simps]| |app RK23_ObjMap_app| lemma RK23_ObjMap_app_0[cat_Kan_cs_simps]: assumes "a = 0" shows "RK23 \\ObjMap\\a\ = \\ObjMap\\0\" unfolding RK23_components assms cat_ordinal_components by simp lemma RK23_ObjMap_app_1[cat_Kan_cs_simps]: assumes "a = 1\<^sub>\" shows "RK23 \\ObjMap\\a\ = \\ObjMap\\1\<^sub>\\" unfolding RK23_components assms cat_ordinal_components by simp lemma RK23_ObjMap_app_2[cat_Kan_cs_simps]: assumes "a = 2\<^sub>\" shows "RK23 \\ObjMap\\a\ = \\ObjMap\\1\<^sub>\\" unfolding RK23_components assms cat_ordinal_components by simp subsubsection\Arrow map\ mk_VLambda RK23_components(2) |vsv RK23_ArrMap_vsv[cat_Kan_cs_intros]| |vdomain RK23_ArrMap_vdomain[cat_Kan_cs_simps]| |app RK23_ArrMap_app| lemma RK23_ArrMap_app_00[cat_Kan_cs_simps]: assumes "f = [0, 0]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\0, 0\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by auto qed lemma RK23_ArrMap_app_01[cat_Kan_cs_simps]: assumes "f = [0, 1\<^sub>\]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\0, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by auto qed lemma RK23_ArrMap_app_02[cat_Kan_cs_simps]: assumes "f = [0, 2\<^sub>\]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\0, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by auto qed lemma RK23_ArrMap_app_11[cat_Kan_cs_simps]: assumes "f = [1\<^sub>\, 1\<^sub>\]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: V_cs_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by auto qed lemma RK23_ArrMap_app_12[cat_Kan_cs_simps]: assumes "f = [1\<^sub>\, 2\<^sub>\]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: omega_of_set cs_intro: nat_omega_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by auto qed lemma RK23_ArrMap_app_22[cat_Kan_cs_simps]: assumes "f = [2\<^sub>\, 2\<^sub>\]\<^sub>\" shows "RK23 \\ArrMap\\f\ = \\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\" proof- from 0123 have f: "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cs_intro: nat_omega_intros cat_ordinal_cs_intros cat_cs_intros assms ) then show ?thesis unfolding RK23_components assms by simp qed subsubsection\\RK23\ is a functor\ lemma cat_RK23_is_functor: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) from ord_of_nat_\ interpret cat_ordinal_2: finite_category \ \cat_ordinal (2\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) from ord_of_nat_\ interpret cat_ordinal_3: finite_category \ \cat_ordinal (3\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms) show ?thesis proof(intro is_functorI') show "vfsequence (RK23 \)" unfolding RK23_def by auto show "vcard (RK23 \) = 4\<^sub>\" unfolding RK23_def by (simp add: nat_omega_simps) show "\\<^sub>\ (RK23 \\ObjMap\) \\<^sub>\ \\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_Kan_cs_simps) fix x assume prems: "x \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" then consider \x = 0\ | \x = 1\<^sub>\\ | \x = 2\<^sub>\\ unfolding cat_ordinal_cs_simps three by auto then show "RK23 \\ObjMap\\x\ \\<^sub>\ \\Obj\" by cases ( cs_concl cs_simp: cat_Kan_cs_simps cat_ordinal_cs_simps omega_of_set cs_intro: cat_cs_intros nat_omega_intros )+ qed (cs_concl cs_intro: cat_Kan_cs_intros) show "RK23 \\ArrMap\\f\ : RK23 \\ObjMap\\a\ \\<^bsub>\\<^esub> RK23 \\ObjMap\\b\" if "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for a b f proof- from 0123 that show ?thesis by (elim cat_ordinal_3_is_arrE; simp only:) ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ qed show "RK23 \\ArrMap\\g \\<^sub>A\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> f\ = RK23 \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> RK23 \\ArrMap\\f\" if "g : b \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> c" and "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for b c g a f using 0123 that by (elim cat_ordinal_3_is_arrE; simp only:; (solves\simp\)?) (*slow*) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps \.cf_ArrMap_Comp[symmetric] cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ show "RK23 \\ArrMap\\cat_ordinal (3\<^sub>\)\CId\\c\\ = \\CId\\RK23 \\ObjMap\\c\\" if "c \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" for c proof- from that consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ unfolding cat_ordinal_components three by auto then show ?thesis by (cases, use 0123 in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps is_functor.cf_ObjMap_CId[symmetric] cs_intro: V_cs_intros cat_cs_intros cat_ordinal_cs_intros )+ qed qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma cat_RK23_is_functor'[cat_Kan_cs_intros]: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = cat_ordinal (3\<^sub>\)" shows "RK23 \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cat_RK23_is_functor) subsubsection\The fundamental property of \RK23\\ lemma cf_comp_RK23_\23[cat_Kan_cs_simps]: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "RK23 \ \\<^sub>C\<^sub>F \23 = \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret RK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \RK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) show ?thesis proof(rule cf_eqI) show "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms) have ObjMap_dom_lhs: "\\<^sub>\ ((RK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\) = 2\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros ) have ObjMap_dom_rhs: "\\<^sub>\ (\\ObjMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) show "(RK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\ = \\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume prems: "a \\<^sub>\ 2\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ by force then show "(RK23 \ \\<^sub>C\<^sub>F \23)\ObjMap\\a\ = \\ObjMap\\a\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: omega_of_set cat_cs_simps cat_ordinal_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros nat_omega_intros )+ qed (cs_concl cs_simp: cs_intro: cat_cs_intros V_cs_intros)+ have ArrMap_dom_lhs: "\\<^sub>\ ((RK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\) = cat_ordinal (2\<^sub>\)\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have ArrMap_dom_rhs: "\\<^sub>\ (\\ArrMap\) = cat_ordinal (2\<^sub>\)\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(RK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\ = \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume prems: "f \\<^sub>\ cat_ordinal (2\<^sub>\)\Arr\" then obtain a b where "f : a \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> b" by auto then show "(RK23 \ \\<^sub>C\<^sub>F \23)\ArrMap\\f\ = \\ArrMap\\f\" by (elim cat_ordinal_2_is_arrE; simp only:) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros )+ qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros)+ qed (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) qed subsection\ \RK_\23\: towards the universal property of the right Kan extension along \\23\ \ subsubsection\Definition and elementary properties\ definition RK_\23 :: "V \ V \ V \ V" where "RK_\23 \ \' \' = [ ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \'\NTMap\\0\ | a = 1\<^sub>\ \ \'\NTMap\\1\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \'\ArrMap\\1\<^sub>\, 2\<^sub>\\\<^sub>\ | a = 2\<^sub>\ \ \'\NTMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Arr\ ), \', RK23 \, cat_ordinal (3\<^sub>\), \'\HomCod\ ]\<^sub>\" text\Components.\ lemma RK_\23_components: shows "RK_\23 \ \' \'\NTMap\ = ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \'\NTMap\\0\ | a = 1\<^sub>\ \ \'\NTMap\\1\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \'\ArrMap\\1\<^sub>\, 2\<^sub>\\\<^sub>\ | a = 2\<^sub>\ \ \'\NTMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Arr\ )" and "RK_\23 \ \' \'\NTDom\ = \'" and "RK_\23 \ \' \'\NTCod\ = RK23 \" and "RK_\23 \ \' \'\NTDGDom\ = cat_ordinal (3\<^sub>\)" and "RK_\23 \ \' \'\NTDGCod\ = \'\HomCod\" unfolding RK_\23_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \' \ assumes \': "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule \') interpretation \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule \) lemmas RK_\23_components' = RK_\23_components[where \'=\' and \=\, unfolded cat_cs_simps] lemmas [cat_Kan_cs_simps] = RK_\23_components'(2-5) end subsubsection\Natural transformation map\ mk_VLambda RK_\23_components(1) |vsv RK_\23_NTMap_vsv[cat_Kan_cs_intros]| |vdomain RK_\23_NTMap_vdomain[cat_Kan_cs_simps]| |app RK_\23_NTMap_app| lemma RK_\23_NTMap_app_0[cat_Kan_cs_simps]: assumes "a = 0" shows "RK_\23 \ \' \'\NTMap\\a\ = \'\NTMap\\0\" using assms unfolding RK_\23_components cat_ordinal_cs_simps by simp lemma (in is_functor) RK_\23_NTMap_app_1[cat_Kan_cs_simps]: assumes "a = 1\<^sub>\" shows "RK_\23 \ \' \'\NTMap\\a\ = \'\NTMap\\1\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \'\ArrMap\\1\<^sub>\, 2\<^sub>\\\<^sub>\" using assms unfolding RK_\23_components cat_ordinal_cs_simps cat_cs_simps by simp lemmas [cat_Kan_cs_simps] = is_functor.RK_\23_NTMap_app_1 lemma RK_\23_NTMap_app_2[cat_Kan_cs_simps]: assumes "a = 2\<^sub>\" shows "RK_\23 \ \' \'\NTMap\\a\ = \'\NTMap\\1\<^sub>\\" using assms unfolding RK_\23_components cat_ordinal_cs_simps by simp subsubsection\\RK_\23\ is a natural transformation\ lemma RK_\23_is_ntcf: assumes "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "RK_\23 \ \' \' : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule assms(1)) interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(2)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \ \\' \\<^sub>C\<^sub>F \23\ \ \' by (rule assms(3)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret RK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \RK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) from 0123 have [cat_cs_simps]: "\\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ = \\CId\\\\ObjMap\\1\<^sub>\\\" by ( cs_concl cs_simp: cat_ordinal_cs_simps is_functor.cf_ObjMap_CId[symmetric] cs_intro: cat_cs_intros ) show ?thesis proof(rule is_ntcfI') show "vfsequence (RK_\23 \ \' \')" unfolding RK_\23_def by simp show "vcard (RK_\23 \ \' \') = 5\<^sub>\" unfolding RK_\23_def by (simp_all add: nat_omega_simps) show "RK_\23 \ \' \'\NTMap\\a\ : \'\ObjMap\\a\ \\<^bsub>\\<^esub> RK23 \\ObjMap\\a\" if "a \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" for a proof- from that consider \a = 0\ | \a = 1\<^sub>\\ | \a = 2\<^sub>\\ unfolding cat_ordinal_cs_simps three by auto from this 0123 show ?thesis by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_ordinal_cs_intros cat_Kan_cs_intros nat_omega_intros )+ qed show "RK_\23 \ \' \'\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \'\ArrMap\\f\ = RK23 \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> RK_\23 \ \' \'\NTMap\\a\" if "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for a b f using that 0123 by (elim cat_ordinal_3_is_arrE, use nothing in \simp_all only:\) (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps \'.cf_ArrMap_Comp[symmetric] \'.HomCod.cat_Comp_assoc \'.ntcf_Comp_commute[symmetric] cat_Kan_cs_simps cs_intro: cat_cs_intros cat_ordinal_cs_intros nat_omega_intros )+ qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma RK_\23_is_ntcf'[cat_Kan_cs_intros]: assumes "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \'" and "\' = RK23 \" and "\' = cat_ordinal (3\<^sub>\)" shows "RK_\23 \ \' \' : \' \\<^sub>C\<^sub>F \': \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4-6) by (rule RK_\23_is_ntcf) subsection\The right Kan extension along \\23\\ lemma \23_is_cat_rKe: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id \ : RK23 \ \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret RK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \RK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) from 0123 have [cat_cs_simps]: "\\ArrMap\\1\<^sub>\, 1\<^sub>\\\<^sub>\ = \\CId\\\\ObjMap\\1\<^sub>\\\" by ( cs_concl cs_simp: cat_ordinal_cs_simps is_functor.cf_ObjMap_CId[symmetric] cs_intro: cat_cs_intros ) show ?thesis proof(intro is_cat_rKeI') fix \' \' assume prems: "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule prems(1)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \ \\' \\<^sub>C\<^sub>F \23\ \ \' by (rule prems(2)) interpret RK_\23: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \ \' \RK23 \\ \RK_\23 \ \' \'\ by (intro RK_\23_is_ntcf prems assms) show "\!\. \ : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" proof(intro ex1I conjI; (elim conjE)?) show "RK_\23 \ \' \' : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (intro RK_\23.is_ntcf_axioms) show "\' = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" proof(rule ntcf_eqI) show "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (intro prems) then have dom_lhs: "\\<^sub>\ (\'\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show rhs: "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\'\NTMap\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ 2\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ unfolding two by auto then show "\'\NTMap\\a\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\\a\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: omega_of_set cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros nat_omega_intros )+ qed (use rhs in \cs_concl cs_simp: cs_intro: V_cs_intros cat_cs_intros\)+ qed simp_all fix \ assume prems': "\ : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" "\' = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" interpret \: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \ \' \RK23 \\ \ by (rule prems'(1)) from prems'(2) have "\'\NTMap\\0\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\\0\" by auto then have [cat_cs_simps]: "\'\NTMap\\0\ = \\NTMap\\0\" by ( cs_prems cs_simp: cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros nat_omega_intros ) from prems'(2) have "\'\NTMap\\1\<^sub>\\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\\1\<^sub>\\" by auto then have [cat_cs_simps]: "\'\NTMap\\1\<^sub>\\ = \\NTMap\\2\<^sub>\\" by ( cs_prems cs_simp: omega_of_set cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros nat_omega_intros ) show "\ = RK_\23 \ \' \'" proof(rule ntcf_eqI) show "\ : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (rule prems'(1)) then have dom_lhs: "\\<^sub>\ (\\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) show "RK_\23 \ \' \' : \' \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) then have dom_rhs: "\\<^sub>\ (RK_\23 \ \' \'\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) from 0123 have 013: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 1\<^sub>\" by (cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros) from 0123 have 123: "[1\<^sub>\, 2\<^sub>\]\<^sub>\ : 1\<^sub>\ \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 2\<^sub>\" by (cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros) from \.ntcf_Comp_commute[OF 123] 013 0123 have [symmetric, cat_Kan_cs_simps]: "\\NTMap\\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \'\ArrMap\ \1\<^sub>\, 2\<^sub>\\\<^sub>\ = \\NTMap\\1\<^sub>\\" by ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps RK23_ArrMap_app_12 cs_intro: cat_cs_intros ) show "\\NTMap\ = RK_\23 \ \' \'\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ 3\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ | \a = 2\<^sub>\\ unfolding three by auto then show "\\NTMap\\a\ = RK_\23 \ \' \'\NTMap\\a\" by (cases; use nothing in \simp_all only:\) (cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps)+ qed auto qed simp_all qed qed (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros)+ qed subsection\ \LK_\23\: towards the universal property of the left Kan extension along \\23\ \ subsubsection\Definition and elementary properties\ definition LK_\23 :: "V \ V \ V \ V" where "LK_\23 \ \' \' = [ ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \'\NTMap\\0\ | a = 1\<^sub>\ \ \'\ArrMap\\0, 1\<^sub>\\\<^sub>\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \'\NTMap\\0\ | a = 2\<^sub>\ \ \'\NTMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Arr\ ), LK23 \, \', cat_ordinal (3\<^sub>\), \'\HomCod\ ]\<^sub>\" text\Components.\ lemma LK_\23_components: shows "LK_\23 \ \' \'\NTMap\ = ( \a\\<^sub>\cat_ordinal (3\<^sub>\)\Obj\. if a = 0 \ \'\NTMap\\0\ | a = 1\<^sub>\ \ \'\ArrMap\\0, 1\<^sub>\\\<^sub>\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \'\NTMap\\0\ | a = 2\<^sub>\ \ \'\NTMap\\1\<^sub>\\ | otherwise \ \\HomCod\\Arr\ )" and "LK_\23 \ \' \'\NTDom\ = LK23 \" and "LK_\23 \ \' \'\NTCod\ = \'" and "LK_\23 \ \' \'\NTDGDom\ = cat_ordinal (3\<^sub>\)" and "LK_\23 \ \' \'\NTDGCod\ = \'\HomCod\" unfolding LK_\23_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \' \ assumes \': "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule \') interpretation \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule \) lemmas LK_\23_components' = LK_\23_components[where \'=\' and \=\, unfolded cat_cs_simps] lemmas [cat_Kan_cs_simps] = LK_\23_components'(2-5) end subsubsection\Natural transformation map\ mk_VLambda LK_\23_components(1) |vsv LK_\23_NTMap_vsv[cat_Kan_cs_intros]| |vdomain LK_\23_NTMap_vdomain[cat_Kan_cs_simps]| |app LK_\23_NTMap_app| lemma LK_\23_NTMap_app_0[cat_Kan_cs_simps]: assumes "a = 0" shows "LK_\23 \ \' \'\NTMap\\a\ = \'\NTMap\\0\" using assms unfolding LK_\23_components cat_ordinal_cs_simps by simp lemma (in is_functor) LK_\23_NTMap_app_1[cat_Kan_cs_simps]: assumes "a = 1\<^sub>\" shows "LK_\23 \ \' \'\NTMap\\a\ = \'\ArrMap\\0, 1\<^sub>\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \'\NTMap\\0\" using assms unfolding LK_\23_components cat_ordinal_cs_simps cat_cs_simps by simp lemmas [cat_Kan_cs_simps] = is_functor.LK_\23_NTMap_app_1 lemma LK_\23_NTMap_app_2[cat_Kan_cs_simps]: assumes "a = 2\<^sub>\" shows "LK_\23 \ \' \'\NTMap\\a\ = \'\NTMap\\1\<^sub>\\" using assms unfolding LK_\23_components cat_ordinal_cs_simps by simp subsubsection\\LK_\23\ is a natural transformation\ lemma LK_\23_is_ntcf: assumes "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "LK_\23 \ \' \' : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule assms(1)) interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(2)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \ \ \\' \\<^sub>C\<^sub>F \23\ \' by (rule assms(3)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret LK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \LK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) show ?thesis proof(rule is_ntcfI') show "vfsequence (LK_\23 \ \' \')" unfolding LK_\23_def by simp show "vcard (LK_\23 \ \' \') = 5\<^sub>\" unfolding LK_\23_def by (simp_all add: nat_omega_simps) show "LK_\23 \ \' \'\NTMap\\a\ : LK23 \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\a\" if "a \\<^sub>\ cat_ordinal (3\<^sub>\)\Obj\" for a proof- from that consider \a = 0\ | \a = 1\<^sub>\\ | \a = 2\<^sub>\\ unfolding cat_ordinal_cs_simps three by auto from this 0123 show "LK_\23 \ \' \'\NTMap\\a\ : LK23 \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\a\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_ordinal_cs_intros cat_Kan_cs_intros nat_omega_intros )+ qed show "LK_\23 \ \' \'\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> LK23 \\ArrMap\\f\ = \'\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> LK_\23 \ \' \'\NTMap\\a\" if "f : a \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b" for a b f using that 0123 by (elim cat_ordinal_3_is_arrE, use nothing in \simp_all only:\) (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps \'.cf_ArrMap_Comp[symmetric] \'.HomCod.cat_Comp_assoc[symmetric] \'.ntcf_Comp_commute cat_Kan_cs_simps cs_intro: cat_cs_intros cat_ordinal_cs_intros nat_omega_intros )+ qed ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros )+ qed lemma LK_\23_is_ntcf'[cat_Kan_cs_intros]: assumes "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = LK23 \" and "\' = \'" and "\' = cat_ordinal (3\<^sub>\)" shows "LK_\23 \ \' \' : \' \\<^sub>C\<^sub>F \': \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4-6) by (rule LK_\23_is_ntcf) subsection\The left Kan extension along \\23\\ lemma \23_is_cat_rKe: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> LK23 \ \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) interpret \23: is_functor \ \cat_ordinal (2\<^sub>\)\ \cat_ordinal (3\<^sub>\)\ \\23\ by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) interpret LK23: is_functor \ \cat_ordinal (3\<^sub>\)\ \ \LK23 \\ by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) show ?thesis proof(intro is_cat_lKeI') fix \' \' assume prems: "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" "\' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \ \' by (rule prems(1)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \ \ \\' \\<^sub>C\<^sub>F \23\ \' by (rule prems(2)) interpret LK_\23: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \ \LK23 \\ \' \LK_\23 \ \' \'\ by (intro LK_\23_is_ntcf prems assms) show "\!\. \ : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \ \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" proof(intro ex1I conjI; (elim conjE)?) show "LK_\23 \ \' \' : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (intro LK_\23.is_ntcf_axioms) show "\' = LK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" proof(rule ntcf_eqI) show "\' : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (intro prems) then have dom_lhs: "\\<^sub>\ (\'\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show rhs: "LK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ ((LK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\'\NTMap\ = (LK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ 2\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ unfolding two by auto then show "\'\NTMap\\a\ = (LK_\23 \ \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: omega_of_set cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros nat_omega_intros )+ qed (use rhs in \cs_concl cs_simp: cs_intro: V_cs_intros cat_cs_intros\)+ qed simp_all fix \ assume prems': "\ : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" interpret \: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \ \LK23 \\ \' \ by (rule prems'(1)) from prems'(2) have "\'\NTMap\\0\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\0\" by auto then have [cat_cs_simps]: "\'\NTMap\\0\ = \\NTMap\\0\" by ( cs_prems cs_simp: cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros nat_omega_intros ) from prems'(2) have "\'\NTMap\\1\<^sub>\\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\1\<^sub>\\" by auto then have [cat_cs_simps]: "\'\NTMap\\1\<^sub>\\ = \\NTMap\\2\<^sub>\\" by ( cs_prems cs_simp: omega_of_set cat_Kan_cs_simps cat_cs_simps cat_ordinal_cs_simps cs_intro: cat_cs_intros nat_omega_intros ) show "\ = LK_\23 \ \' \'" proof(rule ntcf_eqI) show "\ : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (rule prems'(1)) then have dom_lhs: "\\<^sub>\ (\\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) show "LK_\23 \ \' \' : LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) then have dom_rhs: "\\<^sub>\ (LK_\23 \ \' \'\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_cs_simps cat_ordinal_cs_simps) from 0123 have 012: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> 1\<^sub>\" by (cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros) from 0123 have 013: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 1\<^sub>\" by (cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros) from 0123 have 00: "[0, 0]\<^sub>\ = (cat_ordinal (2\<^sub>\))\CId\\0\" by (cs_concl cs_simp: cat_ordinal_cs_simps) from \.ntcf_Comp_commute[OF 013] 013 0123 have [symmetric, cat_Kan_cs_simps]: "\\NTMap\\1\<^sub>\\ = \'\ArrMap\\0, 1\<^sub>\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\0\" by ( cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps 00 LK23_ArrMap_app_01 cs_intro: cat_cs_intros cat_ordinal_cs_intros nat_omega_intros ) show "\\NTMap\ = LK_\23 \ \' \'\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ 3\<^sub>\" then consider \a = 0\ | \a = 1\<^sub>\\ | \a = 2\<^sub>\\ unfolding three by auto then show "\\NTMap\\a\ = LK_\23 \ \' \'\NTMap\\a\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros )+ qed auto qed simp_all qed qed (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros)+ qed subsection\Pointwise Kan extensions along \\23\\ lemma \23_is_cat_pw_rKe: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id \ : RK23 \ \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> \ : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) show ?thesis proof(intro is_cat_pw_rKeI \23_is_cat_rKe[OF assms]) fix a assume prems: "a \\<^sub>\ \\Obj\" show "ntcf_id \ : RK23 \ \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> \ : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C cat_Set \)" proof(intro is_cat_rKe_preservesI \23_is_cat_rKe[OF assms]) from prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \) \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C cat_Set \" proof(intro is_cat_rKeI') show "\23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_ordinal (3\<^sub>\)" by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros) from prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cs_intro: cat_cs_intros cat_Kan_cs_intros) from prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) fix \' \' assume prems': "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \cat_Set \\ \' by (rule prems'(1)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \cat_Set \\ \\' \\<^sub>C\<^sub>F \23\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \\ \' by (rule prems'(2)) show "\!\. \ : \' \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \ \ \' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" proof(intro ex1I conjI; (elim conjE)?) have [cat_Kan_cs_simps]: "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ = RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)" proof(rule cf_eqI) from prems show lhs: "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) from prems show rhs: "RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Kan_cs_intros ) from lhs prems have ObjMap_dom_lhs: "\\<^sub>\ ((Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ObjMap\) = 3\<^sub>\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) from rhs prems have ObjMap_dom_rhs: "\\<^sub>\ ((RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \))\ObjMap\) = 3\<^sub>\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros ) show "(Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ObjMap\ = RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix c assume prems'': "c \\<^sub>\ 3\<^sub>\" with 0123 consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ by force from this prems prems'' 0123 show "(Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ObjMap\\c\ = RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)\ObjMap\\c\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_op_simps cat_Kan_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros )+ qed ( use prems in \ cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros \ )+ from lhs prems have ArrMap_dom_lhs: "\\<^sub>\ ((Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ArrMap\) = cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) from rhs prems have ArrMap_dom_rhs: "\\<^sub>\ ((RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \))\ArrMap\) = cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cs_intro: cat_Kan_cs_intros ) show "(Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ArrMap\ = RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume prems'': "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" then obtain a' b' where "f : a' \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b'" by auto from this 0123 prems show "(Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \)\ArrMap\\f\ = RK23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \)\ArrMap\\f\" by (*slow*) ( elim cat_ordinal_3_is_arrE; use nothing in \simp_all only:\ ) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_op_simps cs_intro: cat_ordinal_cs_intros cat_Kan_cs_intros cat_cs_intros nat_omega_intros )+ qed ( use prems in \cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros\ )+ qed simp_all show "RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' : \' \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (intro RK_\23_is_ntcf') (cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_cs_intros)+ show "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" proof(rule ntcf_eqI) show "\' : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (intro prems') then have dom_lhs: "\\<^sub>\ (\'\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) from prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) : \' \\<^sub>C\<^sub>F \23 \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ ( (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) )\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\'\NTMap\ = ( Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) )\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix c assume prems'': "c \\<^sub>\ 2\<^sub>\" then consider \c = 0\ | \c = 1\<^sub>\\ unfolding two by auto from this prems 0123 show "\'\NTMap\\c\ = ( Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) )\NTMap\\c\" - by (cases; use nothing in \simp_all only:\) + by (cases; use nothing in \simp_all only:\) ( - cs_concl + cs_concl cs_simp: cat_Kan_cs_simps - cat_ordinal_cs_simps - cat_cs_simps - cat_op_simps + cat_ordinal_cs_simps + cat_cs_simps + cat_op_simps + RK_\23_NTMap_app_0 cat_Set_components(1) - cs_intro: - cat_Kan_cs_intros - cat_cs_intros - cat_prod_cs_intros + cs_intro: + cat_Kan_cs_intros + cat_cs_intros + cat_prod_cs_intros \.HomCod.cat_Hom_in_Vset )+ qed (cs_concl cs_simp: cs_intro: cat_cs_intros V_cs_intros)+ qed simp_all fix \ assume prems'': "\ : \' \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23)" interpret \: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \cat_Set \\ \' \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \\ \ by (rule prems''(1)) from prems''(2) have "\'\NTMap\\0\ = (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23))\NTMap\\0\" by auto from this prems 0123 have \'_NTMap_app_0: "\'\NTMap\\0\ = \\NTMap\\0\" by ( cs_prems cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_Kan_cs_simps cat_op_simps \23_ObjMap_app_0 cat_Set_components(1) cs_intro: cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros \.HomCod.cat_Hom_in_Vset ) from 0123 have 01: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> 1\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_ordinal_cs_intros nat_omega_intros ) from prems''(2) have "\'\NTMap\\1\<^sub>\\ = ( Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23) )\NTMap\\1\<^sub>\\" by auto from this prems 0123 have \'_NTMap_app_1: "\'\NTMap\\1\<^sub>\\ = \\NTMap\\2\<^sub>\\" by ( cs_prems cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_Kan_cs_simps cat_op_simps \23_ObjMap_app_1 cat_Set_components(1) cs_intro: cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros \.HomCod.cat_Hom_in_Vset ) from 0123 have 012: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> 1\<^sub>\" by ( cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros ) from 0123 have 013: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 1\<^sub>\" by ( cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros ) from 0123 have 123: "[1\<^sub>\, 2\<^sub>\]\<^sub>\ : 1\<^sub>\ \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 2\<^sub>\" by ( cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros ) from 0123 have 11: "[1\<^sub>\, 1\<^sub>\]\<^sub>\ = (cat_ordinal (2\<^sub>\))\CId\\1\<^sub>\\" by (cs_concl cs_simp: cat_ordinal_cs_simps) from \.ntcf_Comp_commute[OF 123] prems 012 013 have [cat_Kan_cs_simps]: "\'\NTMap\\1\<^sub>\\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \'\ArrMap\\1\<^sub>\, 2\<^sub>\\\<^sub>\ = \\NTMap\\1\<^sub>\\" by (*slow*) ( - cs_prems 1 + cs_prems cs_simp: cat_cs_simps cat_Kan_cs_simps \'_NTMap_app_1[symmetric] is_functor.cf_ObjMap_CId RK23_ArrMap_app_12 11 cs_intro: cat_cs_intros nat_omega_intros ) show "\ = RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \'" proof(rule ntcf_eqI) show \: "\ : \' \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (rule prems''(1)) then have dom_lhs: "\\<^sub>\ (\\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \' : \' \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F RK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ (RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \'\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\\NTMap\ = RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \'\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix c assume "c \\<^sub>\ 3\<^sub>\" then consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ unfolding three by auto from this 0123 show "\\NTMap\\c\ = RK_\23 (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F \) \' \'\NTMap\\c\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_Kan_cs_simps \'_NTMap_app_1 \'_NTMap_app_0 )+ qed (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros V_cs_intros)+ qed simp_all qed qed qed qed qed lemma \23_is_cat_pw_lKe: assumes "\ : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^sub>.\<^sub>p\<^sub>w\<^bsub>\\<^esub> LK23 \ \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C \" proof- interpret \: is_functor \ \cat_ordinal (2\<^sub>\)\ \ \ by (rule assms(1)) from ord_of_nat_\ interpret cat_ordinal_3: finite_category \ \cat_ordinal (3\<^sub>\)\ by (cs_concl cs_intro: cat_ordinal_cs_intros) from 0123 have 002: "[0, 0]\<^sub>\ : 0 \\<^bsub>cat_ordinal (2\<^sub>\)\<^esub> 0" by (cs_concl cs_simp: cat_ordinal_cs_simps cs_intro: cat_cs_intros) show ?thesis proof(intro is_cat_pw_lKeI \23_is_cat_rKe assms, unfold cat_op_simps) fix a assume prems: "a \\<^sub>\ \\Obj\" show "op_ntcf (ntcf_id \) : op_cf (LK23 \) \\<^sub>C\<^sub>F op_cf \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat (cat_ordinal (2\<^sub>\)) \\<^sub>C op_cat (cat_ordinal (3\<^sub>\)) \\<^sub>C (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) : op_cat \ \\\<^sub>C cat_Set \)" proof(intro is_cat_rKe_preservesI) show "op_ntcf (ntcf_id \) : op_cf (LK23 \) \\<^sub>C\<^sub>F op_cf \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf \ : op_cat (cat_ordinal (2\<^sub>\)) \\<^sub>C op_cat (cat_ordinal (3\<^sub>\)) \\<^sub>C op_cat \" proof(cs_intro_step cat_op_intros) show "ntcf_id \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> LK23 \ \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C \" by (intro \23_is_cat_rKe assms) qed simp_all from prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cs_intro: cat_cs_intros) have "op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \) \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\<^sub>C cat_ordinal (3\<^sub>\) \\<^sub>C op_cat (cat_Set \)" proof(intro is_cat_lKeI') show "\23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> cat_ordinal (3\<^sub>\)" by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros) from prems show "op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) from prems show "op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) fix \' \' assume prems': "\' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" "\' : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" interpret \': is_functor \ \cat_ordinal (3\<^sub>\)\ \op_cat (cat_Set \)\ \' by (rule prems'(1)) interpret \': is_ntcf \ \cat_ordinal (2\<^sub>\)\ \op_cat (cat_Set \)\ \op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \\ \\' \\<^sub>C\<^sub>F \23\ \' by (rule prems'(2)) note [unfolded cat_op_simps, cat_cs_intros] = \'.ntcf_NTMap_is_arr' \'.cf_ArrMap_is_arr' show "\!\. \ : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \) \ \' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)" proof(intro ex1I conjI; (elim conjE)?) have [cat_Kan_cs_simps]: "op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ = LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)" proof(rule cf_eqI) from prems show lhs: "op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) from prems show rhs: "LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros cat_cs_intros) from lhs prems have ObjMap_dom_lhs: "\\<^sub>\ ((op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ObjMap\) = 3\<^sub>\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) from rhs prems have ObjMap_dom_rhs: "\\<^sub>\ (LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ObjMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "(op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ObjMap\ = LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix c assume prems'': "c \\<^sub>\ 3\<^sub>\" then consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ unfolding three by auto from this prems 0123 show "(op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ObjMap\\c\ = LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ObjMap\\c\" by (cases; use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros )+ qed ( use prems in \ cs_concl cs_simp: cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros \ )+ from lhs prems have ArrMap_dom_lhs: "\\<^sub>\ ((op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ArrMap\) = cat_ordinal (3\<^sub>\)\Arr\" by ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) from rhs prems have ArrMap_dom_rhs: "\\<^sub>\ (LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ArrMap\) = cat_ordinal (3\<^sub>\)\Arr\" by (cs_concl cs_simp: cat_cs_simps) show "(op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ArrMap\ = LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume "f \\<^sub>\ cat_ordinal (3\<^sub>\)\Arr\" then obtain a' b' where f: "f : a' \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> b'" by auto from f prems 0123 002 show "(op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \)\ArrMap\\f\ = LK23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \)\ArrMap\\f\" by (elim cat_ordinal_3_is_arrE, (simp_all only:)?) ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_op_simps cs_intro: cat_ordinal_cs_intros cat_Kan_cs_intros cat_cs_intros cat_op_intros nat_omega_intros )+ qed ( use prems in \ cs_concl cs_simp: cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros\ )+ qed simp_all show "LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_cs_simps cat_Kan_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) show "\' = LK_\23 ( op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ )" proof(rule ntcf_eqI) show lhs: "\' : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by (rule prems'(2)) from lhs have "\\<^sub>\ (\'\NTMap\) = cat_ordinal (2\<^sub>\)\Obj\" by (cs_concl cs_simp: cat_cs_simps) from prems show rhs: "LK_\23 ( op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ ) : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \23 : cat_ordinal (2\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_Kan_cs_simps cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) from lhs have dom_lhs: "\\<^sub>\ (\'\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) from rhs have dom_rhs: "\\<^sub>\ ((LK_\23 ( op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ ))\NTMap\) = 2\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\'\NTMap\ = ( LK_\23 ( op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ ) )\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs cat_ordinal_cs_simps) fix c assume "c \\<^sub>\ 2\<^sub>\" then consider \c = 0\ | \c = 1\<^sub>\\ unfolding two by auto from this prems 0123 show "\'\NTMap\\c\ = ( LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \) )\NTMap\\c\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cat_cs_simps cat_op_simps \23_ObjMap_app_1 \23_ObjMap_app_0 LK_\23_NTMap_app_0 cat_Set_components(1) cs_intro: cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros \.HomCod.cat_Hom_in_Vset )+ qed (cs_concl cs_simp: cs_intro: V_cs_intros cat_cs_intros)+ qed simp_all fix \ assume prems'': "\ : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" "\' = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)" interpret \: is_ntcf \ \cat_ordinal (3\<^sub>\)\ \op_cat (cat_Set \)\ \op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \\ \' \ by (rule prems''(1)) note [cat_Kan_cs_intros] = \.ntcf_NTMap_is_arr'[unfolded cat_op_simps] from prems''(2) have "\'\NTMap\\0\ = ( \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \) )\NTMap\\0\" by simp from this prems 0123 have \'_NTMap_app_0: "\'\NTMap\\0\ = \\NTMap\\0\" by (*very slow*) ( cs_prems cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cat_cs_simps cat_op_simps cat_Set_components(1) cs_intro: cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros \.HomCod.cat_Hom_in_Vset ) from prems''(2) have "\'\NTMap\\1\<^sub>\\ = ( \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \23 \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \) )\NTMap\\1\<^sub>\\" by simp from this prems 0123 have \'_NTMap_app_1: "\'\NTMap\\1\<^sub>\\ = \\NTMap\\2\<^sub>\\" by (*very slow*) ( cs_prems cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cat_cs_simps cat_op_simps cat_Set_components(1) cs_intro: cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros \.HomCod.cat_Hom_in_Vset )+ from 0123 have 013: "[0, 1\<^sub>\]\<^sub>\ : 0 \\<^bsub>cat_ordinal (3\<^sub>\)\<^esub> 1\<^sub>\" by (cs_concl cs_simp: cs_intro: cat_ordinal_cs_intros nat_omega_intros) from 0123 have 00: "[0, 0]\<^sub>\ = (cat_ordinal (2\<^sub>\))\CId\\0\" by (cs_concl cs_simp: cat_ordinal_cs_simps) from \.ntcf_Comp_commute[OF 013] prems 0123 013 have [cat_Kan_cs_simps]: "\\NTMap\\1\<^sub>\\ = \'\NTMap\\0\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \'\ArrMap\\0, 1\<^sub>\\\<^sub>\" by ( cs_prems cs_simp: cat_ordinal_cs_simps cat_Kan_cs_simps cat_cs_simps cat_op_simps LK23_ArrMap_app_01 cs_intro: cat_ordinal_cs_intros cat_Kan_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros nat_omega_intros cs_simp: 00 \'_NTMap_app_0[symmetric] ) show "\ = LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \'" proof(rule ntcf_eqI) show lhs: "\ : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by (rule prems''(1)) show rhs: "LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \' : op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F LK23 \ \\<^sub>C\<^sub>F \' : cat_ordinal (3\<^sub>\) \\\<^sub>C\<^bsub>\\<^esub> op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_Kan_cs_simps cs_intro: cat_Kan_cs_intros cat_cs_intros ) from lhs have dom_lhs: "\\<^sub>\ (\\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) from rhs have dom_rhs: "\\<^sub>\ (LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \'\NTMap\) = 3\<^sub>\" by (cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps) show "\\NTMap\ = LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \'\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix c assume "c \\<^sub>\ 3\<^sub>\" then consider \c = 0\ | \c = 1\<^sub>\\ | \c = 2\<^sub>\\ unfolding three by auto from this 0123 show "\\NTMap\\c\ = LK_\23 (op_cf Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F \) \' \'\NTMap\\c\" by (cases, use nothing in \simp_all only:\) ( cs_concl cs_simp: cat_ordinal_cs_simps cat_cs_simps cat_Kan_cs_simps cat_op_simps \'_NTMap_app_0 LK_\23_NTMap_app_0 \'_NTMap_app_1 cs_intro: cat_ordinal_cs_intros cat_Kan_cs_intros cat_cs_intros cat_op_intros nat_omega_intros )+ qed (cs_concl cs_simp: cs_intro: cat_Kan_cs_intros V_cs_intros)+ qed simp_all qed qed then have "op_ntcf (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf (ntcf_id \)) : op_cf (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F op_cf \) \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> op_cf ((Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F op_cf (LK23 \))) \\<^sub>C\<^sub>F op_cf (op_cf \23) : op_cat (op_cat (cat_ordinal (2\<^sub>\))) \\<^sub>C op_cat (op_cat (cat_ordinal (3\<^sub>\))) \\<^sub>C op_cat (cat_Set \)" by ( cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_Kan_cs_intros cat_op_intros ) from is_cat_lKe.is_cat_rKe_op[OF this] prems show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf (ntcf_id \) : (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F op_cf (LK23 \)) \\<^sub>C\<^sub>F op_cf \23 \\<^sub>C\<^sub>F\<^sub>.\<^sub>r\<^sub>K\<^sub>e\<^bsub>\\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,a) \\<^sub>C\<^sub>F op_cf \ : op_cat (cat_ordinal (2\<^sub>\)) \\<^sub>C op_cat (cat_ordinal (3\<^sub>\)) \\<^sub>C cat_Set \" by ( cs_prems cs_simp: cat_op_simps cs_intro: cat_Kan_cs_intros cat_cs_intros cat_op_intros ) qed qed qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Universal.thy b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Universal.thy --- a/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Universal.thy +++ b/thys/CZH_Universal_Constructions/czh_ucategories/CZH_UCAT_Universal.thy @@ -1,1276 +1,1131 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Universal arrow\ theory CZH_UCAT_Universal imports CZH_UCAT_Introduction CZH_Elementary_Categories.CZH_ECAT_FUNCT CZH_Elementary_Categories.CZH_ECAT_Set CZH_Elementary_Categories.CZH_ECAT_Hom begin subsection\Background\ text\ The following section is based, primarily, on the elements of the content of Chapter III-1 in \cite{mac_lane_categories_2010}. \ subsection\Universal map\ text\ The universal map is a convenience utility that allows treating a part of the definition of the universal arrow as an arrow in the category \Set\. \ subsubsection\Definition and elementary properties\ definition umap_of :: "V \ V \ V \ V \ V \ V" where "umap_of \ c r u d = [ (\f'\\<^sub>\Hom (\\HomDom\) r d. \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> u), Hom (\\HomDom\) r d, Hom (\\HomCod\) c (\\ObjMap\\d\) ]\<^sub>\" definition umap_fo :: "V \ V \ V \ V \ V \ V" where "umap_fo \ c r u d = umap_of (op_cf \) c r u d" text\Components.\ lemma (in is_functor) umap_of_components: assumes "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" (*do not remove*) shows "umap_of \ c r u d\ArrVal\ = (\f'\\<^sub>\Hom \ r d. \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> u)" and "umap_of \ c r u d\ArrDom\ = Hom \ r d" and "umap_of \ c r u d\ArrCod\ = Hom \ c (\\ObjMap\\d\)" unfolding umap_of_def arr_field_simps by (simp_all add: cat_cs_simps nat_omega_simps) lemma (in is_functor) umap_fo_components: assumes "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "umap_fo \ c r u d\ArrVal\ = (\f'\\<^sub>\Hom \ d r. u \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\)" and "umap_fo \ c r u d\ArrDom\ = Hom \ d r" and "umap_fo \ c r u d\ArrCod\ = Hom \ (\\ObjMap\\d\) c" unfolding umap_fo_def is_functor.umap_of_components[ OF is_functor_op, unfolded cat_op_simps, OF assms ] proof(rule vsv_eqI) fix f' assume "f' \\<^sub>\ \\<^sub>\ (\f'\\<^sub>\Hom \ d r. \\ArrMap\\f'\ \\<^sub>A\<^bsub>op_cat \\<^esub> u)" then have f': "f' : d \\<^bsub>\\<^esub> r" by simp then have \f': "\\ArrMap\\f'\ : \\ObjMap\\d\ \\<^bsub>\\<^esub> \\ObjMap\\r\" by (auto intro: cat_cs_intros) from f' show "(\f'\\<^sub>\Hom \ d r. \\ArrMap\\f'\ \\<^sub>A\<^bsub>op_cat \\<^esub> u)\f'\ = (\f'\\<^sub>\Hom \ d r. u \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\)\f'\" by (simp add: HomCod.op_cat_Comp[OF assms \f']) qed simp_all text\Universal maps for the opposite functor.\ lemma (in is_functor) op_umap_of[cat_op_simps]: "umap_of (op_cf \) = umap_fo \" unfolding umap_fo_def by simp lemma (in is_functor) op_umap_fo[cat_op_simps]: "umap_fo (op_cf \) = umap_of \" unfolding umap_fo_def by (simp add: cat_op_simps) lemmas [cat_op_simps] = is_functor.op_umap_of is_functor.op_umap_fo subsubsection\Arrow value\ lemma umap_of_ArrVal_vsv[cat_cs_intros]: "vsv (umap_of \ c r u d\ArrVal\)" unfolding umap_of_def arr_field_simps by (simp add: nat_omega_simps) lemma umap_fo_ArrVal_vsv[cat_cs_intros]: "vsv (umap_fo \ c r u d\ArrVal\)" unfolding umap_fo_def by (rule umap_of_ArrVal_vsv) lemma (in is_functor) umap_of_ArrVal_vdomain: assumes "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "\\<^sub>\ (umap_of \ c r u d\ArrVal\) = Hom \ r d" unfolding umap_of_components[OF assms] by simp lemmas [cat_cs_simps] = is_functor.umap_of_ArrVal_vdomain lemma (in is_functor) umap_fo_ArrVal_vdomain: assumes "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "\\<^sub>\ (umap_fo \ c r u d\ArrVal\) = Hom \ d r" unfolding umap_fo_components[OF assms] by simp lemmas [cat_cs_simps] = is_functor.umap_fo_ArrVal_vdomain lemma (in is_functor) umap_of_ArrVal_app: assumes "f' : r \\<^bsub>\\<^esub> d" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "umap_of \ c r u d\ArrVal\\f'\ = \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> u" using assms(1) unfolding umap_of_components[OF assms(2)] by simp lemmas [cat_cs_simps] = is_functor.umap_of_ArrVal_app lemma (in is_functor) umap_fo_ArrVal_app: assumes "f' : d \\<^bsub>\\<^esub> r" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "umap_fo \ c r u d\ArrVal\\f'\ = u \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\" proof- from assms have "\\ArrMap\\f'\ : \\ObjMap\\d\ \\<^bsub>\\<^esub> \\ObjMap\\r\" by (auto intro: cat_cs_intros) from this assms(2) have \f'[simp]: "\\ArrMap\\f'\ \\<^sub>A\<^bsub>op_cat \\<^esub> u = u \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\" by (simp add: cat_op_simps) from is_functor_axioms is_functor.umap_of_ArrVal_app[ OF is_functor_op, unfolded cat_op_simps, OF assms ] show ?thesis by (simp add: cat_op_simps cat_cs_simps) qed lemmas [cat_cs_simps] = is_functor.umap_fo_ArrVal_app lemma (in is_functor) umap_of_ArrVal_vrange: assumes "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "\\<^sub>\ (umap_of \ c r u d\ArrVal\) \\<^sub>\ Hom \ c (\\ObjMap\\d\)" proof(intro vsubset_antisym vsubsetI) interpret vsv \umap_of \ c r u d\ArrVal\\ unfolding umap_of_components[OF assms] by simp fix g assume "g \\<^sub>\ \\<^sub>\ (umap_of \ c r u d\ArrVal\)" then obtain f' where g_def: "g = umap_of \ c r u d\ArrVal\\f'\" and f': "f' \\<^sub>\ \\<^sub>\ (umap_of \ c r u d\ArrVal\)" unfolding umap_of_components[OF assms] by auto then have f': "f' : r \\<^bsub>\\<^esub> d" unfolding umap_of_ArrVal_vdomain[OF assms] by simp then have \f': "\\ArrMap\\f'\ : \\ObjMap\\r\ \\<^bsub>\\<^esub> \\ObjMap\\d\" by (auto intro!: cat_cs_intros) have g_def: "g = \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> u" unfolding g_def umap_of_ArrVal_app[OF f' assms].. from \f' assms show "g \\<^sub>\ Hom \ c (\\ObjMap\\d\)" unfolding g_def by (auto intro: cat_cs_intros) qed lemma (in is_functor) umap_fo_ArrVal_vrange: assumes "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "\\<^sub>\ (umap_fo \ c r u d\ArrVal\) \\<^sub>\ Hom \ (\\ObjMap\\d\) c" by ( rule is_functor.umap_of_ArrVal_vrange[ OF is_functor_op, unfolded cat_op_simps, OF assms, folded umap_fo_def ] ) subsubsection\Universal map is an arrow in the category \Set\\ lemma (in is_functor) cf_arr_Set_umap_of: assumes "category \ \" and "category \ \" and r: "r \\<^sub>\ \\Obj\" and d: "d \\<^sub>\ \\Obj\" and u: "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "arr_Set \ (umap_of \ c r u d)" proof(intro arr_SetI) interpret HomDom: category \ \ by (rule assms(1)) interpret HomCod: category \ \ by (rule assms(2)) note umap_of_components = umap_of_components[OF u] from u d have c: "c \\<^sub>\ \\Obj\" and \d: "(\\ObjMap\\d\) \\<^sub>\ \\Obj\" by (auto intro: cat_cs_intros) show "vfsequence (umap_of \ c r u d)" unfolding umap_of_def by simp show "vcard (umap_of \ c r u d) = 3\<^sub>\" unfolding umap_of_def by (simp add: nat_omega_simps) from umap_of_ArrVal_vrange[OF u] show "\\<^sub>\ (umap_of \ c r u d\ArrVal\) \\<^sub>\ umap_of \ c r u d\ArrCod\" unfolding umap_of_components by simp from r d show "umap_of \ c r u d\ArrDom\ \\<^sub>\ Vset \" unfolding umap_of_components by (intro HomDom.cat_Hom_in_Vset) from c \d show "umap_of \ c r u d\ArrCod\ \\<^sub>\ Vset \" unfolding umap_of_components by (intro HomCod.cat_Hom_in_Vset) qed (auto simp: umap_of_components[OF u]) lemma (in is_functor) cf_arr_Set_umap_fo: assumes "category \ \" and "category \ \" and r: "r \\<^sub>\ \\Obj\" and d: "d \\<^sub>\ \\Obj\" and u: "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "arr_Set \ (umap_fo \ c r u d)" proof- from assms(1) have \: "category \ (op_cat \)" by (auto intro: cat_cs_intros) from assms(2) have \: "category \ (op_cat \)" by (auto intro: cat_cs_intros) show ?thesis by ( rule is_functor.cf_arr_Set_umap_of[ OF is_functor_op, unfolded cat_op_simps, OF \ \ r d u ] ) qed lemma (in is_functor) cf_umap_of_is_arr: assumes "category \ \" and "category \ \" and "r \\<^sub>\ \\Obj\" and "d \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "umap_of \ c r u d : Hom \ r d \\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\d\)" proof(intro cat_Set_is_arrI) show "arr_Set \ (umap_of \ c r u d)" by (rule cf_arr_Set_umap_of[OF assms]) qed (simp_all add: umap_of_components[OF assms(5)]) lemma (in is_functor) cf_umap_of_is_arr': assumes "category \ \" and "category \ \" and "r \\<^sub>\ \\Obj\" and "d \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" and "A = Hom \ r d" and "B = Hom \ c (\\ObjMap\\d\)" and "\ = cat_Set \" shows "umap_of \ c r u d : A \\<^bsub>\\<^esub> B" using assms(1-5) unfolding assms(6-8) by (rule cf_umap_of_is_arr) lemmas [cat_cs_intros] = is_functor.cf_umap_of_is_arr' lemma (in is_functor) cf_umap_fo_is_arr: assumes "category \ \" and "category \ \" and "r \\<^sub>\ \\Obj\" and "d \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "umap_fo \ c r u d : Hom \ d r \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\d\) c" proof(intro cat_Set_is_arrI) show "arr_Set \ (umap_fo \ c r u d)" by (rule cf_arr_Set_umap_fo[OF assms]) qed (simp_all add: umap_fo_components[OF assms(5)]) lemma (in is_functor) cf_umap_fo_is_arr': assumes "category \ \" and "category \ \" and "r \\<^sub>\ \\Obj\" and "d \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" and "A = Hom \ d r" and "B = Hom \ (\\ObjMap\\d\) c" and "\ = cat_Set \" shows "umap_fo \ c r u d : A \\<^bsub>\\<^esub> B" using assms(1-5) unfolding assms(6-8) by (rule cf_umap_fo_is_arr) lemmas [cat_cs_intros] = is_functor.cf_umap_fo_is_arr' subsection\Universal arrow: definition and elementary properties\ text\See Chapter III-1 in \cite{mac_lane_categories_2010}.\ definition universal_arrow_of :: "V \ V \ V \ V \ bool" where "universal_arrow_of \ c r u \ ( r \\<^sub>\ \\HomDom\\Obj\ \ u : c \\<^bsub>\\HomCod\\<^esub> \\ObjMap\\r\ \ ( \r' u'. r' \\<^sub>\ \\HomDom\\Obj\ \ u' : c \\<^bsub>\\HomCod\\<^esub> \\ObjMap\\r'\ \ (\!f'. f' : r \\<^bsub>\\HomDom\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\f'\) ) )" definition universal_arrow_fo :: "V \ V \ V \ V \ bool" where "universal_arrow_fo \ c r u \ universal_arrow_of (op_cf \) c r u" text\Rules.\ mk_ide (in is_functor) rf universal_arrow_of_def[where \=\, unfolded cf_HomDom cf_HomCod] |intro universal_arrow_ofI| |dest universal_arrow_ofD[dest]| |elim universal_arrow_ofE[elim]| lemma (in is_functor) universal_arrow_foI: assumes "r \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" and "\r' u'. \ r' \\<^sub>\ \\Obj\; u' : \\ObjMap\\r'\ \\<^bsub>\\<^esub> c \ \ \!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = umap_fo \ c r u r'\ArrVal\\f'\" shows "universal_arrow_fo \ c r u" by ( simp add: is_functor.universal_arrow_ofI [ OF is_functor_op, folded universal_arrow_fo_def, unfolded cat_op_simps, OF assms ] ) lemma (in is_functor) universal_arrow_foD[dest]: assumes "universal_arrow_fo \ c r u" shows "r \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" and "\r' u'. \ r' \\<^sub>\ \\Obj\; u' : \\ObjMap\\r'\ \\<^bsub>\\<^esub> c \ \ \!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = umap_fo \ c r u r'\ArrVal\\f'\" by ( auto simp: is_functor.universal_arrow_ofD [ OF is_functor_op, folded universal_arrow_fo_def, unfolded cat_op_simps, OF assms ] ) lemma (in is_functor) universal_arrow_foE[elim]: assumes "universal_arrow_fo \ c r u" obtains "r \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" and "\r' u'. \ r' \\<^sub>\ \\Obj\; u' : \\ObjMap\\r'\ \\<^bsub>\\<^esub> c \ \ \!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = umap_fo \ c r u r'\ArrVal\\f'\" using assms by (auto simp: universal_arrow_foD) text\Elementary properties.\ lemma (in is_functor) op_cf_universal_arrow_of[cat_op_simps]: "universal_arrow_of (op_cf \) c r u \ universal_arrow_fo \ c r u" unfolding universal_arrow_fo_def .. lemma (in is_functor) op_cf_universal_arrow_fo[cat_op_simps]: "universal_arrow_fo (op_cf \) c r u \ universal_arrow_of \ c r u" unfolding universal_arrow_fo_def cat_op_simps .. lemmas (in is_functor) [cat_op_simps] = is_functor.op_cf_universal_arrow_of is_functor.op_cf_universal_arrow_fo subsection\Uniqueness\ text\ The following properties are related to the uniqueness of the universal arrow. These properties can be inferred from the content of Chapter III-1 in \cite{mac_lane_categories_2010}. \ lemma (in is_functor) cf_universal_arrow_of_ex_is_arr_isomorphism: \\The proof is based on the ideas expressed in the proof of Theorem 5.2 in Chapter Introduction in \cite{hungerford_algebra_2003}.\ assumes "universal_arrow_of \ c r u" and "universal_arrow_of \ c r' u'" obtains f where "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" and "u' = umap_of \ c r u r'\ArrVal\\f\" proof- note ua1 = universal_arrow_ofD[OF assms(1)] note ua2 = universal_arrow_ofD[OF assms(2)] from ua1(1) have \r: "\\CId\\r\ : r \\<^bsub>\\<^esub> r" by (auto intro: cat_cs_intros) from ua1(1) have "\\ArrMap\\\\CId\\r\\ = \\CId\\\\ObjMap\\r\\" by (auto intro: cat_cs_intros) with ua1(1,2) have u_def: "u = umap_of \ c r u r\ArrVal\\\\CId\\r\\" unfolding umap_of_ArrVal_app[OF \r ua1(2)] by (auto simp: cat_cs_simps) from ua2(1) have \r': "\\CId\\r'\ : r' \\<^bsub>\\<^esub> r'" by (auto intro: cat_cs_intros) from ua2(1) have "\\ArrMap\\\\CId\\r'\\ = \\CId\\\\ObjMap\\r'\\" by (auto intro: cat_cs_intros) with ua2(1,2) have u'_def: "u' = umap_of \ c r' u' r'\ArrVal\\\\CId\\r'\\" unfolding umap_of_ArrVal_app[OF \r' ua2(2)] by (auto simp: cat_cs_simps) from \r u_def universal_arrow_ofD(3)[OF assms(1) ua1(1,2)] have eq_CId_rI: "\ f' : r \\<^bsub>\\<^esub> r; u = umap_of \ c r u r\ArrVal\\f'\ \ \ f' = \\CId\\r\" for f' by blast from \r' u'_def universal_arrow_ofD(3)[OF assms(2) ua2(1,2)] have eq_CId_r'I: "\ f' : r' \\<^bsub>\\<^esub> r'; u' = umap_of \ c r' u' r'\ArrVal\\f'\ \ \ f' = \\CId\\r'\" for f' by blast from ua1(3)[OF ua2(1,2)] obtain f where f: "f : r \\<^bsub>\\<^esub> r'" and u'_def: "u' = umap_of \ c r u r'\ArrVal\\f\" and "g : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\g\ \ f = g" for g by metis from ua2(3)[OF ua1(1,2)] obtain f' where f': "f' : r' \\<^bsub>\\<^esub> r" and u_def: "u = umap_of \ c r' u' r\ArrVal\\f'\" and "g : r' \\<^bsub>\\<^esub> r \ u = umap_of \ c r' u' r\ArrVal\\g\ \ f' = g" for g by metis have "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" proof(intro is_arr_isomorphismI is_inverseI) show f: "f : r \\<^bsub>\\<^esub> r'" by (rule f) show f': "f' : r' \\<^bsub>\\<^esub> r" by (rule f') show "f : r \\<^bsub>\\<^esub> r'" by (rule f) from f' have \f': "\\ArrMap\\f'\ : \\ObjMap\\r'\ \\<^bsub>\\<^esub> \\ObjMap\\r\" by (auto intro: cat_cs_intros) from f have \f: "\\ArrMap\\f\ : \\ObjMap\\r\ \\<^bsub>\\<^esub> \\ObjMap\\r'\" by (auto intro: cat_cs_intros) note u'_def' = u'_def[symmetric, unfolded umap_of_ArrVal_app[OF f ua1(2)]] and u_def' = u_def[symmetric, unfolded umap_of_ArrVal_app[OF f' ua2(2)]] show "f' \\<^sub>A\<^bsub>\\<^esub> f = \\CId\\r\" proof(rule eq_CId_rI) from f f' show f'f: "f' \\<^sub>A\<^bsub>\\<^esub> f : r \\<^bsub>\\<^esub> r" by (auto intro: cat_cs_intros) from ua1(2) \f' \f show "u = umap_of \ c r u r\ArrVal\\f' \\<^sub>A\<^bsub>\\<^esub> f\" unfolding umap_of_ArrVal_app[OF f'f ua1(2)] cf_ArrMap_Comp[OF f' f] by (simp add: HomCod.cat_Comp_assoc u'_def' u_def') qed show "f \\<^sub>A\<^bsub>\\<^esub> f' = \\CId\\r'\" proof(rule eq_CId_r'I) from f f' show ff': "f \\<^sub>A\<^bsub>\\<^esub> f' : r' \\<^bsub>\\<^esub> r'" by (auto intro: cat_cs_intros) from ua2(2) \f' \f show "u' = umap_of \ c r' u' r'\ArrVal\\f \\<^sub>A\<^bsub>\\<^esub> f'\" unfolding umap_of_ArrVal_app[OF ff' ua2(2)] cf_ArrMap_Comp[OF f f'] by (simp add: HomCod.cat_Comp_assoc u'_def' u_def') qed qed with u'_def that show ?thesis by auto qed lemma (in is_functor) cf_universal_arrow_fo_ex_is_arr_isomorphism: assumes "universal_arrow_fo \ c r u" and "universal_arrow_fo \ c r' u'" obtains f where "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and "u' = umap_fo \ c r u r'\ArrVal\\f\" by ( elim is_functor.cf_universal_arrow_of_ex_is_arr_isomorphism[ OF is_functor_op, unfolded cat_op_simps, OF assms ] ) lemma (in is_functor) cf_universal_arrow_of_unique: assumes "universal_arrow_of \ c r u" and "universal_arrow_of \ c r' u'" shows "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\f'\" proof- note ua1 = universal_arrow_ofD[OF assms(1)] note ua2 = universal_arrow_ofD[OF assms(2)] from ua1(3)[OF ua2(1,2)] show ?thesis . qed lemma (in is_functor) cf_universal_arrow_fo_unique: assumes "universal_arrow_fo \ c r u" and "universal_arrow_fo \ c r' u'" shows "\!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = umap_fo \ c r u r'\ArrVal\\f'\" proof- note ua1 = universal_arrow_foD[OF assms(1)] note ua2 = universal_arrow_foD[OF assms(2)] from ua1(3)[OF ua2(1,2)] show ?thesis . qed lemma (in is_functor) cf_universal_arrow_of_is_arr_isomorphism: assumes "universal_arrow_of \ c r u" and "universal_arrow_of \ c r' u'" and "f : r \\<^bsub>\\<^esub> r'" and "u' = umap_of \ c r u r'\ArrVal\\f\" shows "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" proof- from assms(3,4) cf_universal_arrow_of_unique[OF assms(1,2)] have eq: "g : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\g\ \ f = g" for g by blast from assms(1,2) obtain f' where iso_f': "f' : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r'" and u'_def: "u' = umap_of \ c r u r'\ArrVal\\f'\" by (auto elim: cf_universal_arrow_of_ex_is_arr_isomorphism) then have f': "f' : r \\<^bsub>\\<^esub> r'" by auto from iso_f' show ?thesis unfolding eq[OF f' u'_def, symmetric]. qed lemma (in is_functor) cf_universal_arrow_fo_is_arr_isomorphism: assumes "universal_arrow_fo \ c r u" and "universal_arrow_fo \ c r' u'" and "f : r' \\<^bsub>\\<^esub> r" and "u' = umap_fo \ c r u r'\ArrVal\\f\" shows "f : r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" by ( rule is_functor.cf_universal_arrow_of_is_arr_isomorphism[ OF is_functor_op, unfolded cat_op_simps, OF assms ] ) subsection\Universal natural transformation\ subsubsection\Definition and elementary properties\ text\ The concept of the universal natural transformation is introduced for the statement of the elements of a variant of Proposition 1 in Chapter III-2 in \cite{mac_lane_categories_2010}. \ definition ntcf_ua_of :: "V \ V \ V \ V \ V \ V" where "ntcf_ua_of \ \ c r u = [ (\d\\<^sub>\\\HomDom\\Obj\. umap_of \ c r u d), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomDom\(r,-), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(c,-) \\<^sub>C\<^sub>F \, \\HomDom\, cat_Set \ ]\<^sub>\" definition ntcf_ua_fo :: "V \ V \ V \ V \ V \ V" where "ntcf_ua_fo \ \ c r u = ntcf_ua_of \ (op_cf \) c r u" text\Components.\ lemma ntcf_ua_of_components: shows "ntcf_ua_of \ \ c r u\NTMap\ = (\d\\<^sub>\\\HomDom\\Obj\. umap_of \ c r u d)" and "ntcf_ua_of \ \ c r u\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomDom\(r,-)" and "ntcf_ua_of \ \ c r u\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(c,-) \\<^sub>C\<^sub>F \" and "ntcf_ua_of \ \ c r u\NTDGDom\ = \\HomDom\" and "ntcf_ua_of \ \ c r u\NTDGCod\ = cat_Set \" unfolding ntcf_ua_of_def nt_field_simps by (simp_all add: nat_omega_simps) lemma ntcf_ua_fo_components: shows "ntcf_ua_fo \ \ c r u\NTMap\ = (\d\\<^sub>\\\HomDom\\Obj\. umap_fo \ c r u d)" and "ntcf_ua_fo \ \ c r u\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat (\\HomDom\)(r,-)" and "ntcf_ua_fo \ \ c r u\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat (\\HomCod\)(c,-) \\<^sub>C\<^sub>F op_cf \" and "ntcf_ua_fo \ \ c r u\NTDGDom\ = op_cat (\\HomDom\)" and "ntcf_ua_fo \ \ c r u\NTDGCod\ = cat_Set \" unfolding ntcf_ua_fo_def ntcf_ua_of_components umap_fo_def cat_op_simps by simp_all context is_functor begin lemmas ntcf_ua_of_components' = ntcf_ua_of_components[where \=\ and \=\, unfolded cat_cs_simps] lemmas [cat_cs_simps] = ntcf_ua_of_components'(2-5) lemma ntcf_ua_fo_components': assumes "c \\<^sub>\ \\Obj\" and "r \\<^sub>\ \\Obj\" shows "ntcf_ua_fo \ \ c r u\NTMap\ = (\d\\<^sub>\\\Obj\. umap_fo \ c r u d)" and [cat_cs_simps]: "ntcf_ua_fo \ \ c r u\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r)" and [cat_cs_simps]: "ntcf_ua_fo \ \ c r u\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \" and [cat_cs_simps]: "ntcf_ua_fo \ \ c r u\NTDGDom\ = op_cat \" and [cat_cs_simps]: "ntcf_ua_fo \ \ c r u\NTDGCod\ = cat_Set \" unfolding ntcf_ua_fo_components cat_cs_simps HomDom.cat_op_cat_cf_Hom_snd[OF assms(2)] HomCod.cat_op_cat_cf_Hom_snd[OF assms(1)] by simp_all end lemmas [cat_cs_simps] = is_functor.ntcf_ua_of_components'(2-5) is_functor.ntcf_ua_fo_components'(2-5) subsubsection\Natural transformation map\ mk_VLambda (in is_functor) ntcf_ua_of_components(1)[where \=\ and \=\, unfolded cf_HomDom] |vsv ntcf_ua_of_NTMap_vsv| |vdomain ntcf_ua_of_NTMap_vdomain| |app ntcf_ua_of_NTMap_app| context is_functor begin context fixes c r assumes r: "r \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" begin mk_VLambda ntcf_ua_fo_components'(1)[OF c r] |vsv ntcf_ua_fo_NTMap_vsv| |vdomain ntcf_ua_fo_NTMap_vdomain| |app ntcf_ua_fo_NTMap_app| end end lemmas [cat_cs_intros] = is_functor.ntcf_ua_fo_NTMap_vsv is_functor.ntcf_ua_of_NTMap_vsv lemmas [cat_cs_simps] = is_functor.ntcf_ua_fo_NTMap_vdomain is_functor.ntcf_ua_fo_NTMap_app is_functor.ntcf_ua_of_NTMap_vdomain is_functor.ntcf_ua_of_NTMap_app lemma (in is_functor) ntcf_ua_of_NTMap_vrange: assumes "category \ \" and "category \ \" and "r \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "\\<^sub>\ (ntcf_ua_of \ \ c r u\NTMap\) \\<^sub>\ cat_Set \\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold ntcf_ua_of_NTMap_vdomain) show "vsv (ntcf_ua_of \ \ c r u\NTMap\)" by (rule ntcf_ua_of_NTMap_vsv) fix d assume prems: "d \\<^sub>\ \\Obj\" with category_cat_Set is_functor_axioms assms show "ntcf_ua_of \ \ c r u\NTMap\\d\ \\<^sub>\ cat_Set \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed subsubsection\Commutativity of the universal maps and \hom\-functions\ lemma (in is_functor) cf_umap_of_cf_hom_commute: assumes "category \ \" and "category \ \" and "c \\<^sub>\ \\Obj\" and "r \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" and "f : a \\<^bsub>\\<^esub> b" shows "umap_of \ c r u b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\ = cf_hom \ [\\CId\\c\, \\ArrMap\\f\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> umap_of \ c r u a" (is \?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf = ?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a\) proof- from is_functor_axioms category_cat_Set assms(1,2,4-6) have b_rf: "?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf : Hom \ r a \\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\b\)" by (cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros) from is_functor_axioms category_cat_Set assms(1,2,4-6) have cf_a: "?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a : Hom \ r a \\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\b\)" by (cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros) show ?thesis proof(rule arr_Set_eqI[of \]) from b_rf show arr_Set_b_rf: "arr_Set \ (?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf)" by (auto dest: cat_Set_is_arrD(1)) from b_rf have dom_lhs: "\\<^sub>\ ((?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf)\ArrVal\) = Hom \ r a" by (cs_concl cs_simp: cat_cs_simps)+ from cf_a show arr_Set_cf_a: "arr_Set \ (?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a)" by (auto dest: cat_Set_is_arrD(1)) from cf_a have dom_rhs: "\\<^sub>\ ((?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a)\ArrVal\) = Hom \ r a" by (cs_concl cs_simp: cat_cs_simps) show "(?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf)\ArrVal\ = (?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix q assume "q : r \\<^bsub>\\<^esub> a" with is_functor_axioms category_cat_Set assms show "(?uof_b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?rf)\ArrVal\\q\ = (?cf \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_a)\ArrVal\\q\" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (use arr_Set_b_rf arr_Set_cf_a in auto) qed (use b_rf cf_a in \cs_concl cs_simp: cat_cs_simps\)+ qed lemma cf_umap_of_cf_hom_unit_commute: assumes "category \ \" and "category \ \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \
\\\<^sub>C\<^bsub>\\<^esub> \" and "\ : cf_id \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g : c' \\<^bsub>\\<^esub> c" and "f : d \\<^bsub>\
\<^esub> d'" shows "umap_of \ c' (\\ObjMap\\c'\) (\\NTMap\\c'\) d' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \
[\\ArrMap\\g\, f]\<^sub>\ = cf_hom \ [g, \\ArrMap\\f\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> umap_of \ c (\\ObjMap\\c\) (\\NTMap\\c\) d" (is \?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf = ?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd\) proof- interpret \: is_ntcf \ \ \ \cf_id \\ \\ \\<^sub>C\<^sub>F \\ \ by (rule assms(5)) from assms have c'd'_\gf: "?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf : Hom \
(\\ObjMap\\c\) d \\<^bsub>cat_Set \\<^esub> Hom \ c' (\\ObjMap\\d'\)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) then have dom_lhs: "\\<^sub>\ ((?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf)\ArrVal\) = Hom \
(\\ObjMap\\c\) d" by (cs_concl cs_simp: cat_cs_simps) from assms have g\f_cd: "?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd : Hom \
(\\ObjMap\\c\) d \\<^bsub>cat_Set \\<^esub> Hom \ c' (\\ObjMap\\d'\)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) then have dom_rhs: "\\<^sub>\ ((?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd)\ArrVal\) = Hom \
(\\ObjMap\\c\) d" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from c'd'_\gf show arr_Set_c'd'_\gf: "arr_Set \ (?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf)" by (auto dest: cat_Set_is_arrD(1)) from g\f_cd show arr_Set_g\f_cd: "arr_Set \ (?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd)" by (auto dest: cat_Set_is_arrD(1)) show "(?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf)\ArrVal\ = (?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix h assume prems: "h : \\ObjMap\\c\ \\<^bsub>\
\<^esub> d" from \.ntcf_Comp_commute[OF assms(6)] assms have [cat_cs_simps]: "\\NTMap\\c\ \\<^sub>A\<^bsub>\\<^esub> g = \\ArrMap\\\\ArrMap\\g\\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\c'\" by (cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) from assms prems show "(?uof_c'd' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\gf)\ArrVal\\h\ = (?g\f \\<^sub>A\<^bsub>cat_Set \\<^esub> ?uof_cd)\ArrVal\\h\" by ( cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros cs_simp: cat_cs_simps ) qed (use arr_Set_c'd'_\gf arr_Set_g\f_cd in auto) qed (use c'd'_\gf g\f_cd in \cs_concl cs_simp: cat_cs_simps\)+ qed subsubsection\Universal natural transformation is a natural transformation\ lemma (in is_functor) cf_ntcf_ua_of_is_ntcf: assumes "r \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" shows "ntcf_ua_of \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof(intro is_ntcfI') let ?ua = \ntcf_ua_of \ \ c r u\ show "vfsequence (ntcf_ua_of \ \ c r u)" unfolding ntcf_ua_of_def by simp show "vcard ?ua = 5\<^sub>\" unfolding ntcf_ua_of_def by (simp add: nat_omega_simps) from assms(1) show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) from is_functor_axioms assms(2) show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) from is_functor_axioms assms show "\\<^sub>\ (?ua\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "?ua\NTMap\\a\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using is_functor_axioms assms that by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) show "?ua\NTMap\\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f\ = (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ua\NTMap\\a\" if "f : a \\<^bsub>\\<^esub> b" for a b f using is_functor_axioms assms that by ( cs_concl cs_simp: cf_umap_of_cf_hom_commute cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) qed (auto simp: ntcf_ua_of_components cat_cs_simps) lemma (in is_functor) cf_ntcf_ua_fo_is_ntcf: assumes "r \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" shows "ntcf_ua_fo \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- from assms(2) have c: "c \\<^sub>\ \\Obj\" by auto show ?thesis by ( rule is_functor.cf_ntcf_ua_of_is_ntcf [ OF is_functor_op, unfolded cat_op_simps, OF assms(1,2), unfolded HomDom.cat_op_cat_cf_Hom_snd[OF assms(1)] HomCod.cat_op_cat_cf_Hom_snd[OF c] ntcf_ua_fo_def[symmetric] ] ) qed subsubsection\Universal natural transformation and universal arrow\ text\ The lemmas in this subsection correspond to variants of elements of Proposition 1 in Chapter III-2 in \cite{mac_lane_categories_2010}. \ lemma (in is_functor) cf_ntcf_ua_of_is_iso_ntcf: assumes "universal_arrow_of \ c r u" shows "ntcf_ua_of \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- have r: "r \\<^sub>\ \\Obj\" and u: "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" and bij: "\r' u'. \ r' \\<^sub>\ \\Obj\; u' : c \\<^bsub>\\<^esub> \\ObjMap\\r'\ \ \ \!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\f'\" by (auto intro!: universal_arrow_ofD[OF assms(1)]) show ?thesis proof(intro is_iso_ntcfI) show "ntcf_ua_of \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (rule cf_ntcf_ua_of_is_ntcf[OF r u]) fix a assume prems: "a \\<^sub>\ \\Obj\" from is_functor_axioms prems r u have [simp]: "umap_of \ c r u a : Hom \ r a \\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\a\)" by (cs_concl cs_intro: cat_cs_intros) then have dom: "\\<^sub>\ (umap_of \ c r u a\ArrVal\) = Hom \ r a" by (cs_concl cs_simp: cat_cs_simps) have "umap_of \ c r u a : Hom \ r a \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\a\)" proof(intro cat_Set_is_arr_isomorphismI, unfold dom) show umof_a: "v11 (umap_of \ c r u a\ArrVal\)" proof(intro vsv.vsv_valeq_v11I, unfold dom in_Hom_iff) fix g f assume prems': "g : r \\<^bsub>\\<^esub> a" "f : r \\<^bsub>\\<^esub> a" "umap_of \ c r u a\ArrVal\\g\ = umap_of \ c r u a\ArrVal\\f\" from is_functor_axioms r u prems'(1) have \g: "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> u : c \\<^bsub>\\<^esub> \\ObjMap\\a\" by (cs_concl cs_intro: cat_cs_intros) from bij[OF prems \g] have unique: "\ f' : r \\<^bsub>\\<^esub> a; \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> u = umap_of \ c r u a\ArrVal\\f'\ \ \ g = f'" for f' by (metis prems'(1) u umap_of_ArrVal_app) from is_functor_axioms prems'(1,2) u have \g_u: "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> u = umap_of \ c r u a\ArrVal\\f\" by (cs_concl cs_simp: prems'(3)[symmetric] cat_cs_simps) show "g = f" by (rule unique[OF prems'(2) \g_u]) qed (auto simp: cat_cs_simps cat_cs_intros) interpret umof_a: v11 \umap_of \ c r u a\ArrVal\\ by (rule umof_a) show "\\<^sub>\ (umap_of \ c r u a\ArrVal\) = Hom \ c (\\ObjMap\\a\)" proof(intro vsubset_antisym) from u show "\\<^sub>\ (umap_of \ c r u a\ArrVal\) \\<^sub>\ Hom \ c (\\ObjMap\\a\)" by (rule umap_of_ArrVal_vrange) show "Hom \ c (\\ObjMap\\a\) \\<^sub>\ \\<^sub>\ (umap_of \ c r u a\ArrVal\)" proof(rule vsubsetI, unfold in_Hom_iff ) fix f assume prems': "f : c \\<^bsub>\\<^esub> \\ObjMap\\a\" from bij[OF prems prems'] obtain f' where f': "f' : r \\<^bsub>\\<^esub> a" and f_def: "f = umap_of \ c r u a\ArrVal\\f'\" by auto from is_functor_axioms prems prems' u f' have "f' \\<^sub>\ \\<^sub>\ (umap_of \ c r u a\ArrVal\)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from this show "f \\<^sub>\ \\<^sub>\ (umap_of \ c r u a\ArrVal\)" unfolding f_def by (rule umof_a.vsv_vimageI2) qed qed qed simp_all from is_functor_axioms prems r u this show "ntcf_ua_of \ \ c r u\NTMap\\a\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ObjMap\\a\" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) qed qed lemmas [cat_cs_intros] = is_functor.cf_ntcf_ua_of_is_iso_ntcf lemma (in is_functor) cf_ntcf_ua_fo_is_iso_ntcf: assumes "universal_arrow_fo \ c r u" shows "ntcf_ua_fo \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- from universal_arrow_foD[OF assms] have r: "r \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" by auto show ?thesis by ( rule is_functor.cf_ntcf_ua_of_is_iso_ntcf [ OF is_functor_op, unfolded cat_op_simps, OF assms, unfolded HomDom.cat_op_cat_cf_Hom_snd[OF r] HomCod.cat_op_cat_cf_Hom_snd[OF c] ntcf_ua_fo_def[symmetric] ] ) qed lemmas [cat_cs_intros] = is_functor.cf_ntcf_ua_fo_is_iso_ntcf lemma (in is_functor) cf_ua_of_if_ntcf_ua_of_is_iso_ntcf: assumes "r \\<^sub>\ \\Obj\" and "u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" and "ntcf_ua_of \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "universal_arrow_of \ c r u" proof(rule universal_arrow_ofI) interpret ua_of_u: is_iso_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \\ \ntcf_ua_of \ \ c r u\ by (rule assms(3)) fix r' u' assume prems: "r' \\<^sub>\ \\Obj\" "u' : c \\<^bsub>\\<^esub> \\ObjMap\\r'\" have "ntcf_ua_of \ \ c r u\NTMap\\r'\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ObjMap\\r'\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ObjMap\\r'\" by (rule is_iso_ntcf.iso_ntcf_is_arr_isomorphism[OF assms(3) prems(1)]) from this is_functor_axioms assms(1-2) prems have uof_r': "umap_of \ c r u r' : Hom \ r r' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> Hom \ c (\\ObjMap\\r'\)" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) note uof_r' = cat_Set_is_arr_isomorphismD[OF uof_r'] interpret uof_r': v11 \umap_of \ c r u r'\ArrVal\\ by (rule uof_r'(2)) from uof_r'.v11_vrange_ex1_eq[ THEN iffD1, unfolded uof_r'(3,4) in_Hom_iff, OF prems(2) ] show "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r u r'\ArrVal\\f'\" by metis qed (intro assms)+ lemma (in is_functor) cf_ua_fo_if_ntcf_ua_fo_is_iso_ntcf: assumes "r \\<^sub>\ \\Obj\" and "u : \\ObjMap\\r\ \\<^bsub>\\<^esub> c" and "ntcf_ua_fo \ \ c r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "universal_arrow_fo \ c r u" proof- from assms(2) have c: "c \\<^sub>\ \\Obj\" by auto show ?thesis by ( rule is_functor.cf_ua_of_if_ntcf_ua_of_is_iso_ntcf [ OF is_functor_op, unfolded cat_op_simps, OF assms(1,2), unfolded HomDom.cat_op_cat_cf_Hom_snd[OF assms(1)] HomCod.cat_op_cat_cf_Hom_snd[OF c] ntcf_ua_fo_def[symmetric], OF assms(3) ] ) qed lemma (in is_functor) cf_universal_arrow_of_if_is_iso_ntcf: assumes "r \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" and "\ : - Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : - \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" + Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "universal_arrow_of \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\)" (is \universal_arrow_of \ c r ?u\) proof- interpret \: is_iso_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \\ \ by (rule assms(3)) show ?thesis proof(intro universal_arrow_ofI assms) from assms(1,2) show u: "?u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) fix r' u' assume prems: "r' \\<^sub>\ \\Obj\" "u' : c \\<^bsub>\\<^esub> \\ObjMap\\r'\" have \r'_ArrVal_app[symmetric, cat_cs_simps]: "\\NTMap\\r'\\ArrVal\\f'\ = \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\r\\ArrVal\\\\CId\\r\\" if "f' : r \\<^bsub>\\<^esub> r'" for f' proof- have "\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f'\ = (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ArrMap\\f'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\" using that by (intro \.ntcf_Comp_commute) then have "\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f']\<^sub>\ = cf_hom \ [\\CId\\c\, \\ArrMap\\f'\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\" using assms(1,2) that prems by (cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) then have "(\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f']\<^sub>\)\ArrVal\\\\CId\\r\\ = (cf_hom \ [\\CId\\c\, \\ArrMap\\f'\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\)\ArrVal\\\\CId\\r\\" by simp from this assms(1,2) u that show ?thesis by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed show "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r ?u r'\ArrVal\\f'\" proof(intro ex1I conjI; (elim conjE)?) from assms prems show "(\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\ : r \\<^bsub>\\<^esub> r'" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_arrow_cs_intros ) with assms(1,2) prems show "u' = umap_of \ c r ?u r'\ArrVal\\(\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\\" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_arrow_cs_intros cat_cs_intros cat_op_intros ) fix f' assume prems': "f' : r \\<^bsub>\\<^esub> r'" "u' = umap_of \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\) r'\ArrVal\\f'\" from prems'(2,1) assms(1,2) have u'_def: "u' = \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\r\\ArrVal\\\\CId\\r\\" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) from prems' show "f' = (\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\" unfolding u'_def \r'_ArrVal_app[OF prems'(1)] by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros cat_op_intros ) qed qed qed lemma (in is_functor) cf_universal_arrow_fo_if_is_iso_ntcf: assumes "r \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "universal_arrow_fo \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\)" by ( rule is_functor.cf_universal_arrow_of_if_is_iso_ntcf [ OF is_functor_op, unfolded cat_op_simps, OF assms(1,2), unfolded HomDom.cat_op_cat_cf_Hom_snd[OF assms(1)] HomCod.cat_op_cat_cf_Hom_snd[OF assms(2)] ntcf_ua_fo_def[symmetric], OF assms(3) ] ) -lemma (in is_functor) cf_universal_arrow_of_if_is_iso_ntcf_if_ge_Limit: - assumes "\ \" - and "\ \\<^sub>\ \" - and "r \\<^sub>\ \\Obj\" - and "c \\<^sub>\ \\Obj\" - and "\ : - Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \ : - \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" - shows "universal_arrow_of \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\)" - (is \universal_arrow_of \ c r ?u\) -proof- - - interpret \: \ \ by (rule assms(1)) - interpret cat_Set_\\: subcategory \ \cat_Set \\ \cat_Set \\ - by (rule subcategory_cat_Set_cat_Set[OF assms(1,2)]) - interpret \: is_iso_ntcf - \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \\ \ - by (rule assms(5)) - interpret \\: category \ \ - by (rule category.cat_category_if_ge_Limit) - (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ - interpret \\: category \ \ - by (rule category.cat_category_if_ge_Limit) - (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ - interpret \\: is_functor \ \ \ \ - by (rule cf_is_functor_if_ge_Limit) - (use assms(2) in \cs_concl cs_simp: cs_intro: cat_cs_intros\)+ - - show ?thesis - proof(intro universal_arrow_ofI assms) - - from assms(3,4) show u: "?u : c \\<^bsub>\\<^esub> \\ObjMap\\r\" - by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) - fix r' u' assume prems: "r' \\<^sub>\ \\Obj\" "u' : c \\<^bsub>\\<^esub> \\ObjMap\\r'\" - have \r'_ArrVal_app[symmetric, cat_cs_simps]: - "\\NTMap\\r'\\ArrVal\\f'\ = - \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\r\\ArrVal\\\\CId\\r\\" - if "f' : r \\<^bsub>\\<^esub> r'" for f' - proof- - have "\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f'\ = - (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F \)\ArrMap\\f'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\" - using that by (intro \.ntcf_Comp_commute) - then have - "\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f']\<^sub>\ = - cf_hom \ [\\CId\\c\, \\ArrMap\\f'\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\" - using assms(3,4) assms(1,2) that prems - by (cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) - then have - "(\\NTMap\\r'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> - cf_hom \ [\\CId\\r\, f']\<^sub>\)\ArrVal\\\\CId\\r\\ = - (cf_hom \ [\\CId\\c\, \\ArrMap\\f'\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> - \\NTMap\\r\)\ArrVal\\\\CId\\r\\" - by simp - from - this assms(3,4,2) u that HomDom.category_axioms HomCod.category_axioms - show ?thesis - by - ( - cs_prems - cs_simp: cat_cs_simps cat_op_simps - cs_intro: - cat_cs_intros - cat_op_intros - cat_prod_cs_intros - cat_Set_\\.subcat_is_arrD - ) - qed - - show "\!f'. f' : r \\<^bsub>\\<^esub> r' \ u' = umap_of \ c r ?u r'\ArrVal\\f'\" - proof(intro ex1I conjI; (elim conjE)?) - from assms prems HomDom.category_axioms show - "(\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\ : r \\<^bsub>\\<^esub> r'" - by - ( - cs_concl - cs_simp: cat_cs_simps cat_op_simps - cs_intro: cat_cs_intros cat_arrow_cs_intros - ) - with assms(3,4) prems show "u' = - umap_of \ c r ?u r'\ArrVal\\(\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\\" - by - ( - cs_concl - cs_simp: cat_cs_simps cat_op_simps - cs_intro: cat_arrow_cs_intros cat_cs_intros cat_op_intros - ) - fix f' assume prems': - "f' : r \\<^bsub>\\<^esub> r'" - "u' = umap_of \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\) r'\ArrVal\\f'\" - from prems'(2,1) assms(3,4) have u'_def: - "u' = \\ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\r\\ArrVal\\\\CId\\r\\" - by - ( - cs_prems - cs_simp: cat_cs_simps cat_op_simps - cs_intro: cat_cs_intros cat_op_intros - ) - from prems' show "f' = (\\NTMap\\r'\)\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\u'\" - unfolding u'_def \r'_ArrVal_app[OF prems'(1)] - by - ( - cs_concl - cs_simp: cat_cs_simps - cs_intro: cat_arrow_cs_intros cat_cs_intros cat_op_intros - ) - - qed - - qed - -qed - -lemma (in is_functor) cf_universal_arrow_fo_if_is_iso_ntcf_if_ge_Limit: - assumes "\ \" - and "\ \\<^sub>\ \" - and "r \\<^sub>\ \\Obj\" - and "c \\<^sub>\ \\Obj\" - and "\ : - Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c) \\<^sub>C\<^sub>F op_cf \ : - op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" - shows "universal_arrow_fo \ c r (\\NTMap\\r\\ArrVal\\\\CId\\r\\)" -proof- - interpret \: \ \ by (rule assms(1)) - interpret \\: is_functor \ \ \ \ - by (rule cf_is_functor_if_ge_Limit) - (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ - show ?thesis - by - ( - rule is_functor.cf_universal_arrow_of_if_is_iso_ntcf_if_ge_Limit - [ - OF is_functor_op, - OF assms(1,2), - unfolded cat_op_simps, - OF assms(3,4), - unfolded - \\.HomDom.cat_op_cat_cf_Hom_snd[OF assms(3)] - \\.HomCod.cat_op_cat_cf_Hom_snd[OF assms(4)] - ntcf_ua_fo_def[symmetric], - OF assms(5) - ] - ) -qed - text\\newpage\ end \ No newline at end of file