diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Comma.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Comma.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Comma.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Comma.thy @@ -1,3371 +1,4763 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Comma categories\ theory CZH_ECAT_Comma imports CZH_ECAT_NTCF CZH_ECAT_Simple begin subsection\Background\ named_theorems cat_comma_cs_simps named_theorems cat_comma_cs_intros subsection\Comma category\ subsubsection\Definition and elementary properties\ text\ See Exercise 1.3.vi in \cite{riehl_category_2016} or Chapter II-6 in \cite{mac_lane_categories_2010}. \ definition cat_comma_Obj :: "V \ V \ V" where "cat_comma_Obj \ \ \ set { [a, b, f]\<^sub>\ | a b f. a \\<^sub>\ \\HomDom\\Obj\ \ b \\<^sub>\ \\HomDom\\Obj\ \ f : \\ObjMap\\a\ \\<^bsub>\\HomCod\\<^esub> \\ObjMap\\b\ }" lemma small_cat_comma_Obj[simp]: "small { [a, b, f]\<^sub>\ | a b f. a \\<^sub>\ \\Obj\ \ b \\<^sub>\ \\Obj\ \ f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\ }" (is \small ?abfs\) proof- define Q where "Q i = (if i = 0 then \\Obj\ else if i = 1\<^sub>\ then \\Obj\ else \\Arr\)" for i have "?abfs \ elts (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" unfolding Q_def proof ( intro subsetI, unfold mem_Collect_eq, elim exE conjE, intro vproductI; simp only: ) fix a b f show "\\<^sub>\ [a, b, f]\<^sub>\ = set {0, 1\<^sub>\, 2\<^sub>\}" by (simp add: three nat_omega_simps) qed (force simp: nat_omega_simps)+ then show "small ?abfs" by (rule down) qed definition cat_comma_Hom :: "V \ V \ V \ V \ V" - where "cat_comma_Hom \ \ abf a'b'f' \ set + where "cat_comma_Hom \ \ A B \ set { - [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ | g h. - abf \\<^sub>\ cat_comma_Obj \ \ \ - a'b'f' \\<^sub>\ cat_comma_Obj \ \ \ - g : abf\0\ \\<^bsub>\\HomDom\\<^esub> a'b'f'\0\ \ - h : abf\1\<^sub>\\ \\<^bsub>\\HomDom\\<^esub> a'b'f'\1\<^sub>\\ \ - a'b'f'\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \\ArrMap\\g\ = - \\ArrMap\\h\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> abf\2\<^sub>\\ + [A, B, [g, h]\<^sub>\]\<^sub>\ | g h. + A \\<^sub>\ cat_comma_Obj \ \ \ + B \\<^sub>\ cat_comma_Obj \ \ \ + g : A\0\ \\<^bsub>\\HomDom\\<^esub> B\0\ \ + h : A\1\<^sub>\\ \\<^bsub>\\HomDom\\<^esub> B\1\<^sub>\\ \ + B\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> \\ArrMap\\g\ = + \\ArrMap\\h\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> A\2\<^sub>\\ }" lemma small_cat_comma_Hom[simp]: "small { - [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ | g h. - abf \\<^sub>\ cat_comma_Obj \ \ \ - a'b'f' \\<^sub>\ cat_comma_Obj \ \ \ - g : abf\0\ \\<^bsub>\\<^esub> a'b'f'\0\ \ - h : abf\1\<^sub>\\ \\<^bsub>\\<^esub> a'b'f'\1\<^sub>\\ \ - a'b'f'\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> abf\2\<^sub>\\ + [A, B, [g, h]\<^sub>\]\<^sub>\ | g h. + A \\<^sub>\ cat_comma_Obj \ \ \ + B \\<^sub>\ cat_comma_Obj \ \ \ + g : A\0\ \\<^bsub>\\<^esub> B\0\ \ + h : A\1\<^sub>\\ \\<^bsub>\\<^esub> B\1\<^sub>\\ \ + B\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> A\2\<^sub>\\ }" (is \small ?abf_a'b'f'_gh\) proof- define Q where "Q i = ( if i = 0 then cat_comma_Obj \ \ else if i = 1\<^sub>\ then cat_comma_Obj \ \ else \\Arr\ \\<^sub>\ \\Arr\ )" for i have "?abf_a'b'f'_gh \ elts (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" unfolding Q_def proof ( intro subsetI, unfold mem_Collect_eq, elim exE conjE, intro vproductI; simp only: ) fix a b f show "\\<^sub>\ [a, b, f]\<^sub>\ = ZFC_in_HOL.set {0, 1\<^sub>\, 2\<^sub>\}" by (simp add: three nat_omega_simps) qed (force simp : nat_omega_simps)+ then show "small ?abf_a'b'f'_gh" by (rule down) qed definition cat_comma_Arr :: "V \ V \ V" where "cat_comma_Arr \ \ \ ( - \\<^sub>\abf\\<^sub>\cat_comma_Obj \ \. \\<^sub>\a'b'f'\\<^sub>\cat_comma_Obj \ \. - cat_comma_Hom \ \ abf a'b'f' + \\<^sub>\A\\<^sub>\cat_comma_Obj \ \. \\<^sub>\B\\<^sub>\cat_comma_Obj \ \. + cat_comma_Hom \ \ A B )" definition cat_comma_composable :: "V \ V \ V" where "cat_comma_composable \ \ \ set { - [[a'b'f', a''b''f'', g'h']\<^sub>\, [abf, a'b'f', gh]\<^sub>\]\<^sub>\ | - abf a'b'f' a''b''f'' g'h' gh. - [a'b'f', a''b''f'', g'h']\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ \ - [abf, a'b'f', gh]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ + [[B, C, G]\<^sub>\, [A, B, F]\<^sub>\]\<^sub>\ | A B C G F. + [B, C, G]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ \ [A, B, F]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ }" lemma small_cat_comma_composable[simp]: shows "small { - [[a'b'f', a''b''f'', g'h']\<^sub>\, [abf, a'b'f', gh]\<^sub>\]\<^sub>\ | - abf a'b'f' a''b''f'' g'h' gh. - [a'b'f', a''b''f'', g'h']\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ \ - [abf, a'b'f', gh]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ + [[B, C, G]\<^sub>\, [A, B, F]\<^sub>\]\<^sub>\ | A B C G F. + [B, C, G]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ \ [A, B, F]\<^sub>\ \\<^sub>\ cat_comma_Arr \ \ }" (is \small ?S\) proof(rule down) show "?S \ elts (cat_comma_Arr \ \ \\<^sub>\ cat_comma_Arr \ \)" by auto qed definition cat_comma :: "V \ V \ V" (\(_ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F _)\ [1000, 1000] 999) where "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ = [ cat_comma_Obj \ \, cat_comma_Arr \ \, (\F\\<^sub>\cat_comma_Arr \ \. F\0\), (\F\\<^sub>\cat_comma_Arr \ \. F\1\<^sub>\\), ( \GF\\<^sub>\cat_comma_composable \ \. [ GF\1\<^sub>\\\0\, GF\0\\1\<^sub>\\, [ GF\0\\2\<^sub>\\\0\ \\<^sub>A\<^bsub>\\HomDom\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\0\, GF\0\\2\<^sub>\\\1\<^sub>\\ \\<^sub>A\<^bsub>\\HomDom\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\1\<^sub>\\ ]\<^sub>\ ]\<^sub>\ ), ( - \abf\\<^sub>\cat_comma_Obj \ \. - [abf, abf, [\\HomDom\\CId\\abf\0\\, \\HomDom\\CId\\abf\1\<^sub>\\\]\<^sub>\]\<^sub>\ + \A\\<^sub>\cat_comma_Obj \ \. + [A, A, [\\HomDom\\CId\\A\0\\, \\HomDom\\CId\\A\1\<^sub>\\\]\<^sub>\]\<^sub>\ ) ]\<^sub>\" text\Components.\ lemma cat_comma_components: shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ = cat_comma_Obj \ \" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\ = cat_comma_Arr \ \" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\ = (\F\\<^sub>\cat_comma_Arr \ \. F\0\)" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\ = (\F\\<^sub>\cat_comma_Arr \ \. F\1\<^sub>\\)" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Comp\ = ( \GF\\<^sub>\cat_comma_composable \ \. [ GF\1\<^sub>\\\0\, GF\0\\1\<^sub>\\, [ GF\0\\2\<^sub>\\\0\ \\<^sub>A\<^bsub>\\HomDom\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\0\, GF\0\\2\<^sub>\\\1\<^sub>\\ \\<^sub>A\<^bsub>\\HomDom\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\1\<^sub>\\ ]\<^sub>\ ]\<^sub>\ )" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\ = ( - \abf\\<^sub>\cat_comma_Obj \ \. - [abf, abf, [\\HomDom\\CId\\abf\0\\, \\HomDom\\CId\\abf\1\<^sub>\\\]\<^sub>\]\<^sub>\ + \A\\<^sub>\cat_comma_Obj \ \. + [A, A, [\\HomDom\\CId\\A\0\\, \\HomDom\\CId\\A\1\<^sub>\\\]\<^sub>\]\<^sub>\ )" unfolding cat_comma_def dg_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 \) lemma cat_comma_Obj_def': "cat_comma_Obj \ \ \ set { [a, b, f]\<^sub>\ | a b f. a \\<^sub>\ \\Obj\ \ b \\<^sub>\ \\Obj\ \ f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\ }" unfolding cat_comma_Obj_def cat_cs_simps by simp lemma cat_comma_Hom_def': - "cat_comma_Hom \ \ abf a'b'f' \ set + "cat_comma_Hom \ \ A B \ set { - [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ | g h. - abf \\<^sub>\ cat_comma_Obj \ \ \ - a'b'f' \\<^sub>\ cat_comma_Obj \ \ \ - g : abf\0\ \\<^bsub>\\<^esub> a'b'f'\0\ \ - h : abf\1\<^sub>\\ \\<^bsub>\\<^esub> a'b'f'\1\<^sub>\\ \ - a'b'f'\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> abf\2\<^sub>\\ + [A, B, [g, h]\<^sub>\]\<^sub>\ | g h. + A \\<^sub>\ cat_comma_Obj \ \ \ + B \\<^sub>\ cat_comma_Obj \ \ \ + g : A\0\ \\<^bsub>\\<^esub> B\0\ \ + h : A\1\<^sub>\\ \\<^bsub>\\<^esub> B\1\<^sub>\\ \ + B\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> A\2\<^sub>\\ }" unfolding cat_comma_Hom_def cat_cs_simps by simp lemma cat_comma_components': shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ = cat_comma_Obj \ \" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\ = cat_comma_Arr \ \" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\ = (\F\\<^sub>\cat_comma_Arr \ \. F\0\)" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\ = (\F\\<^sub>\cat_comma_Arr \ \. F\1\<^sub>\\)" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Comp\ = ( \GF\\<^sub>\cat_comma_composable \ \. [ GF\1\<^sub>\\\0\, GF\0\\1\<^sub>\\, [ GF\0\\2\<^sub>\\\0\ \\<^sub>A\<^bsub>\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\0\, GF\0\\2\<^sub>\\\1\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> GF\1\<^sub>\\\2\<^sub>\\\1\<^sub>\\ ]\<^sub>\ ]\<^sub>\ )" and "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\ = - (\abf\\<^sub>\cat_comma_Obj \ \. [abf, abf, [\\CId\\abf\0\\, \\CId\\abf\1\<^sub>\\\]\<^sub>\]\<^sub>\)" + (\A\\<^sub>\cat_comma_Obj \ \. [A, A, [\\CId\\A\0\\, \\CId\\A\1\<^sub>\\\]\<^sub>\]\<^sub>\)" unfolding cat_comma_components cat_cs_simps by simp_all end subsubsection\Objects\ lemma cat_comma_ObjI[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "A = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" shows "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" using assms(4-6) unfolding cat_comma_Obj_def'[OF assms(1,2)] assms(3) cat_comma_components by simp lemma cat_comma_ObjD[dest]: assumes "[a, b, f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" using assms unfolding cat_comma_components'[OF assms(2,3)] cat_comma_Obj_def'[OF assms(2,3)] by auto lemma cat_comma_ObjE[elim]: assumes "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains a b f where "A = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" using assms unfolding cat_comma_components'[OF assms(2,3)] cat_comma_Obj_def'[OF assms(2,3)] by auto subsubsection\Arrows\ lemma cat_comma_HomI[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + and "F = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" - shows "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" + shows "F \\<^sub>\ cat_comma_Hom \ \ A B" using assms(1,2,6-10) unfolding cat_comma_Hom_def'[OF assms(1,2)] assms(3-5) by ( intro in_set_CollectI exI conjI small_cat_comma_Hom, unfold cat_comma_components'(1,2)[OF assms(1,2), symmetric], ( cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros )+ ) (clarsimp simp: nat_omega_simps)+ lemma cat_comma_HomE[elim]: - assumes "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" + assumes "F \\<^sub>\ cat_comma_Hom \ \ A B" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains a b f a' b' f' g h - where "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where "F = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" using assms(1) by ( unfold cat_comma_components'[OF assms(2,3)] cat_comma_Hom_def'[OF assms(2,3)], elim in_small_setE; (unfold mem_Collect_eq, elim exE conjE cat_comma_ObjE[OF _ assms(2,3)])?, insert that, all\ (unfold cat_comma_components'(1,2)[OF assms(2,3), symmetric], elim cat_comma_ObjE[OF _ assms(2,3)]) | - \ ) (auto simp: nat_omega_simps) lemma cat_comma_HomD[dest]: - assumes "[[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" + assumes "[[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ A B" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" using assms(1) by (force elim!: cat_comma_HomE[OF _ assms(2,3)])+ lemma cat_comma_ArrI[cat_comma_cs_intros]: - assumes "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + assumes "F \\<^sub>\ cat_comma_Hom \ \ A B" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" shows "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" using assms unfolding cat_comma_components cat_comma_Arr_def by (intro vifunionI) lemma cat_comma_ArrE[elim]: assumes "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - obtains abf a'b'f' - where "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + obtains A B + where "F \\<^sub>\ cat_comma_Hom \ \ A B" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" using assms unfolding cat_comma_components cat_comma_Arr_def by auto lemma cat_comma_ArrD[dest]: - assumes "[abf, a'b'f', gh]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + assumes "[A, B, F]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - shows "[abf, a'b'f', gh]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + shows "[A, B, F]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ A B" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" proof- - from assms obtain abf' a'b'f'' - where "[abf, a'b'f', gh]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ abf' a'b'f''" - and "abf' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f'' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from assms obtain C D + where "[A, B, F]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ C D" + and "C \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "D \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (elim cat_comma_ArrE) - moreover from cat_comma_HomE[OF this(1) assms(2,3)] have - "abf = abf'" and "a'b'f' = a'b'f''" + moreover from cat_comma_HomE[OF this(1) assms(2,3)] have "A = C" and "B = D" by auto - ultimately show "[abf, a'b'f', gh]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + ultimately show "[A, B, F]\<^sub>\ \\<^sub>\ cat_comma_Hom \ \ A B" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto qed subsubsection\Domain\ lemma cat_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\)" unfolding cat_comma_components by simp lemma cat_comma_Dom_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" unfolding cat_comma_components by simp lemma cat_comma_Dom_app[cat_comma_cs_simps]: - assumes "F = [abf, a'b'f', gh]\<^sub>\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\F\ = abf" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\ABF\ = A" using assms(2) unfolding assms(1) cat_comma_components by simp lemma cat_comma_Dom_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\) \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" proof(rule vsv.vsv_vrange_vsubset) - fix F assume "F \\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\)" - then have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" by (cs_prems cs_simp: cat_comma_cs_simps) - then obtain abf a'b'f' - where F: "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and abf: "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and a'b'f': "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + fix ABF assume "ABF \\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\)" + then have "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" by (cs_prems cs_simp: cat_comma_cs_simps) + then obtain A B + where ABF: "ABF \\<^sub>\ cat_comma_Hom \ \ A B" + and A: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and B: "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto from this(1) obtain a b f a' b' f' g h - where "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by (elim cat_comma_HomE[OF _ assms(1,2)]) - from F this abf a'b'f' show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\F\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from ABF this A B show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\ABF\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) qed (auto intro: cat_comma_cs_intros) subsubsection\Codomain\ lemma cat_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\)" unfolding cat_comma_components by simp lemma cat_comma_Cod_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" unfolding cat_comma_components by simp lemma cat_comma_Cod_app[cat_comma_cs_simps]: - assumes "F = [abf, a'b'f', gh]\<^sub>\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\F\ = a'b'f'" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\ABF\ = B" using assms(2) unfolding assms(1) cat_comma_components by (simp add: nat_omega_simps) lemma cat_comma_Cod_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\) \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" proof(rule vsv.vsv_vrange_vsubset) - fix F assume "F \\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\)" - then have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" by (cs_prems cs_simp: cat_comma_cs_simps) - then obtain abf a'b'f' - where F: "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and abf: "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and a'b'f': "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + fix ABF assume "ABF \\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\)" + then have "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" by (cs_prems cs_simp: cat_comma_cs_simps) + then obtain A B + where F: "ABF \\<^sub>\ cat_comma_Hom \ \ A B" + and A: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and B: "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto from this(1) obtain a b f a' b' f' g h - where "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by (elim cat_comma_HomE[OF _ assms(1,2)]) - from F this abf a'b'f' show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\F\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from F this A B show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\ABF\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) qed (auto intro: cat_comma_cs_intros) subsubsection\Arrow with a domain and a codomain\ lemma cat_comma_is_arrI[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "F = [abf, a'b'f', gh]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" - and "gh = [g, h]\<^sub>\" + and "ABF = [A, B, F]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" + and "F = [g, h]\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" - shows "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" + shows "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" proof(intro is_arrI) interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) - from assms(7-11) show "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + from assms(7-11) show "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" unfolding assms(3-6) by ( cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) - with assms(7-11) show - "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\F\ = abf" "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\F\ = a'b'f'" + with assms(7-11) show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\\ABF\ = A" "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\\ABF\ = B" unfolding assms(3-6) by (cs_concl cs_simp: cat_comma_cs_simps)+ qed lemma cat_comma_is_arrD[dest]: assumes "[[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\ : [a, b, f]\<^sub>\ \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> [a', b', f']\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" proof- note F_is_arrD = is_arrD[OF assms(1)] note F_cat_comma_ArrD = cat_comma_ArrD[OF F_is_arrD(1) assms(2,3)] show "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by (intro cat_comma_HomD[OF F_cat_comma_ArrD(1) assms(2,3)])+ qed lemma cat_comma_is_arrE[elim]: - assumes "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" + assumes "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains a b f a' b' f' g h - where "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where "ABF = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" proof- note F_is_arrD = is_arrD[OF assms(1)] - from F_is_arrD(1) obtain abf a'b'f' - where "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from F_is_arrD(1) obtain C D + where "ABF \\<^sub>\ cat_comma_Hom \ \ C D" + and "C \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "D \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto from this(1) obtain a b f a' b' f' g h - where F_def: "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where F_def: "ABF = [C, D, [g, h]\<^sub>\]\<^sub>\" + and "C = [a, b, f]\<^sub>\" + and "D = [a', b', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by (elim cat_comma_HomE[OF _ assms(2,3)]) with that show ?thesis by (metis F_is_arrD(1,2,3) cat_comma_Cod_app cat_comma_Dom_app) qed subsubsection\Composition\ lemma cat_comma_composableI: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "GF = [G, F]\<^sub>\" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" - shows "GF \\<^sub>\ cat_comma_composable \ \" + and "ABCGF = [BCG, ABF]\<^sub>\" + and "BCG : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + shows "ABCGF \\<^sub>\ cat_comma_composable \ \" proof- from assms(1,2,5) obtain a b f a' b' f' gh - where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, gh]\<^sub>\" - and "abf = [a, b, f]\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" + where ABF_def: "ABF = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, gh]\<^sub>\" + and "A = [a, b, f]\<^sub>\" + and "B = [a', b', f']\<^sub>\" by auto with assms(1,2,4) obtain a'' b'' f'' g'h' - where G_def: "G = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, g'h']\<^sub>\" - and "a'b'f' = [a', b', f']\<^sub>\" - and "a''b''f'' = [a'', b'', f'']\<^sub>\" + where BCG_def: "BCG = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, g'h']\<^sub>\" + and "B = [a', b', f']\<^sub>\" + and "C = [a'', b'', f'']\<^sub>\" by auto - from is_arrD(1)[OF assms(4)] have "G \\<^sub>\ cat_comma_Arr \ \" + from is_arrD(1)[OF assms(4)] have "BCG \\<^sub>\ cat_comma_Arr \ \" unfolding cat_comma_components'(2)[OF assms(1,2)]. - moreover from is_arrD(1)[OF assms(5)] have "F \\<^sub>\ cat_comma_Arr \ \" + moreover from is_arrD(1)[OF assms(5)] have "ABF \\<^sub>\ cat_comma_Arr \ \" unfolding cat_comma_components'(2)[OF assms(1,2)]. ultimately show ?thesis - unfolding assms(3) F_def G_def cat_comma_composable_def + unfolding assms(3) ABF_def BCG_def cat_comma_composable_def by simp qed lemma cat_comma_composableE[elim]: - assumes "GF \\<^sub>\ cat_comma_composable \ \" + assumes "ABCGF \\<^sub>\ cat_comma_composable \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - obtains G F abf a'b'f' a''b''f'' - where "GF = [G, F]\<^sub>\" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" + obtains BCG ABF A B C + where "ABCGF = [BCG, ABF]\<^sub>\" + and "BCG : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" proof- - from assms(1) obtain abf a'b'f' a''b''f'' g'h' gh - where GF_def: "GF = [[a'b'f', a''b''f'', g'h']\<^sub>\, [abf, a'b'f', gh]\<^sub>\]\<^sub>\" - and g'h': "[a'b'f', a''b''f'', g'h']\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - and gh: "[abf, a'b'f', gh]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + from assms(1) obtain A B C G F + where ABCGF_def: "ABCGF = [[B, C, G]\<^sub>\, [A, B, F]\<^sub>\]\<^sub>\" + and BCG: "[B, C, G]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + and ABF: "[A, B, F]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" unfolding cat_comma_composable_def by (auto simp: cat_comma_components'[OF assms(2,3)]) - note g'h' = cat_comma_ArrD[OF g'h' assms(2,3)] - and gh = cat_comma_ArrD[OF gh assms(2,3)] - from gh(1) assms(2,3) obtain a b f a' b' f' g h - where "[abf, a'b'f', gh]\<^sub>\ = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and abf_def: "abf = [a, b, f]\<^sub>\" - and a'b'f'_def: "a'b'f' = [a', b', f']\<^sub>\" - and gh_def: "gh = [g, h]\<^sub>\" + note BCG = cat_comma_ArrD[OF BCG assms(2,3)] + and ABF = cat_comma_ArrD[OF ABF assms(2,3)] + from ABF(1) assms(2,3) obtain a b f a' b' f' g h + where "[A, B, F]\<^sub>\ = [A, B, [g, h]\<^sub>\]\<^sub>\" + and A_def: "A = [a, b, f]\<^sub>\" + and B_def: "B = [a', b', f']\<^sub>\" + and F_def: "F = [g, h]\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and [cat_comma_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto - with g'h'(1) assms(2,3) obtain a'' b'' f'' g' h' - where g'h'_def: "[a'b'f', a''b''f'', g'h']\<^sub>\ = [a'b'f', a''b''f'', [g', h']\<^sub>\]\<^sub>\" - and a''b''f''_def: "a''b''f'' = [a'', b'', f'']\<^sub>\" - and g'h'_def: "g'h' = [g', h']\<^sub>\" + with BCG(1) assms(2,3) obtain a'' b'' f'' g' h' + where g'h'_def: "[B, C, G]\<^sub>\ = [B, C, [g', h']\<^sub>\]\<^sub>\" + and C_def: "C = [a'', b'', f'']\<^sub>\" + and G_def: "G = [g', h']\<^sub>\" and g': "g' : a' \\<^bsub>\\<^esub> a''" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" and [cat_comma_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto - from gh_def have "gh = [g, h]\<^sub>\" by simp - from assms(2,3) GF_def g h f f' g' h' f'' have - "[a'b'f', a''b''f'', g'h']\<^sub>\ : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" - unfolding GF_def gh_def g'h'_def abf_def a'b'f'_def a''b''f''_def + from F_def have "F = [g, h]\<^sub>\" by simp + from assms(2,3) g h f f' g' h' f'' have + "[B, C, G]\<^sub>\ : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + unfolding ABCGF_def F_def G_def A_def B_def C_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_is_arrI)+ - moreover from assms(2,3) GF_def g h f f' g' h' f'' have - "[abf, a'b'f', gh]\<^sub>\ : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" - unfolding GF_def gh_def g'h'_def abf_def a'b'f'_def a''b''f''_def + moreover from assms(2,3) g h f f' g' h' f'' have + "[A, B, F]\<^sub>\ : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + unfolding ABCGF_def F_def G_def A_def B_def C_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_is_arrI)+ - ultimately show ?thesis using that GF_def by auto + thm that + ultimately show ?thesis using that ABCGF_def by auto qed lemma cat_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Comp\)" unfolding cat_comma_components by auto lemma cat_comma_Comp_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Comp\) = cat_comma_composable \ \" unfolding cat_comma_components by auto lemma cat_comma_Comp_app[cat_comma_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "G = [a'b'f', a''b''f'', [g', h']\<^sub>\]\<^sub>\" - and "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" - shows "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F = [abf, a''b''f'', [g' \\<^sub>A\<^bsub>\\<^esub> g, h' \\<^sub>A\<^bsub>\\<^esub> h]\<^sub>\]\<^sub>\" + and "G = [B, C, [g', h']\<^sub>\]\<^sub>\" + and "F = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "G : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + shows "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F = [A, C, [g' \\<^sub>A\<^bsub>\\<^esub> g, h' \\<^sub>A\<^bsub>\\<^esub> h]\<^sub>\]\<^sub>\" using assms(1,2,5,6) unfolding cat_comma_components'[OF assms(1,2)] assms(3,4) by (*slow*) ( cs_concl cs_simp: omega_of_set V_cs_simps vfsequence_simps cs_intro: nat_omega_intros V_cs_intros cat_comma_composableI TrueI ) lemma cat_comma_Comp_is_arr[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" - shows "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" + and "BCG : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + shows "BCG \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms(1,2,4) obtain a b f a' b' f' g h - where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" - and abf_def: "abf = [a, b, f]\<^sub>\" - and a'b'f'_def: "a'b'f' = [a', b', f']\<^sub>\" + where ABF_def: "ABF = [[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 g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and [symmetric, cat_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto with assms(1,2,3) obtain a'' b'' f'' g' h' - where G_def: "G = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" - and a''b''f''_def: "a''b''f'' = [a'', b'', f'']\<^sub>\" + where BCG_def: "BCG = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" + and C_def: "C = [a'', b'', f'']\<^sub>\" and g': "g' : a' \\<^bsub>\\<^esub> a''" and h': "h' : b' \\<^bsub>\\<^esub> b''" - and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto (*slow*) from g' have \g': "\\ArrMap\\g'\ : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\a''\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) note [cat_cs_simps] = category.cat_assoc_helper[ where \=\ and h=f'' and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ ] category.cat_assoc_helper[ where \=\ and h=f and g=\\\ArrMap\\h\\ and q=\f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\\ ] from assms(1,2,3,4) g h f f' g' h' f'' show ?thesis - unfolding F_def G_def abf_def a'b'f'_def a''b''f''_def + unfolding ABF_def BCG_def A_def B_def C_def by (intro cat_comma_is_arrI[OF assms(1,2)]) ( - cs_concl + cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros )+ qed subsubsection\Identity\ lemma cat_comma_CId_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\)" unfolding cat_comma_components by simp lemma cat_comma_CId_vdomain[cat_comma_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" unfolding cat_comma_components'[OF assms(1,2)] by simp lemma cat_comma_CId_app[cat_comma_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "A = [a, b ,f]\<^sub>\" and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" shows "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\A\ = [A, A, [\\CId\\a\, \\CId\\b\]\<^sub>\]\<^sub>\" proof- from assms(4)[unfolded assms(3), unfolded cat_comma_components'[OF assms(1,2)]] have "[a, b, f]\<^sub>\ \\<^sub>\ cat_comma_Obj \ \". then show ?thesis unfolding cat_comma_components'(6)[OF assms(1,2)] assms(3) by (simp add: nat_omega_simps) qed subsubsection\\Hom\-set\ lemma cat_comma_Hom: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - shows "Hom (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) abf a'b'f' = cat_comma_Hom \ \ abf a'b'f'" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + shows "Hom (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) A B = cat_comma_Hom \ \ A B" proof(intro vsubset_antisym vsubsetI, unfold in_Hom_iff) - fix F assume "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" - with assms(1,2) show "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" + fix ABF assume "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + with assms(1,2) show "ABF \\<^sub>\ cat_comma_Hom \ \ A B" by (elim cat_comma_is_arrE[OF _ assms(1,2)], intro cat_comma_HomI) force+ next - fix F assume "F \\<^sub>\ cat_comma_Hom \ \ abf a'b'f'" - with assms(1,2) show "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" + fix ABF assume "ABF \\<^sub>\ cat_comma_Hom \ \ A B" + with assms(1,2) show "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by (elim cat_comma_HomE[OF _ assms(1,2)], intro cat_comma_is_arrI) force+ qed subsubsection\Comma category is a category\ lemma category_cat_comma[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "category \ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule categoryI') show "vfsequence (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" unfolding cat_comma_def by auto show "vcard (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) = 6\<^sub>\" unfolding cat_comma_def by (simp add: nat_omega_simps) show "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Dom\) \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (rule cat_comma_Dom_vrange[OF assms]) show "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Cod\) \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (rule cat_comma_Cod_vrange[OF assms]) show "(GF \\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Comp\)) \ (\g f b c a. GF = [g, f]\<^sub>\ \ g : b \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> c \ f : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> b)" for GF proof(intro iffI; (elim exE conjE)?; (simp only: cat_comma_Comp_vdomain)?) assume prems: "GF \\<^sub>\ cat_comma_composable \ \" with assms obtain G F abf a'b'f' a''b''f'' where "GF = [G, F]\<^sub>\" and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a''b''f''" and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" by auto with assms show "\g f b c a. GF = [g, f]\<^sub>\ \ g : b \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> c \ f : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> b" by auto qed (use assms in \cs_concl cs_intro: cat_comma_composableI\) from assms show "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (cs_concl cs_simp: cat_comma_cs_simps) - from assms show "g \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> f : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> c" - if "g : b \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> c" and "f : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> b" - for b c g a f + from assms show "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" + if "G : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + for B C G A F using that by (cs_concl cs_intro: cat_comma_cs_intros) from assms show "H \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F = H \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> (G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F)" if "H : C \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> D" and "G : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for C D H B G A F using assms that proof- from that(3) assms obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^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 g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and [cat_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto with that(2) assms obtain a'' b'' f'' g' h' where G_def: "G = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" and C_def: "C = [a'', b'', f'']\<^sub>\" and g': "g' : a' \\<^bsub>\\<^esub> a''" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto (*slow*) with that(1) assms obtain a''' b''' f''' g'' h'' where H_def: "H = [[a'', b'', f'']\<^sub>\, [a''', b''', f''']\<^sub>\, [g'', h'']\<^sub>\]\<^sub>\" and D_def: "D = [a''', b''', f''']\<^sub>\" and g'': "g'' : a'' \\<^bsub>\\<^esub> a'''" and h'': "h'' : b'' \\<^bsub>\\<^esub> b'''" and f''': "f''' : \\ObjMap\\a'''\ \\<^bsub>\\<^esub> \\ObjMap\\b'''\" and [cat_cs_simps]: "f''' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g''\ = \\ArrMap\\h''\ \\<^sub>A\<^bsub>\\<^esub> f''" by auto (*slow*) - note [cat_cs_simps] = + note [cat_cs_simps] = category.cat_assoc_helper[ - where \=\ - and h=f'' - and g=\\\ArrMap\\g'\\ + where \=\ + and h=f'' + and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ ] category.cat_assoc_helper[ - where \=\ - and h=f'' - and g=\\\ArrMap\\g'\\ + where \=\ + and h=f'' + and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ ] category.cat_assoc_helper[ where \=\ and h=f''' and g=\\\ArrMap\\g''\\ and q=\\\ArrMap\\h''\ \\<^sub>A\<^bsub>\\<^esub> f''\ ] from assms that g h f f' g' h' f'' g'' h'' f''' show ?thesis unfolding F_def G_def H_def A_def B_def C_def D_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) qed - show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\a\ : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a" - if "a \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" for a + show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\A\ : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> A" + if "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" for A using that by (elim cat_comma_ObjE[OF _ assms(1)]; (simp only:)?) ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros )+ - show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\b\ \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> f = f" - if "f : a \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> b" for a b f + show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\B\ \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F = F" + if "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for A B F using that by (elim cat_comma_is_arrE[OF _ assms]; (simp only:)?) ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros )+ - show "f \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\b\ = f" - if "f : b \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> c" for b c f + show "F \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\B\ = F" + if "F : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" for B C F using that by (elim cat_comma_is_arrE[OF _ assms]; (simp only:)?) ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros )+ show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ \\<^sub>\ Vset \" proof(intro vsubsetI, elim cat_comma_ObjE[OF _ assms]) fix F a b f assume prems: "F = [a, b, f]\<^sub>\" "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" from prems(2-4) show "F \\<^sub>\ Vset \" unfolding prems(1) by (cs_concl cs_intro: cat_cs_intros V_cs_intros) qed show "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) a b) \\<^sub>\ Vset \" if "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" and "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" and "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" for A B proof- define A0 where "A0 = \\<^sub>\ (\F\\<^sub>\A. F\0\)" define A1 where "A1 = \\<^sub>\ (\F\\<^sub>\A. F\1\<^sub>\\)" define B0 where "B0 = \\<^sub>\ (\F\\<^sub>\B. F\0\)" define B1 where "B1 = \\<^sub>\ (\F\\<^sub>\B. F\1\<^sub>\\)" define A0B0 where "A0B0 = (\\<^sub>\a\\<^sub>\A0. \\<^sub>\b\\<^sub>\B0. Hom \ a b)" define A1B1 where "A1B1 = (\\<^sub>\a\\<^sub>\A1. \\<^sub>\b\\<^sub>\B1. Hom \ a b)" have A0B0: "A0B0 \\<^sub>\ Vset \" unfolding A0B0_def proof(rule \.HomDom.cat_Hom_vifunion_in_Vset; (intro vsubsetI)?) show "A0 \\<^sub>\ Vset \" unfolding A0_def proof(intro vrange_vprojection_in_VsetI that(3)) fix F assume "F \\<^sub>\ A" with that(1) have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto with assms obtain a b f where F_def: "F = [a, b, f]\<^sub>\" by auto show "vsv F" unfolding F_def by auto show "0 \\<^sub>\ \\<^sub>\ F" unfolding F_def by simp qed auto show "B0 \\<^sub>\ Vset \" unfolding B0_def proof(intro vrange_vprojection_in_VsetI that(4)) fix F assume "F \\<^sub>\ B" with that(2) have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto with assms obtain a b f where F_def: "F = [a, b, f]\<^sub>\" by auto show "vsv F" unfolding F_def by auto show "0 \\<^sub>\ \\<^sub>\ F" unfolding F_def by simp qed auto next fix a assume "a \\<^sub>\ A0" with that(1) obtain F where a_def: "a = F\0\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" unfolding A0_def by force with assms obtain b f where "F = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" by auto then show "a \\<^sub>\ \\Obj\" unfolding a_def by simp next fix a assume "a \\<^sub>\ B0" with that(2) obtain F where a_def: "a = F\0\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" unfolding B0_def by force with assms obtain b f where "F = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" by auto then show "a \\<^sub>\ \\Obj\" unfolding a_def by simp qed have A1B1: "A1B1 \\<^sub>\ Vset \" unfolding A1B1_def proof(rule \.HomDom.cat_Hom_vifunion_in_Vset; (intro vsubsetI)?) show "A1 \\<^sub>\ Vset \" unfolding A1_def proof(intro vrange_vprojection_in_VsetI that(3)) fix F assume "F \\<^sub>\ A" with that(1) have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto with assms obtain a b f where F_def: "F = [a, b, f]\<^sub>\" by auto show "vsv F" unfolding F_def by auto show "1\<^sub>\ \\<^sub>\ \\<^sub>\ F" unfolding F_def by (simp add: nat_omega_simps) qed auto show "B1 \\<^sub>\ Vset \" unfolding B1_def proof(intro vrange_vprojection_in_VsetI that(4)) fix F assume "F \\<^sub>\ B" with that(2) have "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto with assms obtain a b f where F_def: "F = [a, b, f]\<^sub>\" by auto show "vsv F" unfolding F_def by auto show "1\<^sub>\ \\<^sub>\ \\<^sub>\ F" unfolding F_def by (simp add: nat_omega_simps) qed auto next fix b assume "b \\<^sub>\ A1" with that(1) obtain F where b_def: "b = F\1\<^sub>\\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" unfolding A1_def by force with assms obtain a f where "F = [a, b, f]\<^sub>\" and "b \\<^sub>\ \\Obj\" by (auto simp: nat_omega_simps) then show "b \\<^sub>\ \\Obj\" unfolding b_def by simp next fix b assume "b \\<^sub>\ B1" with that(2) obtain F where b_def: "b = F\1\<^sub>\\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" unfolding B1_def by force with assms obtain a f where "F = [a, b, f]\<^sub>\" and "b \\<^sub>\ \\Obj\" by (auto simp: nat_omega_simps) then show "b \\<^sub>\ \\Obj\" unfolding b_def by simp qed define Q where "Q i = (if i = 0 then A else if i = 1\<^sub>\ then B else (A0B0 \\<^sub>\ A1B1))" for i have "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) a b) \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof ( intro vsubsetI, elim vifunionE, unfold in_Hom_iff, intro vproductI ballI ) fix abf a'b'f' F assume prems: "abf \\<^sub>\ A" "a'b'f' \\<^sub>\ B" "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" from prems(3) assms obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" and abf_def: "abf = [a, b, f]\<^sub>\" and a'b'f'_def: "a'b'f' = [a', b', f']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto have gh: "[g, h]\<^sub>\ \\<^sub>\ A0B0 \\<^sub>\ A1B1" unfolding A0B0_def A1B1_def proof ( intro ftimesI2 vifunionI, unfold in_Hom_iff A0_def B0_def A1_def B1_def ) from prems(1) show "a \\<^sub>\ \\<^sub>\ (\F\\<^sub>\A. F\0\)" by (intro vsv.vsv_vimageI2'[where a=abf]) (simp_all add: abf_def) from prems(2) show "a' \\<^sub>\ \\<^sub>\ (\F\\<^sub>\B. F\0\)" by (intro vsv.vsv_vimageI2'[where a=a'b'f']) (simp_all add: a'b'f'_def) from prems(1) show "b \\<^sub>\ \\<^sub>\ (\F\\<^sub>\A. F\1\<^sub>\\)" by (intro vsv.vsv_vimageI2'[where a=abf]) (simp_all add: nat_omega_simps abf_def) from prems(2) show "b' \\<^sub>\ \\<^sub>\ (\F\\<^sub>\B. F\1\<^sub>\\)" by (intro vsv.vsv_vimageI2'[where a=a'b'f']) (simp_all add: nat_omega_simps a'b'f'_def) qed (intro g h)+ show "vsv F" unfolding F_def by auto show "\\<^sub>\ F = set {0, 1\<^sub>\, 2\<^sub>\}" by (simp add: F_def three nat_omega_simps) fix i assume "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" then consider \i = 0\ | \i = 1\<^sub>\\ | \i = 2\<^sub>\\ by auto from this prems show "F\i\ \\<^sub>\ Q i" by cases (simp_all add: F_def Q_def gh abf_def a'b'f'_def nat_omega_simps) 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 \" by (cs_concl cs_intro: V_cs_intros) from A0B0 A1B1 assms(1,2) that(3,4) show "Q i \\<^sub>\ Vset \" if "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" for i by (simp_all add: Q_def Limit_ftimes_in_VsetI nat_omega_simps) qed auto ultimately show "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) a b) \\<^sub>\ Vset \" by auto qed qed (auto simp: cat_comma_cs_simps intro: cat_comma_cs_intros) qed subsubsection\Tiny comma category\ lemma tiny_category_cat_comma[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" shows "tiny_category \ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" proof- interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) note \ = \.is_functor_axioms and \ = \.is_functor_axioms interpret category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) show ?thesis proof(intro tiny_categoryI' category_cat_comma) have vrange_\: "\\<^sub>\ (\\ObjMap\) \\<^sub>\ Vset \" by (simp add: vrange_in_VsetI \.tm_cf_ObjMap_in_Vset) moreover have vrange_\: "\\<^sub>\ (\\ObjMap\) \\<^sub>\ Vset \" by (simp add: vrange_in_VsetI \.tm_cf_ObjMap_in_Vset) ultimately have UU_Hom_in_Vset: "(\\<^sub>\a\\<^sub>\\\<^sub>\ (\\ObjMap\). \\<^sub>\b\\<^sub>\\\<^sub>\ (\\ObjMap\). Hom \ a b) \\<^sub>\ Vset \" using \.cf_ObjMap_vrange \.cf_ObjMap_vrange by (auto intro: \.HomCod.cat_Hom_vifunion_in_Vset) define Q where "Q i = ( if i = 0 then \\Obj\ else if i = 1\<^sub>\ then \\Obj\ else (\\<^sub>\a\\<^sub>\\\<^sub>\ (\\ObjMap\). \\<^sub>\b\\<^sub>\\\<^sub>\ (\\ObjMap\). Hom \ a b) )" for i have "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro vsubsetI) fix A assume "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" then obtain a b f where A_def: "A = [a, b, f]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by (elim cat_comma_ObjE[OF _ \ \]) from f have f: "f \\<^sub>\ (\\<^sub>\a\\<^sub>\\\<^sub>\ (\\ObjMap\). \\<^sub>\b\\<^sub>\\\<^sub>\ (\\ObjMap\). Hom \ a b)" by (intro vifunionI, unfold in_Hom_iff) ( simp_all add: a b \.ObjMap.vsv_vimageI2 \.cf_ObjMap_vdomain \.ObjMap.vsv_vimageI2 \.cf_ObjMap_vdomain ) show "A \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro vproductI, unfold Ball_def; (intro allI impI)?) show "\\<^sub>\ A = set {0, 1\<^sub>\, 2\<^sub>\}" unfolding A_def by (simp add: three nat_omega_simps) fix i assume "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" then consider \i = 0\ | \i = 1\<^sub>\\ | \i = 2\<^sub>\\ by auto from this a b f show "A\i\ \\<^sub>\ Q i" unfolding A_def Q_def by cases (simp_all add: nat_omega_simps) qed (auto simp: A_def) 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 this show "Q i \\<^sub>\ Vset \" if "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" for i using that assms(1,2) UU_Hom_in_Vset by ( simp_all add: Q_def \.HomDom.tiny_cat_Obj_in_Vset \.HomDom.tiny_cat_Obj_in_Vset nat_omega_simps ) qed auto ultimately show [simp]: "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ \\<^sub>\ Vset \" by auto define Q where "Q i = ( if i = 0 then \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ else if i = 1\<^sub>\ then \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\ else \\Arr\ \\<^sub>\ \\Arr\ )" for i have "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\ \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro vsubsetI) fix F assume "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" then obtain abf a'b'f' where F: "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" by (auto intro: is_arrI) with assms obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" and abf_def: "abf = [a, b, f]\<^sub>\" and a'b'f'_def: "a'b'f' = [a', b', f']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto from g h have "[g, h]\<^sub>\ \\<^sub>\ \\Arr\ \\<^sub>\ \\Arr\" by auto show "F \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro vproductI, unfold Ball_def; (intro allI impI)?) show "\\<^sub>\ F = set {0, 1\<^sub>\, 2\<^sub>\}" by (simp add: F_def three nat_omega_simps) fix i assume "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" then consider \i = 0\ | \i = 1\<^sub>\\ | \i = 2\<^sub>\\ by auto from this F g h show "F\i\ \\<^sub>\ Q i" unfolding Q_def F_def abf_def[symmetric] a'b'f'_def[symmetric] by cases (auto simp: nat_omega_simps) qed (auto simp: F_def) qed moreover have "(\\<^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 \" by (simp add: three[symmetric] nat_omega_simps) moreover have "\\Arr\ \\<^sub>\ \\Arr\ \\<^sub>\ Vset \" by ( auto intro!: Limit_ftimes_in_VsetI \.\_\ \_def \.HomDom.tiny_cat_Arr_in_Vset \.HomDom.tiny_cat_Arr_in_Vset ) ultimately show "Q i \\<^sub>\ Vset \" if "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" for i using that assms(1,2) UU_Hom_in_Vset by ( simp_all add: Q_def \.HomDom.tiny_cat_Obj_in_Vset \.HomDom.tiny_cat_Obj_in_Vset nat_omega_simps ) qed auto ultimately show "\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\ \\<^sub>\ Vset \" by auto qed (rule \, rule \) qed +subsection\Opposite comma category functor\ + + +subsubsection\Background\ + + +text\ +See \cite{noauthor_wikipedia_2001}\footnote{ +\url{https://en.wikipedia.org/wiki/Opposite_category} +} for background information. +\ + + +subsubsection\Object flip\ + +definition op_cf_commma_obj_flip :: "V \ V \ V" + where "op_cf_commma_obj_flip \ \ = + (\A\\<^sub>\(\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\. [A\1\<^sub>\\, A\0\, A\2\<^sub>\\]\<^sub>\)" + + +text\Elementary properties.\ + +mk_VLambda op_cf_commma_obj_flip_def + |vsv op_cf_commma_obj_flip_vsv[cat_comma_cs_intros]| + |vdomain op_cf_commma_obj_flip_vdomain[cat_comma_cs_simps]| + |app op_cf_commma_obj_flip_app'| + +lemma op_cf_commma_obj_flip_app[cat_comma_cs_simps]: + assumes "A = [a, b, f]\<^sub>\" and "A \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" + shows "op_cf_commma_obj_flip \ \\A\ = [b, a, f]\<^sub>\" + using assms unfolding op_cf_commma_obj_flip_def by (simp add: nat_omega_simps) + +lemma op_cf_commma_obj_flip_v11[cat_comma_cs_intros]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "v11 (op_cf_commma_obj_flip \ \)" +proof(rule vsv.vsv_valeq_v11I, unfold op_cf_commma_obj_flip_vdomain) + fix A B assume prems: + "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + "op_cf_commma_obj_flip \ \\A\ = op_cf_commma_obj_flip \ \\B\" + from prems(1,2) assms obtain a b f a' b' f' + where A_def: "A = [a, b, f]\<^sub>\" + and B_def: "B = [a', b', f']\<^sub>\" + by (elim cat_comma_ObjE[OF _ assms]) + from prems(3,1,2) show "A = B" + by (simp_all add: A_def B_def op_cf_commma_obj_flip_app nat_omega_simps) +qed (auto intro: op_cf_commma_obj_flip_vsv) + +lemma op_cf_commma_obj_flip_vrange[cat_comma_cs_simps]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "\\<^sub>\ (op_cf_commma_obj_flip \ \) = (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Obj\" +proof(intro vsubset_antisym) + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + show "\\<^sub>\ (op_cf_commma_obj_flip \ \) \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Obj\" + proof + ( + intro vsv.vsv_vrange_vsubset op_cf_commma_obj_flip_vsv, + unfold cat_comma_cs_simps + ) + fix A assume "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + then obtain a b f + where A_def: "A = [a, b, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and b: "b \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + by (elim cat_comma_ObjE[OF _ assms]) + from a b f show + "op_cf_commma_obj_flip \ \\A\ \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Obj\" + unfolding A_def + by + ( + cs_concl + cs_simp: cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + show "(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Obj\ \\<^sub>\ \\<^sub>\ (op_cf_commma_obj_flip \ \)" + proof(intro vsubsetI) + fix B assume "B \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Obj\" + then obtain a b f + where B_def: "B = [b, a, f]\<^sub>\" + and b: "b \\<^sub>\ \\Obj\" + and a: "a \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + by + ( + elim cat_comma_ObjE[ + OF _ \.is_functor_op \.is_functor_op, unfolded cat_op_simps + ] + ) + from a b f have B_def: "B = op_cf_commma_obj_flip \ \\a, b, f\\<^sub>\" + by + ( + cs_concl + cs_simp: cat_comma_cs_simps B_def + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from a b f have "[a, b, f]\<^sub>\ \\<^sub>\ \\<^sub>\ (op_cf_commma_obj_flip \ \)" + by + ( + cs_concl + cs_simp: cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + with op_cf_commma_obj_flip_vsv show "B \\<^sub>\ \\<^sub>\ (op_cf_commma_obj_flip \ \)" + unfolding B_def by auto + qed +qed + + +subsubsection\Definition and elementary properties\ + +definition op_cf_comma :: "V \ V \ V" + where "op_cf_comma \ \ = + [ + op_cf_commma_obj_flip \ \, + ( + \ABF\\<^sub>\(\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Arr\. + [ + op_cf_commma_obj_flip \ \\ABF\1\<^sub>\\\, + op_cf_commma_obj_flip \ \\ABF\0\\, + [ABF\2\<^sub>\\\1\<^sub>\\, ABF\2\<^sub>\\\0\<^sub>\\]\<^sub>\ + ]\<^sub>\ + ), + op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \), + (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \) + ]\<^sub>\" + + +text\Components.\ + +lemma op_cf_comma_components: + shows [cat_comma_cs_simps]: + "op_cf_comma \ \\ObjMap\ = op_cf_commma_obj_flip \ \" + and "op_cf_comma \ \\ArrMap\ = + ( + \ABF\\<^sub>\(\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Arr\. + [ + op_cf_commma_obj_flip \ \\ABF\1\<^sub>\\\, + op_cf_commma_obj_flip \ \\ABF\0\\, + [ABF\2\<^sub>\\\1\<^sub>\\, ABF\2\<^sub>\\\0\<^sub>\\]\<^sub>\ + ]\<^sub>\ + )" + and [cat_comma_cs_simps]: + "op_cf_comma \ \\HomDom\ = op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" + and [cat_comma_cs_simps]: + "op_cf_comma \ \\HomCod\ = (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)" + unfolding op_cf_comma_def dghm_field_simps by (simp_all add: nat_omega_simps) + + +subsubsection\Arrow map\ + +mk_VLambda op_cf_comma_components(2) + |vsv op_cf_comma_ArrMap_vsv[cat_comma_cs_intros]| + |vdomain op_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]| + |app op_cf_comma_ArrMap_app'| + +lemma op_cf_comma_ArrMap_app[cat_comma_cs_simps]: + assumes "ABF = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" + and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + shows "op_cf_comma \ \\ArrMap\\ABF\ = + [ + op_cf_commma_obj_flip \ \\a', b', f'\\<^sub>\, + op_cf_commma_obj_flip \ \\a, b, f\\<^sub>\, + [h, g]\<^sub>\ + ]\<^sub>\" + using assms(2) by (simp add: assms(1) op_cf_comma_ArrMap_app' nat_omega_simps) + +lemma op_cf_comma_ArrMap_v11[cat_comma_cs_intros]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "v11 (op_cf_comma \ \\ArrMap\)" +proof + ( + rule vsv.vsv_valeq_v11I, + unfold op_cf_comma_ArrMap_vdomain, + intro op_cf_comma_ArrMap_vsv + ) + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + fix ABF ABF' assume prems: + "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + "ABF' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + "op_cf_comma \ \\ArrMap\\ABF\ = op_cf_comma \ \\ArrMap\\ABF'\" + from prems(1) obtain A B where ABF: "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto + from prems(2) obtain A' B' where ABF': "ABF' : A' \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B'" by auto + from ABF obtain a b f a' b' f' g h + where ABF_def: "ABF = [[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 "g : a \\<^bsub>\\<^esub> a'" + and "h : b \\<^bsub>\\<^esub> b'" + and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" + and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" + by (elim cat_comma_is_arrE[OF _ assms]) + from ABF' obtain a'' b'' f'' a''' b''' f''' g' h' + where ABF'_def: "ABF' = [[a'', b'', f'']\<^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 "g' : a'' \\<^bsub>\\<^esub> a'''" + and "h' : b'' \\<^bsub>\\<^esub> b'''" + and "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" + and "f''' : \\ObjMap\\a'''\ \\<^bsub>\\<^esub> \\ObjMap\\b'''\" + and "f''' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f''" + by (elim cat_comma_is_arrE[OF _ assms]) + from ABF ABF' have abf: + "[a, b, f]\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" + "[a', b', f']\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" + "[a'', b'', f'']\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" + "[a''', b''', f''']\<^sub>\ \\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" + unfolding ABF_def ABF'_def A_def B_def A'_def B'_def by auto + note v11_injective = v11.v11_injective[ + OF op_cf_commma_obj_flip_v11, OF assms, unfolded cat_comma_cs_simps + ] + from + prems(3,1,2) assms + op_cf_commma_obj_flip_v11 + v11_injective[OF abf(1,3)] + v11_injective[OF abf(2,4)] + show "ABF = ABF'" + by + ( + simp_all add: + ABF_def ABF'_def op_cf_comma_ArrMap_app' nat_omega_simps + ) +qed + +lemma op_cf_comma_ArrMap_is_arr: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + shows "op_cf_comma \ \\ArrMap\\ABF\ : + op_cf_commma_obj_flip \ \\B\ \\<^bsub>(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\<^esub> + op_cf_commma_obj_flip \ \\A\" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms(3) obtain a b f a' b' f' g h + where ABF_def: "ABF = [[a, b, f]\<^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 g: "g : a \\<^bsub>\\<^esub> a'" + and h: "h : b \\<^bsub>\\<^esub> b'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" + and f'g_hf: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" + by (elim cat_comma_is_arrE[OF _ assms(1,2)]) + from g h f f' f'g_hf show ?thesis + unfolding ABF_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) +qed + +lemma op_cf_comma_ArrMap_is_arr': + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" + and "A' = op_cf_commma_obj_flip \ \\B\" + and "B' = op_cf_commma_obj_flip \ \\A\" + shows "op_cf_comma \ \\ArrMap\\ABF\ : A' \\<^bsub>(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\<^esub> B'" + using assms(1-3) unfolding assms(4,5) by (intro op_cf_comma_ArrMap_is_arr) + +lemma op_cf_comma_ArrMap_vrange[cat_comma_cs_simps]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "\\<^sub>\ (op_cf_comma \ \\ArrMap\) = (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Arr\" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + interpret op_\\: category \ \(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\ + by (cs_concl cs_intro: cat_comma_cs_intros cat_op_intros) + show ?thesis + proof(intro vsubset_antisym) + show "\\<^sub>\ (op_cf_comma \ \\ArrMap\) \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Arr\" + proof + ( + intro vsv.vsv_vrange_vsubset op_cf_comma_ArrMap_vsv, + unfold cat_comma_cs_simps + ) + fix ABF assume prems: "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + then obtain A B where ABF: "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto + from op_cf_comma_ArrMap_is_arr[OF assms this] show + "op_cf_comma \ \\ArrMap\\ABF\ \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Arr\" + by auto + qed + show "(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Arr\ \\<^sub>\ \\<^sub>\ (op_cf_comma \ \\ArrMap\)" + proof(intro vsubsetI) + fix ABF assume prems: "ABF \\<^sub>\ (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\Arr\" + then obtain A B where ABF: "ABF : A \\<^bsub>(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\<^esub> B" + by auto + then obtain a b f a' b' f' g h + where ABF_def: "ABF = [[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 g: "g : a' \\<^bsub>\\<^esub> a" + and h: "h : b' \\<^bsub>\\<^esub> b" + and f: "f : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\a\" + and f': "f' : \\ObjMap\\b'\ \\<^bsub>\\<^esub> \\ObjMap\\a'\" + and f'g_hf: "f' \\<^sub>A\<^bsub>op_cat \\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>op_cat \\<^esub> f" + by + ( + elim cat_comma_is_arrE[ + OF _ \.is_functor_op \.is_functor_op, unfolded cat_op_simps + ] + ) + from f'g_hf g h f f' have gf'_fh: + "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f' = f \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\h\" + by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros) + with g h f f' have + "[[b', a', f']\<^sub>\, [b, a, f]\<^sub>\, [h, g]\<^sub>\]\<^sub>\ \\<^sub>\ \\<^sub>\ (op_cf_comma \ \\ArrMap\)" + "ABF = op_cf_comma \ \\ArrMap\\[b', a', f']\<^sub>\, [b, a, f]\<^sub>\, [h, g]\<^sub>\\\<^sub>\" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps ABF_def + cs_intro: cat_cs_intros cat_comma_cs_intros + )+ + with op_cf_comma_ArrMap_vsv show "ABF \\<^sub>\ \\<^sub>\ (op_cf_comma \ \\ArrMap\)" + by auto + qed + qed +qed + + +subsubsection\Opposite comma category functor is an isomorphism of categories\ + +lemma op_cf_comma_is_iso_functor: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "op_cf_comma \ \ : + op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + show ?thesis + proof(intro is_iso_functorI' is_functorI') + show "vfsequence (op_cf_comma \ \)" + unfolding op_cf_comma_def by simp + show "vcard (op_cf_comma \ \) = 4\<^sub>\" + unfolding op_cf_comma_def by (simp add: nat_omega_simps) + from assms show "op_cf_comma \ \\ArrMap\\ABF\ : + op_cf_comma \ \\ObjMap\\A\ \\<^bsub>(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\<^esub> + op_cf_comma \ \\ObjMap\\B\" + if "ABF : A \\<^bsub>op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\<^esub> B" for A B ABF + using that + unfolding cat_op_simps + by + ( + cs_concl + cs_intro: op_cf_comma_ArrMap_is_arr' cs_simp: cat_comma_cs_simps + ) + show + "op_cf_comma \ \\ArrMap\\G \\<^sub>A\<^bsub>op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\<^esub> F\ = + op_cf_comma \ \\ArrMap\\G\ \\<^sub>A\<^bsub>(op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\<^esub> + op_cf_comma \ \\ArrMap\\F\" + if "G : B \\<^bsub>op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\<^esub> C" + and "F : A \\<^bsub>op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\<^esub> B" + for B C G A F + proof- + note G = that(1)[unfolded cat_op_simps] + note F = that(2)[unfolded cat_op_simps] + from assms G obtain a b f a' b' f' g h + where G_def: "G = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" + and C_def: "C = [a, b, f]\<^sub>\" + and B_def: "B = [a', b', f']\<^sub>\" + and g: "g : a \\<^bsub>\\<^esub> a'" + and h: "h : b \\<^bsub>\\<^esub> b'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" + and [symmetric, cat_comma_cs_simps]: + "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" + by auto + with assms F obtain a'' b'' f'' g' h' + where F_def: "F = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" + and A_def: "A = [a'', b'', f'']\<^sub>\" + and g': "g' : a' \\<^bsub>\\<^esub> a''" + and h': "h' : b' \\<^bsub>\\<^esub> b''" + and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" + and [cat_comma_cs_simps]: + "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" + by auto (*slow*) + note [cat_comma_cs_simps] = + category.cat_assoc_helper[ + where \=\ and h=f'' and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ + ] + from assms that g h f f' g' h' f' f'' show ?thesis + unfolding cat_op_simps G_def C_def B_def F_def A_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + show + "op_cf_comma \ \\ArrMap\\op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\CId\\C\\ = + (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)\CId\\op_cf_comma \ \\ObjMap\\C\\" + if "C \\<^sub>\ op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)\Obj\" for C + proof- + from that[unfolded cat_op_simps] assms obtain a b f + where C_def: "C = [a, b, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and b: "b \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + by auto + from a b f that show ?thesis + unfolding cat_op_simps C_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + qed + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_op_intros + )+ +qed + +lemma op_cf_comma_is_iso_functor'[cat_comma_cs_intros]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\' = op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" + and "\' = (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)" + shows "op_cf_comma \ \ : \' \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" + using assms(1,2) unfolding assms(3,4) by (rule op_cf_comma_is_iso_functor) + +lemma op_cf_comma_is_functor: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "op_cf_comma \ \ : + op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)" + by (rule is_iso_functorD(1)[OF op_cf_comma_is_iso_functor[OF assms]]) + +lemma op_cf_comma_is_functor'[cat_comma_cs_intros]: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\' = op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \)" + and "\' = (op_cf \) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (op_cf \)" + shows "op_cf_comma \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + using assms(1,2) unfolding assms(3,4) by (rule op_cf_comma_is_functor) + + + subsection\Projections for a comma category\ subsubsection\Definitions and elementary properties\ text\See Chapter II-6 in \cite{mac_lane_categories_2010}.\ definition cf_comma_proj_left :: "V \ V \ V" (\(_ \<^sub>C\<^sub>F\ _)\ [1000, 1000] 999) where "\ \<^sub>C\<^sub>F\ \ = [ (\a\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\. a\0\), (\f\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\. f\2\<^sub>\\\0\), \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \, \\HomDom\ ]\<^sub>\" definition cf_comma_proj_right :: "V \ V \ V" (\(_ \\<^sub>C\<^sub>F _)\ [1000, 1000] 999) where "\ \\<^sub>C\<^sub>F \ = [ (\a\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\. a\1\<^sub>\\), (\f\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\. f\2\<^sub>\\\1\<^sub>\\), \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \, \\HomDom\ ]\<^sub>\" text\Components.\ lemma cf_comma_proj_left_components: shows "\ \<^sub>C\<^sub>F\ \\ObjMap\ = (\a\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\. a\0\)" and "\ \<^sub>C\<^sub>F\ \\ArrMap\ = (\f\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\. f\2\<^sub>\\\0\)" and "\ \<^sub>C\<^sub>F\ \\HomDom\ = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" and "\ \<^sub>C\<^sub>F\ \\HomCod\ = \\HomDom\" unfolding cf_comma_proj_left_def dghm_field_simps by (simp_all add: nat_omega_simps) lemma cf_comma_proj_right_components: shows "\ \\<^sub>C\<^sub>F \\ObjMap\ = (\a\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\. a\1\<^sub>\\)" and "\ \\<^sub>C\<^sub>F \\ArrMap\ = (\f\\<^sub>\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\. f\2\<^sub>\\\1\<^sub>\\)" and "\ \\<^sub>C\<^sub>F \\HomDom\ = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" and "\ \\<^sub>C\<^sub>F \\HomCod\ = \\HomDom\" unfolding cf_comma_proj_right_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 cf_comma_proj_left_components' = cf_comma_proj_left_components[of \ \, unfolded \.cf_HomDom] lemmas cf_comma_proj_right_components' = cf_comma_proj_right_components[of \ \, unfolded \.cf_HomDom] lemmas [cat_comma_cs_simps] = cf_comma_proj_left_components'(3,4) cf_comma_proj_right_components'(3,4) end subsubsection\Object map\ mk_VLambda cf_comma_proj_left_components(1) |vsv cf_comma_proj_left_ObjMap_vsv[cat_comma_cs_intros]| |vdomain cf_comma_proj_left_ObjMap_vdomain[cat_comma_cs_simps]| mk_VLambda cf_comma_proj_right_components(1) |vsv cf_comma_proj_right_ObjMap_vsv[cat_comma_cs_intros]| |vdomain cf_comma_proj_right_ObjMap_vdomain[cat_comma_cs_simps]| lemma cf_comma_proj_left_ObjMap_app[cat_comma_cs_simps]: assumes "A = [a, b, f]\<^sub>\" and "[a, b, f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" shows "\ \<^sub>C\<^sub>F\ \\ObjMap\\A\ = a" using assms(2) unfolding assms(1) cf_comma_proj_left_components by simp lemma cf_comma_proj_right_ObjMap_app[cat_comma_cs_simps]: assumes "A = [a, b, f]\<^sub>\" and "[a, b, f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" shows "\ \\<^sub>C\<^sub>F \\ObjMap\\A\ = b" using assms(2) unfolding assms(1) cf_comma_proj_right_components by (simp add: nat_omega_simps) lemma cf_comma_proj_left_ObjMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\ \\ObjMap\) \\<^sub>\ \\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps) fix A assume prems: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" with assms obtain a b f where A_def: "A = [a, b, f]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" by auto from assms prems a show "\ \<^sub>C\<^sub>F\ \\ObjMap\\A\ \\<^sub>\ \\Obj\" unfolding A_def by (cs_concl cs_simp: cat_comma_cs_simps) qed (auto intro: cat_comma_cs_intros) lemma cf_comma_proj_right_ObjMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ \\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps) fix A assume prems: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" with assms obtain a b f where A_def: "A = [a, b, f]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" by auto from assms prems b show "\ \\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^sub>\ \\Obj\" unfolding A_def by (cs_concl cs_simp: cat_comma_cs_simps) qed (auto intro: cat_comma_cs_intros) subsubsection\Arrow map\ mk_VLambda cf_comma_proj_left_components(2) |vsv cf_comma_proj_left_ArrMap_vsv[cat_comma_cs_intros]| |vdomain cf_comma_proj_left_ArrMap_vdomain[cat_comma_cs_simps]| mk_VLambda cf_comma_proj_right_components(2) |vsv cf_comma_proj_right_ArrMap_vsv[cat_comma_cs_intros]| |vdomain cf_comma_proj_right_ArrMap_vdomain[cat_comma_cs_simps]| lemma cf_comma_proj_left_ArrMap_app[cat_comma_cs_simps]: - assumes "A = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "[abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - shows "\ \<^sub>C\<^sub>F\ \\ArrMap\\A\ = g" + assumes "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" and "[A, B, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + shows "\ \<^sub>C\<^sub>F\ \\ArrMap\\ABF\ = g" using assms(2) unfolding assms(1) cf_comma_proj_left_components by (simp add: nat_omega_simps) lemma cf_comma_proj_right_ArrMap_app[cat_comma_cs_simps]: - assumes "A = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "[abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - shows "\ \\<^sub>C\<^sub>F \\ArrMap\\A\ = h" + assumes "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "[A, B, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + shows "\ \\<^sub>C\<^sub>F \\ArrMap\\ABF\ = h" using assms(2) unfolding assms(1) cf_comma_proj_right_components by (simp add: nat_omega_simps) lemma cf_comma_proj_left_ArrMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\ \\ArrMap\) \\<^sub>\ \\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps) from assms interpret category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) fix F assume prems: "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - then obtain abf a'b'f' where "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" by auto + then obtain A B where "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto with assms obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" by auto from assms prems g show "\ \<^sub>C\<^sub>F\ \\ArrMap\\F\ \\<^sub>\ \\Arr\" unfolding F_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) qed (auto intro: cat_comma_cs_intros) lemma cf_comma_proj_right_ArrMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (\ \\<^sub>C\<^sub>F \\ArrMap\) \\<^sub>\ \\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_comma_cs_simps) from assms interpret category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) fix F assume prems: "F \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" - then obtain abf a'b'f' where F: "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> a'b'f'" by auto + then obtain A B where F: "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto with assms obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\" and h: "h : b \\<^bsub>\\<^esub> b'" by auto from assms prems h show "\ \\<^sub>C\<^sub>F \\ArrMap\\F\ \\<^sub>\ \\Arr\" unfolding F_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) qed (auto intro: cat_comma_cs_intros) subsubsection\Projections for a comma category are functors\ lemma cf_comma_proj_left_is_functor: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \<^sub>C\<^sub>F\ \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) show ?thesis proof(rule is_functorI') show "vfsequence (\ \<^sub>C\<^sub>F\ \)" unfolding cf_comma_proj_left_def by auto show "vcard (\ \<^sub>C\<^sub>F\ \) = 4\<^sub>\" unfolding cf_comma_proj_left_def by (simp add: nat_omega_simps) from assms show "\\<^sub>\ (\ \<^sub>C\<^sub>F\ \\ObjMap\) \\<^sub>\ \\Obj\" by (rule cf_comma_proj_left_ObjMap_vrange) show "\ \<^sub>C\<^sub>F\ \\ArrMap\\F\ : \ \<^sub>C\<^sub>F\ \\ObjMap\\A\ \\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ \\ObjMap\\B\" if "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for A B F proof- from assms that obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^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 g: "g : a \\<^bsub>\\<^esub> a'" by auto from that g show "\ \<^sub>C\<^sub>F\ \\ArrMap\\F\ : \ \<^sub>C\<^sub>F\ \\ObjMap\\A\ \\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ \\ObjMap\\B\" unfolding F_def A_def B_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) qed show "\ \<^sub>C\<^sub>F\ \\ArrMap\\G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F\ = \ \<^sub>C\<^sub>F\ \\ArrMap\\G\ \\<^sub>A\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ \\ArrMap\\F\" if "G : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for B C G A F proof- from assms that(2) obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^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 g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and [cat_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto with that(1) assms obtain a'' b'' f'' g' h' where G_def: "G = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" and C_def: "C = [a'', b'', f'']\<^sub>\" and g': "g' : a' \\<^bsub>\\<^esub> a''" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto (*slow*) note [cat_cs_simps] = category.cat_assoc_helper [ where \=\ and h=f'' and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ ] category.cat_assoc_helper [ where \=\ and h=f and g=\\\ArrMap\\h\\ and q=\f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\\ ] from assms that g g' h h' f f' f'' show ?thesis unfolding F_def G_def A_def B_def C_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) qed show "\ \<^sub>C\<^sub>F\ \\ArrMap\\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\A\\ = \\CId\\\ \<^sub>C\<^sub>F\ \\ObjMap\\A\\" if "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" for A proof- from assms that obtain a b f where A_def: "A = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from assms that this(2-4) show ?thesis unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) qed qed ( use assms in \ cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros \ )+ qed lemma cf_comma_proj_left_is_functor'[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" shows "\ \<^sub>C\<^sub>F\ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_left_is_functor) lemma cf_comma_proj_right_is_functor: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) show ?thesis proof(rule is_functorI') show "vfsequence (\ \\<^sub>C\<^sub>F \)" unfolding cf_comma_proj_right_def by auto show "vcard (\ \\<^sub>C\<^sub>F \) = 4\<^sub>\" unfolding cf_comma_proj_right_def by (simp add: nat_omega_simps) from assms show "\\<^sub>\ (\ \\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ \\Obj\" by (rule cf_comma_proj_right_ObjMap_vrange) show "\ \\<^sub>C\<^sub>F \\ArrMap\\F\ : \ \\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \\ObjMap\\B\" if "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for A B F proof- from assms that obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^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 h: "h : b \\<^bsub>\\<^esub> b'" by auto from that h show "\ \\<^sub>C\<^sub>F \\ArrMap\\F\ : \ \\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \\ObjMap\\B\" unfolding F_def A_def B_def by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) qed show "\ \\<^sub>C\<^sub>F \\ArrMap\\G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> F\ = \ \\<^sub>C\<^sub>F \\ArrMap\\G\ \\<^sub>A\<^bsub>\\<^esub> \ \\<^sub>C\<^sub>F \\ArrMap\\F\" if "G : B \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" for B C G A F proof- from assms that(2) obtain a b f a' b' f' g h where F_def: "F = [[a, b, f]\<^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 g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b \\<^bsub>\\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" and [cat_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by auto with that(1) assms obtain a'' b'' f'' g' h' where G_def: "G = [[a', b', f']\<^sub>\, [a'', b'', f'']\<^sub>\, [g', h']\<^sub>\]\<^sub>\" and C_def: "C = [a'', b'', f'']\<^sub>\" and g': "g' : a' \\<^bsub>\\<^esub> a''" and h': "h' : b' \\<^bsub>\\<^esub> b''" and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> \\ObjMap\\b''\" and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g'\ = \\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'" by auto (*slow*) note [cat_cs_simps] = category.cat_assoc_helper [ where \=\ and h=f'' and g=\\\ArrMap\\g'\\ and q=\\\ArrMap\\h'\ \\<^sub>A\<^bsub>\\<^esub> f'\ ] category.cat_assoc_helper [ where \=\ and h=f and g=\\\ArrMap\\h\\ and q=\f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\\ ] from assms that g g' h h' f f' f'' show ?thesis unfolding F_def G_def A_def B_def C_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) qed show "\ \\<^sub>C\<^sub>F \\ArrMap\\\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\CId\\A\\ = \\CId\\\ \\<^sub>C\<^sub>F \\ObjMap\\A\\" if "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" for A proof- from assms that obtain a b f where A_def: "A = [a, b, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto from assms that this(2-4) show ?thesis unfolding A_def by ( cs_concl cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros ) qed qed ( use assms in \ cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros \ )+ qed lemma cf_comma_proj_right_is_functor'[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" shows "\ \\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_right_is_functor) +subsubsection\Opposite projections for a comma category\ + +lemma op_cf_comma_proj_left: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "op_cf (\ \<^sub>C\<^sub>F\ \) = (op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + show "op_cf (\ \<^sub>C\<^sub>F\ \) = (op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \" + proof(rule cf_eqI) + show "op_cf (\ \<^sub>C\<^sub>F\ \) : op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by + ( + cs_concl cs_simp: cs_intro: + cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + then have ObjMap_dom_lhs: "\\<^sub>\ (op_cf (\ \<^sub>C\<^sub>F\ \)\ObjMap\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and ArrMap_dom_lhs: "\\<^sub>\ (op_cf (\ \<^sub>C\<^sub>F\ \)\ArrMap\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps cat_op_simps)+ + show "(op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \ : + op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros) + then have ObjMap_dom_rhs: + "\\<^sub>\ (((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\) = + \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and ArrMap_dom_rhs: + "\\<^sub>\ (((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ArrMap\) = + \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_cs_simps cat_op_simps)+ + show + "op_cf (\ \<^sub>C\<^sub>F\ \)\ObjMap\ = + ((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + fix A assume "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + with assms obtain a b f + where A_def: "A = [a, b, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and b: "b \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + by auto + from a b f show + "op_cf (\ \<^sub>C\<^sub>F\ \)\ObjMap\\A\ = + ((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\\A\" + unfolding A_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + ( + cs_concl + cs_simp: cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + )+ + show + "op_cf (\ \<^sub>C\<^sub>F\ \)\ArrMap\ = + ((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>D\<^sub>G\<^sub>H\<^sub>M op_cf_comma \ \)\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix ABF assume "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + then obtain A B where "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto + with assms obtain a b f a' b' f' g h + where ABF_def: "ABF = [[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 g: "g : a \\<^bsub>\\<^esub> a'" + and h: "h : b \\<^bsub>\\<^esub> b'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" + and [symmetric, cat_cs_simps]: + "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" + by auto + from g h f f' show "op_cf (\ \<^sub>C\<^sub>F\ \)\ArrMap\\ABF\ = + ((op_cf \) \\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ArrMap\\ABF\" + unfolding ABF_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + ( + cs_concl + cs_simp: cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + )+ + qed simp_all +qed + +lemma op_cf_comma_proj_right: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "op_cf (\ \\<^sub>C\<^sub>F \) = (op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret \\: category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + show "op_cf (\ \\<^sub>C\<^sub>F \) = (op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \" + proof(rule cf_eqI) + show "op_cf (\ \\<^sub>C\<^sub>F \) : op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by + ( + cs_concl cs_simp: cs_intro: + cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + then have ObjMap_dom_lhs: "\\<^sub>\ (op_cf (\ \\<^sub>C\<^sub>F \)\ObjMap\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and ArrMap_dom_lhs: "\\<^sub>\ (op_cf (\ \\<^sub>C\<^sub>F \)\ArrMap\) = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps cat_op_simps)+ + show "(op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \ : + op_cat (\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros) + then have ObjMap_dom_rhs: + "\\<^sub>\ (((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\) = + \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and ArrMap_dom_rhs: + "\\<^sub>\ (((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ArrMap\) = + \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_cs_simps cat_op_simps)+ + show + "op_cf (\ \\<^sub>C\<^sub>F \)\ObjMap\ = + ((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + fix A assume prems: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + with assms obtain a b f + where A_def: "A = [a, b, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and b: "b \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + by auto + from a b f show + "op_cf (\ \\<^sub>C\<^sub>F \)\ObjMap\\A\ = + ((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ObjMap\\A\" + unfolding A_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + ( + cs_concl + cs_simp: cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + )+ + show + "op_cf (\ \\<^sub>C\<^sub>F \)\ArrMap\ = + ((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>D\<^sub>G\<^sub>H\<^sub>M op_cf_comma \ \)\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix ABF assume prems: "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Arr\" + then obtain A B where ABF: "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\<^esub> B" by auto + with assms obtain a b f a' b' f' g h + where ABF_def: "ABF = [[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 g: "g : a \\<^bsub>\\<^esub> a'" + and h: "h : b \\<^bsub>\\<^esub> b'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" + and [symmetric, cat_cs_simps]: + "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = \\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" + by auto + from g h f f' show "op_cf (\ \\<^sub>C\<^sub>F \)\ArrMap\\ABF\ = + ((op_cf \) \<^sub>C\<^sub>F\ (op_cf \) \\<^sub>C\<^sub>F op_cf_comma \ \)\ArrMap\\ABF\" + unfolding ABF_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + ( + cs_concl + cs_simp: cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + )+ + qed simp_all +qed + + subsubsection\Projections for a tiny comma category\ lemma cf_comma_proj_left_is_tm_functor: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" shows "\ \<^sub>C\<^sub>F\ \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" proof(intro is_tm_functorI') interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) show \_\\: "\ \<^sub>C\<^sub>F\ \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) interpret \_\\: is_functor \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ \ \\ \<^sub>C\<^sub>F\ \\ by (rule \_\\) interpret \\: tiny_category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_small_cs_intros cat_comma_cs_intros) show "\ \<^sub>C\<^sub>F\ \\ObjMap\ \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\ \<^sub>C\<^sub>F\ \\ObjMap\) \\<^sub>\ Vset \" proof- note \_\\.cf_ObjMap_vrange moreover have "\\Obj\ \\<^sub>\ Vset \" by (intro cat_small_cs_intros) ultimately show ?thesis by auto qed qed (auto simp: cf_comma_proj_left_components intro: cat_small_cs_intros) show "\ \<^sub>C\<^sub>F\ \\ArrMap\ \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\ \<^sub>C\<^sub>F\ \\ArrMap\) \\<^sub>\ Vset \" proof- note \_\\.cf_ArrMap_vrange moreover have "\\Arr\ \\<^sub>\ Vset \" by (intro cat_small_cs_intros) ultimately show ?thesis by auto qed qed (auto simp: cf_comma_proj_left_components intro: cat_small_cs_intros) qed lemma cf_comma_proj_left_is_tm_functor'[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\\ = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" shows "\ \<^sub>C\<^sub>F\ \ : \\ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_left_is_tm_functor) lemma cf_comma_proj_right_is_tm_functor: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" proof(intro is_tm_functorI') interpret \: is_tm_functor \ \ \ \ by (rule assms(1)) interpret \: is_tm_functor \ \ \ \ by (rule assms(2)) show \_\\: "\ \\<^sub>C\<^sub>F \ : \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) interpret \_\\: is_functor \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ \ \\ \\<^sub>C\<^sub>F \\ by (rule \_\\) interpret \\: tiny_category \ \\ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_small_cs_intros cat_comma_cs_intros) show "\ \\<^sub>C\<^sub>F \\ObjMap\ \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\ \\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ Vset \" proof- note \_\\.cf_ObjMap_vrange moreover have "\\Obj\ \\<^sub>\ Vset \" by (intro cat_small_cs_intros) ultimately show ?thesis by auto qed qed (auto simp: cf_comma_proj_right_components intro: cat_small_cs_intros) show "\ \\<^sub>C\<^sub>F \\ArrMap\ \\<^sub>\ Vset \" proof(rule vbrelation.vbrelation_Limit_in_VsetI) show "\\<^sub>\ (\ \\<^sub>C\<^sub>F \\ArrMap\) \\<^sub>\ Vset \" proof- note \_\\.cf_ArrMap_vrange moreover have "\\Arr\ \\<^sub>\ Vset \" by (intro cat_small_cs_intros) ultimately show ?thesis by auto qed qed (auto simp: cf_comma_proj_right_components intro: cat_small_cs_intros) qed lemma cf_comma_proj_right_is_tm_functor'[cat_comma_cs_intros]: assumes "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" and "\\ = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" shows "\ \\<^sub>C\<^sub>F \ : \\ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3) by (rule cf_comma_proj_right_is_tm_functor) subsection\Comma categories constructed from a functor and an object\ subsubsection\Definitions and elementary properties\ text\See Chapter II-6 in \cite{mac_lane_categories_2010}.\ definition cat_cf_obj_comma :: "V \ V \ V" (\(_ \<^sub>C\<^sub>F\ _)\ [1000, 1000] 999) where "\ \<^sub>C\<^sub>F\ b \ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) (\\HomCod\) b)" definition cat_obj_cf_comma :: "V \ V \ V" (\(_ \\<^sub>C\<^sub>F _)\ [1000, 1000] 999) where "b \\<^sub>C\<^sub>F \ \ (cf_const (cat_1 0 0) (\\HomCod\) b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" text\Alternative forms of the definitions.\ lemma (in is_functor) cat_cf_obj_comma_def: "\ \<^sub>C\<^sub>F\ b = \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)" unfolding cat_cf_obj_comma_def cf_HomCod .. lemma (in is_functor) cat_obj_cf_comma_def: "b \\<^sub>C\<^sub>F \ = (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \" unfolding cat_obj_cf_comma_def cf_HomCod .. subsubsection\Objects\ lemma (in is_functor) cat_cf_obj_comma_ObjI[cat_comma_cs_intros]: assumes "A = [a, 0, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" shows "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" using assms(2,3) unfolding assms(1) by ( cs_concl cs_simp: cat_cs_simps cat_cf_obj_comma_def cs_intro: cat_cs_intros vempty_is_zet cat_comma_ObjI ) lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_ObjI lemma (in is_functor) cat_obj_cf_comma_ObjI[cat_comma_cs_intros]: assumes "A = [0, a, f]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" shows "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" using assms(2,3) unfolding assms(1) by ( cs_concl cs_simp: cat_cs_simps cat_obj_cf_comma_def cs_intro: vempty_is_zet cat_cs_intros cat_comma_ObjI ) lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_ObjI lemma (in is_functor) cat_cf_obj_comma_ObjD[dest]: assumes "[a, b', f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" and "b \\<^sub>\ \\Obj\" - shows "a \\<^sub>\ \\Obj\" - and "b' = 0" - and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" + shows "a \\<^sub>\ \\Obj\" and "b' = 0" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" proof- from assms(2) have "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) note obj = cat_comma_ObjD[ OF assms(1)[unfolded cat_cf_obj_comma_def] is_functor_axioms this ] from obj[unfolded cat_1_components] have [cat_cs_simps]: "b' = 0" by simp moreover have "cf_const (cat_1 0 0) \ b\ObjMap\\b'\ = b" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show "a \\<^sub>\ \\Obj\" "b' = 0" "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" using obj by auto qed lemmas [dest] = is_functor.cat_cf_obj_comma_ObjD[rotated 1] lemma (in is_functor) cat_obj_cf_comma_ObjD[dest]: assumes "[b', a, f]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" and "b \\<^sub>\ \\Obj\" - shows "a \\<^sub>\ \\Obj\" - and "b' = 0" - and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" + shows "a \\<^sub>\ \\Obj\" and "b' = 0" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" proof- from assms(2) have "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) note obj = cat_comma_ObjD[ OF assms(1)[unfolded cat_obj_cf_comma_def] this is_functor_axioms ] from obj[unfolded cat_1_components] have [cat_cs_simps]: "b' = 0" by simp moreover have "cf_const (cat_1 0 0) \ b\ObjMap\\b'\ = b" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show "a \\<^sub>\ \\Obj\" "b' = 0" "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" using obj by auto qed lemmas [dest] = is_functor.cat_obj_cf_comma_ObjD[rotated 1] lemma (in is_functor) cat_cf_obj_comma_ObjE[elim]: assumes "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" and "b \\<^sub>\ \\Obj\" - obtains a f where "A = [a, 0, f]\<^sub>\" - and "a \\<^sub>\ \\Obj\" - and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" + obtains a f + where "A = [a, 0, f]\<^sub>\" + and "a \\<^sub>\ \\Obj\" + and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" proof- from assms(2) have "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from assms(1)[unfolded cat_cf_obj_comma_def] is_functor_axioms this obtain a b' f where "A = [a, b', f]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ cat_1 0 0\Obj\" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ObjMap\\b'\" by auto moreover from b' have [cat_cs_simps]: "b' = 0" unfolding cat_1_components by auto moreover have "cf_const (cat_1 0 0) \ b\ObjMap\\b'\ = b" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show ?thesis using that by auto qed lemmas [elim] = is_functor.cat_cf_obj_comma_ObjE[rotated 1] lemma (in is_functor) cat_obj_cf_comma_ObjE[elim]: assumes "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" and "b \\<^sub>\ \\Obj\" - obtains a f where "A = [0, a, f]\<^sub>\" - and "a \\<^sub>\ \\Obj\" - and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" + obtains a f + where "A = [0, a, f]\<^sub>\" + and "a \\<^sub>\ \\Obj\" + and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" proof- from assms(2) have "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from assms(1)[unfolded cat_obj_cf_comma_def] is_functor_axioms this obtain a b' f where A_def: "A = [b', a, f]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ cat_1 0 0\Obj\" and f: "f : cf_const (cat_1 0 0) \ b\ObjMap\\b'\ \\<^bsub>\\<^esub> \\ObjMap\\a\" by auto moreover from b' have [cat_cs_simps]: "b' = 0" unfolding cat_1_components by auto moreover have "cf_const (cat_1 0 0) \ b\ObjMap\\b'\ = b" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) ultimately show ?thesis using that by auto qed lemmas [elim] = is_functor.cat_obj_cf_comma_ObjE[rotated 1] subsubsection\Arrows\ lemma (in is_functor) cat_cf_obj_comma_ArrI[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "F = [abf, a'b'f', [g, 0]\<^sub>\]\<^sub>\" - and "abf = [a, 0, f]\<^sub>\" - and "a'b'f' = [a', 0, f']\<^sub>\" + and "F = [A, B, [g, 0]\<^sub>\]\<^sub>\" + and "A = [a, 0, f]\<^sub>\" + and "B = [a', 0, f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" shows "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" unfolding cat_cf_obj_comma_def proof(intro cat_comma_ArrI cat_comma_HomI) show "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms(1) show const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from vempty_is_zet show 0: "0 : 0 \\<^bsub>cat_1 0 0\<^esub> 0" by (cs_concl cs_simp: cat_cs_simps cat_1_CId_app cs_intro: cat_cs_intros) from assms(6) show "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ObjMap\\0\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(7) show "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ObjMap\\0\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from 0 assms(6) show "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = cf_const (cat_1 0 0) \ b\ArrMap\\0\ \\<^sub>A\<^bsub>\\<^esub> f" by (cs_concl cs_simp: cat_cs_simps assms(8) cs_intro: cat_cs_intros) - from const assms(5,6) show - "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (dghm_const (cat_1 []\<^sub>\ []\<^sub>\) \ b (\\CId\\b\))\Obj\" + from const assms(5,6) show "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" by (fold cat_cf_obj_comma_def) (cs_concl cs_simp: assms(3) cs_intro: cat_cs_intros cat_comma_cs_intros) - from const assms(5,7) show - "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (dghm_const (cat_1 []\<^sub>\ []\<^sub>\) \ b (\\CId\\b\))\Obj\" + from const assms(5,7) show "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" by (fold cat_cf_obj_comma_def) (cs_concl cs_simp: assms(4) cs_intro: cat_cs_intros cat_comma_cs_intros) qed (intro assms)+ lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_ArrI lemma (in is_functor) cat_obj_cf_comma_ArrI[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "F = [abf, a'b'f', [0, g]\<^sub>\]\<^sub>\" - and "abf = [0, a, f]\<^sub>\" - and "a'b'f' = [0, a', f']\<^sub>\" + and "F = [A, B, [0, g]\<^sub>\]\<^sub>\" + and "A = [0, a, f]\<^sub>\" + and "B = [0, a', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\ " and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" shows "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" unfolding cat_obj_cf_comma_def proof(intro cat_comma_ArrI cat_comma_HomI) show "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms(1) show const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from vempty_is_zet show 0: "0 : 0 \\<^bsub>cat_1 0 0\<^esub> 0" by (cs_concl cs_simp: cat_1_CId_app cs_intro: cat_cs_intros) from assms(6) show "f : cf_const (cat_1 0 0) \ b\ObjMap\\0\ \\<^bsub>\\<^esub> \\ObjMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(7) show "f' : cf_const (cat_1 0 0) \ b\ObjMap\\0\ \\<^bsub>\\<^esub> \\ObjMap\\a'\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from 0 assms(7) show "f' \\<^sub>A\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ArrMap\\0\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f" by (cs_concl cs_simp: cat_cs_simps assms(8) cs_intro: cat_cs_intros) - from const assms(5,6) show - "abf \\<^sub>\ (dghm_const (cat_1 0 0) \ b (\\CId\\b\)) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from const assms(5,6) show "A \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (fold cat_obj_cf_comma_def) (cs_concl cs_simp: assms(3) cs_intro: cat_cs_intros cat_comma_cs_intros) - from const assms(5,7) show - "a'b'f' \\<^sub>\ (dghm_const (cat_1 []\<^sub>\ []\<^sub>\) \ b (\\CId\\b\)) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + from const assms(5,7) show "B \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by (fold cat_obj_cf_comma_def) (cs_concl cs_simp: assms(4) cs_intro: cat_cs_intros cat_comma_cs_intros) qed (intro assms)+ lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_ArrI lemma (in is_functor) cat_cf_obj_comma_ArrE[elim]: assumes "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" and "b \\<^sub>\ \\Obj\" - obtains abf a'b'f' a f a' f' g - where "F = [abf, a'b'f', [g, 0]\<^sub>\]\<^sub>\" - and "abf = [a, 0, f]\<^sub>\" - and "a'b'f' = [a', 0, f']\<^sub>\" + obtains A B a f a' f' g + where "F = [A, B, [g, 0]\<^sub>\]\<^sub>\" + and "A = [a, 0, f]\<^sub>\" + and "B = [a', 0, f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" proof- from cat_comma_ArrE[OF assms(1)[unfolded cat_cf_obj_comma_def]] - obtain abf a'b'f' - where F: "F \\<^sub>\ cat_comma_Hom \ (cf_const (cat_1 0 0) \ b) abf a'b'f'" - and abf: "abf \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" - and a'b'f': "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" + obtain A B + where F: "F \\<^sub>\ cat_comma_Hom \ (cf_const (cat_1 0 0) \ b) A B" + and A: "A \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" + and B: "B \\<^sub>\ \ \<^sub>C\<^sub>F\\<^sub>C\<^sub>F (cf_const (cat_1 0 0) \ b)\Obj\" by auto from assms(2) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from F obtain a b'' f a' b' f' g h - where F_def: "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and abf_def: "abf = [a, b'', f]\<^sub>\" - and a'b'f'_def: "a'b'f' = [a', b', f']\<^sub>\" + where F_def: "F = [A, B, [g, h]\<^sub>\]\<^sub>\" + and A_def: "A = [a, b'', f]\<^sub>\" + and B_def: "B = [a', b', f']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and h: "h : b'' \\<^bsub>cat_1 0 0\<^esub> b'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ObjMap\\b''\" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ObjMap\\b'\" and f_def: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = cf_const (cat_1 0 0) \ b\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f" by (elim cat_comma_HomE[OF _ is_functor_axioms const]) blast note hb'b'' = cat_1_is_arrD[OF h] - from F_def have F_def: "F = [abf, a'b'f', [g, 0]\<^sub>\]\<^sub>\" + from F_def have F_def: "F = [A, B, [g, 0]\<^sub>\]\<^sub>\" unfolding hb'b'' by simp - from abf_def have abf_def: "abf = [a, 0, f]\<^sub>\" + from A_def have A_def: "A = [a, 0, f]\<^sub>\" unfolding hb'b'' by simp - from a'b'f'_def have a'b'f'_def: "a'b'f' = [a', 0, f']\<^sub>\" + from B_def have B_def: "B = [a', 0, f']\<^sub>\" unfolding hb'b'' by simp from f have f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f' have f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f_def f f' g h have f_def: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from - that F_def abf_def a'b'f'_def g f f' f_def - a'b'f'[folded cat_cf_obj_comma_def] abf[folded cat_cf_obj_comma_def] + that F_def A_def B_def g f f' f_def + B[folded cat_cf_obj_comma_def] A[folded cat_cf_obj_comma_def] show ?thesis by blast qed lemmas [elim] = is_functor.cat_cf_obj_comma_ArrE[rotated 1] lemma (in is_functor) cat_obj_cf_comma_ArrE[elim]: assumes "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" and "b \\<^sub>\ \\Obj\" - obtains baf b'a'f' a f a' f' g - where "F = [baf, b'a'f', [0, g]\<^sub>\]\<^sub>\" - and "baf = [0, a, f]\<^sub>\" - and "b'a'f' = [0, a', f']\<^sub>\" + obtains A B a f a' f' g + where "F = [A, B, [0, g]\<^sub>\]\<^sub>\" + and "A = [0, a, f]\<^sub>\" + and "B = [0, a', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" - and "baf \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" - and "b'a'f' \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" proof- from cat_comma_ArrE[OF assms(1)[unfolded cat_obj_cf_comma_def]] - obtain baf b'a'f' - where F: "F \\<^sub>\ cat_comma_Hom (cf_const (cat_1 0 0) \ b) \ baf b'a'f'" - and baf: "baf \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" - and b'a'f': "b'a'f' \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + obtain A B + where F: "F \\<^sub>\ cat_comma_Hom (cf_const (cat_1 0 0) \ b) \ A B" + and A: "A \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" + and B: "B \\<^sub>\ (cf_const (cat_1 0 0) \ b) \<^sub>C\<^sub>F\\<^sub>C\<^sub>F \\Obj\" by auto from assms(2) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from F obtain a b'' f a' b' f' h g - where F_def: "F = [baf, b'a'f', [h, g]\<^sub>\]\<^sub>\" - and baf_def: "baf = [b', a, f]\<^sub>\" - and b'a'f'_def: "b'a'f' = [b'', a', f']\<^sub>\" + where F_def: "F = [A, B, [h, g]\<^sub>\]\<^sub>\" + and A_def: "A = [b', a, f]\<^sub>\" + and B_def: "B = [b'', a', f']\<^sub>\" and h: "h : b' \\<^bsub>cat_1 0 0\<^esub> b''" and g: "g : a \\<^bsub>\\<^esub> a'" and f: "f : cf_const (cat_1 0 0) \ b\ObjMap\\b'\ \\<^bsub>\\<^esub> \\ObjMap\\a\" and f': "f' : cf_const (cat_1 0 0) \ b\ObjMap\\b''\ \\<^bsub>\\<^esub> \\ObjMap\\a'\" and f'_def: "f' \\<^sub>A\<^bsub>\\<^esub> cf_const (cat_1 0 0) \ b\ArrMap\\h\ = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f" by (elim cat_comma_HomE[OF _ const is_functor_axioms]) blast note hb'b'' = cat_1_is_arrD[OF h] - from F_def have F_def: "F = [baf, b'a'f', [0, g]\<^sub>\]\<^sub>\" + from F_def have F_def: "F = [A, B, [0, g]\<^sub>\]\<^sub>\" unfolding hb'b'' by simp - from baf_def have baf_def: "baf = [0, a, f]\<^sub>\" - unfolding hb'b'' by simp - from b'a'f'_def have b'a'f'_def: "b'a'f' = [0, a', f']\<^sub>\" - unfolding hb'b'' by simp + from A_def have A_def: "A = [0, a, f]\<^sub>\" unfolding hb'b'' by simp + from B_def have B_def: "B = [0, a', f']\<^sub>\" unfolding hb'b'' by simp from f have f: "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f' have f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from f'_def f f' g h have f'_def[symmetric]: "f' = \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f" unfolding hb'b'' by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from - that F_def baf_def b'a'f'_def g f f' f'_def - baf[folded cat_obj_cf_comma_def] b'a'f'[folded cat_obj_cf_comma_def] + that F_def A_def B_def g f f' f'_def + A[folded cat_obj_cf_comma_def] B[folded cat_obj_cf_comma_def] show ?thesis by blast qed lemmas [elim] = is_functor.cat_obj_cf_comma_ArrE lemma (in is_functor) cat_cf_obj_comma_ArrD[dest]: assumes "[[a, b', f]\<^sub>\, [a', b'', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" and "b \\<^sub>\ \\Obj\" shows "b' = 0" and "b'' = 0" and "h = 0" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" and "[a, b', f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" and "[a', b'', f']\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" using cat_cf_obj_comma_ArrE[OF assms] by auto lemmas [dest] = is_functor.cat_cf_obj_comma_ArrD[rotated 1] lemma (in is_functor) cat_obj_cf_comma_ArrD[dest]: assumes "[[b', a, f]\<^sub>\, [b'', a', f']\<^sub>\, [h, g]\<^sub>\]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" and "b \\<^sub>\ \\Obj\" shows "b' = 0" and "b'' = 0" and "h = 0" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" and "[b', a, f]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" and "[b'', a', f']\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" using cat_obj_cf_comma_ArrE[OF assms] by auto lemmas [dest] = is_functor.cat_obj_cf_comma_ArrD subsubsection\Domain\ lemma cat_cf_obj_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\ b\Dom\)" unfolding cat_cf_obj_comma_def cat_comma_components by simp lemma cat_cf_obj_comma_Dom_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\ b\Dom\) = \ \<^sub>C\<^sub>F\ b\Arr\" unfolding cat_cf_obj_comma_def cat_comma_components by simp lemma cat_cf_obj_comma_Dom_app[cat_comma_cs_simps]: - assumes "F = [abf, a'b'f', gh]\<^sub>\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" - shows "\ \<^sub>C\<^sub>F\ b\Dom\\F\ = abf" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + shows "\ \<^sub>C\<^sub>F\ b\Dom\\ABF\ = A" using assms(2) unfolding assms(1) cat_cf_obj_comma_def cat_comma_components by simp lemma (in is_functor) cat_cf_obj_comma_Dom_vrange: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\ b\Dom\) \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Dom_vrange[ OF is_functor_axioms const, folded cat_cf_obj_comma_def ] ) qed lemma cat_obj_cf_comma_Dom_vsv[cat_comma_cs_intros]: "vsv (b \\<^sub>C\<^sub>F \\Dom\)" unfolding cat_obj_cf_comma_def cat_comma_components by simp lemma cat_obj_cf_comma_Dom_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (b \\<^sub>C\<^sub>F \\Dom\) = b \\<^sub>C\<^sub>F \\Arr\" unfolding cat_obj_cf_comma_def cat_comma_components by simp lemma cat_obj_cf_comma_Dom_app[cat_comma_cs_simps]: - assumes "F = [baf, b'a'f', gh]\<^sub>\" and "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" - shows "b \\<^sub>C\<^sub>F \\Dom\\F\ = baf" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + shows "b \\<^sub>C\<^sub>F \\Dom\\ABF\ = A" using assms(2) unfolding assms(1) cat_obj_cf_comma_def cat_comma_components by simp lemma (in is_functor) cat_obj_cf_comma_Dom_vrange: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (b \\<^sub>C\<^sub>F \\Dom\) \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Dom_vrange[ OF const is_functor_axioms, folded cat_obj_cf_comma_def ] ) qed subsubsection\Codomain\ lemma cat_cf_obj_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\ b\Cod\)" unfolding cat_cf_obj_comma_def cat_comma_components by simp lemma cat_cf_obj_comma_Cod_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\ b\Cod\) = \ \<^sub>C\<^sub>F\ b\Arr\" unfolding cat_cf_obj_comma_def cat_comma_components by simp lemma cat_cf_obj_comma_Cod_app[cat_comma_cs_simps]: - assumes "F = [abf, a'b'f', gh]\<^sub>\" and "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" - shows "\ \<^sub>C\<^sub>F\ b\Cod\\F\ = a'b'f'" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + shows "\ \<^sub>C\<^sub>F\ b\Cod\\ABF\ = B" using assms(2) unfolding assms(1) cat_cf_obj_comma_def cat_comma_components by (simp add: nat_omega_simps) lemma (in is_functor) cat_cf_obj_comma_Cod_vrange: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\ b\Cod\) \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Cod_vrange[ OF is_functor_axioms const, folded cat_cf_obj_comma_def ] ) qed lemma cat_obj_cf_comma_Cod_vsv[cat_comma_cs_intros]: "vsv (b \\<^sub>C\<^sub>F \\Cod\)" unfolding cat_obj_cf_comma_def cat_comma_components by simp lemma cat_obj_cf_comma_Cod_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (b \\<^sub>C\<^sub>F \\Cod\) = b \\<^sub>C\<^sub>F \\Arr\" unfolding cat_obj_cf_comma_def cat_comma_components by simp lemma cat_obj_cf_comma_Cod_app[cat_comma_cs_simps]: - assumes "F = [baf, b'a'f', gh]\<^sub>\" and "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" - shows "b \\<^sub>C\<^sub>F \\Cod\\F\ = b'a'f'" + assumes "ABF = [A, B, F]\<^sub>\" and "ABF \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + shows "b \\<^sub>C\<^sub>F \\Cod\\ABF\ = B" using assms(2) unfolding assms(1) cat_obj_cf_comma_def cat_comma_components by (simp add: nat_omega_simps) lemma (in is_functor) cat_obj_cf_comma_Cod_vrange: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (b \\<^sub>C\<^sub>F \\Dom\) \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Dom_vrange[ OF const is_functor_axioms, folded cat_obj_cf_comma_def ] ) qed subsubsection\Arrow with a domain and a codomain\ lemma (in is_functor) cat_cf_obj_comma_is_arrI[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "F = [abf, a'b'f', gh]\<^sub>\" - and "abf = [a, 0, f]\<^sub>\" - and "a'b'f' = [a', 0, f']\<^sub>\" - and "gh = [g, 0]\<^sub>\" + and "ABF = [A, B, F]\<^sub>\" + and "A = [a, 0, f]\<^sub>\" + and "B = [a', 0, f']\<^sub>\" + and "F = [g, 0]\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" - shows "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a'b'f'" + shows "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> B" proof(intro is_arrI) - from assms(1,6,7,8) show "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + from assms(1,6,7,8) show "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" by (cs_concl cs_simp: assms(2,3,4,5,9) cs_intro: cat_comma_cs_intros) - with assms(2) show "\ \<^sub>C\<^sub>F\ b\Dom\\F\ = abf" "\ \<^sub>C\<^sub>F\ b\Cod\\F\ = a'b'f'" + with assms(2) show "\ \<^sub>C\<^sub>F\ b\Dom\\ABF\ = A" "\ \<^sub>C\<^sub>F\ b\Cod\\ABF\ = B" by (cs_concl cs_simp: cat_comma_cs_simps)+ qed lemmas [cat_comma_cs_intros] = is_functor.cat_cf_obj_comma_is_arrI lemma (in is_functor) cat_obj_cf_comma_is_arrI[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "F = [baf, b'a'f', gh]\<^sub>\" - and "baf = [0, a, f]\<^sub>\" - and "b'a'f' = [0, a', f']\<^sub>\" - and "gh = [0, g]\<^sub>\" + and "ABF = [A, B, F]\<^sub>\" + and "A = [0, a, f]\<^sub>\" + and "B = [0, a', f']\<^sub>\" + and "F = [0, g]\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" - shows "F : baf \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b'a'f'" + shows "ABF : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" proof(intro is_arrI) - from assms(1,6,7,8) show "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + from assms(1,6,7,8) show "ABF \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" by (cs_concl cs_simp: assms(2,3,4,5,9) cs_intro: cat_comma_cs_intros) - with assms(2) show "b \\<^sub>C\<^sub>F \\Dom\\F\ = baf" "b \\<^sub>C\<^sub>F \\Cod\\F\ = b'a'f'" + with assms(2) show "b \\<^sub>C\<^sub>F \\Dom\\ABF\ = A" "b \\<^sub>C\<^sub>F \\Cod\\ABF\ = B" by (cs_concl cs_simp: cat_comma_cs_simps)+ qed lemmas [cat_comma_cs_intros] = is_functor.cat_obj_cf_comma_is_arrI lemma (in is_functor) cat_cf_obj_comma_is_arrD[dest]: assumes "[[a, b', f]\<^sub>\, [a', b'', f']\<^sub>\, [g, h]\<^sub>\]\<^sub>\ : [a, b', f]\<^sub>\ \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> [a', b'', f']\<^sub>\" and "b \\<^sub>\ \\Obj\" - shows "b' = []\<^sub>\" - and "b'' = []\<^sub>\" - and "h = []\<^sub>\" + shows "b' = 0" + and "b'' = 0" + and "h = 0" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" and "[a, b', f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" and "[a', b'', f']\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" by (intro cat_cf_obj_comma_ArrD[OF is_arrD(1)[OF assms(1)] assms(2)])+ lemma (in is_functor) cat_obj_cf_comma_is_arrD[dest]: assumes "[[b', a, f]\<^sub>\, [b'', a', f']\<^sub>\, [h, g]\<^sub>\]\<^sub>\ : [b', a, f]\<^sub>\ \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> [b'', a', f']\<^sub>\" and "b \\<^sub>\ \\Obj\" shows "b' = 0" and "b'' = 0" and "h = 0" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" and "[b', a, f]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" and "[b'', a', f']\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" by (intro cat_obj_cf_comma_ArrD[OF is_arrD(1)[OF assms(1)] assms(2)])+ lemmas [dest] = is_functor.cat_obj_cf_comma_is_arrD lemma (in is_functor) cat_cf_obj_comma_is_arrE[elim]: - assumes "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a'b'f'" and "b \\<^sub>\ \\Obj\" + assumes "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> B" and "b \\<^sub>\ \\Obj\" obtains a f a' f' g - where "F = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" - and "abf = [a, 0, f]\<^sub>\" - and "a'b'f' = [a', 0, f']\<^sub>\" + where "ABF = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" + and "A = [a, 0, f]\<^sub>\" + and "B = [a', 0, f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" - and "abf \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" - and "a'b'f' \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and "B \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" proof- - note F = is_arrD[OF assms(1)] - from F(1) obtain abf' a'b'f'' a f a' f' g - where F_def: "F = [abf', a'b'f'', [g, 0]\<^sub>\]\<^sub>\" - and abf'_def: "abf' = [a, 0, f]\<^sub>\" - and a'b'f''_def: "a'b'f'' = [a', 0, f']\<^sub>\" + note ABF = is_arrD[OF assms(1)] + from ABF(1) obtain C D a f a' f' g + where ABF_def: "ABF = [C, D, [g, 0]\<^sub>\]\<^sub>\" + and C_def: "C = [a, 0, f]\<^sub>\" + and D_def: "D = [a', 0, f']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" and f_def: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" - and abf': "abf' \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" - and a'b'f'': "a'b'f'' \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and C: "C \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + and D: "D \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" by (elim cat_cf_obj_comma_ArrE[OF _ assms(2)]) - from F(2) assms(2) abf'_def a'b'f''_def g f f' f_def have "abf' = abf" - unfolding F_def + from ABF(2) assms(2) C_def D_def g f f' f_def have [simp]: "C = A" + unfolding ABF_def by (cs_prems cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) - from F(3) assms(2) abf'_def a'b'f''_def g f f' f_def have "a'b'f'' = a'b'f'" - unfolding F_def + from ABF(3) assms(2) C_def D_def g f f' f_def have [simp]: "D = B" + unfolding ABF_def by (cs_prems cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) - from that F_def abf'_def a'b'f''_def g f f' f_def abf' a'b'f'' show ?thesis - unfolding \abf' = abf\ \a'b'f'' = a'b'f'\ by auto + from that ABF_def C_def D_def g f f' f_def C D show ?thesis by auto qed lemmas [elim] = is_functor.cat_cf_obj_comma_is_arrE lemma (in is_functor) cat_obj_cf_comma_is_arrE[elim]: - assumes "F : baf \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b'a'f'" - and "b \\<^sub>\ \\Obj\" + assumes "ABF : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" and "b \\<^sub>\ \\Obj\" obtains a f a' f' g - where "F = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" - and "baf = [0, a, f]\<^sub>\" - and "b'a'f' = [0, a', f']\<^sub>\" + where "ABF = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" + and "A = [0, a, f]\<^sub>\" + and "B = [0, a', f']\<^sub>\" and "g : a \\<^bsub>\\<^esub> a'" and "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" - and "baf \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" - and "b'a'f' \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and "B \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" proof- - note F = is_arrD[OF assms(1)] - from F(1) obtain baf' b'a'f'' a f a' f' g - where F_def: "F = [baf', b'a'f'', [0, g]\<^sub>\]\<^sub>\" - and baf'_def: "baf' = [0, a, f]\<^sub>\" - and b'a'f''_def: "b'a'f'' = [0, a', f']\<^sub>\" + note ABF = is_arrD[OF assms(1)] + from ABF(1) obtain C D a f a' f' g + where ABF_def: "ABF = [C, D, [0, g]\<^sub>\]\<^sub>\" + and C_def: "C = [0, a, f]\<^sub>\" + and D_def: "D = [0, a', f']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and f: "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" and f'_def: "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> f = f'" - and baf': "baf' \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" - and b'a'f'': "b'a'f'' \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and C: "C \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + and D: "D \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" by (elim cat_obj_cf_comma_ArrE[OF _ assms(2)]) - from F(2) assms(2) baf'_def b'a'f''_def g f f' f'_def have "baf' = baf" - unfolding F_def + from ABF(2) assms(2) C_def D_def g f f' f'_def have [simp]: "C = A" + unfolding ABF_def by (cs_prems cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) - from F(3) assms(2) baf'_def b'a'f''_def g f f' f'_def have "b'a'f'' = b'a'f'" - unfolding F_def + from ABF(3) assms(2) C_def D_def g f f' f'_def have [simp]: "D = B" + unfolding ABF_def by (cs_prems cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) - from that F_def baf'_def b'a'f''_def g f f' f'_def baf' b'a'f'' show ?thesis - unfolding \baf' = baf\ \b'a'f'' = b'a'f'\ by auto + from that ABF_def C_def D_def g f f' f'_def C D show ?thesis by auto qed lemmas [elim] = is_functor.cat_obj_cf_comma_is_arrE subsubsection\Composition\ lemma cat_cf_obj_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\ b\Comp\)" unfolding cat_cf_obj_comma_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma cat_obj_cf_comma_Comp_vsv[cat_comma_cs_intros]: "vsv (b \\<^sub>C\<^sub>F \\Comp\)" unfolding cat_obj_cf_comma_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma (in is_functor) cat_cf_obj_comma_Comp_app[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" - and "G = [a'b'f', a''b''f'', [g', h']\<^sub>\]\<^sub>\" - and "F = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a'b'f'" - shows "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> F = [abf, a''b''f'', [g' \\<^sub>A\<^bsub>\\<^esub> g, 0]\<^sub>\]\<^sub>\" + and "BCG = [B, C, [g', h']\<^sub>\]\<^sub>\" + and "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "BCG : B \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> C" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> B" + shows "BCG \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> ABF = [A, C, [g' \\<^sub>A\<^bsub>\\<^esub> g, 0]\<^sub>\]\<^sub>\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from assms(4) obtain a f a' f' g - where G_def: "G = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" + where BCG_def: "BCG = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" by (elim cat_cf_obj_comma_is_arrE[OF _ assms(1)]) from assms(5) obtain a f a' f' g - where F_def: "F = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" + where ABF_def: "ABF = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" by (elim cat_cf_obj_comma_is_arrE[OF _ assms(1)]) - from assms(2)[unfolded G_def] assms(3)[unfolded F_def] have [cat_cs_simps]: + from assms(2)[unfolded BCG_def] assms(3)[unfolded ABF_def] have [cat_cs_simps]: "h' = 0" "h = 0" by simp_all have "h' \\<^sub>A\<^bsub>cat_1 0 0\<^esub> h = 0" by (cs_concl cs_simp: cat_cs_simps) show ?thesis by ( rule cat_comma_Comp_app [ OF is_functor_axioms const assms(2,3) assms(4)[unfolded cat_cf_obj_comma_def] assms(5)[unfolded cat_cf_obj_comma_def], folded cat_cf_obj_comma_def, unfolded cat_cs_simps ] ) qed lemma (in is_functor) cat_obj_cf_comma_Comp_app[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" - and "G = [b'a'f', b''a''f'', [h', g']\<^sub>\]\<^sub>\" - and "F = [baf, b'a'f', [h, g]\<^sub>\]\<^sub>\" - and "G : b'a'f' \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b''a''f''" - and "F : baf \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b'a'f'" - shows "G \\<^sub>A\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> F = [baf, b''a''f'', [0, g' \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\]\<^sub>\" + and "BCG = [B, C, [h', g']\<^sub>\]\<^sub>\" + and "ABF = [A, B, [h, g]\<^sub>\]\<^sub>\" + and "BCG : B \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> C" + and "ABF : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" + shows "BCG \\<^sub>A\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> ABF = [A, C, [0, g' \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\]\<^sub>\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cs_intro: vempty_is_zet cat_cs_intros) from assms(4) obtain a f a' f' g - where G_def: "G = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" + where BCG_def: "BCG = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" by (elim cat_obj_cf_comma_is_arrE[OF _ assms(1)]) from assms(5) obtain a f a' f' g - where F_def: "F = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" + where ABF_def: "ABF = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, g]\<^sub>\]\<^sub>\" by (elim cat_obj_cf_comma_is_arrE[OF _ assms(1)]) - from assms(2)[unfolded G_def] assms(3)[unfolded F_def] have [cat_cs_simps]: + from assms(2)[unfolded BCG_def] assms(3)[unfolded ABF_def] have [cat_cs_simps]: "h' = 0" "h = 0" by simp_all have "h' \\<^sub>A\<^bsub>cat_1 0 0\<^esub> h = 0" by (cs_concl cs_simp: cat_cs_simps) show ?thesis by ( rule cat_comma_Comp_app [ OF const is_functor_axioms assms(2,3) assms(4)[unfolded cat_obj_cf_comma_def] assms(5)[unfolded cat_obj_cf_comma_def], folded cat_obj_cf_comma_def, unfolded cat_cs_simps ] ) qed lemma (in is_functor) cat_cf_obj_comma_Comp_is_arr[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "G : a'b'f' \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a''b''f''" - and "F : abf \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a'b'f'" - shows "G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> F : abf \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> a''b''f''" + and "BCG : B \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> C" + and "ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> B" + shows "BCG \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> ABF : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> C" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Comp_is_arr [ OF is_functor_axioms const assms(2)[unfolded cat_cf_obj_comma_def] assms(3)[unfolded cat_cf_obj_comma_def], folded cat_cf_obj_comma_def ] ) qed lemma (in is_functor) cat_obj_cf_comma_Comp_is_arr[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" - and "G : b'a'f' \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b''a''f''" - and "F : baf \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b'a'f'" - shows "G \\<^sub>A\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> F : baf \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> b''a''f''" + and "BCG : B \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> C" + and "ABF : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" + shows "BCG \\<^sub>A\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> ABF : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> C" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_Comp_is_arr [ OF const is_functor_axioms assms(2)[unfolded cat_obj_cf_comma_def] assms(3)[unfolded cat_obj_cf_comma_def], folded cat_obj_cf_comma_def ] ) qed subsubsection\Identity\ lemma cat_cf_obj_comma_CId_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\ b\CId\)" unfolding cat_cf_obj_comma_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma cat_obj_cf_comma_CId_vsv[cat_comma_cs_intros]: "vsv (b \\<^sub>C\<^sub>F \\CId\)" unfolding cat_obj_cf_comma_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma (in is_functor) cat_cf_obj_comma_CId_vdomain[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\ b\CId\) = \ \<^sub>C\<^sub>F\ b\Obj\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule cat_comma_CId_vdomain[ OF is_functor_axioms const, folded cat_cf_obj_comma_def ] ) qed lemma (in is_functor) cat_obj_cf_comma_CId_vdomain[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (b \\<^sub>C\<^sub>F \\CId\) = b \\<^sub>C\<^sub>F \\Obj\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show "\\<^sub>\ (b \\<^sub>C\<^sub>F \\CId\) = b \\<^sub>C\<^sub>F \\Obj\" by ( rule cat_comma_CId_vdomain[ OF const is_functor_axioms, folded cat_obj_cf_comma_def ] ) qed lemma (in is_functor) cat_cf_obj_comma_CId_app[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" and "A = [a, b', f]\<^sub>\" and "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" shows "\ \<^sub>C\<^sub>F\ b\CId\\A\ = [A, A, [\\CId\\a\, 0]\<^sub>\]\<^sub>\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from assms(3,2) have b'_def: "b' = 0" by (auto elim: cat_cf_obj_comma_ObjE[OF _ assms(1)]) have [cat_cs_simps]: "cat_1 0 0\CId\\b'\ = 0" unfolding cat_1_components b'_def by simp show ?thesis by ( rule cat_comma_CId_app [ OF is_functor_axioms const assms(2,3)[unfolded cat_cf_obj_comma_def], unfolded cat_cf_obj_comma_def[symmetric] cat_cs_simps ] ) qed lemma (in is_functor) cat_obj_cf_comma_CId_app[cat_comma_cs_simps]: assumes "b \\<^sub>\ \\Obj\" and "A = [b', a, f]\<^sub>\" and "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" shows "b \\<^sub>C\<^sub>F \\CId\\A\ = [A, A, [0, \\CId\\a\]\<^sub>\]\<^sub>\" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) from assms(3,2) have b'_def: "b' = 0" by (auto elim: cat_obj_cf_comma_ObjE[OF _ assms(1)]) have [cat_cs_simps]: "cat_1 0 0\CId\\b'\ = 0" unfolding cat_1_components b'_def by simp show ?thesis by ( rule cat_comma_CId_app [ OF const is_functor_axioms assms(2,3)[unfolded cat_obj_cf_comma_def], unfolded cat_obj_cf_comma_def[symmetric] cat_cs_simps ] ) qed subsubsection\ Comma categories constructed from a functor and an object are categories \ lemma (in is_functor) category_cat_cf_obj_comma[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" shows "category \ (\ \<^sub>C\<^sub>F\ b)" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule category_cat_comma[ OF is_functor_axioms const, folded cat_cf_obj_comma_def ] ) qed lemmas [cat_comma_cs_intros] = is_functor.category_cat_cf_obj_comma lemma (in is_functor) category_cat_obj_cf_comma[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" shows "category \ (b \\<^sub>C\<^sub>F \)" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: vempty_is_zet cat_cs_intros) show ?thesis by ( rule category_cat_comma[ OF const is_functor_axioms, folded cat_obj_cf_comma_def ] ) qed lemmas [cat_comma_cs_intros] = is_functor.category_cat_obj_cf_comma subsubsection\Tiny comma categories constructed from a functor and an object\ lemma (in is_tm_functor) tiny_category_cat_cf_obj_comma[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" shows "tiny_category \ (\ \<^sub>C\<^sub>F\ b)" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by ( cs_concl cs_intro: vempty_is_zet cat_small_cs_intros cat_cs_intros ) show ?thesis by ( rule tiny_category_cat_comma[ OF is_tm_functor_axioms const, folded cat_cf_obj_comma_def ] ) qed lemma (in is_tm_functor) tiny_category_cat_obj_cf_comma[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" shows "tiny_category \ (b \\<^sub>C\<^sub>F \)" proof- from assms(1) have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by ( cs_concl cs_intro: vempty_is_zet cat_small_cs_intros cat_cs_intros ) show ?thesis by ( rule tiny_category_cat_comma[ OF const is_tm_functor_axioms, folded cat_obj_cf_comma_def ] ) qed subsection\ +Opposite comma category functors for the comma categories +constructed from a functor and an object +\ + + +subsubsection\Definitions and elementary properties\ + +definition op_cf_obj_comma :: "V \ V \ V" + where "op_cf_obj_comma \ b = + op_cf_comma \ (cf_const (cat_1 0 0) (\\HomCod\) b)" + +definition op_obj_cf_comma :: "V \ V \ V" + where "op_obj_cf_comma b \ = + op_cf_comma (cf_const (cat_1 0 0) (\\HomCod\) b) \" + + +text\Alternative forms of the definitions.\ + +lemma (in is_functor) op_cf_obj_comma_def: + "op_cf_obj_comma \ b = op_cf_comma \ (cf_const (cat_1 0 0) \ b)" + unfolding op_cf_obj_comma_def cat_cs_simps by simp + +lemma (in is_functor) op_obj_cf_comma_def: + "op_obj_cf_comma b \ = op_cf_comma (cf_const (cat_1 0 0) \ b) \" + unfolding op_obj_cf_comma_def cat_cs_simps by simp + + +subsubsection\Object map\ + +lemma op_cf_obj_comma_ObjMap_vsv[cat_comma_cs_intros]: + "vsv (op_cf_obj_comma \ b\ObjMap\)" + unfolding op_cf_obj_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) + +lemma op_obj_cf_comma_ObjMap_vsv[cat_comma_cs_intros]: + "vsv (op_obj_cf_comma b \\ObjMap\)" + unfolding op_obj_cf_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) + +lemma (in is_functor) op_cf_obj_comma_ObjMap_vdomain[cat_comma_cs_simps]: + "\\<^sub>\ (op_cf_obj_comma \ b\ObjMap\) = \ \<^sub>C\<^sub>F\ b\Obj\" + unfolding op_cf_obj_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cat_cf_obj_comma_def[symmetric]) + +lemma (in is_functor) op_obj_cf_comma_ObjMap_vdomain[cat_comma_cs_simps]: + "\\<^sub>\ (op_obj_cf_comma b \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + unfolding op_obj_cf_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cat_obj_cf_comma_def[symmetric]) + +lemma (in is_functor) op_cf_obj_comma_ObjMap_app[cat_comma_cs_simps]: + assumes "A = [a, 0, f]\<^sub>\" and "b \\<^sub>\ \\Obj\" and "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + shows "op_cf_obj_comma \ b\ObjMap\\A\ = [0, a, f]\<^sub>\" +proof- + have a: "a \\<^sub>\ \\Obj\" and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" + by (intro cat_cf_obj_comma_ObjD[OF assms(3)[unfolded assms(1)] assms(2)])+ + from assms(2) a f show ?thesis + using assms(2) + unfolding assms(1) op_cf_obj_comma_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros + ) +qed + +lemma (in is_functor) op_obj_cf_comma_ObjMap_app[cat_comma_cs_simps]: + assumes "A = [0, a, f]\<^sub>\" and "b \\<^sub>\ \\Obj\" and "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + shows "op_obj_cf_comma b \ \ObjMap\\A\ = [a, 0, f]\<^sub>\" +proof- + have a: "a \\<^sub>\ \\Obj\" and f: "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" + by (intro cat_obj_cf_comma_ObjD[OF assms(3)[unfolded assms(1)] assms(2)])+ + from assms(2) a f show ?thesis + using assms(2) + unfolding assms(1) op_obj_cf_comma_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros + ) +qed + + +subsubsection\Arrow map\ + +lemma op_cf_obj_comma_ArrMap_vsv[cat_comma_cs_intros]: + "vsv (op_cf_obj_comma \ b\ArrMap\)" + unfolding op_cf_obj_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) + +lemma op_obj_cf_comma_ArrMap_vsv[cat_comma_cs_intros]: + "vsv (op_obj_cf_comma b \\ArrMap\)" + unfolding op_obj_cf_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_comma_cs_intros) + +lemma (in is_functor) op_cf_obj_comma_ArrMap_vdomain[cat_comma_cs_simps]: + "\\<^sub>\ (op_cf_obj_comma \ b\ArrMap\) = \ \<^sub>C\<^sub>F\ b\Arr\" + unfolding op_cf_obj_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cat_cf_obj_comma_def[symmetric]) + +lemmas [cat_comma_cs_simps] = is_functor.op_cf_obj_comma_ArrMap_vdomain + +lemma (in is_functor) op_obj_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]: + "\\<^sub>\ (op_obj_cf_comma b \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + unfolding op_obj_cf_comma_def + by (cs_concl cs_simp: cat_comma_cs_simps cat_obj_cf_comma_def[symmetric]) + +lemmas [cat_comma_cs_simps] = is_functor.op_obj_cf_comma_ArrMap_vdomain + +lemma (in is_functor) op_cf_obj_comma_ArrMap_app[cat_comma_cs_simps]: + assumes "ABF = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [g, 0]\<^sub>\]\<^sub>\" + and "b \\<^sub>\ \\Obj\" + and "ABF \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + shows "op_cf_obj_comma \ b\ArrMap\\ABF\ = [[0, a', f']\<^sub>\, [0, a, f]\<^sub>\, [0, g]\<^sub>\]\<^sub>\" +proof- + from assms(3) have g: "g : a \\<^bsub>\\<^esub> a'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" + and [cat_comma_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\ = f" + by (intro cat_cf_obj_comma_ArrD[OF assms(3)[unfolded assms(1)] assms(2)])+ + from assms(2) g f f' show ?thesis + unfolding assms(1) op_cf_obj_comma_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_1_CId_app + cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_1_is_arrI + ) +qed + +lemmas [cat_comma_cs_simps] = is_functor.op_cf_obj_comma_ArrMap_app + +lemma (in is_functor) op_obj_cf_comma_ArrMap_app[cat_comma_cs_simps]: + assumes "ABF = [[0, a, f]\<^sub>\, [0, a', f']\<^sub>\, [0, h]\<^sub>\]\<^sub>\" + and "b \\<^sub>\ \\Obj\" + and "ABF \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + shows "op_obj_cf_comma b \\ArrMap\\ABF\ = [[a', 0, f']\<^sub>\, [a, 0, f]\<^sub>\, [h, 0]\<^sub>\]\<^sub>\" +proof- + from assms(3) have h: "h : a \\<^bsub>\\<^esub> a'" + and f: "f : b \\<^bsub>\\<^esub> \\ObjMap\\a\" + and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" + and [cat_comma_cs_simps]: "\\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f = f'" + by (intro cat_obj_cf_comma_ArrD[OF assms(3)[unfolded assms(1)] assms(2)])+ + from assms(2) h f f' show ?thesis + unfolding assms(1) op_obj_cf_comma_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_1_CId_app + cs_intro: V_cs_intros cat_cs_intros cat_comma_cs_intros cat_1_is_arrI + ) +qed + +lemmas [cat_comma_cs_simps] = is_functor.op_obj_cf_comma_ArrMap_app + + +subsubsection\ +Opposite comma category functors for the comma categories +constructed from a functor and an object are isomorphisms of categories +\ + +lemma (in is_functor) op_cf_obj_comma_is_iso_functor: + assumes "b \\<^sub>\ \\Obj\" + shows "op_cf_obj_comma \ b : op_cat (\ \<^sub>C\<^sub>F\ b) \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F (op_cf \)" +proof- + from assms have cf_const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros) + note cat_obj_cf_comma_def = + is_functor.cat_obj_cf_comma_def[ + OF is_functor_op, unfolded cat_op_simps + ] + show ?thesis + by + ( + rule op_cf_comma_is_iso_functor + [ + OF is_functor_axioms cf_const, + folded cat_cf_obj_comma_def op_cf_obj_comma_def, + unfolded cat_op_simps, + folded cat_obj_cf_comma_def + ] + ) +qed + +lemma (in is_functor) op_cf_obj_comma_is_iso_functor'[cat_comma_cs_intros]: + assumes "b \\<^sub>\ \\Obj\" + and "\' = op_cat (\ \<^sub>C\<^sub>F\ b)" + and "\' = b \\<^sub>C\<^sub>F (op_cf \)" + shows "op_cf_obj_comma \ b : \' \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" + using assms(1) unfolding assms(2,3) by (rule op_cf_obj_comma_is_iso_functor) + +lemmas [cat_comma_cs_intros] = is_functor.op_cf_obj_comma_is_iso_functor' + +lemma (in is_functor) op_cf_obj_comma_is_functor: + assumes "b \\<^sub>\ \\Obj\" + shows "op_cf_obj_comma \ b : op_cat (\ \<^sub>C\<^sub>F\ b) \\\<^sub>C\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F (op_cf \)" + by (rule is_iso_functorD(1)[OF op_cf_obj_comma_is_iso_functor[OF assms]]) + +lemma (in is_functor) op_cf_obj_comma_is_functor'[cat_comma_cs_intros]: + assumes "b \\<^sub>\ \\Obj\" + and "\' = op_cat (\ \<^sub>C\<^sub>F\ b)" + and "\' = b \\<^sub>C\<^sub>F (op_cf \)" + shows "op_cf_obj_comma \ b : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + using assms(1) unfolding assms(2,3) by (rule op_cf_obj_comma_is_functor) + +lemmas [cat_comma_cs_intros] = is_functor.op_cf_obj_comma_is_functor' + +lemma (in is_functor) op_obj_cf_comma_is_iso_functor: + assumes "b \\<^sub>\ \\Obj\" + shows "op_obj_cf_comma b \ : op_cat (b \\<^sub>C\<^sub>F \) \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> (op_cf \) \<^sub>C\<^sub>F\ b" +proof- + from assms have cf_const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros) + note cat_cf_obj_comma_def = + is_functor.cat_cf_obj_comma_def[ + OF is_functor_op, unfolded cat_op_simps + ] + show ?thesis + by + ( + rule op_cf_comma_is_iso_functor + [ + OF cf_const is_functor_axioms, + folded cat_obj_cf_comma_def op_obj_cf_comma_def, + unfolded cat_op_simps, + folded cat_cf_obj_comma_def + ] + ) +qed + +lemma (in is_functor) op_obj_cf_comma_is_iso_functor'[cat_comma_cs_intros]: + assumes "b \\<^sub>\ \\Obj\" + and "\' = op_cat (b \\<^sub>C\<^sub>F \)" + and "\' = (op_cf \) \<^sub>C\<^sub>F\ b" + shows "op_obj_cf_comma b \ : \' \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" + using assms(1) unfolding assms(2,3) by (rule op_obj_cf_comma_is_iso_functor) + +lemmas [cat_comma_cs_intros] = is_functor.op_obj_cf_comma_is_iso_functor' + +lemma (in is_functor) op_obj_cf_comma_is_functor: + assumes "b \\<^sub>\ \\Obj\" + shows "op_obj_cf_comma b \ : op_cat (b \\<^sub>C\<^sub>F \) \\\<^sub>C\<^bsub>\\<^esub> (op_cf \) \<^sub>C\<^sub>F\ b" + by (rule is_iso_functorD(1)[OF op_obj_cf_comma_is_iso_functor[OF assms]]) + +lemma (in is_functor) op_obj_cf_comma_is_functor'[cat_comma_cs_intros]: + assumes "b \\<^sub>\ \\Obj\" + and "\' = op_cat (b \\<^sub>C\<^sub>F \)" + and "\' = (op_cf \) \<^sub>C\<^sub>F\ b" + shows "op_obj_cf_comma b \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + using assms(1) unfolding assms(2,3) by (rule op_obj_cf_comma_is_functor) + + + +subsection\ Projections for comma categories constructed from a functor and an object \ subsubsection\Definitions and elementary properties\ definition cf_cf_obj_comma_proj :: "V \ V \ V" (\(_ \<^sub>C\<^sub>F\\<^sub>O _)\ [1000, 1000] 999) where "\ \<^sub>C\<^sub>F\\<^sub>O b \ \ \<^sub>C\<^sub>F\ (cf_const (cat_1 0 0) (\\HomCod\) b)" definition cf_obj_cf_comma_proj :: "V \ V \ V" (\(_ \<^sub>O\\<^sub>C\<^sub>F _)\ [1000, 1000] 999) where "b \<^sub>O\\<^sub>C\<^sub>F \ \ (cf_const (cat_1 0 0) (\\HomCod\) b) \\<^sub>C\<^sub>F \" text\Alternative forms of the definitions.\ lemma (in is_functor) cf_cf_obj_comma_proj_def: "\ \<^sub>C\<^sub>F\\<^sub>O b = \ \<^sub>C\<^sub>F\ (cf_const (cat_1 0 0) \ b)" unfolding cf_cf_obj_comma_proj_def cf_HomCod.. lemma (in is_functor) cf_obj_cf_comma_proj_def: "b \<^sub>O\\<^sub>C\<^sub>F \ = (cf_const (cat_1 0 0) \ b) \\<^sub>C\<^sub>F \" unfolding cf_obj_cf_comma_proj_def cf_HomCod.. text\Components.\ lemma (in is_functor) cf_cf_obj_comma_proj_components[cat_comma_cs_simps]: shows "\ \<^sub>C\<^sub>F\\<^sub>O b\HomDom\ = \ \<^sub>C\<^sub>F\ b" and "\ \<^sub>C\<^sub>F\\<^sub>O b\HomCod\ = \" unfolding cf_cf_obj_comma_proj_def cf_comma_proj_left_components cat_cf_obj_comma_def[symmetric] cat_cs_simps by simp_all lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_components lemma (in is_functor) cf_obj_cf_comma_proj_components[cat_comma_cs_simps]: shows "b \<^sub>O\\<^sub>C\<^sub>F \\HomDom\ = b \\<^sub>C\<^sub>F \" and "b \<^sub>O\\<^sub>C\<^sub>F \\HomCod\ = \" unfolding cf_obj_cf_comma_proj_def cf_comma_proj_right_components cat_obj_cf_comma_def[symmetric] cat_cs_simps by simp_all lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_components subsubsection\Object map\ lemma cf_cf_obj_comma_proj_ObjMap_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>O b\ObjMap\)" unfolding cf_cf_obj_comma_proj_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma cf_obj_cf_comma_proj_ObjMap_vsv[cat_comma_cs_intros]: "vsv (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\)" unfolding cf_obj_cf_comma_proj_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma (in is_functor) cf_cf_obj_comma_proj_ObjMap_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>O b\ObjMap\) = \ \<^sub>C\<^sub>F\ b\Obj\" unfolding cf_cf_obj_comma_proj_def cf_comma_proj_left_ObjMap_vdomain unfolding cf_cf_obj_comma_proj_def[symmetric] cf_comma_proj_left_components[symmetric] cat_comma_cs_simps by simp lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_vdomain lemma (in is_functor) cf_obj_cf_comma_proj_ObjMap_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" unfolding cf_obj_cf_comma_proj_def cf_comma_proj_right_ObjMap_vdomain unfolding cf_obj_cf_comma_proj_def[symmetric] cf_comma_proj_right_components[symmetric] cat_comma_cs_simps by simp lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ObjMap_vdomain lemma (in is_functor) cf_cf_obj_comma_proj_ObjMap_app[cat_comma_cs_simps]: assumes "A = [a, b', f]\<^sub>\" and "[a, b', f]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" shows "\ \<^sub>C\<^sub>F\\<^sub>O b\ObjMap\\A\ = a" by ( rule cf_comma_proj_left_ObjMap_app[ OF assms(1) assms(2)[unfolded cat_cf_obj_comma_def], folded cf_cf_obj_comma_proj_def ] ) lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_app lemma (in is_functor) cf_obj_cf_comma_proj_ObjMap_app[cat_comma_cs_simps]: assumes "A = [b', a, f]\<^sub>\" and "[b', a, f]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" shows "b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\\A\ = a" by ( rule cf_comma_proj_right_ObjMap_app[ OF assms(1) assms(2)[unfolded cat_obj_cf_comma_def], folded cf_obj_cf_comma_proj_def ] ) lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ObjMap_app subsubsection\Arrow map\ lemma cf_cf_obj_comma_proj_ArrMap_vsv[cat_comma_cs_intros]: "vsv (\ \<^sub>C\<^sub>F\\<^sub>O b\ArrMap\)" unfolding cf_cf_obj_comma_proj_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma cf_obj_cf_comma_proj_ArrMap_vsv[cat_comma_cs_intros]: "vsv (b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\)" unfolding cf_obj_cf_comma_proj_def by (cs_concl cs_intro: cat_comma_cs_intros) lemma (in is_functor) cf_cf_obj_comma_proj_ArrMap_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>O b\ArrMap\) = \ \<^sub>C\<^sub>F\ b\Arr\" unfolding cf_cf_obj_comma_proj_def cf_comma_proj_left_ArrMap_vdomain unfolding cf_cf_obj_comma_proj_def[symmetric] cf_comma_proj_left_components[symmetric] cat_comma_cs_simps by simp lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ObjMap_vdomain lemma (in is_functor) cf_obj_cf_comma_proj_ArrMap_vdomain[cat_comma_cs_simps]: "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" unfolding cf_obj_cf_comma_proj_def cf_comma_proj_right_ArrMap_vdomain unfolding cf_obj_cf_comma_proj_def[symmetric] cf_comma_proj_right_components[symmetric] cat_comma_cs_simps by simp lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ArrMap_vdomain lemma (in is_functor) cf_cf_obj_comma_proj_ArrMap_app[cat_comma_cs_simps]: - assumes "A = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "[abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" - shows "\ \<^sub>C\<^sub>F\\<^sub>O b\ArrMap\\A\ = g" + assumes "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "[A, B, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + shows "\ \<^sub>C\<^sub>F\\<^sub>O b\ArrMap\\ABF\ = g" by ( rule cf_comma_proj_left_ArrMap_app[ OF assms(1) assms(2)[unfolded cat_cf_obj_comma_def], folded cf_cf_obj_comma_proj_def ] ) lemmas [cat_comma_cs_simps] = is_functor.cf_cf_obj_comma_proj_ArrMap_app lemma (in is_functor) cf_obj_cf_comma_proj_ArrMap_app[cat_comma_cs_simps]: - assumes "A = [abf, a'b'f', [g, h]\<^sub>\]\<^sub>\" - and "[abf, a'b'f', [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" - shows "b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\\A\ = h" + assumes "ABF = [A, B, [g, h]\<^sub>\]\<^sub>\" + and "[A, B, [g, h]\<^sub>\]\<^sub>\ \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + shows "b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\\ABF\ = h" by ( rule cf_comma_proj_right_ArrMap_app[ OF assms(1) assms(2)[unfolded cat_obj_cf_comma_def], folded cf_obj_cf_comma_proj_def ] ) lemmas [cat_comma_cs_simps] = is_functor.cf_obj_cf_comma_proj_ArrMap_app subsubsection\Projections for a comma category are functors\ lemma (in is_functor) cf_cf_obj_comma_proj_is_functor: assumes "b \\<^sub>\ \\Obj\" shows "\ \<^sub>C\<^sub>F\\<^sub>O b : \ \<^sub>C\<^sub>F\ b \\\<^sub>C\<^bsub>\\<^esub> \" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_cs_intros) show ?thesis by ( rule cf_comma_proj_left_is_functor[ OF is_functor_axioms const, folded cf_cf_obj_comma_proj_def cat_cf_obj_comma_def ] ) qed lemma (in is_functor) cf_cf_obj_comma_proj_is_functor'[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" and "\' = \ \<^sub>C\<^sub>F\ b" shows "\ \<^sub>C\<^sub>F\\<^sub>O b : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_cf_obj_comma_proj_is_functor) lemmas [cat_comma_cs_intros] = is_functor.cf_cf_obj_comma_proj_is_functor' lemma (in is_functor) cf_obj_cf_comma_proj_is_functor: assumes "b \\<^sub>\ \\Obj\" shows "b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_cs_intros) show ?thesis by ( rule cf_comma_proj_right_is_functor[ OF const is_functor_axioms, folded cf_obj_cf_comma_proj_def cat_obj_cf_comma_def ] ) qed lemma (in is_functor) cf_obj_cf_comma_proj_is_functor'[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" and "\' = b \\<^sub>C\<^sub>F \" shows "b \<^sub>O\\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_obj_cf_comma_proj_is_functor) lemmas [cat_comma_cs_intros] = is_functor.cf_obj_cf_comma_proj_is_functor' +subsubsection\ +Opposite projections for comma categories constructed from a functor +and an object +\ + +lemma (in is_functor) op_cf_cf_obj_comma_proj: + assumes "b \\<^sub>\ \\Obj\" + shows "op_cf (\ \<^sub>C\<^sub>F\\<^sub>O b) = b \<^sub>O\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ b" +proof- + from assms have cf_const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros) + show ?thesis + by + ( + rule op_cf_comma_proj_left + [ + OF is_functor_axioms cf_const, + unfolded cat_op_simps, + folded + cf_cf_obj_comma_proj_def + op_cf_obj_comma_def + is_functor.cf_obj_cf_comma_proj_def[ + OF is_functor_op, unfolded cat_op_simps + ] + ] + ) +qed + +lemma (in is_functor) op_cf_obj_cf_comma_proj: + assumes "b \\<^sub>\ \\Obj\" + shows "op_cf (b \<^sub>O\\<^sub>C\<^sub>F \) = (op_cf \) \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F op_obj_cf_comma b \" +proof- + from assms have cf_const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros) + show ?thesis + by + ( + rule op_cf_comma_proj_right + [ + OF cf_const is_functor_axioms, + unfolded cat_op_simps, + folded + cf_obj_cf_comma_proj_def + op_obj_cf_comma_def + is_functor.cf_cf_obj_comma_proj_def[ + OF is_functor_op, unfolded cat_op_simps + ] + ] + ) +qed + + subsubsection\Projections for a tiny comma category\ lemma (in is_tm_functor) cf_cf_obj_comma_proj_is_tm_functor: assumes "b \\<^sub>\ \\Obj\" shows "\ \<^sub>C\<^sub>F\\<^sub>O b : \ \<^sub>C\<^sub>F\ b \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_small_cs_intros cat_cs_intros) show ?thesis by ( rule cf_comma_proj_left_is_tm_functor[ OF is_tm_functor_axioms const, folded cf_cf_obj_comma_proj_def cat_cf_obj_comma_def ] ) qed lemma (in is_tm_functor) cf_cf_obj_comma_proj_is_tm_functor'[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" and "\b = \ \<^sub>C\<^sub>F\ b" shows "\ \<^sub>C\<^sub>F\\<^sub>O b : \b \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_cf_obj_comma_proj_is_tm_functor) lemmas [cat_comma_cs_intros] = is_tm_functor.cf_cf_obj_comma_proj_is_tm_functor' lemma (in is_tm_functor) cf_obj_cf_comma_proj_is_tm_functor: assumes "b \\<^sub>\ \\Obj\" shows "b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" proof- from assms have const: "cf_const (cat_1 0 0) \ b : cat_1 0 0 \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_small_cs_intros cat_cs_intros) show ?thesis by ( rule cf_comma_proj_right_is_tm_functor[ OF const is_tm_functor_axioms, folded cf_obj_cf_comma_proj_def cat_obj_cf_comma_def ] ) qed lemma (in is_tm_functor) cf_obj_cf_comma_proj_is_tm_functor'[cat_comma_cs_intros]: assumes "b \\<^sub>\ \\Obj\" and "\' = b \\<^sub>C\<^sub>F \" shows "b \<^sub>O\\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^sub>.\<^sub>t\<^sub>m\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_obj_cf_comma_proj_is_tm_functor) lemmas [cat_comma_cs_intros] = is_tm_functor.cf_obj_cf_comma_proj_is_tm_functor' subsection\Comma functors\ subsubsection\Definition and elementary properties\ text\See Theorem 1 in Chapter X-3 in \cite{mac_lane_categories_2010}.\ -definition cf_cf_arr_comma :: "V \ V \ V" +definition cf_arr_cf_comma :: "V \ V \ V" (\(_ \<^sub>A\\<^sub>C\<^sub>F _)\ [1000, 1000] 999) where "g \<^sub>A\\<^sub>C\<^sub>F \ = [ (\A\\<^sub>\(\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \\Obj\. [0, A\1\<^sub>\\, A\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\), ( \F\\<^sub>\(\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \\Arr\. [ [0, F\0\\1\<^sub>\\, F\0\\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\, [0, F\1\<^sub>\\\1\<^sub>\\, F\1\<^sub>\\\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\, F\2\<^sub>\\ ]\<^sub>\ ), (\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \, (\\HomCod\\Dom\\g\) \\<^sub>C\<^sub>F \ ]\<^sub>\" +definition cf_cf_arr_comma :: "V \ V \ V" + (\(_ \<^sub>C\<^sub>F\\<^sub>A _)\ [1000, 1000] 999) + where "\ \<^sub>C\<^sub>F\\<^sub>A g = + [ + (\A\\<^sub>\\ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\)\Obj\. [A\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> A\2\<^sub>\\]\<^sub>\), + ( + \F\\<^sub>\\ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\)\Arr\. + [ + [F\0\\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> F\0\\2\<^sub>\\]\<^sub>\, + [F\1\<^sub>\\\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> F\1\<^sub>\\\2\<^sub>\\]\<^sub>\, + F\2\<^sub>\\ + ]\<^sub>\ + ), + \ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\), + \ \<^sub>C\<^sub>F\ (\\HomCod\\Cod\\g\) + ]\<^sub>\" + text\Components.\ -lemma cf_cf_arr_comma_components: +lemma cf_arr_cf_comma_components: shows "g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\ = (\A\\<^sub>\(\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \\Obj\. [0, A\1\<^sub>\\, A\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\)" and "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\ = ( \F\\<^sub>\(\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \\Arr\. [ [0, F\0\\1\<^sub>\\, F\0\\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\, [0, F\1\<^sub>\\\1\<^sub>\\, F\1\<^sub>\\\2\<^sub>\\ \\<^sub>A\<^bsub>\\HomCod\\<^esub> g]\<^sub>\, F\2\<^sub>\\ ]\<^sub>\ )" and "g \<^sub>A\\<^sub>C\<^sub>F \\HomDom\ = (\\HomCod\\Cod\\g\) \\<^sub>C\<^sub>F \" and "g \<^sub>A\\<^sub>C\<^sub>F \\HomCod\ = (\\HomCod\\Dom\\g\) \\<^sub>C\<^sub>F \" + unfolding cf_arr_cf_comma_def dghm_field_simps + by (simp_all add: nat_omega_simps) + +lemma cf_cf_arr_comma_components: + shows "\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\ = + (\A\\<^sub>\\ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\)\Obj\. [A\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> A\2\<^sub>\\]\<^sub>\)" + and "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\ = + ( + \F\\<^sub>\\ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\)\Arr\. + [ + [F\0\\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> F\0\\2\<^sub>\\]\<^sub>\, + [F\1\<^sub>\\\0\, 0, g \\<^sub>A\<^bsub>\\HomCod\\<^esub> F\1\<^sub>\\\2\<^sub>\\]\<^sub>\, + F\2\<^sub>\\ + ]\<^sub>\ + )" + and "\ \<^sub>C\<^sub>F\\<^sub>A g\HomDom\ = \ \<^sub>C\<^sub>F\ (\\HomCod\\Dom\\g\)" + and "\ \<^sub>C\<^sub>F\\<^sub>A g\HomCod\ = \ \<^sub>C\<^sub>F\ (\\HomCod\\Cod\\g\)" unfolding cf_cf_arr_comma_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_functor begin -lemma cf_cf_arr_comma_components': +lemma cf_arr_cf_comma_components': assumes "g : c \\<^bsub>\\<^esub> c'" shows "g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\ = (\A\\<^sub>\c' \\<^sub>C\<^sub>F \\Obj\. [0, A\1\<^sub>\\, A\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\)" and "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\ = ( \F\\<^sub>\c' \\<^sub>C\<^sub>F \\Arr\. [ [0, F\0\\1\<^sub>\\, F\0\\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\, [0, F\1\<^sub>\\\1\<^sub>\\, F\1\<^sub>\\\2\<^sub>\\ \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\, F\2\<^sub>\\ ]\<^sub>\ )" and [cat_comma_cs_simps]: "g \<^sub>A\\<^sub>C\<^sub>F \\HomDom\ = c' \\<^sub>C\<^sub>F \" and [cat_comma_cs_simps]: "g \<^sub>A\\<^sub>C\<^sub>F \\HomCod\ = c \\<^sub>C\<^sub>F \" using assms + unfolding cf_arr_cf_comma_components + by (simp_all add: cat_cs_simps) + +lemma cf_cf_arr_comma_components': + assumes "g : c \\<^bsub>\\<^esub> c'" + shows "\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\ = (\A\\<^sub>\\ \<^sub>C\<^sub>F\ c\Obj\. [A\0\, 0, g \\<^sub>A\<^bsub>\\<^esub> A\2\<^sub>\\]\<^sub>\)" + and "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\ = + ( + \F\\<^sub>\\ \<^sub>C\<^sub>F\ c\Arr\. + [ + [F\0\\0\, 0, g \\<^sub>A\<^bsub>\\<^esub> F\0\\2\<^sub>\\]\<^sub>\, + [F\1\<^sub>\\\0\, 0, g \\<^sub>A\<^bsub>\\<^esub> F\1\<^sub>\\\2\<^sub>\\]\<^sub>\, + F\2\<^sub>\\ + ]\<^sub>\ + )" + and [cat_comma_cs_simps]: "\ \<^sub>C\<^sub>F\\<^sub>A g\HomDom\ = \ \<^sub>C\<^sub>F\ c" + and [cat_comma_cs_simps]: "\ \<^sub>C\<^sub>F\\<^sub>A g\HomCod\ = \ \<^sub>C\<^sub>F\ c'" + using assms unfolding cf_cf_arr_comma_components by (simp_all add: cat_cs_simps) end +lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_components'(3,4) lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_components'(3,4) subsubsection\Object map\ +mk_VLambda cf_arr_cf_comma_components(1)[unfolded VLambda_vid_on[symmetric]] + |vsv cf_arr_cf_comma_ObjMap_vsv[cat_comma_cs_intros]| + mk_VLambda cf_cf_arr_comma_components(1)[unfolded VLambda_vid_on[symmetric]] |vsv cf_cf_arr_comma_ObjMap_vsv[cat_comma_cs_intros]| context is_functor begin context fixes g c c' assumes g: "g : c \\<^bsub>\\<^esub> c'" begin mk_VLambda + cf_arr_cf_comma_components'(1)[OF g, unfolded VLambda_vid_on[symmetric]] + |vdomain cf_arr_cf_comma_ObjMap_vdomain[cat_comma_cs_simps]| + +mk_VLambda cf_cf_arr_comma_components'(1)[OF g, unfolded VLambda_vid_on[symmetric]] |vdomain cf_cf_arr_comma_ObjMap_vdomain[cat_comma_cs_simps]| end end +lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ObjMap_vdomain lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ObjMap_vdomain -lemma (in is_functor) cf_cf_arr_comma_ObjMap_app[cat_comma_cs_simps]: +lemma (in is_functor) cf_arr_cf_comma_ObjMap_app[cat_comma_cs_simps]: assumes "A = [a', b', f']\<^sub>\" and "A \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" and "g : c \\<^bsub>\\<^esub> c'" shows "g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ = [a', b', f' \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\" proof- from assms have b': "b' \\<^sub>\ \\Obj\" and f: "f' : c' \\<^bsub>\\<^esub> \\ObjMap\\b'\" and a'_def: "a' = 0" by auto from assms(2) show ?thesis - unfolding cf_cf_arr_comma_components'[OF assms(3)] assms(1) + unfolding cf_arr_cf_comma_components'[OF assms(3)] assms(1) by (simp add: nat_omega_simps a'_def) qed -lemma (in is_functor) cf_cf_arr_comma_ObjMap_vrange: +lemma (in is_functor) cf_cf_arr_comma_ObjMap_app[cat_comma_cs_simps]: + assumes "A = [a', b', f']\<^sub>\" and "A \\<^sub>\ \ \<^sub>C\<^sub>F\ c\Obj\" and "g : c \\<^bsub>\\<^esub> c'" + shows "\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\\A\ = [a', b', g \\<^sub>A\<^bsub>\\<^esub> f']\<^sub>\" +proof- + from assms have b'_def: "b' = 0" + and f: "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> c" + and a': "a' \\<^sub>\ \\Obj\" + by auto + from assms(2) show ?thesis + unfolding cf_cf_arr_comma_components'[OF assms(3)] assms(1) + by (simp add: nat_omega_simps b'_def) +qed + +lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ObjMap_app +lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ObjMap_app + +lemma (in is_functor) cf_arr_cf_comma_ObjMap_vrange: assumes "g : c \\<^bsub>\\<^esub> c'" shows "\\<^sub>\ (g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" proof ( rule vsv.vsv_vrange_vsubset, - unfold cf_cf_arr_comma_ObjMap_vdomain[OF assms] + unfold cf_arr_cf_comma_ObjMap_vdomain[OF assms] ) fix A assume "A \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" with assms is_functor_axioms obtain a f - where A_def: "A = [[]\<^sub>\, a, f]\<^sub>\" + where A_def: "A = [0, a, f]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\a\" by auto from assms a f show "g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" by ( cs_concl cs_simp: cat_comma_cs_simps A_def cs_intro: cat_cs_intros cat_comma_cs_intros ) qed (cs_concl cs_intro: cat_comma_cs_intros) +lemma (in is_functor) cf_cf_arr_comma_ObjMap_vrange: + assumes "g : c \\<^bsub>\\<^esub> c'" + shows "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\) \\<^sub>\ \ \<^sub>C\<^sub>F\ c'\Obj\" +proof + ( + rule vsv.vsv_vrange_vsubset, + unfold cf_cf_arr_comma_ObjMap_vdomain[OF assms] + ) + fix A assume "A \\<^sub>\ \ \<^sub>C\<^sub>F\ c\Obj\" + with assms is_functor_axioms obtain a f + where A_def: "A = [a, 0, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + by auto + from assms a f show "\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\\A\ \\<^sub>\ \ \<^sub>C\<^sub>F\ c'\Obj\" + by + ( + cs_concl + cs_simp: cat_comma_cs_simps A_def + cs_intro: cat_cs_intros cat_comma_cs_intros + ) +qed (cs_concl cs_intro: cat_comma_cs_intros) + subsubsection\Arrow map\ +mk_VLambda cf_arr_cf_comma_components(2) + |vsv cf_arr_cf_comma_ArrMap_vsv[cat_comma_cs_intros]| + mk_VLambda cf_cf_arr_comma_components(2) |vsv cf_cf_arr_comma_ArrMap_vsv[cat_comma_cs_intros]| context is_functor begin context fixes g c c' assumes g: "g : c \\<^bsub>\\<^esub> c'" begin mk_VLambda + cf_arr_cf_comma_components'(2)[OF g, unfolded VLambda_vid_on[symmetric]] + |vdomain cf_arr_cf_comma_ArrMap_vdomain[cat_comma_cs_simps]| + +mk_VLambda cf_cf_arr_comma_components'(2)[OF g, unfolded VLambda_vid_on[symmetric]] |vdomain cf_cf_arr_comma_ArrMap_vdomain[cat_comma_cs_simps]| end end +lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ArrMap_vdomain lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ArrMap_vdomain -lemma (in is_functor) cf_cf_arr_comma_ArrMap_app[cat_comma_cs_simps]: +lemma (in is_functor) cf_arr_cf_comma_ArrMap_app[cat_comma_cs_simps]: assumes "A = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [h, k]\<^sub>\]\<^sub>\" and "[[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [h, k]\<^sub>\]\<^sub>\ : [a, b, f]\<^sub>\ \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> [a', b', f']\<^sub>\" and "g : c \\<^bsub>\\<^esub> c'" shows "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\A\ = [[a, b, f \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\, [a', b', f' \\<^sub>A\<^bsub>\\<^esub> g]\<^sub>\, [h, k]\<^sub>\]\<^sub>\" proof- from assms(3) have c': "c' \\<^sub>\ \\Obj\" by auto from cat_obj_cf_comma_is_arrD(1,2)[OF assms(2)[unfolded cat_comma_cs_simps] c'] is_arrD(1)[OF assms(2)] show ?thesis + unfolding assms(1) cf_arr_cf_comma_components'[OF assms(3)] + by (simp_all add: nat_omega_simps) +qed + +lemmas [cat_comma_cs_simps] = is_functor.cf_arr_cf_comma_ArrMap_app + +lemma (in is_functor) cf_cf_arr_comma_ArrMap_app[cat_comma_cs_simps]: + assumes "A = [[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [h, k]\<^sub>\]\<^sub>\" + and "[[a, b, f]\<^sub>\, [a', b', f']\<^sub>\, [h, k]\<^sub>\]\<^sub>\ : + [a, b, f]\<^sub>\ \\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> [a', b', f']\<^sub>\" + and "g : c \\<^bsub>\\<^esub> c'" + shows "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\A\ = + [[a, b, g \\<^sub>A\<^bsub>\\<^esub> f]\<^sub>\, [a', b', g \\<^sub>A\<^bsub>\\<^esub> f']\<^sub>\, [h, k]\<^sub>\]\<^sub>\" +proof- + from assms(3) have c: "c \\<^sub>\ \\Obj\" by auto + from + cat_cf_obj_comma_is_arrD(1,2)[OF assms(2)[unfolded cat_comma_cs_simps] c] + is_arrD(1)[OF assms(2)] + show ?thesis unfolding assms(1) cf_cf_arr_comma_components'[OF assms(3)] by (simp_all add: nat_omega_simps) qed +lemmas [cat_comma_cs_simps] = is_functor.cf_cf_arr_comma_ArrMap_app + subsubsection\Comma functors are functors\ +lemma (in is_functor) cf_arr_cf_comma_is_functor: + assumes "g : c \\<^bsub>\\<^esub> c'" + shows "g \<^sub>A\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> c \\<^sub>C\<^sub>F \" +proof(rule is_functorI') + show "vfsequence (g \<^sub>A\\<^sub>C\<^sub>F \)" unfolding cf_arr_cf_comma_def by simp + from assms show "category \ (c' \\<^sub>C\<^sub>F \)" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms show "category \ (c \\<^sub>C\<^sub>F \)" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + show "vcard (g \<^sub>A\\<^sub>C\<^sub>F \) = 4\<^sub>\" + unfolding cf_arr_cf_comma_def by (simp_all add: nat_omega_simps) + from assms show "\\<^sub>\ (g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" + by (intro cf_arr_cf_comma_ObjMap_vrange) + show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\ : + g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\B\" + if "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" for A B F + proof- + from assms that 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 assms that k f f' show ?thesis + unfolding F_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps f'_def[symmetric] + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + qed + show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\G \\<^sub>A\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> F\ = + g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\G\ \\<^sub>A\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\" + if "G : B \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" for B C G A F + proof- + from that(2) assms obtain b f b' f' k + where F_def: "F = [[0, b, f]\<^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 + with that(1) assms obtain b'' f'' k' + where G_def: "G = [[0, b', f']\<^sub>\, [0, b'', f'']\<^sub>\, [0, k']\<^sub>\]\<^sub>\" + and C_def: "C = [0, b'', f'']\<^sub>\" + and k': "k' : b' \\<^bsub>\\<^esub> b''" + and f'': "f'' : c' \\<^bsub>\\<^esub> \\ObjMap\\b''\" + and f''_def: "\\ArrMap\\k'\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" + by auto (*slow*) + from assms that k f f' f'' k' show ?thesis + unfolding F_def G_def A_def B_def C_def + by (*slow*) + ( + cs_concl + cs_simp: + cat_cs_simps cat_comma_cs_simps + f''_def[symmetric] f'_def[symmetric] + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + qed + show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\c' \\<^sub>C\<^sub>F \\CId\\C\\ = c \\<^sub>C\<^sub>F \\CId\\g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\C\\" + if "C \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" for C + proof- + from that assms obtain a f + where C_def: "C = [0, a, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\a\" + by auto + from a assms f show + "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\c' \\<^sub>C\<^sub>F \\CId\\C\\ = c \\<^sub>C\<^sub>F \\CId\\g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\C\\" + unfolding C_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + qed +qed + ( + use assms in + \ + cs_concl + cs_simp: cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros + \ + )+ + lemma (in is_functor) cf_cf_arr_comma_is_functor: assumes "g : c \\<^bsub>\\<^esub> c'" - shows "g \<^sub>A\\<^sub>C\<^sub>F \ : c' \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> c \\<^sub>C\<^sub>F \" -proof- - show ?thesis - proof(rule is_functorI') - show "vfsequence (g \<^sub>A\\<^sub>C\<^sub>F \)" unfolding cf_cf_arr_comma_def by simp - from assms show "category \ (c' \\<^sub>C\<^sub>F \)" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - from assms show "category \ (c \\<^sub>C\<^sub>F \)" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - show "vcard (g \<^sub>A\\<^sub>C\<^sub>F \) = 4\<^sub>\" - unfolding cf_cf_arr_comma_def by (simp_all add: nat_omega_simps) - from assms show "\\<^sub>\ (g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\) \\<^sub>\ c \\<^sub>C\<^sub>F \\Obj\" - by (intro cf_cf_arr_comma_ObjMap_vrange) - show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\ : - g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ \\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\B\" - if "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" for A B F - proof- - from assms that 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 assms that k f f' show ?thesis - unfolding F_def A_def B_def - by - ( - cs_concl - cs_simp: cat_cs_simps cat_comma_cs_simps f'_def[symmetric] - cs_intro: cat_cs_intros cat_comma_cs_intros - ) - qed - show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\G \\<^sub>A\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> F\ = - g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\G\ \\<^sub>A\<^bsub>c \\<^sub>C\<^sub>F \\<^esub> g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\" - if "G : B \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> C" and "F : A \\<^bsub>c' \\<^sub>C\<^sub>F \\<^esub> B" for B C G A F - proof- - from that(2) assms obtain b f b' f' k - where F_def: "F = [[0, b, f]\<^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 - with that(1) assms obtain b'' f'' k' - where G_def: "G = [[0, b', f']\<^sub>\, [0, b'', f'']\<^sub>\, [0, k']\<^sub>\]\<^sub>\" - and C_def: "C = [0, b'', f'']\<^sub>\" - and k': "k' : b' \\<^bsub>\\<^esub> b''" - and f'': "f'' : c' \\<^bsub>\\<^esub> \\ObjMap\\b''\" - and f''_def: "\\ArrMap\\k'\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" - by auto (*slow*) - from assms that k f f' f'' k' show ?thesis - unfolding F_def G_def A_def B_def C_def - by (*slow*) - ( - cs_concl - cs_simp: - cat_cs_simps cat_comma_cs_simps - f''_def[symmetric] f'_def[symmetric] - cs_intro: cat_cs_intros cat_comma_cs_intros - ) - qed - show "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\c' \\<^sub>C\<^sub>F \\CId\\C\\ = c \\<^sub>C\<^sub>F \\CId\\g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\C\\" - if "C \\<^sub>\ c' \\<^sub>C\<^sub>F \\Obj\" for C - proof- - from that assms obtain a f - where C_def: "C = [0, a, f]\<^sub>\" - and a: "a \\<^sub>\ \\Obj\" - and f: "f : c' \\<^bsub>\\<^esub> \\ObjMap\\a\" - by auto - from a assms f show - "g \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\c' \\<^sub>C\<^sub>F \\CId\\C\\ = c \\<^sub>C\<^sub>F \\CId\\g \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\C\\" - unfolding C_def - by - ( - cs_concl - cs_simp: cat_cs_simps cat_comma_cs_simps - cs_intro: cat_cs_intros cat_comma_cs_intros - ) - qed + shows "\ \<^sub>C\<^sub>F\\<^sub>A g : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ c'" +proof(rule is_functorI') + from assms have c: "c \\<^sub>\ \\Obj\" by auto + show "vfsequence (\ \<^sub>C\<^sub>F\\<^sub>A g)" unfolding cf_cf_arr_comma_def by simp + from assms show "category \ (\ \<^sub>C\<^sub>F\ c')" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms show "category \ (\ \<^sub>C\<^sub>F\ c)" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + show "vcard (\ \<^sub>C\<^sub>F\\<^sub>A g) = 4\<^sub>\" + unfolding cf_cf_arr_comma_def by (simp_all add: nat_omega_simps) + from assms show "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\) \\<^sub>\ \ \<^sub>C\<^sub>F\ c'\Obj\" + by (intro cf_cf_arr_comma_ObjMap_vrange) + show "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\F\ : + \ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\\A\ \\<^bsub>\ \<^sub>C\<^sub>F\ c'\<^esub> \ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\\B\" + if "F : A \\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> B" for A B F + proof- + from assms that obtain a f a' f' h + where F_def: "F = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [h, 0]\<^sub>\]\<^sub>\" + and A_def: "A = [a, 0, f]\<^sub>\" + and B_def: "B = [a', 0, f']\<^sub>\" + and h: "h : a \\<^bsub>\\<^esub> a'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> c" + and f'_def: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\h\ = f" + by auto + from assms that h f f' show ?thesis + unfolding F_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps f'_def + cs_intro: cat_cs_intros cat_comma_cs_intros + ) qed - ( - use assms in - \ + show "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\G \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> F\ = + \ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\G\ \\<^sub>A\<^bsub>\ \<^sub>C\<^sub>F\ c'\<^esub> \ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\F\" + if "G : B \\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> C" and "F : A \\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> B" for B C G A F + proof- + from that(2) assms obtain a f a' f' h + where F_def: "F = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [h, 0]\<^sub>\]\<^sub>\" + and A_def: "A = [a, 0, f]\<^sub>\" + and B_def: "B = [a', 0, f']\<^sub>\" + and h: "h : a \\<^bsub>\\<^esub> a'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> c" + and [cat_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\h\ = f" + by auto + with that(1) assms obtain a'' f'' h' + where G_def: "G = [[a', 0, f']\<^sub>\, [a'', 0, f'']\<^sub>\, [h', 0]\<^sub>\]\<^sub>\" + and C_def: "C = [a'', 0, f'']\<^sub>\" + and h': "h' : a' \\<^bsub>\\<^esub> a''" + and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> c" + and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\h'\ = f'" + by auto (*slow*) + note [cat_cs_simps] = category.cat_assoc_helper[ + where \=\, where h=f'' and g=\\\ArrMap\\h'\\ and q=f' + ] + from assms that c h f f' f'' h' show ?thesis + unfolding F_def G_def A_def B_def C_def + by + ( cs_concl - cs_simp: cat_comma_cs_simps + cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros - \ - )+ + ) + qed + show "\ \<^sub>C\<^sub>F\\<^sub>A g\ArrMap\\\ \<^sub>C\<^sub>F\ c\CId\\C\\ = \ \<^sub>C\<^sub>F\ c'\CId\\\ \<^sub>C\<^sub>F\\<^sub>A g\ObjMap\\C\\" + if "C \\<^sub>\ \ \<^sub>C\<^sub>F\ c\Obj\" for C + proof- + from that assms obtain a f + where C_def: "C = [a, 0, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + by auto + from a c assms f show ?thesis + unfolding C_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros + ) + qed qed + ( + use assms in + \ + cs_concl + cs_simp: cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros + \ + )+ + +lemma (in is_functor) cf_arr_cf_comma_is_functor'[cat_comma_cs_intros]: + assumes "g : c \\<^bsub>\\<^esub> c'" and "\' = c' \\<^sub>C\<^sub>F \" and "\' = c \\<^sub>C\<^sub>F \" + shows "g \<^sub>A\\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + using assms(1) unfolding assms(2,3) by (rule cf_arr_cf_comma_is_functor(1)) + +lemmas [cat_comma_cs_intros] = is_functor.cf_arr_cf_comma_is_functor' lemma (in is_functor) cf_cf_arr_comma_is_functor'[cat_comma_cs_intros]: - assumes "g : c \\<^bsub>\\<^esub> c'" and "\' = c' \\<^sub>C\<^sub>F \" and "\' = c \\<^sub>C\<^sub>F \" - shows "g \<^sub>A\\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + assumes "g : c \\<^bsub>\\<^esub> c'" and "\' = \ \<^sub>C\<^sub>F\ c" and "\' = \ \<^sub>C\<^sub>F\ c'" + shows "\ \<^sub>C\<^sub>F\\<^sub>A g : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1) unfolding assms(2,3) by (rule cf_cf_arr_comma_is_functor(1)) lemmas [cat_comma_cs_intros] = is_functor.cf_cf_arr_comma_is_functor' -lemma (in is_functor) cf_cf_arr_comma_CId: +lemma (in is_functor) cf_arr_cf_comma_CId: assumes "b \\<^sub>\ \\Obj\" shows "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \ = cf_id (b \\<^sub>C\<^sub>F \)" -proof- - - show ?thesis - proof(rule cf_eqI) - from vempty_is_zet assms show "cf_id (b \\<^sub>C\<^sub>F \) : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F \" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - from vempty_is_zet assms show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F \" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - from assms have ObjMap_dom_lhs: - "\\<^sub>\ ((\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" - by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) - from assms have ObjMap_dom_rhs: - "\\<^sub>\ (dghm_id (b \\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) - show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\ = cf_id (b \\<^sub>C\<^sub>F \)\ObjMap\" - proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) - fix A assume prems: "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" - with assms obtain a' f' - where A_def: "A = [0, a', f']\<^sub>\" - and a': "a' \\<^sub>\ \\Obj\" - and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" - by auto - from prems assms vempty_is_zet a' f' show - "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ = cf_id (b \\<^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 - ) - qed (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)+ - - from assms have ArrMap_dom_lhs: - "\\<^sub>\ ((\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" - by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) - from assms have ArrMap_dom_rhs: - "\\<^sub>\ (dghm_id (b \\<^sub>C\<^sub>F \)\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" - by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) - - show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\ = cf_id (b \\<^sub>C\<^sub>F \)\ArrMap\" - proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) - fix F assume prems: "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" - then obtain A B where F: "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" by (auto dest: is_arrI) - from assms F 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 "\\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" - by auto - from assms prems F h f' f'' show - "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\ = cf_id (b \\<^sub>C\<^sub>F \)\ArrMap\\F\" - unfolding F_def A_def B_def - by - ( - cs_concl - cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros - ) - qed (cs_concl cs_intro: cat_comma_cs_intros cat_cs_intros)+ - - qed simp_all - -qed +proof(rule cf_eqI) + from vempty_is_zet assms show "cf_id (b \\<^sub>C\<^sub>F \) : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from vempty_is_zet assms show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> b \\<^sub>C\<^sub>F \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms have ObjMap_dom_lhs: + "\\<^sub>\ ((\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) + from assms have ObjMap_dom_rhs: + "\\<^sub>\ (cf_id (b \\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\ = cf_id (b \\<^sub>C\<^sub>F \)\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + fix A assume prems: "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + with assms obtain a' f' + where A_def: "A = [0, a', f']\<^sub>\" + and a': "a' \\<^sub>\ \\Obj\" + and f': "f' : b \\<^bsub>\\<^esub> \\ObjMap\\a'\" + by auto + from prems assms vempty_is_zet a' f' show + "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ObjMap\\A\ = cf_id (b \\<^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 + ) + qed (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)+ + from assms have ArrMap_dom_lhs: + "\\<^sub>\ ((\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) + from assms have ArrMap_dom_rhs: + "\\<^sub>\ (cf_id (b \\<^sub>C\<^sub>F \)\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + show "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\ = cf_id (b \\<^sub>C\<^sub>F \)\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix F assume prems: "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + then obtain A B where F: "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" by (auto dest: is_arrI) + from assms F 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 "\\ArrMap\\h\ \\<^sub>A\<^bsub>\\<^esub> f' = f''" + by auto + from assms prems F h f' f'' show + "(\\CId\\b\) \<^sub>A\\<^sub>C\<^sub>F \\ArrMap\\F\ = cf_id (b \\<^sub>C\<^sub>F \)\ArrMap\\F\" + unfolding F_def A_def B_def + by + ( + cs_concl + cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros + ) + qed (cs_concl cs_intro: cat_comma_cs_intros cat_cs_intros)+ +qed simp_all + +lemma (in is_functor) cf_cf_arr_comma_CId: + assumes "b \\<^sub>\ \\Obj\" + shows "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\) = cf_id (\ \<^sub>C\<^sub>F\ b)" +proof(rule cf_eqI) + from vempty_is_zet assms show "cf_id (\ \<^sub>C\<^sub>F\ b) : \ \<^sub>C\<^sub>F\ b \\\<^sub>C\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ b" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from vempty_is_zet assms show "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\) : \ \<^sub>C\<^sub>F\ b \\\<^sub>C\<^bsub>\\<^esub> \ \<^sub>C\<^sub>F\ b" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms have ObjMap_dom_lhs: + "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ObjMap\) = \ \<^sub>C\<^sub>F\ b\Obj\" + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) + from assms have ObjMap_dom_rhs: + "\\<^sub>\ (cf_id (\ \<^sub>C\<^sub>F\ b)\ObjMap\) = \ \<^sub>C\<^sub>F\ b\Obj\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + show "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ObjMap\ = cf_id (\ \<^sub>C\<^sub>F\ b)\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + fix A assume prems: "A \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Obj\" + with assms obtain a' f' + where A_def: "A = [a', 0, f']\<^sub>\" + and a': "a' \\<^sub>\ \\Obj\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" + by auto + from prems assms vempty_is_zet a' f' show + "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ObjMap\\A\ = cf_id (\ \<^sub>C\<^sub>F\ b)\ObjMap\\A\" + unfolding A_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cs_intro: cat_cs_intros + ) + qed (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros)+ + from assms have ArrMap_dom_lhs: + "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ArrMap\) = \ \<^sub>C\<^sub>F\ b\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps cs_intro: cat_cs_intros) + from assms have ArrMap_dom_rhs: + "\\<^sub>\ (cf_id (\ \\<^sub>C\<^sub>F b)\ArrMap\) = \ \\<^sub>C\<^sub>F b\Arr\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + show "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ArrMap\ = cf_id (\ \<^sub>C\<^sub>F\ b)\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix F assume prems: "F \\<^sub>\ \ \<^sub>C\<^sub>F\ b\Arr\" + then obtain A B where F: "F : A \\<^bsub>\ \<^sub>C\<^sub>F\ b\<^esub> B" by (auto dest: is_arrI) + from assms F obtain a' f' a'' f'' k + where F_def: "F = [[a', 0, f']\<^sub>\, [a'', 0, f'']\<^sub>\, [k, 0]\<^sub>\]\<^sub>\" + and A_def: "A = [a', 0, f']\<^sub>\" + and B_def: "B = [a'', 0, f'']\<^sub>\" + and k: "k : a' \\<^bsub>\\<^esub> a''" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> b" + and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> b" + and [cat_cs_simps]: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\k\ = f'" + by auto + from assms prems F k f' f'' show + "\ \<^sub>C\<^sub>F\\<^sub>A (\\CId\\b\)\ArrMap\\F\ = cf_id (\ \<^sub>C\<^sub>F\ b)\ArrMap\\F\" + unfolding F_def A_def B_def + by + ( + cs_concl + cs_simp: cat_comma_cs_simps cat_cs_simps cs_intro: cat_cs_intros + ) + qed + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_comma_cs_intros cat_cs_intros + )+ +qed simp_all subsubsection\Comma functors and projections\ lemma (in is_functor) - cf_cf_comp_cf_obj_cf_comma_proj_cf_cf_arr_comma[cat_comma_cs_simps]: + cf_cf_comp_cf_obj_cf_comma_proj_cf_arr_cf_comma[cat_comma_cs_simps]: assumes "f : a \\<^bsub>\\<^esub> b" shows "a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \ = b \<^sub>O\\<^sub>C\<^sub>F \" -proof- - - show ?thesis - proof(rule cf_eqI) - from assms vempty_is_zet show "b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - from assms show - "a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" - by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) - from assms have ObjMap_dom_lhs: - "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" +proof(rule cf_eqI) + from assms vempty_is_zet show "b \<^sub>O\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms show + "a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \ : b \\<^sub>C\<^sub>F \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms have ObjMap_dom_lhs: + "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms have ObjMap_dom_rhs: "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by (cs_concl cs_simp: cat_comma_cs_simps) + show "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\ = b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + from assms show "vsv (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\)" + by (cs_concl cs_intro: cat_comma_cs_intros) + fix A assume prems: "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" + with assms 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 prems assms b' f' show + "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\\A\ = b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\\A\" + unfolding A_def by - ( + ( cs_concl - cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros ) - from assms have ObjMap_dom_rhs: "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" - by (cs_concl cs_simp: cat_comma_cs_simps) - show "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\ = b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\" - proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) - from assms show "vsv (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\)" - by (cs_concl cs_intro: cat_comma_cs_intros) - fix A assume prems: "A \\<^sub>\ b \\<^sub>C\<^sub>F \\Obj\" - with assms 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 prems assms b' f' show - "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\\A\ = b \<^sub>O\\<^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 + qed + ( + use assms vempty_is_zet in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ + ) + from assms have ArrMap_dom_lhs: + "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by ( - use assms vempty_is_zet in - \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros ) - from assms have ArrMap_dom_lhs: - "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + from assms vempty_is_zet have ArrMap_dom_rhs: + "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" + by (cs_concl cs_simp: cat_comma_cs_simps) + from assms vempty_is_zet have ArrMap_dom_lhs: + "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms have ArrMap_dom_rhs: "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps) + show "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\ = b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix F assume "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" + then obtain A B where F: "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" + by (auto dest: 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' : 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 vempty_is_zet h assms f' f'' F show + "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\\F\ = b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\\F\" + unfolding F_def A_def B_def + by (*slow*) + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric] + cs_intro: cat_cs_intros cat_comma_cs_intros + )+ + qed + ( + use assms vempty_is_zet in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ + ) +qed simp_all + +lemma (in is_functor) + cf_cf_comp_cf_cf_obj_comma_proj_cf_cf_arr_comma[cat_comma_cs_simps]: + assumes "f : a \\<^bsub>\\<^esub> b" + shows "\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f = \ \<^sub>C\<^sub>F\\<^sub>O a" +proof(rule cf_eqI) + from assms vempty_is_zet show "\ \<^sub>C\<^sub>F\\<^sub>O a : \ \<^sub>C\<^sub>F\ a \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms show "\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f : \ \<^sub>C\<^sub>F\ a \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros) + from assms have ObjMap_dom_lhs: + "\\<^sub>\ ((\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ObjMap\) = \ \<^sub>C\<^sub>F\ a\Obj\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms have ObjMap_dom_rhs: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>O a\ObjMap\) = \ \<^sub>C\<^sub>F\ a\Obj\" + by (cs_concl cs_simp: cat_comma_cs_simps) + show "(\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ObjMap\ = \ \<^sub>C\<^sub>F\\<^sub>O a\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + from assms show "vsv (\ \<^sub>C\<^sub>F\\<^sub>O a\ObjMap\)" + by (cs_concl cs_intro: cat_comma_cs_intros) + fix A assume prems: "A \\<^sub>\ \ \<^sub>C\<^sub>F\ a\Obj\" + with assms obtain a' f' + where A_def: "A = [a', 0, f']\<^sub>\" + and b': "a' \\<^sub>\ \\Obj\" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> a" + by auto + from prems assms b' f' show + "(\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ObjMap\\A\ = \ \<^sub>C\<^sub>F\\<^sub>O a\ObjMap\\A\" + unfolding A_def by ( cs_concl - cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + cs_simp: cat_cs_simps cat_comma_cs_simps + cs_intro: cat_cs_intros cat_comma_cs_intros ) - from assms vempty_is_zet have ArrMap_dom_rhs: - "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ObjMap\) = b \\<^sub>C\<^sub>F \\Obj\" - by (cs_concl cs_simp: cat_comma_cs_simps) - from assms vempty_is_zet have ArrMap_dom_lhs: - "\\<^sub>\ ((a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" + qed + ( + use assms vempty_is_zet in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ + ) + from assms vempty_is_zet have ArrMap_dom_lhs: + "\\<^sub>\ ((\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ArrMap\) = \ \<^sub>C\<^sub>F\ a\Arr\" + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms have ArrMap_dom_rhs: "\\<^sub>\ (\ \<^sub>C\<^sub>F\\<^sub>O a\ArrMap\) = \ \<^sub>C\<^sub>F\ a\Arr\" + by (cs_concl cs_simp: cat_comma_cs_simps) + show "(\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ArrMap\ = \ \<^sub>C\<^sub>F\\<^sub>O a\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + fix F assume "F \\<^sub>\ \ \<^sub>C\<^sub>F\ a\Arr\" + then obtain A B where F: "F : A \\<^bsub>\ \<^sub>C\<^sub>F\ a\<^esub> B" by (auto dest: is_arrI) + with assms obtain a' f' a'' f'' k + where F_def: "F = [[a', 0, f']\<^sub>\, [a'', 0, f'']\<^sub>\, [k, 0]\<^sub>\]\<^sub>\" + and A_def: "A = [a', 0, f']\<^sub>\" + and B_def: "B = [a'', 0, f'']\<^sub>\" + and k: "k : a' \\<^bsub>\\<^esub> a''" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> a" + and f'': "f'' : \\ObjMap\\a''\ \\<^bsub>\\<^esub> a" + and f'_def: "f'' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\k\ = f'" + by auto + from vempty_is_zet k assms f' f'' F show + "(\ \<^sub>C\<^sub>F\\<^sub>O b \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>A f)\ArrMap\\F\ = \ \<^sub>C\<^sub>F\\<^sub>O a\ArrMap\\F\" + unfolding F_def A_def B_def + by (*slow*) + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps f'_def + cs_intro: cat_cs_intros cat_comma_cs_intros + )+ + qed + ( + use assms vempty_is_zet in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ + ) +qed simp_all + + +subsubsection\Opposite comma functors\ + +lemma (in is_functor) cf_op_cf_obj_comma_cf_arr_cf_comma: + assumes "g : c \\<^bsub>\\<^esub> c'" + shows + "op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g) = + g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c" +proof(rule cf_eqI) + from assms interpret \c: category \ \\ \<^sub>C\<^sub>F\ c\ + by + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms have c: "c \\<^sub>\ \\Obj\" by auto + from assms show "op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g) : + op_cat (\ \<^sub>C\<^sub>F\ c) \\\<^sub>C\<^bsub>\\<^esub> c' \\<^sub>C\<^sub>F (op_cf \)" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + then have ObjMap_dom_lhs: + "\\<^sub>\ ((op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ObjMap\) = + (op_cat (\ \<^sub>C\<^sub>F\ c))\Obj\" + and ArrMap_dom_lhs: + "\\<^sub>\ ((op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ArrMap\) = + (op_cat (\ \<^sub>C\<^sub>F\ c))\Arr\" + by (cs_concl cs_simp: cat_cs_simps)+ + from assms show + "g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c : + op_cat (\ \<^sub>C\<^sub>F\ c) \\\<^sub>C\<^bsub>\\<^esub> c' \\<^sub>C\<^sub>F (op_cf \)" + by + ( + cs_concl + cs_simp: cat_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + then have ObjMap_dom_rhs: + "\\<^sub>\ ((g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c)\ObjMap\) = + (op_cat (\ \<^sub>C\<^sub>F\ c))\Obj\" + and ArrMap_dom_rhs: + "\\<^sub>\ ((g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c)\ArrMap\) = + (op_cat (\ \<^sub>C\<^sub>F\ c))\Arr\" + by (cs_concl cs_simp: cat_cs_simps)+ + show + "(op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ObjMap\ = + (g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c)\ObjMap\" + proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs cat_op_simps) + fix A assume "A \\<^sub>\ \ \<^sub>C\<^sub>F\ c\Obj\" + with assms obtain a f + where A_def: "A = [a, 0, f]\<^sub>\" + and a: "a \\<^sub>\ \\Obj\" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + by auto + from assms a f show + "(op_cf_obj_comma \ c' \\<^sub>D\<^sub>G\<^sub>H\<^sub>M op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ObjMap\\A\ = + (g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>D\<^sub>G\<^sub>H\<^sub>M op_cf_obj_comma \ c)\ObjMap\\A\" + unfolding A_def by ( cs_concl - cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros ) - from assms have ArrMap_dom_rhs: - "\\<^sub>\ (b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\) = b \\<^sub>C\<^sub>F \\Arr\" - by (cs_concl cs_simp: cat_comma_cs_simps) - show "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\ = b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\" - proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) - fix F assume "F \\<^sub>\ b \\<^sub>C\<^sub>F \\Arr\" - then obtain A B where F: "F : A \\<^bsub>b \\<^sub>C\<^sub>F \\<^esub> B" - by (auto dest: 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' : 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 vempty_is_zet h assms f' f'' F show - "(a \<^sub>O\\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F f \<^sub>A\\<^sub>C\<^sub>F \)\ArrMap\\F\ = b \<^sub>O\\<^sub>C\<^sub>F \\ArrMap\\F\" - unfolding F_def A_def B_def - by (*slow*) - ( - cs_concl - cs_simp: cat_cs_simps cat_comma_cs_simps f''_def[symmetric] - cs_intro: cat_cs_intros cat_comma_cs_intros - )+ - qed - ( - use assms vempty_is_zet in - \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros\ - ) - qed simp_all - -qed + qed + ( + use assms in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros\ + )+ + show + "(op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ArrMap\ = + (g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c)\ArrMap\" + proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs cat_op_simps) + fix F assume "F \\<^sub>\ \ \<^sub>C\<^sub>F\ c\Arr\" + then obtain A B where F: "F : A \\<^bsub>\ \<^sub>C\<^sub>F\ c\<^esub> B" by auto + with assms c obtain a f a' f' h + where F_def: "F = [[a, 0, f]\<^sub>\, [a', 0, f']\<^sub>\, [h, 0]\<^sub>\]\<^sub>\" + and A_def: "A = [a, 0, f]\<^sub>\" + and B_def: "B = [a', 0, f']\<^sub>\" + and h: "h : a \\<^bsub>\\<^esub> a'" + and f: "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> c" + and f': "f' : \\ObjMap\\a'\ \\<^bsub>\\<^esub> c" + and [cat_comma_cs_simps]: "f' \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\h\ = f" + by auto + from F assms h f f' c show + "(op_cf_obj_comma \ c' \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>A g))\ArrMap\\F\ = + (g \<^sub>A\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c)\ArrMap\\F\" + unfolding F_def A_def B_def + by + ( + cs_concl + cs_simp: cat_cs_simps cat_comma_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + qed + ( + use assms in + \cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros\ + )+ +qed simp_all 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,2042 +1,2076 @@ (* 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 "\ = \
" 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) +subsubsection\Further properties\ + +lemma (in is_functor) cf_cf_comp_cf_const: + assumes "category \ \" and "a \\<^sub>\ \\Obj\" + shows "cf_const \ \ a \\<^sub>C\<^sub>F \ = cf_const \ \ a" +proof(rule cf_smcf_eqI) + interpret \: category \ \ by (rule assms(1)) + from assms(2) show "cf_const \ \ a \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros) + from assms(2) show "cf_const \ \ a : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: cat_cs_intros) + from assms(2) have CId_a: "\\CId\\a\ : a \\<^bsub>\\<^esub> a" + by (cs_concl cs_intro: cat_cs_intros) + from assms(2) have CId_CId_a: "\\CId\\a\ \\<^sub>A\<^bsub>\\<^esub> \\CId\\a\ = \\CId\\a\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + from is_semifunctor.smcf_smcf_comp_smcf_const[ + OF cf_is_semifunctor \.cat_semicategory, + unfolded slicing_simps, + OF CId_a CId_CId_a + ] + show "cf_smcf (cf_const \ \ a \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) = cf_smcf (cf_const \ \ a)" + by (cs_prems cs_simp: slicing_simps slicing_commute) +qed simp_all + +lemma (in is_functor) cf_cf_comp_cf_const'[cat_cs_simps]: + assumes "category \ \" + and "a \\<^sub>\ \\Obj\" + and "f = \\CId\\a\" + shows "dghm_const \ \ a f \\<^sub>C\<^sub>F \ = cf_const \ \ a" + using assms(1,2) unfolding assms(3) by (rule cf_cf_comp_cf_const) + +lemmas [cat_cs_simps] = is_functor.cf_cf_comp_cf_const' + + 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_NTCF.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_NTCF.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_NTCF.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_NTCF.thy @@ -1,2364 +1,2390 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Natural transformation\ theory CZH_ECAT_NTCF imports CZH_Foundations.CZH_SMC_NTSMCF CZH_ECAT_Functor begin subsection\Background\ named_theorems ntcf_cs_simps named_theorems ntcf_cs_intros lemmas [cat_cs_simps] = dg_shared_cs_simps lemmas [cat_cs_intros] = dg_shared_cs_intros subsubsection\Slicing\ definition ntcf_ntsmcf :: "V \ V" where "ntcf_ntsmcf \ = [ \\NTMap\, cf_smcf (\\NTDom\), cf_smcf (\\NTCod\), cat_smc (\\NTDGDom\), cat_smc (\\NTDGCod\) ]\<^sub>\" text\Components.\ lemma ntcf_ntsmcf_components: shows [slicing_simps]: "ntcf_ntsmcf \\NTMap\ = \\NTMap\" and [slicing_commute]: "ntcf_ntsmcf \\NTDom\ = cf_smcf (\\NTDom\)" and [slicing_commute]: "ntcf_ntsmcf \\NTCod\ = cf_smcf (\\NTCod\)" and [slicing_commute]: "ntcf_ntsmcf \\NTDGDom\ = cat_smc (\\NTDGDom\)" and [slicing_commute]: "ntcf_ntsmcf \\NTDGCod\ = cat_smc (\\NTDGCod\)" unfolding ntcf_ntsmcf_def nt_field_simps by (auto simp: nat_omega_simps) subsection\Definition and elementary properties\ text\ The definition of a natural transformation that is used in this work is is similar to the definition that can be found in Chapter I-4 in \cite{mac_lane_categories_2010}. \ locale is_ntcf = \ \ + vfsequence \ + NTDom: is_functor \ \ \ \ + NTCod: is_functor \ \ \ \ for \ \ \ \ \ \ + assumes ntcf_length[cat_cs_simps]: "vcard \ = 5\<^sub>\" and ntcf_is_ntsmcf[slicing_intros]: "ntcf_ntsmcf \ : cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" and ntcf_NTDom[cat_cs_simps]: "\\NTDom\ = \" and ntcf_NTCod[cat_cs_simps]: "\\NTCod\ = \" and ntcf_NTDGDom[cat_cs_simps]: "\\NTDGDom\ = \" and ntcf_NTDGCod[cat_cs_simps]: "\\NTDGCod\ = \" syntax "_is_ntcf" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ :/ _ \\<^sub>C\<^sub>F _ :/ _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_ntcf \ \ \ \ \ \" abbreviation all_ntcfs :: "V \ V" where "all_ntcfs \ \ set {\. \\ \ \ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" abbreviation ntcfs :: "V \ V \ V \ V" where "ntcfs \ \ \ \ set {\. \\ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" abbreviation these_ntcfs :: "V \ V \ V \ V \ V \ V" where "these_ntcfs \ \ \ \ \ \ set {\. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" lemmas [cat_cs_simps] = is_ntcf.ntcf_length is_ntcf.ntcf_NTDom is_ntcf.ntcf_NTCod is_ntcf.ntcf_NTDGDom is_ntcf.ntcf_NTDGCod lemma (in is_ntcf) ntcf_is_ntsmcf': assumes "\' = cf_smcf \" and "\' = cf_smcf \" and "\' = cat_smc \" and "\' = cat_smc \" shows "ntcf_ntsmcf \ : \' \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \' : \' \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms(1-4) by (rule ntcf_is_ntsmcf) lemmas [slicing_intros] = is_ntcf.ntcf_is_ntsmcf' text\Rules.\ lemma (in is_ntcf) is_ntcf_axioms'[cat_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_ntcf_axioms) mk_ide rf is_ntcf_def[unfolded is_ntcf_axioms_def] |intro is_ntcfI| |dest is_ntcfD[dest]| |elim is_ntcfE[elim]| lemmas [cat_cs_intros] = is_ntcfD(3,4) lemma is_ntcfI': assumes "\ \" and "vfsequence \" and "vcard \ = 5\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\\NTDom\ = \" and "\\NTCod\ = \" and "\\NTDGDom\ = \" and "\\NTDGCod\ = \" and "vsv (\\NTMap\)" and "\\<^sub>\ (\\NTMap\) = \\Obj\" and "\a. a \\<^sub>\ \\Obj\ \ \\NTMap\\a\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\a\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" shows "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (intro is_ntcfI is_ntsmcfI', unfold ntcf_ntsmcf_components slicing_simps) ( simp_all add: assms nat_omega_simps ntcf_ntsmcf_def is_functorD(6)[OF assms(4)] is_functorD(6)[OF assms(5)] ) lemma is_ntcfD': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \" and "vfsequence \" and "vcard \ = 5\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\\NTDom\ = \" and "\\NTCod\ = \" and "\\NTDGDom\ = \" and "\\NTDGCod\ = \" and "vsv (\\NTMap\)" and "\\<^sub>\ (\\NTMap\) = \\Obj\" and "\a. a \\<^sub>\ \\Obj\ \ \\NTMap\\a\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\a\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" by ( simp_all add: is_ntcfD(2-10)[OF assms] is_ntsmcfD'[OF is_ntcfD(6)[OF assms], unfolded slicing_simps] ) lemma is_ntcfE': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" obtains "\ \" and "vfsequence \" and "vcard \ = 5\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\\NTDom\ = \" and "\\NTCod\ = \" and "\\NTDGDom\ = \" and "\\NTDGCod\ = \" and "vsv (\\NTMap\)" and "\\<^sub>\ (\\NTMap\) = \\Obj\" and "\a. a \\<^sub>\ \\Obj\ \ \\NTMap\\a\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\a\" and "\a b f. f : a \\<^bsub>\\<^esub> b \ \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" using assms by (simp add: is_ntcfD') text\Slicing.\ context is_ntcf begin interpretation ntsmcf: is_ntsmcf \ \cat_smc \\ \cat_smc \\ \cf_smcf \\ \cf_smcf \\ \ntcf_ntsmcf \\ by (rule ntcf_is_ntsmcf) lemmas_with [unfolded slicing_simps]: ntcf_NTMap_vsv = ntsmcf.ntsmcf_NTMap_vsv and ntcf_NTMap_vdomain[cat_cs_simps] = ntsmcf.ntsmcf_NTMap_vdomain and ntcf_NTMap_is_arr = ntsmcf.ntsmcf_NTMap_is_arr and ntcf_NTMap_is_arr'[cat_cs_intros] = ntsmcf.ntsmcf_NTMap_is_arr' sublocale NTMap: vsv \\\NTMap\\ rewrites "\\<^sub>\ (\\NTMap\) = \\Obj\" by (rule ntcf_NTMap_vsv) (simp add: cat_cs_simps) lemmas_with [unfolded slicing_simps]: ntcf_NTMap_app_in_Arr[cat_cs_intros] = ntsmcf.ntsmcf_NTMap_app_in_Arr and ntcf_NTMap_vrange_vifunion = ntsmcf.ntsmcf_NTMap_vrange_vifunion and ntcf_NTMap_vrange = ntsmcf.ntsmcf_NTMap_vrange and ntcf_NTMap_vsubset_Vset = ntsmcf.ntsmcf_NTMap_vsubset_Vset and ntcf_NTMap_in_Vset = ntsmcf.ntsmcf_NTMap_in_Vset and ntcf_is_ntsmcf_if_ge_Limit = ntsmcf.ntsmcf_is_ntsmcf_if_ge_Limit lemmas_with [unfolded slicing_simps]: ntcf_Comp_commute[cat_cs_intros] = ntsmcf.ntsmcf_Comp_commute and ntcf_Comp_commute' = ntsmcf.ntsmcf_Comp_commute' and ntcf_Comp_commute'' = ntsmcf.ntsmcf_Comp_commute'' end lemmas [cat_cs_simps] = is_ntcf.ntcf_NTMap_vdomain lemmas [cat_cs_intros] = is_ntcf.ntcf_NTMap_is_arr' ntsmcf_hcomp_NTMap_vsv text\Elementary properties.\ lemma ntcf_eqI: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" and "\\NTMap\ = \'\NTMap\" and "\ = \'" and "\ = \'" and "\ = \'" and "\ = \'" shows "\ = \'" proof- interpret L: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret R: is_ntcf \ \' \' \' \' \' by (rule assms(2)) show ?thesis proof(rule vsv_eqI) have dom: "\\<^sub>\ \ = 5\<^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(4-7) have sup: "\\NTDom\ = \'\NTDom\" "\\NTCod\ = \'\NTCod\" "\\NTDGDom\ = \'\NTDGDom\" "\\NTDGCod\ = \'\NTDGCod\" by (simp_all add: cat_cs_simps) show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = \'\a\" for a by (unfold dom, elim_in_numeral, insert assms(3) sup) (auto simp: nt_field_simps) qed auto qed lemma ntcf_ntsmcf_eqI: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" and "\ = \'" and "\ = \'" and "\ = \'" and "\ = \'" and "ntcf_ntsmcf \ = ntcf_ntsmcf \'" shows "\ = \'" proof(rule ntcf_eqI[of \]) from assms(7) have "ntcf_ntsmcf \\NTMap\ = ntcf_ntsmcf \'\NTMap\" by simp then show "\\NTMap\ = \'\NTMap\" unfolding slicing_simps by simp_all from assms(3-6) show "\ = \'" "\ = \'" "\ = \'" "\ = \'" by simp_all qed (auto simp: assms(1,2)) lemma (in is_ntcf) ntcf_def: "\ = [\\NTMap\, \\NTDom\, \\NTCod\, \\NTDGDom\, \\NTDGCod\]\<^sub>\" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ \ = 5\<^sub>\" by (cs_concl cs_simp: cat_cs_simps V_cs_simps) have dom_rhs: "\\<^sub>\ [\\NTMap\, \\NTDGDom\, \\NTDGCod\, \\NTDom\, \\NTCod\]\<^sub>\ = 5\<^sub>\" by (simp add: nat_omega_simps) then show "\\<^sub>\ \ = \\<^sub>\ [\\NTMap\, \\NTDom\, \\NTCod\, \\NTDGDom\, \\NTDGCod\]\<^sub>\" unfolding dom_lhs dom_rhs by (simp add: nat_omega_simps) show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = [\\NTMap\, \\NTDom\, \\NTCod\, \\NTDGDom\, \\NTDGCod\]\<^sub>\\a\" for a by (unfold dom_lhs, elim_in_numeral, unfold nt_field_simps) (simp_all add: nat_omega_simps) qed (auto simp: vsv_axioms) lemma (in is_ntcf) ntcf_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "\ \\<^sub>\ Vset \" proof- interpret \: \ \ by (rule assms(1)) note [cat_cs_intros] = ntcf_NTMap_in_Vset NTDom.cf_in_Vset NTCod.cf_in_Vset NTDom.HomDom.cat_in_Vset NTDom.HomCod.cat_in_Vset from assms(2) show ?thesis by (subst ntcf_def) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros) qed lemma (in is_ntcf) ntcf_is_ntcf_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_ntcfI) show "ntcf_ntsmcf \ : cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by (rule is_ntsmcf.ntsmcf_is_ntsmcf_if_ge_Limit[OF ntcf_is_ntsmcf assms]) qed ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros assms NTDom.cf_is_functor_if_ge_Limit NTCod.cf_is_functor_if_ge_Limit )+ lemma small_all_ntcfs[simp]: "small {\. \\ \ \ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" proof(cases \\ \\) case True from is_ntcf.ntcf_in_Vset show ?thesis by (intro down[of _ \Vset (\ + \)\]) (auto simp: True \.\_Limit_\\ \.\_\_\\ \.intro \.\_\_\\) next case False then have "{\. \\ \ \ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \} = {}" by auto then show ?thesis by simp qed lemma small_ntcfs[simp]: "small {\. \\ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" by (rule down[of _ \set {\. \\ \ \ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}\]) auto lemma small_these_ntcfs[simp]: "small {\. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}" by (rule down[of _ \set {\. \\ \ \ \. \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \}\]) auto text\Further elementary results.\ lemma these_ntcfs_iff: (*not simp*) "\ \\<^sub>\ these_ntcfs \ \ \ \ \ \ \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto subsection\Opposite natural transformation\ text\See section 1.5 in \cite{bodo_categories_1970}.\ definition op_ntcf :: "V \ V" where "op_ntcf \ = [ \\NTMap\, op_cf (\\NTCod\), op_cf (\\NTDom\), op_cat (\\NTDGDom\), op_cat (\\NTDGCod\) ]\<^sub>\" text\Components.\ lemma op_ntcf_components[cat_op_simps]: shows "op_ntcf \\NTMap\ = \\NTMap\" and "op_ntcf \\NTDom\ = op_cf (\\NTCod\)" and "op_ntcf \\NTCod\ = op_cf (\\NTDom\)" and "op_ntcf \\NTDGDom\ = op_cat (\\NTDGDom\)" and "op_ntcf \\NTDGCod\ = op_cat (\\NTDGCod\)" unfolding op_ntcf_def nt_field_simps by (auto simp: nat_omega_simps) text\Slicing.\ lemma ntcf_ntsmcf_op_ntcf[slicing_commute]: "op_ntsmcf (ntcf_ntsmcf \) = ntcf_ntsmcf (op_ntcf \)" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (op_ntsmcf (ntcf_ntsmcf \)) = 5\<^sub>\" unfolding op_ntsmcf_def by (auto simp: nat_omega_simps) have dom_rhs: "\\<^sub>\ (ntcf_ntsmcf (op_ntcf \)) = 5\<^sub>\" unfolding ntcf_ntsmcf_def by (auto simp: nat_omega_simps) show "\\<^sub>\ (op_ntsmcf (ntcf_ntsmcf \)) = \\<^sub>\ (ntcf_ntsmcf (op_ntcf \))" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (op_ntsmcf (ntcf_ntsmcf \)) \ op_ntsmcf (ntcf_ntsmcf \)\a\ = ntcf_ntsmcf (op_ntcf \)\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold nt_field_simps ntcf_ntsmcf_def op_ntcf_def op_ntsmcf_def ) (auto simp: nat_omega_simps slicing_commute[symmetric]) qed (auto simp: ntcf_ntsmcf_def op_ntsmcf_def) text\Elementary properties.\ lemma op_ntcf_vsv[cat_op_intros]: "vsv (op_ntcf \)" unfolding op_ntcf_def by auto subsubsection\Further properties\ lemma (in is_ntcf) is_ntcf_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" proof(rule is_ntcfI, unfold cat_op_simps) show "vfsequence (op_ntcf \)" by (simp add: op_ntcf_def) show "vcard (op_ntcf \) = 5\<^sub>\" by (simp add: op_ntcf_def nat_omega_simps) qed ( use is_ntcf_axioms in \ cs_concl cs_simp: cat_cs_simps slicing_commute[symmetric] cs_intro: cat_cs_intros cat_op_intros smc_op_intros slicing_intros \ )+ lemma (in is_ntcf) is_ntcf_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_ntcf_op) lemmas [cat_op_intros] = is_ntcf.is_ntcf_op' lemma (in is_ntcf) ntcf_op_ntcf_op_ntcf[cat_op_simps]: "op_ntcf (op_ntcf \) = \" proof(rule ntcf_eqI[of \ \ \ \ \ _ \ \ \ \], unfold cat_op_simps) interpret op: is_ntcf \ \op_cat \\ \op_cat \\ \op_cf \\ \op_cf \\ \op_ntcf \\ by (rule is_ntcf_op) from op.is_ntcf_op show "op_ntcf (op_ntcf \) : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (simp add: cat_op_simps) qed (auto simp: cat_cs_intros) lemmas ntcf_op_ntcf_op_ntcf[cat_op_simps] = is_ntcf.ntcf_op_ntcf_op_ntcf lemma eq_op_ntcf_iff[cat_op_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" shows "op_ntcf \ = op_ntcf \' \ \ = \'" proof interpret L: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret R: is_ntcf \ \' \' \' \' \' by (rule assms(2)) assume prems: "op_ntcf \ = op_ntcf \'" show "\ = \'" proof(rule ntcf_eqI[OF assms]) from prems L.ntcf_op_ntcf_op_ntcf R.ntcf_op_ntcf_op_ntcf show "\\NTMap\ = \'\NTMap\" by metis+ from prems L.ntcf_op_ntcf_op_ntcf R.ntcf_op_ntcf_op_ntcf have "\\NTDom\ = \'\NTDom\" and "\\NTCod\ = \'\NTCod\" and "\\NTDGDom\ = \'\NTDGDom\" and "\\NTDGCod\ = \'\NTDGCod\" by metis+ then show "\ = \'" "\ = \'" "\ = \'" "\ = \'" by (auto simp: cat_cs_simps) qed qed auto subsection\Vertical composition of natural transformations\ subsubsection\Definition and elementary properties\ text\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ abbreviation (input) ntcf_vcomp :: "V \ V \ V" (infixl \\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\ 55) where "ntcf_vcomp \ ntsmcf_vcomp" lemmas [cat_cs_simps] = ntsmcf_vcomp_components(2-5) text\Slicing.\ lemma ntcf_ntsmcf_ntcf_vcomp[slicing_commute]: "ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" unfolding ntsmcf_vcomp_def ntcf_ntsmcf_def cat_smc_def nt_field_simps dg_field_simps by (simp add: nat_omega_simps) subsubsection\Natural transformation map\ lemma ntcf_vcomp_NTMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \\Obj\" proof- interpret \: is_ntcf \ \ \ \ \ \ using assms by auto show ?thesis by ( rule ntsmcf_vcomp_NTMap_vdomain [ OF \.ntcf_is_ntsmcf, of \ntcf_ntsmcf \\, unfolded slicing_commute slicing_simps ] ) qed lemma ntcf_vcomp_NTMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = \\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" proof- interpret \: is_ntcf \ \ \ \ \ \ using assms by clarsimp interpret \: is_ntcf \ \ \ \ \ \ using assms by clarsimp show ?thesis by ( rule ntsmcf_vcomp_NTMap_app [ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed lemma ntcf_vcomp_NTMap_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_ntcf \ \ \ \ \ \ using assms by auto interpret \: is_ntcf \ \ \ \ \ \ using assms by auto show ?thesis by ( rule ntsmcf_vcomp_NTMap_vrange[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] ) qed subsubsection\Further properties\ lemma ntcf_vcomp_composable_commute[cat_cs_simps]: \\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and [intro]: "f : a \\<^bsub>\\<^esub> b" shows "(\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\) \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis by ( rule ntsmcf_vcomp_composable_commute[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps, OF assms(3) ] ) qed lemma ntcf_vcomp_is_ntcf[cat_cs_intros]: \\see Chapter II-4 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis proof(intro is_ntcfI) show "vfsequence (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by (simp add: ntsmcf_vcomp_def) show "vcard (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = 5\<^sub>\" unfolding ntsmcf_vcomp_def by (simp add: nat_omega_simps) show "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by ( rule ntsmcf_vcomp_is_ntsmcf[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] ) qed (auto simp: ntsmcf_vcomp_components(1) cat_cs_simps cat_cs_intros) qed lemma ntcf_vcomp_assoc[cat_cs_simps]: \\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^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- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) show ?thesis proof(rule ntcf_eqI[of \]) from ntsmcf_vcomp_assoc[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] have "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" by simp then 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\" unfolding slicing_simps . qed (auto intro: cat_cs_intros) qed subsubsection\ The opposite of the vertical composition of natural transformations \ lemma op_ntcf_ntcf_vcomp[cat_op_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \" proof- interpret \: is_ntcf \ \ \ \ \ \ using assms(1) by auto interpret \: is_ntcf \ \ \ \ \ \ using assms(2) by auto show ?thesis proof(rule sym, rule ntcf_eqI[of \]) from op_ntsmcf_ntsmcf_vcomp [ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] have "ntcf_ntsmcf (op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \)\NTMap\ = ntcf_ntsmcf (op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" by simp then show "(op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \)\NTMap\ = op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\" unfolding slicing_simps . qed (auto intro: cat_cs_intros cat_op_intros) qed subsection\Horizontal composition of natural transformations\ subsubsection\Definition and elementary properties\ text\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ abbreviation (input) ntcf_hcomp :: "V \ V \ V" (infixl \\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\ 55) where "ntcf_hcomp \ ntsmcf_hcomp" lemmas [cat_cs_simps] = ntsmcf_hcomp_components(2-5) text\Slicing.\ lemma ntcf_ntsmcf_ntcf_hcomp[slicing_commute]: "ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" proof(rule vsv_eqI) show "vsv (ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \)" unfolding ntsmcf_hcomp_def by auto show "vsv (ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))" unfolding ntcf_ntsmcf_def by auto have dom_lhs: "\\<^sub>\ (ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \) = 5\<^sub>\" unfolding ntsmcf_hcomp_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) = 5\<^sub>\" unfolding ntcf_ntsmcf_def by (simp add: nat_omega_simps) show "\\<^sub>\ (ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \) = \\<^sub>\ (ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))" unfolding dom_lhs dom_rhs .. fix a assume "a \\<^sub>\ \\<^sub>\ (ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \)" then show "(ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \)\a\ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\a\" unfolding dom_lhs by (elim_in_numeral; fold nt_field_simps) (simp_all add: ntsmcf_hcomp_components slicing_simps slicing_commute) qed subsubsection\Natural transformation map\ lemma ntcf_hcomp_NTMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \\Obj\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis unfolding ntsmcf_hcomp_components by (simp add: cat_cs_simps) qed lemma ntcf_hcomp_NTMap_app[cat_cs_simps]: assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = \'\ArrMap\\\\NTMap\\a\\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\\\ObjMap\\a\\" proof- interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) from assms(3) show ?thesis unfolding ntsmcf_hcomp_components by (simp add: cat_cs_simps) qed lemma ntcf_hcomp_NTMap_vrange: assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis by ( rule ntsmcf_hcomp_NTMap_vrange[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] ) qed subsubsection\Further properties\ lemma ntcf_hcomp_composable_commute: \\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f : a \\<^bsub>\\<^esub> b" shows "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> (\' \\<^sub>C\<^sub>F \)\ArrMap\\f\ = (\' \\<^sub>C\<^sub>F \)\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\" (is \?\\b \\<^sub>A\<^bsub>\\<^esub> ?\'\f = ?\'\f \\<^sub>A\<^bsub>\\<^esub> ?\\a\) proof- interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis by ( rule ntsmcf_hcomp_composable_commute[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute, OF assms(3) ] ) qed lemma ntcf_hcomp_is_ntcf: \\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis proof(intro is_ntcfI) show "vfsequence (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" unfolding ntsmcf_hcomp_def by (simp add: nat_omega_simps) show "vcard (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = 5\<^sub>\" unfolding ntsmcf_hcomp_def by (simp add: nat_omega_simps) show "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : cf_smcf (\' \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf (\' \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by ( rule ntsmcf_hcomp_is_ntsmcf[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] ) qed (auto simp: ntsmcf_hcomp_components(1) cat_cs_simps intro: cat_cs_intros) qed lemma ntcf_hcomp_is_ntcf'[cat_cs_intros]: assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \' \\<^sub>C\<^sub>F \" and "\' = \' \\<^sub>C\<^sub>F \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3,4) by (rule ntcf_hcomp_is_ntcf) lemma ntcf_hcomp_associativ[cat_cs_simps]: assumes "\ : \'' \\<^sub>C\<^sub>F \'' : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^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- interpret \: is_ntcf \ \ \
\'' \'' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) show ?thesis proof(rule ntcf_eqI[of \]) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : \'' \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \'' \\<^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 ntsmcf_hcomp_assoc[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_commute ] have "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" by simp then 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\" unfolding slicing_simps . qed (auto intro: cat_cs_intros) qed subsubsection\ The opposite of the horizontal composition of natural transformations \ lemma op_ntcf_ntcf_hcomp[cat_op_simps]: assumes "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \" proof- interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis proof(rule sym, rule ntcf_eqI[of \]) from op_ntsmcf_ntsmcf_hcomp[ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf, unfolded slicing_simps slicing_commute ] have "ntcf_ntsmcf (op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \)\NTMap\ = ntcf_ntsmcf (op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" by simp then show "(op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \)\NTMap\ = op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\" unfolding slicing_simps . have "\ \\<^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 (rule ntcf_hcomp_is_ntcf[OF assms]) from is_ntcf.is_ntcf_op[OF this] show "op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : op_cf \' \\<^sub>C\<^sub>F op_cf \ \\<^sub>C\<^sub>F op_cf \' \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" unfolding cat_op_simps . qed (auto intro: cat_op_intros cat_cs_intros) qed subsection\Interchange law\ lemma ntcf_comp_interchange_law: \\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "((\' \\<^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 \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \': is_ntcf \ \ \ \' \' \' by (rule assms(3)) interpret \': is_ntcf \ \ \ \' \' \' by (rule assms(4)) show ?thesis proof(rule ntcf_eqI) from ntsmcf_comp_interchange_law [ OF \.ntcf_is_ntsmcf \.ntcf_is_ntsmcf \'.ntcf_is_ntsmcf \'.ntcf_is_ntsmcf ] have "( (ntcf_ntsmcf \' \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \') \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F (ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \) )\NTMap\ = ( (ntcf_ntsmcf \' \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (ntcf_ntsmcf \' \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \) )\NTMap\" by simp then show "(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \' \\<^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 (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \))\NTMap\" unfolding slicing_simps slicing_commute . qed (auto intro: cat_cs_intros) qed subsection\Identity natural transformation\ subsubsection\Definition and elementary properties\ text\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ definition ntcf_id :: "V \ V" where "ntcf_id \ = [\\HomCod\\CId\ \\<^sub>\ \\ObjMap\, \, \, \\HomDom\, \\HomCod\]\<^sub>\" text\Components.\ lemma ntcf_id_components: shows "ntcf_id \\NTMap\ = \\HomCod\\CId\ \\<^sub>\ \\ObjMap\" and [dg_shared_cs_simps, cat_cs_simps]: "ntcf_id \\NTDom\ = \" and [dg_shared_cs_simps, cat_cs_simps]: "ntcf_id \\NTCod\ = \" and [dg_shared_cs_simps, cat_cs_simps]: "ntcf_id \\NTDGDom\ = \\HomDom\" and [dg_shared_cs_simps, cat_cs_simps]: "ntcf_id \\NTDGCod\ = \\HomCod\" unfolding ntcf_id_def nt_field_simps by (simp_all add: nat_omega_simps) lemma (in is_functor) is_functor_ntcf_id_components: shows "ntcf_id \\NTMap\ = \\CId\ \\<^sub>\ \\ObjMap\" and "ntcf_id \\NTDom\ = \" and "ntcf_id \\NTCod\ = \" and "ntcf_id \\NTDGDom\ = \" and "ntcf_id \\NTDGCod\ = \" unfolding ntcf_id_components by (simp_all add: cat_cs_simps) subsubsection\Natural transformation map\ lemma (in is_functor) ntcf_id_NTMap_vdomain[cat_cs_simps]: "\\<^sub>\ (ntcf_id \\NTMap\) = \\Obj\" using cf_ObjMap_vrange unfolding is_functor_ntcf_id_components by (auto simp: cat_cs_simps) lemmas [cat_cs_simps] = is_functor.ntcf_id_NTMap_vdomain lemma (in is_functor) ntcf_id_NTMap_app_vdomain[cat_cs_simps]: assumes [simp]: "a \\<^sub>\ \\Obj\" shows "ntcf_id \\NTMap\\a\ = \\CId\\\\ObjMap\\a\\" unfolding is_functor_ntcf_id_components by (rule vsv_vcomp_at) (auto simp: cf_ObjMap_vrange cat_cs_simps cat_cs_intros) lemmas [cat_cs_simps] = is_functor.ntcf_id_NTMap_app_vdomain lemma (in is_functor) ntcf_id_NTMap_vsv[cat_cs_intros]: "vsv (ntcf_id \\NTMap\)" unfolding is_functor_ntcf_id_components by (auto intro: vsv_vcomp) lemmas [cat_cs_intros] = is_functor.ntcf_id_NTMap_vsv lemma (in is_functor) ntcf_id_NTMap_vrange: "\\<^sub>\ (ntcf_id \\NTMap\) \\<^sub>\ \\Arr\" proof(rule vsubsetI) interpret vsv \ntcf_id \\NTMap\\ by (rule ntcf_id_NTMap_vsv) fix f assume "f \\<^sub>\ \\<^sub>\ (ntcf_id \\NTMap\)" then obtain a where f_def: "f = ntcf_id \\NTMap\\a\" and a: "a \\<^sub>\ \\<^sub>\ (ntcf_id \\NTMap\)" using vrange_atD by metis then have "a \\<^sub>\ \\Obj\" and "f = \\CId\\\\ObjMap\\a\\" by (auto simp: cat_cs_simps) then show "f \\<^sub>\ \\Arr\" by (auto dest: cf_ObjMap_app_in_HomCod_Obj HomCod.cat_CId_is_arr) qed subsubsection\Further properties\ lemma (in is_functor) cf_ntcf_id_is_ntcf[cat_cs_intros]: "ntcf_id \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_ntcfI, unfold is_functor_ntcf_id_components(2,3,4,5)) show "ntcf_ntsmcf (ntcf_id \) : cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" proof ( rule is_ntsmcfI, unfold slicing_simps slicing_commute is_functor_ntcf_id_components(2,3,4,5) ) show "ntsmcf_tdghm (ntcf_ntsmcf (ntcf_id \)) : smcf_dghm (cf_smcf \) \\<^sub>D\<^sub>G\<^sub>H\<^sub>M smcf_dghm (cf_smcf \) : smc_dg (cat_smc \) \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg (cat_smc \)" by ( rule is_tdghmI, unfold slicing_simps slicing_commute is_functor_ntcf_id_components(2,3,4,5) ) ( auto simp: cat_cs_simps cat_cs_intros nat_omega_simps ntsmcf_tdghm_def cf_is_semifunctor intro: slicing_intros ) fix f a b assume "f : a \\<^bsub>\\<^esub> b" with is_functor_axioms show "ntcf_id \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> ntcf_id \\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: ntcf_ntsmcf_def nat_omega_simps intro: slicing_intros) qed (auto simp: ntcf_id_def nat_omega_simps intro: cat_cs_intros) lemma (in is_functor) cf_ntcf_id_is_ntcf': assumes "\' = \" and "\' = \" shows "ntcf_id \ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding assms by (rule cf_ntcf_id_is_ntcf) lemmas [cat_cs_intros] = is_functor.cf_ntcf_id_is_ntcf' lemma (in is_ntcf) ntcf_ntcf_vcomp_ntcf_id_left_left[cat_cs_simps]: \\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = \" proof(rule ntcf_eqI[of \]) interpret id: is_ntcf \ \ \ \ \ \ntcf_id \\ by (rule NTCod.cf_ntcf_id_is_ntcf) show "(ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = \\NTMap\" proof(rule vsv_eqI) show [simp]: "\\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \\<^sub>\ (\\NTMap\)" unfolding ntsmcf_vcomp_components by (simp add: cat_cs_simps) fix a assume "a \\<^sub>\ \\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\)" then have "a \\<^sub>\ \\Obj\" by (simp add: cat_cs_simps) then show "(ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = \\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: ntsmcf_vcomp_components) qed (auto intro: cat_cs_intros) lemmas [cat_cs_simps] = is_ntcf.ntcf_ntcf_vcomp_ntcf_id_left_left lemma (in is_ntcf) ntcf_ntcf_vcomp_ntcf_id_right_left[cat_cs_simps]: \\See Chapter II-4 in \cite{mac_lane_categories_2010}.\ "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ = \" proof(rule ntcf_eqI[of \]) interpret id: is_ntcf \ \ \ \ \ \ntcf_id \\ by (rule NTDom.cf_ntcf_id_is_ntcf) show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\ = \\NTMap\" proof(rule vsv_eqI) show [simp]: "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\) = \\<^sub>\ (\\NTMap\)" unfolding ntsmcf_vcomp_components by (simp add: cat_cs_simps) fix a assume "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\)" then have "a \\<^sub>\ \\Obj\" by (simp add: cat_cs_simps) then show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\ = \\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: ntsmcf_vcomp_components) qed (auto intro: cat_cs_intros) lemmas [cat_cs_simps] = is_ntcf.ntcf_ntcf_vcomp_ntcf_id_right_left lemma (in is_ntcf) ntcf_ntcf_hcomp_ntcf_id_left_left[cat_cs_simps]: \\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ "ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = \" proof(rule ntcf_eqI) interpret id: is_ntcf \ \ \ \cf_id \\ \cf_id \\ \ntcf_id (cf_id \)\ by ( simp add: NTDom.HomCod.cat_cf_id_is_functor is_functor.cf_ntcf_id_is_ntcf ) show "ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : cf_id \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F cf_id \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "(ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = \\NTMap\" proof(rule vsv_eqI) fix a assume "a \\<^sub>\ \\<^sub>\ ((ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\)" then have a: "a \\<^sub>\ \\Obj\" unfolding ntcf_hcomp_NTMap_vdomain[OF is_ntcf_axioms] by simp with is_ntcf_axioms show "(ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = \\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: ntsmcf_hcomp_components(1) cat_cs_simps) qed (auto simp: cat_cs_simps intro: cat_cs_intros) lemmas [cat_cs_simps] = is_ntcf.ntcf_ntcf_hcomp_ntcf_id_left_left lemma (in is_ntcf) ntcf_ntcf_hcomp_ntcf_id_right_left[cat_cs_simps]: \\See Chapter II-5 in \cite{mac_lane_categories_2010}.\ "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \) = \" proof(rule ntcf_eqI[of \]) interpret id: is_ntcf \ \ \ \cf_id \\ \cf_id \\ \ntcf_id (cf_id \)\ by ( simp add: NTDom.HomDom.cat_cf_id_is_functor is_functor.cf_ntcf_id_is_ntcf ) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \) : \ \\<^sub>C\<^sub>F cf_id \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F cf_id \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \))\NTMap\ = \\NTMap\" proof(rule vsv_eqI) fix a assume "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \))\NTMap\)" then have a: "a \\<^sub>\ \\Obj\" unfolding ntcf_hcomp_NTMap_vdomain[OF id.is_ntcf_axioms] by simp with is_ntcf_axioms show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \))\NTMap\\a\ = \\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: ntsmcf_hcomp_components(1) cat_cs_simps) qed (auto simp: cat_cs_simps cat_cs_intros) lemmas [cat_cs_simps] = is_ntcf.ntcf_ntcf_hcomp_ntcf_id_right_left subsubsection\The opposite identity natural transformation\ lemma (in is_functor) cf_ntcf_id_op_cf: "ntcf_id (op_cf \) = op_ntcf (ntcf_id \)" proof(rule ntcf_eqI) show ntcfid_op: "ntcf_id (op_cf \) : op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (simp add: is_functor.cf_ntcf_id_is_ntcf local.is_functor_op) show "ntcf_id (op_cf \)\NTMap\ = op_ntcf (ntcf_id \)\NTMap\" by (rule vsv_eqI, unfold cat_op_simps) ( auto simp: cat_op_simps cat_cs_simps ntcf_id_components(1) intro: vsv_vcomp ) qed (auto intro: cat_op_intros cat_cs_intros) subsubsection\Identity natural transformation of a composition of functors\ lemma ntcf_id_cf_comp: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id (\ \\<^sub>C\<^sub>F \) = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" proof(rule ntcf_eqI) from assms show \\: "ntcf_id (\ \\<^sub>C\<^sub>F \) : \ \\<^sub>C\<^sub>F \ \\<^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) interpret \\: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \\ \\<^sub>C\<^sub>F \\ \ntcf_id (\ \\<^sub>C\<^sub>F \)\ by (rule \\) from assms show \_\: "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : \ \\<^sub>C\<^sub>F \ \\<^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) interpret \_\: is_ntcf \ \ \ \\ \\<^sub>C\<^sub>F \\ \\ \\<^sub>C\<^sub>F \\ \ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \\ by (rule \_\) show "ntcf_id (\ \\<^sub>C\<^sub>F \)\NTMap\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\" proof(rule vsv_eqI, unfold \\.ntcf_NTMap_vdomain \_\.ntcf_NTMap_vdomain) fix a assume "a \\<^sub>\ \\Obj\" with assms show "ntcf_id (\ \\<^sub>C\<^sub>F \)\NTMap\\a\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed auto qed auto lemmas [cat_cs_simps] = ntcf_id_cf_comp[symmetric] subsection\Composition of a natural transformation and a functor\ subsubsection\Definition and elementary properties\ abbreviation (input) ntcf_cf_comp :: "V \ V \ V" (infixl "\\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F" 55) where "ntcf_cf_comp \ tdghm_dghm_comp" text\Slicing.\ lemma ntsmcf_tdghm_ntsmcf_smcf_comp[slicing_commute]: "ntcf_ntsmcf \ \\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F\<^sub>-\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf \ = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" unfolding ntcf_ntsmcf_def cf_smcf_def cat_smc_def tdghm_dghm_comp_def dghm_comp_def ntsmcf_tdghm_def smcf_dghm_def smc_dg_def dg_field_simps dghm_field_simps nt_field_simps by (simp add: nat_omega_simps) (*slow*) subsubsection\Natural transformation map\ mk_VLambda (in is_functor) tdghm_dghm_comp_components(1)[where \=\, unfolded cf_HomDom] |vdomain ntcf_cf_comp_NTMap_vdomain[cat_cs_simps]| |app ntcf_cf_comp_NTMap_app[cat_cs_simps]| lemmas [cat_cs_simps] = is_functor.ntcf_cf_comp_NTMap_vdomain is_functor.ntcf_cf_comp_NTMap_app lemma ntcf_cf_comp_NTMap_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show ?thesis unfolding tdghm_dghm_comp_components by (auto simp: cat_cs_simps intro: cat_cs_intros) qed subsubsection\ Opposite of the composition of a natural transformation and a functor \ lemma op_ntcf_ntcf_cf_comp[cat_op_simps]: "op_ntcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) = op_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf \" unfolding tdghm_dghm_comp_def dghm_comp_def op_ntcf_def op_cf_def op_cat_def dg_field_simps dghm_field_simps nt_field_simps by (simp add: nat_omega_simps) (*slow*) subsubsection\ Composition of a natural transformation and a functor is a natural transformation \ lemma ntcf_cf_comp_is_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^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> \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_ntcfI) show "vfsequence (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" unfolding tdghm_dghm_comp_def by (simp add: nat_omega_simps) from assms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "vcard (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) = 5\<^sub>\" unfolding tdghm_dghm_comp_def by (simp add: nat_omega_simps) from assms show "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \) : cf_smcf (\ \\<^sub>C\<^sub>F \) \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: slicing_intros smc_cs_intros cat_cs_intros ) qed (auto simp: tdghm_dghm_comp_components(1) cat_cs_simps) qed -lemma ntcf_cf_comp_is_functor'[cat_cs_intros]: +lemma ntcf_cf_comp_is_ntcf'[cat_cs_intros]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \\<^sub>C\<^sub>F \" and "\' = \ \\<^sub>C\<^sub>F \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3,4) by (simp add: ntcf_cf_comp_is_ntcf) subsubsection\Further properties\ lemma ntcf_cf_comp_ntcf_cf_comp_assoc: assumes "\ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^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>-\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F \)" proof- interpret \: is_ntcf \ \ \
\ \' \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) show ?thesis proof(rule ntcf_ntsmcf_eqI) from assms show "(\ \\<^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>C\<^sub>F \ \\<^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) show "\ \\<^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\<^sub>F \' \\<^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 assms show "ntcf_ntsmcf ((\ \\<^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 \) = ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F \))" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: slicing_intros ntsmcf_smcf_comp_ntsmcf_smcf_comp_assoc ) qed simp_all qed lemma (in is_ntcf) ntcf_ntcf_cf_comp_cf_id[cat_cs_simps]: "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_id \ = \" proof(rule ntcf_ntsmcf_eqI) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_id \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_id \) = ntcf_ntsmcf \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: cat_cs_intros slicing_intros smc_cs_simps ) qed simp_all lemmas [cat_cs_simps] = is_ntcf.ntcf_ntcf_cf_comp_cf_id lemma ntcf_vcomp_ntcf_cf_comp[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^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>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \" proof(rule ntcf_ntsmcf_eqI) from assms show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^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 assms show "ntcf_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^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_ntsmcf (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)" unfolding slicing_commute[symmetric] by (intro ntsmcf_vcomp_ntsmcf_smcf_comp) (cs_concl cs_intro: slicing_intros) qed (use assms in \cs_concl cs_intro: cat_cs_intros\)+ subsection\Composition of a functor and a natural transformation\ subsubsection\Definition and elementary properties\ abbreviation (input) cf_ntcf_comp :: "V \ V \ V" (infixl "\\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F" 55) where "cf_ntcf_comp \ dghm_tdghm_comp" text\Slicing.\ lemma ntcf_ntsmcf_cf_ntcf_comp[slicing_commute]: "cf_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F ntcf_ntsmcf \ = ntcf_ntsmcf (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" unfolding ntcf_ntsmcf_def cf_smcf_def cat_smc_def dghm_tdghm_comp_def dghm_comp_def ntsmcf_tdghm_def smcf_dghm_def smc_dg_def dg_field_simps dghm_field_simps nt_field_simps by (simp add: nat_omega_simps) (*slow*) subsubsection\Natural transformation map\ mk_VLambda (in is_ntcf) dghm_tdghm_comp_components(1)[where \=\, unfolded ntcf_NTDGDom] |vdomain cf_ntcf_comp_NTMap_vdomain| |app cf_ntcf_comp_NTMap_app| lemmas [cat_cs_simps] = is_ntcf.cf_ntcf_comp_NTMap_vdomain is_ntcf.cf_ntcf_comp_NTMap_app lemma cf_ntcf_comp_NTMap_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis unfolding dghm_tdghm_comp_components by (auto simp: cat_cs_simps intro: cat_cs_intros) qed subsubsection\ Opposite of the composition of a functor and a natural transformation \ lemma op_ntcf_cf_ntcf_comp[cat_op_simps]: "op_ntcf (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = op_cf \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F op_ntcf \" unfolding dghm_tdghm_comp_def dghm_comp_def op_ntcf_def op_cf_def op_cat_def dg_field_simps dghm_field_simps nt_field_simps by (simp add: nat_omega_simps) (*slow*) subsubsection\ Composition of a functor and a natural transformation is a natural transformation \ lemma cf_ntcf_comp_is_ntcf: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^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> \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_ntcfI) show "vfsequence (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" unfolding dghm_tdghm_comp_def by simp from assms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show "\ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "vcard (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = 5\<^sub>\" unfolding dghm_tdghm_comp_def by (simp add: nat_omega_simps) from assms show "ntcf_ntsmcf (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) : cf_smcf (\ \\<^sub>C\<^sub>F \) \\<^sub>S\<^sub>M\<^sub>C\<^sub>F cf_smcf (\ \\<^sub>C\<^sub>F \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: slicing_intros smc_cs_intros ) qed (auto simp: dghm_tdghm_comp_components(1) cat_cs_simps) qed lemma cf_ntcf_comp_is_functor'[cat_cs_intros]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \\<^sub>C\<^sub>F \" and "\' = \ \\<^sub>C\<^sub>F \" shows "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3,4) by (simp add: cf_ntcf_comp_is_ntcf) subsubsection\Further properties\ lemma cf_comp_cf_ntcf_comp_assoc: assumes "\ : \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "(\ \\<^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>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" proof(rule ntcf_ntsmcf_eqI) interpret \: is_ntcf \ \ \ \ \' \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_functor \ \ \
\ by (rule assms(3)) from assms show "(\ \\<^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\<^sub>F \ \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by (cs_concl cs_intro: cat_cs_intros) from assms show "\ \\<^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>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^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 assms show "ntcf_ntsmcf (\ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = ntcf_ntsmcf (\ \\<^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 \))" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: slicing_intros smcf_comp_smcf_ntsmcf_comp_assoc ) qed simp_all lemma (in is_ntcf) ntcf_cf_ntcf_comp_cf_id[cat_cs_simps]: "cf_id \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = \" proof(rule ntcf_ntsmcf_eqI) show "cf_id \ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^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) show "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "ntcf_ntsmcf (smcf_id \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) = ntcf_ntsmcf \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: cat_cs_intros slicing_intros smc_cs_simps ) qed simp_all lemmas [cat_cs_simps] = is_ntcf.ntcf_cf_ntcf_comp_cf_id lemma cf_ntcf_comp_ntcf_cf_comp_assoc: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^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 \ = \ \\<^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 \)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \
\ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) show ?thesis by (rule ntcf_ntsmcf_eqI) ( use assms in \ cs_concl cs_simp: cat_cs_simps slicing_commute[symmetric] cs_intro: cat_cs_intros slicing_intros smcf_ntsmcf_comp_ntsmcf_smcf_comp_assoc \ )+ qed lemma ntcf_cf_comp_ntcf_id[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" proof(rule ntcf_eqI) from assms have dom_lhs: "\\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms have dom_rhs: "\\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with assms show "(ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\\a\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: cat_cs_intros) qed (use assms in \cs_concl cs_intro: cat_cs_intros\)+ lemma cf_comp_cf_const_right[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>\ \\Obj\" shows "\ \\<^sub>C\<^sub>F cf_const \ \ \ = cf_const \ \ (\\ObjMap\\\\)" proof(rule cf_eqI) interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_functor \ \ \ \ by (rule assms(3)) from assms(4) show "\ \\<^sub>C\<^sub>F cf_const \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms(4) show "cf_const \ \ (\\ObjMap\\\\) : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(4) have ObjMap_dom_lhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F cf_const \ \ \)\ObjMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(4) have ObjMap_dom_rhs: "\\<^sub>\ (cf_const \ \ (\\ObjMap\\\\)\ObjMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "(\ \\<^sub>C\<^sub>F cf_const \ \ \)\ObjMap\ = cf_const \ \ (\\ObjMap\\\\)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with assms(4) show "(\ \\<^sub>C\<^sub>F cf_const \ \ \)\ObjMap\\a\ = cf_const \ \ (\\ObjMap\\\\)\ObjMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: assms(4) cat_cs_intros) from assms(4) have ArrMap_dom_lhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F cf_const \ \ \)\ArrMap\) = \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(4) have ArrMap_dom_rhs: "\\<^sub>\ (cf_const \ \ (\\ObjMap\\\\)\ArrMap\) = \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(\ \\<^sub>C\<^sub>F cf_const \ \ \)\ArrMap\ = cf_const \ \ (\\ObjMap\\\\)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix a assume "a \\<^sub>\ \\Arr\" with assms(4) show "(\ \\<^sub>C\<^sub>F cf_const \ \ \)\ArrMap\\a\ = cf_const \ \ (\\ObjMap\\\\)\ArrMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: assms(4) cat_cs_intros) qed simp_all lemma cf_ntcf_comp_ntcf_vcomp: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^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>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(3)) show "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^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>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)" by (rule ntcf_ntsmcf_eqI) ( use assms in \ cs_concl cs_simp: smc_cs_simps slicing_commute[symmetric] cs_intro: cat_cs_intros slicing_intros smcf_ntsmcf_comp_ntsmcf_vcomp \ )+ qed subsection\Constant natural transformation\ subsubsection\Definition and elementary properties\ text\See Chapter III in \cite{mac_lane_categories_2010}.\ definition ntcf_const :: "V \ V \ V \ V" where "ntcf_const \ \ f = [ vconst_on (\\Obj\) f, cf_const \ \ (\\Dom\\f\), cf_const \ \ (\\Cod\\f\), \, \ ]\<^sub>\" text\Components.\ lemma ntcf_const_components: shows "ntcf_const \ \ f\NTMap\ = vconst_on (\\Obj\) f" and "ntcf_const \ \ f\NTDom\ = cf_const \ \ (\\Dom\\f\)" and "ntcf_const \ \ f\NTCod\ = cf_const \ \ (\\Cod\\f\)" and "ntcf_const \ \ f\NTDGDom\ = \" and "ntcf_const \ \ f\NTDGCod\ = \" unfolding ntcf_const_def nt_field_simps by (auto simp: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda ntcf_const_components(1)[folded VLambda_vconst_on] |vsv ntcf_const_ObjMap_vsv[cat_cs_intros]| |vdomain ntcf_const_ObjMap_vdomain[cat_cs_simps]| |app ntcf_const_ObjMap_app[cat_cs_simps]| lemma ntcf_const_NTMap_ne_vrange: assumes "\\Obj\ \ 0" shows "\\<^sub>\ (ntcf_const \ \ f\NTMap\) = set {f}" using assms unfolding ntcf_const_components by simp lemma ntcf_const_NTMap_vempty_vrange: assumes "\\Obj\ = 0" shows "\\<^sub>\ (ntcf_const \ \ f\NTMap\) = 0" using assms unfolding ntcf_const_components by simp subsubsection\Constant natural transformation is a natural transformation\ lemma ntcf_const_is_ntcf: assumes "category \ \" and "category \ \" and "f : a \\<^bsub>\\<^esub> b" shows "ntcf_const \ \ f : cf_const \ \ a \\<^sub>C\<^sub>F cf_const \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) show ?thesis proof(intro is_ntcfI') show "vfsequence (ntcf_const \ \ f)" unfolding ntcf_const_def by auto show "cf_const \ \ a : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule cf_const_is_functor) from assms(3) show "a \\<^sub>\ \\Obj\" by (simp add: cat_cs_intros) qed (auto simp: cat_cs_intros) from assms(3) show const_b_is_functor: "cf_const \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: cf_const_is_functor cat_cs_intros) show "vcard (ntcf_const \ \ f) = 5\<^sub>\" unfolding ntcf_const_def by (simp add: nat_omega_simps) show "ntcf_const \ \ f\NTMap\\a'\ : cf_const \ \ a\ObjMap\\a'\ \\<^bsub>\\<^esub> cf_const \ \ b\ObjMap\\a'\" if "a' \\<^sub>\ \\Obj\" for a' by (simp add: that assms(3) ntcf_const_components(1) dghm_const_ObjMap_app) from assms(3) show "ntcf_const \ \ f\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> cf_const \ \ a\ArrMap\\f'\ = cf_const \ \ b \ArrMap\\f'\ \\<^sub>A\<^bsub>\\<^esub> ntcf_const \ \ f\NTMap\\a'\" if "f' : a' \\<^bsub>\\<^esub> b'" for f' a' b' using that dghm_const_ArrMap_app by (auto simp: ntcf_const_components cat_cs_intros cat_cs_simps) qed (use assms(3) in \auto simp: ntcf_const_components\) qed lemma ntcf_const_is_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "f : a \\<^bsub>\\<^esub> b" and "\ = cf_const \ \ a" and "\ = cf_const \ \ b" and "\' = \" and "\' = \" shows "ntcf_const \ \ f : \ \\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1-3) unfolding assms(4-7) by (rule ntcf_const_is_ntcf) subsubsection\Opposite constant natural transformation\ lemma op_ntcf_ntcf_const[cat_op_simps]: "op_ntcf (ntcf_const \ \ f) = ntcf_const (op_cat \) (op_cat \) f" unfolding nt_field_simps dghm_field_simps dg_field_simps dghm_const_def ntcf_const_def op_cat_def op_cf_def op_ntcf_def by (simp_all add: nat_omega_simps) subsubsection\Further properties\ lemma ntcf_const_ntcf_vcomp[cat_cs_simps]: assumes "category \ \" and "category \ \" and "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" shows "ntcf_const \ \ g \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f = ntcf_const \ \ (g \\<^sub>A\<^bsub>\\<^esub> f)" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) from assms(3,4) have gf: "g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" by (simp add: cat_cs_intros) note \\g = ntcf_const_is_ntcf[OF assms(1,2,3)] and \\f = ntcf_const_is_ntcf[OF assms(1,2,4)] show ?thesis proof(rule ntcf_eqI) from ntcf_const_is_ntcf[OF assms(1,2,3)] ntcf_const_is_ntcf[OF assms(1,2,4)] show "ntcf_const \ \ g \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f : cf_const \ \ a \\<^sub>C\<^sub>F cf_const \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule ntcf_vcomp_is_ntcf) show "ntcf_const \ \ (g \\<^sub>A\<^bsub>\\<^esub> f) : cf_const \ \ a \\<^sub>C\<^sub>F cf_const \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule ntcf_const_is_ntcf[OF assms(1,2) gf]) show "(ntcf_const \ \ g \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f)\NTMap\ = ntcf_const \ \ (g \\<^sub>A\<^bsub>\\<^esub> f)\NTMap\" unfolding ntcf_const_components proof(rule vsv_eqI, unfold ntcf_vcomp_NTMap_vdomain[OF \\f]) fix a assume prems: "a \\<^sub>\ \\Obj\" then show "(ntcf_const \ \ g \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f)\NTMap\\a\ = vconst_on (\\Obj\) (g \\<^sub>A\<^bsub>\\<^esub> f)\a\" unfolding ntcf_vcomp_NTMap_app[OF \\g \\f prems] by (simp add: ntcf_const_components) qed (simp_all add: ntsmcf_vcomp_components) qed auto qed lemma ntcf_id_cf_const[cat_cs_simps]: assumes "category \ \" and "category \ \" and "c \\<^sub>\ \\Obj\" shows "ntcf_id (cf_const \ \ c) = ntcf_const \ \ (\\CId\\c\)" proof(rule ntcf_eqI) interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) from assms show "ntcf_const \ \ (\\CId\\c\) : cf_const \ \ c \\<^sub>C\<^sub>F cf_const \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: ntcf_const_is_ntcf) interpret const_c: is_functor \ \ \ \cf_const \ \ c\ by (rule cf_const_is_functor) (auto simp: assms(3) cat_cs_intros) show "ntcf_id (cf_const \ \ c) : cf_const \ \ c \\<^sub>C\<^sub>F cf_const \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule const_c.cf_ntcf_id_is_ntcf) show "ntcf_id (cf_const \ \ c)\NTMap\ = ntcf_const \ \ (\\CId\\c\)\NTMap\" proof(rule vsv_eqI, unfold ntcf_const_components) show "vsv (ntcf_id (cf_const \ \ c)\NTMap\)" unfolding ntcf_id_components by (auto simp: cat_cs_simps intro: vsv_vcomp) qed (auto simp: cat_cs_simps) qed simp_all lemma ntcf_cf_comp_cf_const_right[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "category \ \" and "b \\<^sub>\ \\Obj\" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_const \ \ b = ntcf_const \ \ (\\NTMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) show ?thesis proof(rule ntcf_eqI) from assms(3) show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_const \ \ b : cf_const \ \ (\\ObjMap\\b\) \\<^sub>C\<^sub>F cf_const \ \ (\\ObjMap\\b\) : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) show "ntcf_const \ \ (\\NTMap\\b\) : cf_const \ \ (\\ObjMap\\b\) \\<^sub>C\<^sub>F cf_const \ \ (\\ObjMap\\b\) : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) have dom_lhs: "\\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_const \ \ b)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3) have dom_rhs: "\\<^sub>\ (ntcf_const \ \ (\\NTMap\\b\)\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 cf_const \ \ b)\NTMap\ = ntcf_const \ \ (\\NTMap\\b\)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with assms(3) show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F cf_const \ \ b)\NTMap\\a\ = ntcf_const \ \ (\\NTMap\\b\)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: cat_cs_intros) qed simp_all qed lemma cf_ntcf_comp_ntcf_id[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ = ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule ntcf_eqI) show "\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ : \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) have dom_lhs: "\\<^sub>\ ((\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have dom_rhs: "\\<^sub>\ ((ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\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 ntcf_id \)\NTMap\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" then show "(\ \\<^sub>C\<^sub>F\<^sub>-\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\ = (ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (cs_concl cs_intro: cat_cs_intros) qed simp_all qed +lemma (in is_functor) cf_ntcf_cf_comp_ntcf_const[cat_cs_simps]: + assumes "category \ \" and "f : a \\<^bsub>\\<^esub> b" + shows "ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ = ntcf_const \ \ f" +proof(rule ntcf_eqI) + interpret \: category \ \ by (rule assms(1)) + from assms(2) show "ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ : + cf_const \ \ a \\<^sub>C\<^sub>F cf_const \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + then have dom_lhs: "\\<^sub>\ ((ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\) = \\Obj\" + by (cs_concl cs_simp: cat_cs_simps) + from assms(2) show + "ntcf_const \ \ f : cf_const \ \ a \\<^sub>C\<^sub>F cf_const \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + then have dom_rhs: "\\<^sub>\ (ntcf_const \ \ f\NTMap\) = \\Obj\" + by (cs_concl cs_simp: cat_cs_simps) + show "(ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\ = ntcf_const \ \ f\NTMap\" + proof(rule vsv_eqI, unfold dom_lhs dom_rhs) + fix a assume "a \\<^sub>\ \\Obj\" + then show + "(ntcf_const \ \ f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \)\NTMap\\a\ = ntcf_const \ \ f\NTMap\\a\" + by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) + qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) +qed simp_all + +lemmas [cat_cs_simps] = is_functor.cf_ntcf_cf_comp_ntcf_const + subsection\Natural isomorphism\ text\See Chapter I-4 in \cite{mac_lane_categories_2010}.\ locale is_iso_ntcf = is_ntcf + assumes iso_ntcf_is_arr_isomorphism[cat_arrow_cs_intros]: "a \\<^sub>\ \\Obj\ \ \\NTMap\\a\ : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" syntax "_is_iso_ntcf" :: "V \ V \ V \ V \ V \ V \ bool" (\(_ : _ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o _ : _ \\\<^sub>C\ _)\ [51, 51, 51, 51, 51] 51) translations "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_iso_ntcf \ \ \ \ \ \" lemma (in is_iso_ntcf) iso_ntcf_is_arr_isomorphism': assumes "a \\<^sub>\ \\Obj\" and "A = \\ObjMap\\a\" and "B = \\ObjMap\\a\" shows "\\NTMap\\a\ : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> B" using assms by (auto intro: cat_arrow_cs_intros) lemmas [cat_arrow_cs_intros] = is_iso_ntcf.iso_ntcf_is_arr_isomorphism' lemma (in is_iso_ntcf) iso_ntcf_is_arr_isomorphism'': assumes "a \\<^sub>\ \\Obj\" and "A = \\ObjMap\\a\" and "B = \\ObjMap\\a\" and "F = \\NTMap\\a\" and "\' = \" shows "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms by (auto intro: cat_arrow_cs_intros) text\Rules.\ lemma (in is_iso_ntcf) is_iso_ntcf_axioms'[cat_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_iso_ntcf_axioms) mk_ide rf is_iso_ntcf_def[unfolded is_iso_ntcf_axioms_def] |intro is_iso_ntcfI| |dest is_iso_ntcfD[dest]| |elim is_iso_ntcfE[elim]| lemmas [ntcf_cs_intros] = is_iso_ntcfD(1) subsection\Inverse natural transformation\ subsubsection\Definition and elementary properties\ definition inv_ntcf :: "V \ V" where "inv_ntcf \ = [ (\a\\<^sub>\\\NTDGDom\\Obj\. SOME g. is_inverse (\\NTDGCod\) g (\\NTMap\\a\)), \\NTCod\, \\NTDom\, \\NTDGDom\, \\NTDGCod\ ]\<^sub>\" text\Slicing.\ lemma inv_ntcf_components: shows "inv_ntcf \\NTMap\ = (\a\\<^sub>\\\NTDGDom\\Obj\. SOME g. is_inverse (\\NTDGCod\) g (\\NTMap\\a\))" and [cat_cs_simps]: "inv_ntcf \\NTDom\ = \\NTCod\" and [cat_cs_simps]: "inv_ntcf \\NTCod\ = \\NTDom\" and [cat_cs_simps]: "inv_ntcf \\NTDGDom\ = \\NTDGDom\" and [cat_cs_simps]: "inv_ntcf \\NTDGCod\ = \\NTDGCod\" unfolding inv_ntcf_def nt_field_simps by (simp_all add: nat_omega_simps) text\Components.\ lemma (in is_iso_ntcf) is_iso_ntcf_inv_ntcf_components[cat_cs_simps]: "inv_ntcf \\NTDom\ = \" "inv_ntcf \\NTCod\ = \" "inv_ntcf \\NTDGDom\ = \" "inv_ntcf \\NTDGCod\ = \" unfolding inv_ntcf_components by (simp_all add: cat_cs_simps) subsubsection\Natural transformation map\ lemma inv_ntcf_NTMap_vsv[cat_cs_intros]: "vsv (inv_ntcf \\NTMap\)" unfolding inv_ntcf_components by auto lemma (in is_iso_ntcf) iso_ntcf_inv_ntcf_NTMap_app_is_inverse[cat_cs_intros]: assumes "a \\<^sub>\ \\Obj\" shows "is_inverse \ (inv_ntcf \\NTMap\\a\) (\\NTMap\\a\)" proof- from assms is_iso_ntcf_axioms have "\g. is_inverse \ g (\\NTMap\\a\)" by auto from assms someI2_ex[OF this] show "is_inverse \ (inv_ntcf \\NTMap\\a\) (\\NTMap\\a\)" unfolding inv_ntcf_components by (simp add: cat_cs_simps) qed lemma (in is_iso_ntcf) iso_ntcf_inv_ntcf_NTMap_app_is_the_inverse[cat_cs_intros]: assumes "a \\<^sub>\ \\Obj\" shows "inv_ntcf \\NTMap\\a\ = (\\NTMap\\a\)\\<^sub>C\<^bsub>\\<^esub>" proof- have "is_inverse \ (inv_ntcf \\NTMap\\a\) (\\NTMap\\a\)" by (rule iso_ntcf_inv_ntcf_NTMap_app_is_inverse[OF assms]) from NTDom.HomCod.cat_is_inverse_eq_the_inverse[OF this] show ?thesis . qed lemmas [cat_cs_simps] = is_iso_ntcf.iso_ntcf_inv_ntcf_NTMap_app_is_the_inverse lemma (in is_ntcf) inv_ntcf_NTMap_vdomain[cat_cs_simps]: "\\<^sub>\ (inv_ntcf \\NTMap\) = \\Obj\" unfolding inv_ntcf_components by (simp add: cat_cs_simps) lemmas [cat_cs_simps] = is_ntcf.inv_ntcf_NTMap_vdomain lemma (in is_iso_ntcf) inv_ntcf_NTMap_vrange: "\\<^sub>\ (inv_ntcf \\NTMap\) \\<^sub>\ \\Arr\" proof(rule vsubsetI) interpret inv_\: vsv \inv_ntcf \\NTMap\\ by (rule inv_ntcf_NTMap_vsv) fix f assume "f \\<^sub>\ \\<^sub>\ (inv_ntcf \\NTMap\)" then obtain a where f_def: "f = inv_ntcf \\NTMap\\a\" and "a \\<^sub>\ \\<^sub>\ (inv_ntcf \\NTMap\)" by (blast elim: inv_\.vrange_atE) then have "a \\<^sub>\ \\Obj\" by (simp add: cat_cs_simps) then have "is_inverse \ f (\\NTMap\\a\)" unfolding f_def by (intro iso_ntcf_inv_ntcf_NTMap_app_is_inverse) then show "f \\<^sub>\ \\Arr\" by auto qed subsubsection\Opposite natural isomorphism\ lemma (in is_iso_ntcf) is_iso_ntcf_op: "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" proof- from is_iso_ntcf_axioms have "op_ntcf \ : op_cf \ \\<^sub>C\<^sub>F op_cf \ : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> op_cat \" by (cs_concl cs_intro: cat_cs_intros cat_op_intros) then show ?thesis by (intro is_iso_ntcfI) (auto simp: cat_op_simps cat_arrow_cs_intros) qed lemma (in is_iso_ntcf) is_iso_ntcf_op'[cat_op_intros]: assumes "\' = op_cf \" and "\' = op_cf \" and "\' = op_cat \" and "\' = op_cat \" shows "op_ntcf \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule is_iso_ntcf_op) lemmas is_iso_ntcf_op[cat_op_intros] = is_iso_ntcf.is_iso_ntcf_op subsection\A natural isomorphism is an isomorphism in the category \Funct\\ text\ The results that are presented in this subsection can be found in nLab (see \cite{noauthor_nlab_nodate}\footnote{\url{ https://ncatlab.org/nlab/show/natural+isomorphism }}). \ lemma is_arr_isomorphism_is_iso_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ntcf_id \" and "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ntcf_id \" shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_iso_ntcfI) fix a assume prems: "a \\<^sub>\ \\Obj\" show "\\NTMap\\a\ : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" proof(rule is_arr_isomorphismI) show "is_inverse \ (\\NTMap\\a\) (\\NTMap\\a\)" proof(rule is_inverseI) from prems have "\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\" by (simp add: ntcf_vcomp_NTMap_app[OF assms(2,1) prems]) also have "\ = ntcf_id \\NTMap\\a\" unfolding assms(4) by simp also from prems \.NTDom.ntcf_id_NTMap_app_vdomain have "\ = \\CId\\\\ObjMap\\a\\" unfolding ntcf_id_components by auto finally show "\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = \\CId\\\\ObjMap\\a\\". from prems have "\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\" by (simp add: ntcf_vcomp_NTMap_app[OF assms(1,2) prems]) also have "\ = ntcf_id \\NTMap\\a\" unfolding assms(3) by simp also from prems \.NTCod.ntcf_id_NTMap_app_vdomain have "\ = \\CId\\\\ObjMap\\a\\" unfolding ntcf_id_components by auto finally show "\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\ = \\CId\\\\ObjMap\\a\\". qed (auto simp: prems cat_cs_intros) qed (auto simp: prems cat_cs_intros) qed (auto simp: cat_cs_intros) qed lemma iso_ntcf_is_arr_isomorphism: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows [ntcf_cs_intros]: "inv_ntcf \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ = ntcf_id \" and "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ntcf_id \" proof- interpret is_iso_ntcf \ \ \ \ \ \ by (rule assms(1)) define m where "m a = inv_ntcf \\NTMap\\a\" for a have is_inverse[intro]: "a \\<^sub>\ \\Obj\ \ is_inverse \ (m a) (\\NTMap\\a\)" for a unfolding m_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have [dest, intro, simp]: "a \\<^sub>\ \\Obj\ \ m a : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" for a proof- assume prems: "a \\<^sub>\ \\Obj\" from prems have "\\NTMap\\a\ : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" by (auto intro: cat_cs_intros cat_arrow_cs_intros) with is_inverse[OF prems] show "m a : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" by ( meson NTDom.HomCod.cat_is_inverse_is_arr_isomorphism is_arr_isomorphismD ) qed have [intro]: "f : a \\<^bsub>\\<^esub> b \ m b \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> m a" for f a b proof- assume prems: "f : a \\<^bsub>\\<^esub> b" then have ma: "m a : \\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\a\" and mb: "m b : \\ObjMap\\b\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\b\" and \f: "\\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and \a: "\\NTMap\\a\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\a\" and \f: "\\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and \b: "\\NTMap\\b\ : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by (auto intro: cat_cs_intros) then have \b\f: "\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by (auto intro: cat_cs_intros) from prems have inv_ma: "is_inverse \ (m a) (\\NTMap\\a\)" and inv_mb: "is_inverse \ (\\NTMap\\b\) (m b)" by (auto simp: is_inverse_sym) from mb have mb_\b: "m b \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\ = \\CId\\\\ObjMap\\b\\" by (auto intro: is_inverse_Comp_CId_right[OF inv_mb]) from prems have \a_ma: "\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> m a = \\CId\\\\ObjMap\\a\\" using \a inv_ma ma by (meson is_inverse_Comp_CId_right) from \f have "m b \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = m b \\<^sub>A\<^bsub>\\<^esub> (\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> m a))" unfolding \a_ma by (cs_concl cs_simp: cat_cs_simps) also have "\ = m b \\<^sub>A\<^bsub>\\<^esub> ((\\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\) \\<^sub>A\<^bsub>\\<^esub> m a)" by ( metis ma \f \a NTDom.HomCod.cat_Comp_assoc is_arr_isomorphismD(1) ) also from prems have "\ = m b \\<^sub>A\<^bsub>\\<^esub> ((\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\) \\<^sub>A\<^bsub>\\<^esub> m a)" by (metis ntcf_Comp_commute) also have "\ = (m b \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\)) \\<^sub>A\<^bsub>\\<^esub> m a" by ( metis \b\f ma mb NTDom.HomCod.cat_Comp_assoc is_arr_isomorphismD(1) ) also from \f \b mb NTDom.HomCod.cat_Comp_assoc have "\ = ((m b \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\) \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\) \\<^sub>A\<^bsub>\\<^esub> m a" by (metis is_arr_isomorphismD(1)) also from \f have "\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> m a" unfolding mb_\b by (simp add: cat_cs_simps) finally show "m b \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> m a" by simp qed show \: "inv_ntcf \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_iso_ntcfI is_ntcfI', unfold m_def[symmetric]) show "vfsequence (inv_ntcf \)" unfolding inv_ntcf_def by simp show "vcard (inv_ntcf \) = 5\<^sub>\" unfolding inv_ntcf_def by (simp add: nat_omega_simps) qed (auto simp: cat_cs_simps intro: cat_cs_intros) interpret \: is_iso_ntcf \ \ \ \ \ \inv_ntcf \\ by (rule \) show \\: "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ = ntcf_id \" proof(rule ntcf_eqI) from NTCod.cf_ntcf_id_is_ntcf show "ntcf_id \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \)\NTMap\ = ntcf_id \\NTMap\" proof(rule vsv_eqI) fix a assume "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \)\NTMap\)" then have "a \\<^sub>\ \\Obj\" unfolding ntcf_vcomp_NTMap_vdomain[OF \.is_ntcf_axioms] by simp then show "(\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \)\NTMap\\a\ = ntcf_id \\NTMap\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_arrow_cs_intros ) qed ( auto simp: ntsmcf_vcomp_components(1) cat_cs_simps intro: cat_cs_intros ) qed (auto intro: cat_cs_intros) show \\: "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ntcf_id \" proof(rule ntcf_eqI) show "(inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\ = ntcf_id \\NTMap\" proof(rule vsv_eqI) show "\\<^sub>\ ((inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\) = \\<^sub>\ (ntcf_id \\NTMap\)" by (simp add: ntsmcf_vcomp_components(1) cat_cs_simps) fix a assume "a \\<^sub>\ \\<^sub>\ ((inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\)" then have "a \\<^sub>\ \\Obj\" unfolding ntsmcf_vcomp_components by (simp add: cat_cs_simps) then show "(inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)\NTMap\\a\ = ntcf_id \\NTMap\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_arrow_cs_intros ) qed ( auto simp: ntsmcf_vcomp_components(1) ntcf_id_components(1) cat_cs_simps intro: vsv_vcomp ) qed (auto intro: cat_cs_intros) qed subsubsection\Vertical composition of natural isomorphisms\ lemma ntcf_vcomp_is_iso_ntcf[cat_cs_intros]: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_arr_isomorphism_is_iso_ntcf) note inv_ntcf_\ = iso_ntcf_is_arr_isomorphism[OF assms(1)] and inv_ntcf_\ = iso_ntcf_is_arr_isomorphism[OF assms(2)] note [cat_cs_simps] = inv_ntcf_\(2,3) inv_ntcf_\(2,3) from assms show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) from inv_ntcf_\(1) inv_ntcf_\(1) show "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) from assms inv_ntcf_\(1) inv_ntcf_\(1) have "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \) = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \" by ( cs_concl cs_simp: ntcf_vcomp_assoc cs_intro: cat_cs_intros ntcf_cs_intros ) also from assms have "\ = ntcf_id \" by (cs_concl cs_simp: cat_cs_simps cs_intro: ntcf_cs_intros) finally show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \) = ntcf_id \" by simp from assms inv_ntcf_\(1) inv_ntcf_\(1) have "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by ( cs_concl cs_simp: ntcf_vcomp_assoc cs_intro: cat_cs_intros ntcf_cs_intros ) also from assms have "\ = ntcf_id \" by (cs_concl cs_simp: cat_cs_simps cs_intro: ntcf_cs_intros) finally show "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = ntcf_id \" by simp qed subsubsection\Horizontal composition of natural isomorphisms\ lemma ntcf_hcomp_is_iso_ntcf: assumes "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_arr_isomorphism_is_iso_ntcf) note inv_ntcf_\ = iso_ntcf_is_arr_isomorphism[OF assms(1)] and inv_ntcf_\ = iso_ntcf_is_arr_isomorphism[OF assms(2)] note [cat_cs_simps] = inv_ntcf_\(2,3) inv_ntcf_\(2,3) from assms show "\ \\<^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 ntcf_cs_intros) from inv_ntcf_\(1) inv_ntcf_\(1) show "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ : \' \\<^sub>C\<^sub>F \ \\<^sub>C\<^sub>F \' \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) from assms inv_ntcf_\(1) inv_ntcf_\(1) have "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \) = ntcf_id \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" by ( cs_concl cs_simp: ntcf_comp_interchange_law[symmetric] cat_cs_simps cs_intro: ntcf_cs_intros ) also from assms have "\ = ntcf_id (\' \\<^sub>C\<^sub>F \)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros ntcf_cs_intros) finally show "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \) = ntcf_id (\' \\<^sub>C\<^sub>F \)" by simp from assms inv_ntcf_\(1) inv_ntcf_\(1) have "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = ntcf_id \' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \" by ( cs_concl cs_simp: ntcf_comp_interchange_law[symmetric] cat_cs_simps cs_intro: ntcf_cs_intros ) also from assms have "\ = ntcf_id (\' \\<^sub>C\<^sub>F \)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros ntcf_cs_intros) finally show "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) = ntcf_id (\' \\<^sub>C\<^sub>F \)" by simp qed lemma ntcf_hcomp_is_iso_ntcf'[ntcf_cs_intros]: assumes "\ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \' \\<^sub>C\<^sub>F \" and "\'' = \' \\<^sub>C\<^sub>F \" shows "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \'' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1,2) unfolding assms(3,4) by (rule ntcf_hcomp_is_iso_ntcf) subsubsection\An identity natural transformation is a natural isomorphism\ lemma (in is_functor) cf_ntcf_id_is_iso_ntcf: "ntcf_id \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- have "ntcf_id \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: cat_cs_intros) moreover then have "ntcf_id \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id \ = ntcf_id \" by (cs_concl cs_simp: cat_cs_simps) ultimately show ?thesis by (auto intro: is_arr_isomorphism_is_iso_ntcf) qed lemma (in is_functor) cf_ntcf_id_is_iso_ntcf'[ntcf_cs_intros]: assumes "\' = \" and "\' = \" shows "ntcf_id \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding assms by (rule cf_ntcf_id_is_iso_ntcf) lemmas [ntcf_cs_intros] = is_functor.cf_ntcf_id_is_iso_ntcf' subsection\Functor isomorphism\ subsubsection\Definition and elementary properties\ text\See subsection 1.5 in \cite{bodo_categories_1970}.\ locale iso_functor = fixes \ \ \ assumes iso_cf_is_iso_ntcf: "\\ \ \. \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" notation iso_functor (infixl "\\<^sub>C\<^sub>F\" 50) text\Rules.\ lemma iso_functorI: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" using assms unfolding iso_functor_def by auto lemma iso_functorD[dest]: assumes "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" shows "\\ \ \. \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms unfolding iso_functor_def by auto lemma iso_functorE[elim]: assumes "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" obtains \ \ \ where "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms unfolding iso_functor_def by auto subsubsection\A functor isomorphism is an equivalence relation\ lemma iso_functor_refl: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" proof(rule iso_functorI) from assms show "ntcf_id \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: ntcf_cs_intros) qed lemma iso_functor_sym[sym]: assumes "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" proof- from assms obtain \ \ \ where \: "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from iso_ntcf_is_arr_isomorphism(1)[OF \] show "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" by (auto simp: iso_functorI) qed lemma iso_functor_trans[trans, intro]: assumes "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" and "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" shows "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" proof- from assms(1) obtain \ \ \ where \: "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto moreover from assms(2) obtain \' \' \ where \: "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" by auto ultimately have "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by blast+ then have eq: "\' = \" "\' = \" by auto from \ have \: "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" unfolding eq . from ntcf_vcomp_is_iso_ntcf[OF \ \] show ?thesis by (rule iso_functorI) qed subsubsection\Opposite functor isomorphism\ lemma (in iso_functor) iso_functor_op: "op_cf \ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> op_cf \" proof- from iso_functor_axioms obtain \ \ \ where "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by auto from is_iso_ntcf_op[OF this] have "op_cf \ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> op_cf \" by (auto simp: iso_functorI) then show "op_cf \ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> op_cf \" by (rule iso_functor_sym) qed lemmas iso_functor_op[cat_op_intros] = iso_functor.iso_functor_op text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_PCategory.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_PCategory.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_PCategory.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_PCategory.thy @@ -1,5829 +1,5829 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Product category\ theory CZH_ECAT_PCategory imports CZH_ECAT_NTCF CZH_ECAT_Small_Category CZH_Foundations.CZH_SMC_PSemicategory begin subsection\Background\ text\See Chapter II-3 in \cite{mac_lane_categories_2010}.\ named_theorems cat_prod_cs_simps named_theorems cat_prod_cs_intros subsection\Product category: definition and elementary properties\ definition cat_prod :: "V \ (V \ V) \ V" where "cat_prod I \ = [ (\\<^sub>\i\\<^sub>\I. \ i\Obj\), (\\<^sub>\i\\<^sub>\I. \ i\Arr\), (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). (\i\\<^sub>\I. \ i\Dom\\f\i\\)), (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). (\i\\<^sub>\I. \ i\Cod\\f\i\\)), ( \gf\\<^sub>\composable_arrs (dg_prod I \). (\i\\<^sub>\I. vpfst gf\i\ \\<^sub>A\<^bsub>\ i\<^esub> vpsnd gf\i\) ), (\a\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Obj\). (\i\\<^sub>\I. \ i\CId\\a\i\\)) ]\<^sub>\" syntax "_PCATEGORY" :: "pttrn \ V \ (V \ V) \ V" ("(3\\<^sub>C_\\<^sub>\_./ _)" [0, 0, 10] 10) translations "\\<^sub>Ci\\<^sub>\I. \" \ "CONST cat_prod I (\i. \)" text\Components.\ lemma cat_prod_components: shows "(\\<^sub>Ci\\<^sub>\I. \ i)\Obj\ = (\\<^sub>\i\\<^sub>\I. \ i\Obj\)" and "(\\<^sub>Ci\\<^sub>\I. \ i)\Arr\ = (\\<^sub>\i\\<^sub>\I. \ i\Arr\)" and "(\\<^sub>Ci\\<^sub>\I. \ i)\Dom\ = (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). (\i\\<^sub>\I. \ i\Dom\\f\i\\))" and "(\\<^sub>Ci\\<^sub>\I. \ i)\Cod\ = (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). (\i\\<^sub>\I. \ i\Cod\\f\i\\))" and "(\\<^sub>Ci\\<^sub>\I. \ i)\Comp\ = ( \gf\\<^sub>\composable_arrs (dg_prod I \). (\i\\<^sub>\I. vpfst gf\i\ \\<^sub>A\<^bsub>\ i\<^esub> vpsnd gf\i\) )" and "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\ = (\a\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Obj\). (\i\\<^sub>\I. \ i\CId\\a\i\\))" unfolding cat_prod_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_cat_prod[slicing_commute]: "smc_prod I (\i. cat_smc (\ i)) = cat_smc (\\<^sub>Ci\\<^sub>\I. \ i)" unfolding dg_prod_def cat_smc_def cat_prod_def smc_prod_def dg_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ :: "V \ V" and \ :: V begin lemmas_with [ where \=\\i. cat_smc (\ i)\, unfolded slicing_simps slicing_commute ]: cat_prod_ObjI = smc_prod_ObjI and cat_prod_ObjD = smc_prod_ObjD and cat_prod_ObjE = smc_prod_ObjE and cat_prod_Obj_cong = smc_prod_Obj_cong and cat_prod_ArrI = smc_prod_ArrI and cat_prod_ArrD = smc_prod_ArrD and cat_prod_ArrE = smc_prod_ArrE and cat_prod_Arr_cong = smc_prod_Arr_cong and cat_prod_Dom_vsv[cat_cs_intros] = smc_prod_Dom_vsv and cat_prod_Dom_vdomain[cat_cs_simps] = smc_prod_Dom_vdomain and cat_prod_Dom_app = smc_prod_Dom_app and cat_prod_Dom_app_component_app[cat_cs_simps] = smc_prod_Dom_app_component_app and cat_prod_Cod_vsv[cat_cs_intros] = smc_prod_Cod_vsv and cat_prod_Cod_app = smc_prod_Cod_app and cat_prod_Cod_vdomain[cat_cs_simps] = smc_prod_Cod_vdomain and cat_prod_Cod_app_component_app[cat_cs_simps] = smc_prod_Cod_app_component_app and cat_prod_Comp = smc_prod_Comp and cat_prod_Comp_vdomain[cat_cs_simps] = smc_prod_Comp_vdomain and cat_prod_Comp_app = smc_prod_Comp_app and cat_prod_Comp_app_component[cat_cs_simps] = smc_prod_Comp_app_component and cat_prod_Comp_app_vdomain = smc_prod_Comp_app_vdomain and cat_prod_vunion_Obj_in_Obj = smc_prod_vunion_Obj_in_Obj and cat_prod_vdiff_vunion_Obj_in_Obj = smc_prod_vdiff_vunion_Obj_in_Obj and cat_prod_vunion_Arr_in_Arr = smc_prod_vunion_Arr_in_Arr and cat_prod_vdiff_vunion_Arr_in_Arr = smc_prod_vdiff_vunion_Arr_in_Arr end subsection\Local assumptions for a product category\ locale pcategory_base = \ \ for \ I \ + assumes pcat_categories: "i \\<^sub>\ I \ category \ (\ i)" and pcat_index_in_Vset[cat_cs_intros]: "I \\<^sub>\ Vset \" lemma (in pcategory_base) pcat_categories'[cat_prod_cs_intros]: assumes "i \\<^sub>\ I" and "\' = \" shows "category \' (\ i)" using assms(1) unfolding assms(2) by (rule pcat_categories) text\Rules.\ lemma (in pcategory_base) pcategory_base_axioms'[cat_prod_cs_intros]: assumes "\' = \" and "I' = I" shows "pcategory_base \' I' \" unfolding assms by (rule pcategory_base_axioms) mk_ide rf pcategory_base_def[unfolded pcategory_base_axioms_def] |intro pcategory_baseI| |dest pcategory_baseD[dest]| |elim pcategory_baseE[elim]| lemma pcategory_base_psemicategory_baseI: assumes "psemicategory_base \ I (\i. cat_smc (\ i))" and "\i. i \\<^sub>\ I \ category \ (\ i)" shows "pcategory_base \ I \" proof- interpret psemicategory_base \ I \\i. cat_smc (\ i)\ by (rule assms(1)) show ?thesis by (intro pcategory_baseI) (auto simp: assms(2) psmc_index_in_Vset psmc_Obj_in_Vset psmc_Arr_in_Vset) qed text\Product category is a product semicategory.\ context pcategory_base begin lemma pcat_psemicategory_base: "psemicategory_base \ I (\i. cat_smc (\ i))" proof(intro psemicategory_baseI) from pcat_index_in_Vset show "I \\<^sub>\ Vset \" by auto qed (auto simp: category.cat_semicategory cat_prod_cs_intros) interpretation psmc: psemicategory_base \ I \\i. cat_smc (\ i)\ by (rule pcat_psemicategory_base) lemmas_with [unfolded slicing_simps slicing_commute]: pcat_Obj_in_Vset = psmc.psmc_Obj_in_Vset and pcat_Arr_in_Vset = psmc.psmc_Arr_in_Vset and pcat_smc_prod_Obj_in_Vset = psmc.psmc_smc_prod_Obj_in_Vset and pcat_smc_prod_Arr_in_Vset = psmc.psmc_smc_prod_Arr_in_Vset and cat_prod_Dom_app_in_Obj[cat_cs_intros] = psmc.smc_prod_Dom_app_in_Obj and cat_prod_Cod_app_in_Obj[cat_cs_intros] = psmc.smc_prod_Cod_app_in_Obj and cat_prod_is_arrI = psmc.smc_prod_is_arrI and cat_prod_is_arrD[dest] = psmc.smc_prod_is_arrD and cat_prod_is_arrE[elim] = psmc.smc_prod_is_arrE end lemma cat_prod_dg_prod_is_arr: "g : b \\<^bsub>dg_prod I \\<^esub> c \ g : b \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> c" unfolding is_arr_def cat_prod_def smc_prod_def dg_prod_def dg_field_simps by (simp add: nat_omega_simps) lemma smc_prod_composable_arrs_dg_prod: "composable_arrs (dg_prod I \) = composable_arrs (\\<^sub>Ci\\<^sub>\I. \ i)" unfolding composable_arrs_def cat_prod_dg_prod_is_arr by simp text\Elementary properties.\ lemma (in pcategory_base) pcat_vsubset_index_pcategory_base: assumes "J \\<^sub>\ I" shows "pcategory_base \ J \" proof(intro pcategory_baseI) show "category \ (\ i)" if "i \\<^sub>\ J" for i using that assms by (auto intro: cat_prod_cs_intros) from assms show "J \\<^sub>\ Vset \" by (simp add: vsubset_in_VsetI cat_cs_intros) qed auto subsubsection\Identity\ lemma cat_prod_CId_vsv[cat_cs_intros]: "vsv ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\)" unfolding cat_prod_components by auto lemma cat_prod_CId_vdomain[cat_cs_simps]: "\\<^sub>\ ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\) = (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" unfolding cat_prod_components by simp lemma cat_prod_CId_app: assumes "a \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" shows "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\a\ = (\i\\<^sub>\I. \ i\CId\\a\i\\)" using assms unfolding cat_prod_components by simp lemma cat_prod_CId_app_component[cat_cs_simps]: assumes "a \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" and "i \\<^sub>\ I" shows "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\a\\i\ = \ i\CId\\a\i\\" using assms unfolding cat_prod_components by simp lemma (in pcategory_base) cat_prod_CId_vrange: "\\<^sub>\ ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. \ i\Arr\)" proof(intro vsubsetI) interpret CId: vsv \((\\<^sub>Ci\\<^sub>\I. \ i)\CId\)\ by (rule cat_prod_CId_vsv) fix f assume "f \\<^sub>\ \\<^sub>\ ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\)" then obtain a where f_def: "f = ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\)\a\" and "a \\<^sub>\ \\<^sub>\ ((\\<^sub>Ci\\<^sub>\I. \ i)\CId\)" by (blast dest: CId.vrange_atD) then have a: "a \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" unfolding cat_prod_components by simp show "f \\<^sub>\ (\\<^sub>\i\\<^sub>\I. \ i\Arr\)" unfolding f_def cat_prod_CId_app[OF a] proof(rule VLambda_in_vproduct) fix i assume prems: "i \\<^sub>\ I" interpret \: category \ \\ i\ by (simp add: \i \\<^sub>\ I\ cat_cs_intros cat_prod_cs_intros) from prems a have "a\i\ \\<^sub>\ \ i\Obj\" unfolding cat_prod_components by auto with is_arrD(1) show "\ i\CId\\a\i\\ \\<^sub>\ \ i\Arr\" by (auto intro: cat_cs_intros) qed qed subsubsection\A product \\\-category is a tiny \\\-category\ lemma (in pcategory_base) pcat_tiny_category_cat_prod: assumes "\ \" and "\ \\<^sub>\ \" shows "tiny_category \ (\\<^sub>Ci\\<^sub>\I. \ i)" proof- interpret \: \ \ by (rule assms(1)) show ?thesis proof(intro tiny_categoryI, (unfold slicing_simps)?) show \: "tiny_semicategory \ (cat_smc (\\<^sub>Ci\\<^sub>\I. \ i))" unfolding slicing_commute[symmetric] by ( intro psemicategory_base.psmc_tiny_semicategory_smc_prod; (rule assms pcat_psemicategory_base)? ) interpret \: tiny_semicategory \ \cat_smc (\\<^sub>Ci\\<^sub>\I. \ i)\ by (rule \) show "vfsequence (\\<^sub>Ci\\<^sub>\I. \ i)" unfolding cat_prod_def by auto show "vcard (\\<^sub>Ci\\<^sub>\I. \ i) = 6\<^sub>\" unfolding cat_prod_def by (simp add: nat_omega_simps) show CId: "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\a\ : a \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> a" if a: "a \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" for a proof(rule cat_prod_is_arrI) have [cat_cs_intros]: "a\i\ \\<^sub>\ \ i\Obj\" if i: "i \\<^sub>\ I" for i by (rule cat_prod_ObjD(3)[OF a i]) from that show "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\a\\i\ : a\i\ \\<^bsub>\ i\<^esub> a\i\" if "i \\<^sub>\ I" for i by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros that ) qed (use that in \auto simp: cat_prod_components cat_prod_CId_app that\) show "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> f = f" if "f : a \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> b" for f a b proof(rule cat_prod_Arr_cong) note f = \.smc_is_arrD[unfolded slicing_simps, OF that] note a = f(2) and b = f(3) and f = f(1) from CId[OF b] have CId_b: "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ : b \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> b" by simp from \.smc_Comp_is_arr[unfolded slicing_simps, OF this that] show "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> f \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by (simp add: cat_cs_intros) from that show "f \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by auto fix i assume prems: "i \\<^sub>\ I" interpret \i: category \ \\ i\ by (simp add: prems cat_prod_cs_intros) from prems cat_prod_is_arrD(7)[OF that] have fi: "f\i\ : a\i\ \\<^bsub>\ i\<^esub> b\i\" by auto from prems show "((\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> f)\i\ = f\i\" unfolding cat_prod_Comp_app_component[OF CId_b that prems] unfolding cat_prod_CId_app[OF b] by (auto intro: \i.cat_CId_left_left[OF fi]) qed show "f \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ = f" if "f : b \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> c" for f b c proof(rule cat_prod_Arr_cong) note f = \.smc_is_arrD[unfolded slicing_simps, OF that] note b = f(2) and c = f(3) and f = f(1) from CId[OF b] have CId_b: "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ : b \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> b" by simp from \.smc_Comp_is_arr[unfolded slicing_simps, OF that this] show "f \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\ \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by (simp add: cat_cs_intros) from that show "f \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by auto fix i assume prems: "i \\<^sub>\ I" interpret \i: category \ \\ i\ by (simp add: prems cat_prod_cs_intros) from prems cat_prod_is_arrD[OF that] have fi: "f\i\ : b\i\ \\<^bsub>\ i\<^esub> c\i\" by simp from prems show "(f \\<^sub>A\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\b\)\i\ = f\i\" unfolding cat_prod_Comp_app_component[OF that CId_b prems] unfolding cat_prod_CId_app[OF b] by (auto intro: \i.cat_CId_right_left[OF fi]) qed qed (auto simp: cat_cs_intros cat_cs_simps intro: cat_cs_intros) qed subsection\Further local assumptions for product categories\ subsubsection\Definition and elementary properties\ locale pcategory = pcategory_base \ I \ for \ I \ + assumes pcat_Obj_vsubset_Vset: "J \\<^sub>\ I \ (\\<^sub>Ci\\<^sub>\J. \ i)\Obj\ \\<^sub>\ Vset \" and pcat_Hom_vifunion_in_Vset: "\ J \\<^sub>\ I; A \\<^sub>\ (\\<^sub>Ci\\<^sub>\J. \ i)\Obj\; B \\<^sub>\ (\\<^sub>Ci\\<^sub>\J. \ i)\Obj\; A \\<^sub>\ Vset \; B \\<^sub>\ Vset \ \ \ (\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (\\<^sub>Ci\\<^sub>\J. \ i) a b) \\<^sub>\ Vset \" text\Rules.\ lemma (in pcategory) pcategory_axioms'[cat_prod_cs_intros]: assumes "\' = \" and "I' = I" shows "pcategory \' I' \" unfolding assms by (rule pcategory_axioms) mk_ide rf pcategory_def[unfolded pcategory_axioms_def] |intro pcategoryI| |dest pcategoryD[dest]| |elim pcategoryE[elim]| lemmas [cat_prod_cs_intros] = pcategoryD(1) lemma pcategory_psemicategoryI: assumes "psemicategory \ I (\i. cat_smc (\ i))" and "\i. i \\<^sub>\ I \ category \ (\ i)" shows "pcategory \ I \" proof- interpret psemicategory \ I \\i. cat_smc (\ i)\ by (rule assms(1)) note [unfolded slicing_simps slicing_commute, cat_cs_intros] = psmc_Obj_vsubset_Vset psmc_Hom_vifunion_in_Vset show ?thesis by (intro pcategoryI pcategory_base_psemicategory_baseI) (auto simp: assms(2) smc_prod_cs_intros intro!: cat_cs_intros) qed text\Product category is a product semicategory.\ context pcategory begin lemma pcat_psemicategory: "psemicategory \ I (\i. cat_smc (\ i))" proof(intro psemicategoryI, unfold slicing_simps slicing_commute) show "psemicategory_base \ I (\i. cat_smc (\ i))" by (rule pcat_psemicategory_base) qed (auto intro!: pcat_Obj_vsubset_Vset pcat_Hom_vifunion_in_Vset) interpretation psmc: psemicategory \ I \\i. cat_smc (\ i)\ by (rule pcat_psemicategory) lemmas_with [unfolded slicing_simps slicing_commute]: pcat_Obj_vsubset_Vset' = psmc.psmc_Obj_vsubset_Vset' and pcat_Hom_vifunion_in_Vset' = psmc.psmc_Hom_vifunion_in_Vset' and pcat_cat_prod_vunion_is_arr = psmc.psmc_smc_prod_vunion_is_arr and pcat_cat_prod_vdiff_vunion_is_arr = psmc.psmc_smc_prod_vdiff_vunion_is_arr lemmas_with [unfolded slicing_simps slicing_commute]: pcat_cat_prod_vunion_Comp = psmc.psmc_smc_prod_vunion_Comp and pcat_cat_prod_vdiff_vunion_Comp = psmc.psmc_smc_prod_vdiff_vunion_Comp end text\Elementary properties.\ lemma (in pcategory) pcat_vsubset_index_pcategory: assumes "J \\<^sub>\ I" shows "pcategory \ J \" proof(intro pcategoryI pcategory_psemicategoryI) show "cat_prod J' \\Obj\ \\<^sub>\ Vset \" if \J' \\<^sub>\ J\ for J' proof- from that assms have "J' \\<^sub>\ I" by simp then show "cat_prod J' \\Obj\ \\<^sub>\ Vset \" by (rule pcat_Obj_vsubset_Vset) qed fix A B J' assume prems: "J' \\<^sub>\ J" "A \\<^sub>\ (\\<^sub>Ci\\<^sub>\J'. \ i)\Obj\" "B \\<^sub>\ (\\<^sub>Ci\\<^sub>\J'. \ i)\Obj\" "A \\<^sub>\ Vset \" "B \\<^sub>\ Vset \" show "(\\<^sub>\a\\<^sub>\A. \\<^sub>\b\\<^sub>\B. Hom (\\<^sub>Ci\\<^sub>\J'. \ i) a b) \\<^sub>\ Vset \" proof- from prems(1) assms have "J' \\<^sub>\ I" by simp from pcat_Hom_vifunion_in_Vset[OF this prems(2-5)] show ?thesis. qed qed (rule pcat_vsubset_index_pcategory_base[OF assms]) subsubsection\A product \\\-category is an \\\-category\ lemma (in pcategory) pcat_category_cat_prod: "category \ (\\<^sub>Ci\\<^sub>\I. \ i)" proof- interpret tiny_category \\ + \\ \\\<^sub>Ci\\<^sub>\I. \ i\ by (intro pcat_tiny_category_cat_prod) (auto simp: \_\_\\ \.intro \_Limit_\\ \_\_\\) show ?thesis by (rule category_if_category) ( auto intro!: pcat_Hom_vifunion_in_Vset pcat_Obj_vsubset_Vset intro: cat_cs_intros ) qed subsection\Local assumptions for a finite product category\ subsubsection\Definition and elementary properties\ locale finite_pcategory = pcategory_base \ I \ for \ I \ + assumes fin_pcat_index_vfinite: "vfinite I" text\Rules.\ lemma (in finite_pcategory) finite_pcategory_axioms[cat_prod_cs_intros]: assumes "\' = \" and "I' = I" shows "finite_pcategory \' I' \" unfolding assms by (rule finite_pcategory_axioms) mk_ide rf finite_pcategory_def[unfolded finite_pcategory_axioms_def] |intro finite_pcategoryI| |dest finite_pcategoryD[dest]| |elim finite_pcategoryE[elim]| lemmas [cat_prod_cs_intros] = finite_pcategoryD(1) lemma finite_pcategory_finite_psemicategoryI: assumes "finite_psemicategory \ I (\i. cat_smc (\ i))" and "\i. i \\<^sub>\ I \ category \ (\ i)" shows "finite_pcategory \ I \" proof- interpret finite_psemicategory \ I \\i. cat_smc (\ i)\ by (rule assms(1)) show ?thesis by ( intro assms finite_pcategoryI pcategory_base_psemicategory_baseI finite_psemicategoryD(1)[OF assms(1)] fin_psmc_index_vfinite ) qed subsubsection\ Local assumptions for a finite product semicategory and local assumptions for an arbitrary product semicategory \ sublocale finite_pcategory \ pcategory \ I \ proof- interpret finite_psemicategory \ I \\i. cat_smc (\ i)\ proof(intro finite_psemicategoryI psemicategory_baseI) fix i assume "i \\<^sub>\ I" then interpret \i: category \ \\ i\ by (simp add: pcat_categories) show "semicategory \ (cat_smc (\ i))" by (simp add: \i.cat_semicategory) qed (auto intro!: cat_cs_intros fin_pcat_index_vfinite) show "pcategory \ I \" by (intro pcategory_psemicategoryI) (simp_all add: pcat_categories psemicategory_axioms) qed subsection\Binary union and complement\ lemma (in pcategory) pcat_cat_prod_vunion_CId: assumes "vdisjnt J K" and "J \\<^sub>\ I" and "K \\<^sub>\ I" and "a \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\Obj\" and "b \\<^sub>\ (\\<^sub>Cj\\<^sub>\K. \ j)\Obj\" shows "(\\<^sub>Cj\\<^sub>\J. \ j)\CId\\a\ \\<^sub>\ (\\<^sub>Cj\\<^sub>\K. \ j)\CId\\b\ = (\\<^sub>Ci\\<^sub>\J \\<^sub>\ K. \ i)\CId\\a \\<^sub>\ b\" proof- interpret J\: pcategory \ J \ using assms(2) by (simp add: pcat_vsubset_index_pcategory) interpret K\: pcategory \ K \ using assms(3) by (simp add: pcat_vsubset_index_pcategory) interpret JK\: pcategory \ \J \\<^sub>\ K\ \ using assms(2,3) by (simp add: pcat_vsubset_index_pcategory) interpret J\': category \ \cat_prod J \\ by (rule J\.pcat_category_cat_prod) interpret K\': category \ \cat_prod K \\ by (rule K\.pcat_category_cat_prod) interpret JK\': category \ \cat_prod (J \\<^sub>\ K) \\ by (rule JK\.pcat_category_cat_prod) from assms(4) have CId_a: "cat_prod J \\CId\\a\ : a \\<^bsub>(\\<^sub>Cj\\<^sub>\J. \ j)\<^esub> a" by (auto intro: cat_cs_intros) from assms(5) have CId_b: "cat_prod K \\CId\\b\ : b \\<^bsub>(\\<^sub>Ck\\<^sub>\K. \ k)\<^esub> b" by (auto intro: cat_cs_intros) have CId_a_CId_b: "cat_prod J \\CId\\a\ \\<^sub>\ cat_prod K \\CId\\b\ : a \\<^sub>\ b \\<^bsub>cat_prod (J \\<^sub>\ K) \\<^esub> a \\<^sub>\ b" by (rule pcat_cat_prod_vunion_is_arr[OF assms(1-3) CId_a CId_b]) from CId_a have a: "a \\<^sub>\ cat_prod J \\Obj\" by (auto intro: cat_cs_intros) from CId_b have b: "b \\<^sub>\ cat_prod K \\Obj\" by (auto intro: cat_cs_intros) from CId_a_CId_b have ab: "a \\<^sub>\ b \\<^sub>\ cat_prod (J \\<^sub>\ K) \\Obj\" by (auto intro: cat_cs_intros) note CId_aD = J\.cat_prod_is_arrD[OF CId_a] and CId_bD = K\.cat_prod_is_arrD[OF CId_b] show ?thesis proof(rule cat_prod_Arr_cong[of _ \J \\<^sub>\ K\ \]) from CId_a_CId_b show "cat_prod J \\CId\\a\ \\<^sub>\ cat_prod K \\CId\\b\ \\<^sub>\ cat_prod (J \\<^sub>\ K) \\Arr\" by auto from ab show "cat_prod (J \\<^sub>\ K) \\CId\\a \\<^sub>\ b\ \\<^sub>\ cat_prod (J \\<^sub>\ K) \\Arr\" by (auto intro: JK\'.cat_is_arrD(1) cat_cs_intros) fix i assume "i \\<^sub>\ J \\<^sub>\ K" then consider (iJ) \i \\<^sub>\ J\ | (iK) \i \\<^sub>\ K\ by auto then show "(cat_prod J \\CId\\a\ \\<^sub>\ cat_prod K \\CId\\b\)\i\ = cat_prod (J \\<^sub>\ K) \\CId\\a \\<^sub>\ b\\i\" by cases ( auto simp: assms(1) CId_aD(1-4) CId_bD(1-4) cat_prod_CId_app[OF ab] cat_prod_CId_app[OF a] cat_prod_CId_app[OF b] ) qed qed lemma (in pcategory) pcat_cat_prod_vdiff_vunion_CId: assumes "J \\<^sub>\ I" and "a \\<^sub>\ (\\<^sub>Cj\\<^sub>\I -\<^sub>\ J. \ j)\Obj\" and "b \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\Obj\" shows "(\\<^sub>Cj\\<^sub>\I -\<^sub>\ J. \ j)\CId\\a\ \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\CId\\b\ = (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\a \\<^sub>\ b\" by ( vdiff_of_vunion' rule: pcat_cat_prod_vunion_CId assms: assms(2-3) subset: assms(1) ) subsection\Projection\ subsubsection\Definition and elementary properties\ text\See Chapter II-3 in \cite{mac_lane_categories_2010}.\ definition cf_proj :: "V \ (V \ V) \ V \ V" (\\\<^sub>C\) where "\\<^sub>C I \ i = [ (\a\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Obj\). a\i\), (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). f\i\), (\\<^sub>Ci\\<^sub>\I. \ i), \ i ]\<^sub>\" text\Components.\ lemma cf_proj_components: shows "\\<^sub>C I \ i\ObjMap\ = (\a\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Obj\). a\i\)" and "\\<^sub>C I \ i\ArrMap\ = (\f\\<^sub>\(\\<^sub>\i\\<^sub>\I. \ i\Arr\). f\i\)" and "\\<^sub>C I \ i\HomDom\ = (\\<^sub>Ci\\<^sub>\I. \ i)" and "\\<^sub>C I \ i\HomCod\ = \ i" unfolding cf_proj_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing\ lemma cf_smcf_cf_proj[slicing_commute]: "\\<^sub>S\<^sub>M\<^sub>C I (\i. cat_smc (\ i)) i = cf_smcf (\\<^sub>C I \ i)" unfolding cat_smc_def cf_smcf_def smcf_proj_def cf_proj_def cat_prod_def smc_prod_def dg_prod_def dg_field_simps dghm_field_simps by (simp add: nat_omega_simps) context pcategory begin interpretation psmc: psemicategory \ I \\i. cat_smc (\ i)\ by (rule pcat_psemicategory) lemmas_with [unfolded slicing_simps slicing_commute]: pcat_cf_proj_is_semifunctor = psmc.psmc_smcf_proj_is_semifunctor end subsubsection\Projection functor is a functor\ lemma (in pcategory) pcat_cf_proj_is_functor: assumes "i \\<^sub>\ I" shows "\\<^sub>C I \ i : (\\<^sub>Ci\\<^sub>\I. \ i) \\\<^sub>C\<^bsub>\\<^esub> \ i" proof(intro is_functorI) interpret \: category \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ by (simp add: pcat_category_cat_prod) show "vfsequence (\\<^sub>C I \ i)" unfolding cf_proj_def by simp show "category \ (\\<^sub>Ci\\<^sub>\I. \ i)" by (simp add: \.category_axioms) show "vcard (\\<^sub>C I \ i) = 4\<^sub>\" unfolding cf_proj_def by (simp add: nat_omega_simps) show "\\<^sub>C I \ i\ArrMap\\(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\\ = \ i\CId\\\\<^sub>C I \ i\ObjMap\\c\\" if "c \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" for c proof- interpret \i: category \ \\ i\ by (auto intro: assms cat_prod_cs_intros) from that have "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\ : c \\<^bsub>(\\<^sub>Ci\\<^sub>\I. \ i)\<^esub> c" by (simp add: \.cat_CId_is_arr) then have "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\ \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by (auto intro: cat_cs_intros) with assms have "\\<^sub>C I \ i\ArrMap\\(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\\ = (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\\i\" unfolding cf_proj_components cat_prod_components by simp also from assms have "\ = \ i\CId\\c\i\\" unfolding cat_prod_CId_app[OF that] by simp also from that have "\ = \ i\CId\\\\<^sub>C I \ i\ObjMap\\c\\" unfolding cf_proj_components cat_prod_components by simp finally show "\\<^sub>C I \ i\ArrMap\\(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\c\\ = \ i\CId\\\\<^sub>C I \ i\ObjMap\\c\\" by simp qed qed ( auto simp: assms cf_proj_components pcat_cf_proj_is_semifunctor cat_prod_cs_intros ) lemma (in pcategory) pcat_cf_proj_is_functor': assumes "i \\<^sub>\ I" and "\ = (\\<^sub>Ci\\<^sub>\I. \ i)" and "\
= \ i" shows "\\<^sub>C I \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1) unfolding assms(2,3) by (rule pcat_cf_proj_is_functor) lemmas [cat_cs_intros] = pcategory.pcat_cf_proj_is_functor' subsection\Category product universal property functor\ subsubsection\Definition and elementary properties\ text\ The functor that is presented in this section is used in the proof of the universal property of the product category later in this work. \ definition cf_up :: "V \ (V \ V) \ V \ (V \ V) \ V" where "cf_up I \ \ \ = [ (\a\\<^sub>\\\Obj\. (\i\\<^sub>\I. \ i\ObjMap\\a\)), (\f\\<^sub>\\\Arr\. (\i\\<^sub>\I. \ i\ArrMap\\f\)), \, (\\<^sub>Ci\\<^sub>\I. \ i) ]\<^sub>\" text\Components.\ lemma cf_up_components: shows "cf_up I \ \ \\ObjMap\ = (\a\\<^sub>\\\Obj\. (\i\\<^sub>\I. \ i\ObjMap\\a\))" and "cf_up I \ \ \\ArrMap\ = (\f\\<^sub>\\\Arr\. (\i\\<^sub>\I. \ i\ArrMap\\f\))" and "cf_up I \ \ \\HomDom\ = \" and "cf_up I \ \ \\HomCod\ = (\\<^sub>Ci\\<^sub>\I. \ i)" unfolding cf_up_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smcf_dghm_cf_up[slicing_commute]: "smcf_up I (\i. cat_smc (\ i)) (cat_smc \) (\i. cf_smcf (\ i)) = cf_smcf (cf_up I \ \ \)" unfolding cat_smc_def cf_smcf_def cf_up_def smcf_up_def cat_prod_def smc_prod_def dg_prod_def dg_field_simps dghm_field_simps by (simp add: nat_omega_simps) context fixes \ \ :: "V \ V" and \ :: V begin lemmas_with [ where \=\\i. cat_smc (\ i)\ and \=\\i. cf_smcf (\ i)\ and \ = \cat_smc \\, unfolded slicing_simps slicing_commute ]: cf_up_ObjMap_vdomain[simp] = smcf_up_ObjMap_vdomain and cf_up_ObjMap_app = smcf_up_ObjMap_app and cf_up_ObjMap_app_vdomain[simp] = smcf_up_ObjMap_app_vdomain and cf_up_ObjMap_app_component = smcf_up_ObjMap_app_component and cf_up_ArrMap_vdomain[simp] = smcf_up_ArrMap_vdomain and cf_up_ArrMap_app = smcf_up_ArrMap_app and cf_up_ArrMap_app_vdomain[simp] = smcf_up_ArrMap_app_vdomain and cf_up_ArrMap_app_component = smcf_up_ArrMap_app_component lemma cf_up_ObjMap_vrange: assumes "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" shows "\\<^sub>\ (cf_up I \ \ \\ObjMap\) \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" proof ( rule smcf_up_ObjMap_vrange[ where \=\\i. cat_smc (\ i)\ and \=\\i. cf_smcf (\ i)\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute ] ) fix i assume "i \\<^sub>\ I" then interpret is_functor \ \ \\ i\ \\ i\ by (rule assms) show "cf_smcf (\ i) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\ i)" by (rule cf_is_semifunctor) qed lemma cf_up_ObjMap_app_vrange: assumes "a \\<^sub>\ \\Obj\" and "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" shows " \\<^sub>\ (cf_up I \ \ \\ObjMap\\a\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. \ i\Obj\)" proof ( rule smcf_up_ObjMap_app_vrange[ where \=\\i. cat_smc (\ i)\ and \=\\i. cf_smcf (\ i)\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute ] ) show "a \\<^sub>\ \\Obj\" by (rule assms) fix i assume "i \\<^sub>\ I" then interpret is_functor \ \ \\ i\ \\ i\ by (rule assms(2)) show "cf_smcf (\ i) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\ i)" by (rule cf_is_semifunctor) qed lemma cf_up_ArrMap_vrange: assumes "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" shows "\\<^sub>\ (cf_up I \ \ \\ArrMap\) \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" proof ( rule smcf_up_ArrMap_vrange[ where \=\\i. cat_smc (\ i)\ and \=\\i. cf_smcf (\ i)\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute ] ) fix i assume "i \\<^sub>\ I" then interpret is_functor \ \ \\ i\ \\ i\ by (rule assms) show "cf_smcf (\ i) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\ i)" by (rule cf_is_semifunctor) qed lemma cf_up_ArrMap_app_vrange: assumes "a \\<^sub>\ \\Arr\" and "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" shows " \\<^sub>\ (cf_up I \ \ \\ArrMap\\a\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. \ i\Arr\)" proof ( rule smcf_up_ArrMap_app_vrange [ where \=\\i. cat_smc (\ i)\ and \=\\i. cf_smcf (\ i)\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute ] ) fix i assume "i \\<^sub>\ I" then interpret is_functor \ \ \\ i\ \\ i\ by (rule assms(2)) show "cf_smcf (\ i) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\ i)" by (rule cf_is_semifunctor) qed (rule assms) end context pcategory begin interpretation psmc: psemicategory \ I \\i. cat_smc (\ i)\ by (rule pcat_psemicategory) lemmas_with [unfolded slicing_simps slicing_commute]: pcat_smcf_comp_smcf_proj_smcf_up = psmc.psmc_Comp_smcf_proj_smcf_up and pcat_smcf_up_eq_smcf_proj = psmc.psmc_smcf_up_eq_smcf_proj end subsubsection\Category product universal property functor is a functor\ lemma (in pcategory) pcat_cf_up_is_functor: assumes "category \ \" and "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" shows "cf_up I \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)" proof- interpret \: category \ \ by (simp add: assms(1)) interpret \: category \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ by (rule pcat_category_cat_prod) show ?thesis proof(intro is_functorI) show "vfsequence (cf_up I \ \ \)" unfolding cf_up_def by simp show "vcard (cf_up I \ \ \) = 4\<^sub>\" unfolding cf_up_def by (simp add: nat_omega_simps) show "cf_smcf (cf_up I \ \ \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\\<^sub>Ci\\<^sub>\I. \ i)" unfolding slicing_commute[symmetric] by (rule psemicategory.psmc_smcf_up_is_semifunctor) ( auto simp: assms(2) pcat_psemicategory is_functor.cf_is_semifunctor slicing_intros ) show "cf_up I \ \ \\ArrMap\\\\CId\\c\\ = (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\cf_up I \ \ \\ObjMap\\c\\" if "c \\<^sub>\ \\Obj\" for c proof(rule cat_prod_Arr_cong) from that is_arrD(1) have CId_c: "\\CId\\c\ \\<^sub>\ \\Arr\" by (auto intro: cat_cs_intros) from CId_c cf_up_ArrMap_vrange[OF assms(2), simplified] show "cf_up I \ \ \\ArrMap\\\\CId\\c\\ \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" unfolding cf_up_components by force have cf_up_\_c: "cf_up I \ \ \\ObjMap\\c\ \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Obj\" unfolding cat_prod_components proof(intro vproductI ballI) fix i assume prems: "i \\<^sub>\ I" interpret \: is_functor \ \ \\ i\ \\ i\ by (simp add: prems assms(2)) from that show "cf_up I \ \ \\ObjMap\\c\\i\ \\<^sub>\ \ i\Obj\" unfolding cf_up_ObjMap_app_component[OF that prems] by (auto intro: cat_cs_intros) qed (simp_all add: cf_up_ObjMap_app that cf_up_ObjMap_app[OF that]) from \.cat_CId_is_arr[OF this] show "(\\<^sub>Ci\\<^sub>\I. \ i)\CId\\cf_up I \ \ \\ObjMap\\c\\ \\<^sub>\ (\\<^sub>Ci\\<^sub>\I. \ i)\Arr\" by auto fix i assume prems: "i \\<^sub>\ I" interpret \: is_functor \ \ \\ i\ \\ i\ by (simp add: prems assms(2)) from cf_up_\_c prems show "cf_up I \ \ \\ArrMap\\\\CId\\c\\\i\ = (\\<^sub>Ci\\<^sub>\I. \ i)\CId\\cf_up I \ \ \\ObjMap\\c\\\i\" unfolding cf_up_ArrMap_app_component[OF CId_c prems] cat_prod_components by ( simp add: that cf_up_ObjMap_app_component[OF that prems] \.cf_ObjMap_CId ) qed qed (auto simp: cf_up_components cat_cs_intros) qed subsubsection\Further properties\ lemma (in pcategory) pcat_Comp_cf_proj_cf_up: assumes "category \ \" and "\i. i \\<^sub>\ I \ \ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" and "i \\<^sub>\ I" shows "\ i = \\<^sub>C I \ i \\<^sub>C\<^sub>F (cf_up I \ \ \)" proof- interpret \: is_functor \ \ \\ i\ \\ i\ by (rule assms(2)[OF assms(3)]) interpret \: is_functor \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ \\ i\ \\\<^sub>C I \ i\ by (simp add: assms(3) pcat_cf_proj_is_functor) interpret up: is_functor \ \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ \cf_up I \ \ \\ by (simp add: assms(2) \.HomDom.category_axioms pcat_cf_up_is_functor) show ?thesis proof(rule cf_smcf_eqI) show "\\<^sub>C I \ i \\<^sub>C\<^sub>F cf_up I \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" by (auto intro: cat_cs_intros) from assms show "cf_smcf (\ i) = cf_smcf (\\<^sub>C I \ i \\<^sub>C\<^sub>F cf_up I \ \ \)" unfolding slicing_simps slicing_commute[symmetric] by ( intro pcat_smcf_comp_smcf_proj_smcf_up[ where \=\\i. cf_smcf (\ i)\, unfolded slicing_commute[symmetric] ] ) (auto simp: is_functor.cf_is_semifunctor) qed (auto intro: cat_cs_intros) qed lemma (in pcategory) pcat_cf_up_eq_cf_proj: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)" and "\i. i \\<^sub>\ I \ \ i = \\<^sub>C I \ i \\<^sub>C\<^sub>F \" shows "cf_up I \ \ \ = \" proof(rule cf_smcf_eqI) interpret \: is_functor \ \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ \ by (rule assms(1)) show "cf_up I \ \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)" proof(rule pcat_cf_up_is_functor) fix i assume prems: "i \\<^sub>\ I" then interpret \: is_functor \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ \\ i\ \\\<^sub>C I \ i\ by (rule pcat_cf_proj_is_functor) show "\ i : \ \\\<^sub>C\<^bsub>\\<^esub> \ i" unfolding assms(2)[OF prems] by (auto intro: cat_cs_intros) qed (auto intro: cat_cs_intros) show "\ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\I. \ i)" by (rule assms(1)) from assms show "cf_smcf (cf_up I \ \ \) = cf_smcf \" unfolding slicing_commute[symmetric] by (intro pcat_smcf_up_eq_smcf_proj) (auto simp: slicing_commute) qed simp_all subsection\Prodfunctor with respect to a fixed argument\ text\ A prodfunctor is a functor whose domain is a product category. It is a generalization of the concept of the bifunctor, as presented in Chapter II-3 in \cite{mac_lane_categories_2010}. \ definition prodfunctor_proj :: "V \ V \ (V \ V) \ V \ V \ V \ V" where "prodfunctor_proj \ I \ \
J c = [ (\b\\<^sub>\(\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)\Obj\. \\ObjMap\\b \\<^sub>\ c\), (\f\\<^sub>\(\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)\Arr\. \\ArrMap\\f \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\CId\\c\\), (\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i), \
]\<^sub>\" syntax "_PPRODFUNCTOR_PROJ" :: "V \ pttrn \ V \ V \ (V \ V) \ V \ V \ V" (\(_\<^bsub>(3\\<^sub>C_\\<^sub>\_-\<^sub>\_./_),_\<^esub>/'(/-,_/'))\ [51, 51, 51, 51, 51, 51, 51] 51) translations "\\<^bsub>\\<^sub>Ci\\<^sub>\I-\<^sub>\J. \,\
\<^esub>(-,c)" \ "CONST prodfunctor_proj \ I (\i. \) \
J c" text\Components.\ lemma prodfunctor_proj_components: shows "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c))\ObjMap\ = (\b\\<^sub>\(\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)\Obj\. \\ObjMap\\b \\<^sub>\ c\)" and "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c))\ArrMap\ = (\f\\<^sub>\(\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)\Arr\. \\ArrMap\\f \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\CId\\c\\)" and "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c))\HomDom\ = (\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)" and "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c))\HomCod\ = \
" unfolding prodfunctor_proj_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda prodfunctor_proj_components(1) |vsv prodfunctor_proj_ObjMap_vsv[cat_cs_intros]| |vdomain prodfunctor_proj_ObjMap_vdomain[cat_cs_simps]| |app prodfunctor_proj_ObjMap_app[cat_cs_simps]| subsubsection\Arrow map\ mk_VLambda prodfunctor_proj_components(2) |vsv prodfunctor_proj_ArrMap_vsv[cat_cs_intros]| |vdomain prodfunctor_proj_ArrMap_vdomain[cat_cs_simps]| |app prodfunctor_proj_ArrMap_app[cat_cs_simps]| subsubsection\Prodfunctor with respect to a fixed argument is a functor\ lemma (in pcategory) pcat_prodfunctor_proj_is_functor: assumes "\ : (\\<^sub>Ci\\<^sub>\I. \ i) \\\<^sub>C\<^bsub>\\<^esub> \
" and "c \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\Obj\" and "J \\<^sub>\ I" shows "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c)) : (\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i) \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret is_functor \ \(\\<^sub>Ci\\<^sub>\I. \ i)\ \
\ by (rule assms(1)) interpret \: pcategory \ J \ using assms(3) by (intro pcat_vsubset_index_pcategory) auto interpret J_\: category \ \\\<^sub>Ci\\<^sub>\J. \ i\ by (rule \.pcat_category_cat_prod) interpret IJ: pcategory \ \I -\<^sub>\ J\ \ using assms(3) by (intro pcat_vsubset_index_pcategory) auto interpret IJ_\: category \ \\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i\ by (rule IJ.pcat_category_cat_prod) let ?IJ\ = \(\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)\ from assms(2) have "c \\<^sub>\ (\\<^sub>\j\\<^sub>\J. \ j\Obj\)" unfolding cat_prod_components by simp then have "(\\<^sub>\j\\<^sub>\J. \ j\Obj\) \ 0" by (auto intro!: cat_cs_intros) show ?thesis proof(intro is_functorI', unfold prodfunctor_proj_components) show "vfsequence (prodfunctor_proj \ I \ \
J c)" unfolding prodfunctor_proj_def by simp show "vcard (prodfunctor_proj \ I \ \
J c) = 4\<^sub>\" unfolding prodfunctor_proj_def by (simp add: nat_omega_simps) show "\\<^sub>\ (\b\\<^sub>\?IJ\\Obj\. \\ObjMap\\b \\<^sub>\ c\) \\<^sub>\ \
\Obj\" proof(intro vsubsetI) fix x assume "x \\<^sub>\ \\<^sub>\ (\b\\<^sub>\?IJ\\Obj\. \\ObjMap\\b \\<^sub>\ c\)" then obtain b where x_def: "x = \\ObjMap\\b \\<^sub>\ c\" and b: "b \\<^sub>\ ?IJ\\Obj\" by auto have "b \\<^sub>\ c \\<^sub>\ cat_prod I \\Obj\" proof(rule cat_prod_vdiff_vunion_Obj_in_Obj) show "b \\<^sub>\ ?IJ\\Obj\" by (rule b) qed (intro assms(2,3))+ then show "x \\<^sub>\ \
\Obj\" unfolding x_def by (auto intro: cat_cs_intros) qed show is_arr: "(\f\\<^sub>\?IJ\\Arr\. \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\)\f\ : (\b\\<^sub>\?IJ\\Obj\. \\ObjMap\\b \\<^sub>\ c\)\a\ \\<^bsub>\
\<^esub> (\b\\<^sub>\?IJ\\Obj\. \\ObjMap\\b \\<^sub>\ c\)\b\" (is \?V_f: ?V_a \\<^bsub>\
\<^esub> ?V_b\) if "f : a \\<^bsub>?IJ\\<^esub> b" for f a b proof- let ?fc = \f \\<^sub>\ cat_prod J \\CId\\c\\ have "?fc : a \\<^sub>\ c \\<^bsub>cat_prod I \\<^esub> b \\<^sub>\ c" proof(rule pcat_cat_prod_vdiff_vunion_is_arr) show "f : a \\<^bsub>?IJ\\<^esub> b" by (rule that) qed (auto simp: assms cat_cs_intros) then have "\\ArrMap\\?fc\ : \\ObjMap\\a \\<^sub>\ c\ \\<^bsub>\
\<^esub> \\ObjMap\\b \\<^sub>\ c\" by (auto intro: cat_cs_intros) moreover from that have "f \\<^sub>\ ?IJ\\Arr\" "a \\<^sub>\ ?IJ\\Obj\" "b \\<^sub>\ ?IJ\\Obj\" by (auto intro: cat_cs_intros) ultimately show ?thesis by simp qed show "(\f\\<^sub>\?IJ\\Arr\. \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\)\g \\<^sub>A\<^bsub>?IJ\\<^esub> f\ = (\f\\<^sub>\?IJ\\Arr\. \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\)\g\ \\<^sub>A\<^bsub>\
\<^esub> (\f\\<^sub>\?IJ\\Arr\. \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\)\f\" if "g : b' \\<^bsub>?IJ\\<^esub> c'" and "f : a' \\<^bsub>?IJ\\<^esub> b'" for g b' c' f a' proof- from that have gf: "g \\<^sub>A\<^bsub>?IJ\\<^esub> f : a' \\<^bsub>?IJ\\<^esub> c'" by (auto intro: cat_cs_intros) from assms(2) have CId_c: "cat_prod J \\CId\\c\ : c \\<^bsub>cat_prod J \\<^esub> c" by (auto intro: cat_cs_intros) then have [simp]: "cat_prod J \\CId\\c\ \\<^sub>A\<^bsub>cat_prod J \\<^esub> cat_prod J \\CId\\c\ = cat_prod J \\CId\\c\" by (auto simp: cat_cs_simps) from assms(3) that(1) CId_c have g_CId_c: "g \\<^sub>\ cat_prod J \\CId\\c\ : b' \\<^sub>\ c \\<^bsub>cat_prod I \\<^esub> c' \\<^sub>\ c" by (rule pcat_cat_prod_vdiff_vunion_is_arr) from assms(3) that(2) CId_c have f_CId_c: "f \\<^sub>\ cat_prod J \\CId\\c\ : a' \\<^sub>\ c \\<^bsub>cat_prod I \\<^esub> b' \\<^sub>\ c" by (rule pcat_cat_prod_vdiff_vunion_is_arr) have "\\ArrMap\\(g \\<^sub>A\<^bsub>?IJ\\<^esub> f) \\<^sub>\ cat_prod J \\CId\\c\\ = \\ArrMap\\g \\<^sub>\ cat_prod J \\CId\\c\\ \\<^sub>A\<^bsub>\
\<^esub> \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\" unfolding pcat_cat_prod_vdiff_vunion_Comp[ OF assms(3) that(1) CId_c that(2) CId_c, simplified ] by (intro cf_ArrMap_Comp[OF g_CId_c f_CId_c]) moreover from gf have "g \\<^sub>A\<^bsub>?IJ\\<^esub> f \\<^sub>\ ?IJ\\Arr\" by auto moreover from that have "g \\<^sub>\ ?IJ\\Arr\" "f \\<^sub>\ ?IJ\\Arr\" by auto ultimately show ?thesis by simp qed show "(\f\\<^sub>\?IJ\\Arr\. \\ArrMap\\f \\<^sub>\ cat_prod J \\CId\\c\\)\?IJ\\CId\\c'\\ = \
\CId\\(\b\\<^sub>\?IJ\\Obj\. \\ObjMap\\b \\<^sub>\ c\)\c'\\" if "c' \\<^sub>\ ?IJ\\Obj\" for c' proof- have "?IJ\\CId\\c'\ \\<^sub>\ cat_prod J \\CId\\c\ = cat_prod I \\CId\\c' \\<^sub>\ c\" unfolding pcat_cat_prod_vdiff_vunion_CId[OF assms(3) that assms(2)] .. moreover from assms(3) that assms(2) have "c' \\<^sub>\ c \\<^sub>\ cat_prod I \\Obj\" by (rule cat_prod_vdiff_vunion_Obj_in_Obj) ultimately have "\\ArrMap\\?IJ\\CId\\c'\ \\<^sub>\ cat_prod J \\CId\\c\\ = \
\CId\\\\ObjMap\\c' \\<^sub>\ c\\" by (auto intro: cat_cs_intros) moreover from that have CId_c': "?IJ\\CId\\c'\ \\<^sub>\ ?IJ\\Arr\" by (auto dest!: IJ_\.cat_CId_is_arr) ultimately show ?thesis by (simp add: that) qed qed (auto intro: cat_cs_intros) qed lemma (in pcategory) pcat_prodfunctor_proj_is_functor': assumes "\ : (\\<^sub>Ci\\<^sub>\I. \ i) \\\<^sub>C\<^bsub>\\<^esub> \
" and "c \\<^sub>\ (\\<^sub>Cj\\<^sub>\J. \ j)\Obj\" and "J \\<^sub>\ I" and "\' = (\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i)" and "\' = \
" shows "(\\<^bsub>\\<^sub>Ci\\<^sub>\I -\<^sub>\ J. \ i,\
\<^esub>(-,c)) : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1-3) unfolding assms(4,5) by (rule pcat_prodfunctor_proj_is_functor) lemmas [cat_cs_intros] = pcategory.pcat_prodfunctor_proj_is_functor' subsection\Singleton category\ subsubsection\Slicing\ context fixes \ :: V begin lemmas_with [where \=\cat_smc \\, unfolded slicing_simps slicing_commute]: cat_singleton_ObjI = smc_singleton_ObjI and cat_singleton_ObjE = smc_singleton_ObjE and cat_singleton_ArrI = smc_singleton_ArrI and cat_singleton_ArrE = smc_singleton_ArrE end context category begin interpretation smc: semicategory \ \cat_smc \\ by (rule cat_semicategory) lemmas_with [unfolded slicing_simps slicing_commute]: cat_finite_psemicategory_cat_singleton = smc.smc_finite_psemicategory_smc_singleton and cat_singleton_is_arrI = smc.smc_singleton_is_arrI and cat_singleton_is_arrD = smc.smc_singleton_is_arrD and cat_singleton_is_arrE = smc.smc_singleton_is_arrE end subsubsection\Identity\ lemma cat_singleton_CId_app: assumes "set {\j, a\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {j}. \)\Obj\" shows "(\\<^sub>Ci\\<^sub>\set {j}. \)\CId\\set {\j, a\}\ = set {\j, \\CId\\a\\}" using assms unfolding cat_prod_components VLambda_vsingleton by simp subsubsection\Singleton category is a category\ lemma (in category) cat_finite_pcategory_cat_singleton: assumes "j \\<^sub>\ Vset \" shows "finite_pcategory \ (set {j}) (\i. \)" by ( auto intro: assms category_axioms finite_pcategory_finite_psemicategoryI cat_finite_psemicategory_cat_singleton ) lemma (in category) cat_category_cat_singleton: assumes "j \\<^sub>\ Vset \" shows "category \ (\\<^sub>Ci\\<^sub>\set {j}. \)" proof- interpret finite_pcategory \ \set {j}\ \\i. \\ using assms by (rule cat_finite_pcategory_cat_singleton) show ?thesis by (rule pcat_category_cat_prod) qed subsection\Singleton functor\ subsubsection\Definition and elementary properties\ definition cf_singleton :: "V \ V \ V" where "cf_singleton j \ = [ (\a\\<^sub>\\\Obj\. set {\j, a\}), (\f\\<^sub>\\\Arr\. set {\j, f\}), \, (\\<^sub>Ci\\<^sub>\set {j}. \) ]\<^sub>\" text\Components.\ lemma cf_singleton_components: shows "cf_singleton j \\ObjMap\ = (\a\\<^sub>\\\Obj\. set {\j, a\})" and "cf_singleton j \\ArrMap\ = (\f\\<^sub>\\\Arr\. set {\j, f\})" and "cf_singleton j \\HomDom\ = \" and "cf_singleton j \\HomCod\ = (\\<^sub>Ci\\<^sub>\set {j}. \)" unfolding cf_singleton_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cf_smcf_cf_singleton[slicing_commute]: "smcf_singleton j (cat_smc \)= cf_smcf (cf_singleton j \)" unfolding smcf_singleton_def cf_singleton_def slicing_simps slicing_commute by ( simp add: nat_omega_simps dghm_field_simps dg_field_simps cat_smc_def cf_smcf_def ) context fixes \ :: V begin lemmas_with [where \=\cat_smc \\, unfolded slicing_simps slicing_commute]: cf_singleton_ObjMap_vsv[cat_cs_intros] = smcf_singleton_ObjMap_vsv and cf_singleton_ObjMap_vdomain[cat_cs_simps] = smcf_singleton_ObjMap_vdomain and cf_singleton_ObjMap_vrange = smcf_singleton_ObjMap_vrange and cf_singleton_ObjMap_app[cat_prod_cs_simps] = smcf_singleton_ObjMap_app and cf_singleton_ArrMap_vsv[cat_cs_intros] = smcf_singleton_ArrMap_vsv and cf_singleton_ArrMap_vdomain[cat_cs_simps] = smcf_singleton_ArrMap_vdomain and cf_singleton_ArrMap_vrange = smcf_singleton_ArrMap_vrange and cf_singleton_ArrMap_app[cat_prod_cs_simps] = smcf_singleton_ArrMap_app end subsubsection\Singleton functor is an isomorphism of categories\ lemma (in category) cat_cf_singleton_is_functor: assumes "j \\<^sub>\ Vset \" shows "cf_singleton j \ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\set {j}. \)" proof(intro is_iso_functorI is_functorI) from assms show smcf_singleton: "cf_smcf (cf_singleton j \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_smc (\\<^sub>Ci\\<^sub>\set {j}. \)" unfolding slicing_commute[symmetric] by (intro semicategory.smc_smcf_singleton_is_iso_semifunctor) (auto intro: smc_cs_intros slicing_intros) show "vfsequence (cf_singleton j \)" unfolding cf_singleton_def by simp show "vcard (cf_singleton j \) = 4\<^sub>\" unfolding cf_singleton_def by (simp add: nat_omega_simps) show "cf_smcf (cf_singleton j \) : cat_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc (\\<^sub>Ci\\<^sub>\set {j}. \)" by (intro is_iso_semifunctor.axioms(1) smcf_singleton) show "cf_singleton j \\ArrMap\\\\CId\\c\\ = (\\<^sub>Ci\\<^sub>\set {j}. \)\CId\\cf_singleton j \\ObjMap\\c\\" if "c \\<^sub>\ \\Obj\" for c proof- from that have CId_c: "\\CId\\c\ : c \\<^bsub>\\<^esub> c" by (auto simp: cat_cs_intros) have "set {\j, c\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {j}. \)\Obj\" by (simp add: cat_singleton_ObjI that) with that have "(\\<^sub>Ci\\<^sub>\set {j}. \)\CId\\cf_singleton j \\ObjMap\\c\\ = set {\j, \\CId\\c\\}" by (simp add: cf_singleton_ObjMap_app cat_singleton_CId_app) moreover from CId_c have "cf_singleton j \\ArrMap\\\\CId\\c\\ = set {\j, \\CId\\c\\}" by (auto simp: cf_singleton_ArrMap_app cat_cs_intros) ultimately show ?thesis by simp qed qed ( auto simp: cat_cs_intros assms cat_category_cat_singleton cf_singleton_components ) subsection\Product of two categories\ subsubsection\Definition and elementary properties.\ text\See Chapter II-3 in \cite{mac_lane_categories_2010}.\ definition cat_prod_2 :: "V \ V \ V" (infixr \\\<^sub>C\ 80) where "\ \\<^sub>C \ \ cat_prod (2\<^sub>\) (\i. if i = 0 then \ else \)" text\Slicing.\ lemma cat_smc_cat_prod_2[slicing_commute]: "cat_smc \ \\<^sub>S\<^sub>M\<^sub>C cat_smc \ = cat_smc (\ \\<^sub>C \)" unfolding cat_prod_2_def smc_prod_2_def slicing_commute[symmetric] if_distrib by simp context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemmas_with [ where \=\cat_smc \\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute, OF \.cat_semicategory \.cat_semicategory ]: cat_prod_2_ObjI = smc_prod_2_ObjI and cat_prod_2_ObjI'[cat_prod_cs_intros] = smc_prod_2_ObjI' and cat_prod_2_ObjE = smc_prod_2_ObjE and cat_prod_2_ArrI = smc_prod_2_ArrI and cat_prod_2_ArrI'[cat_prod_cs_intros] = smc_prod_2_ArrI' and cat_prod_2_ArrE = smc_prod_2_ArrE and cat_prod_2_is_arrI = smc_prod_2_is_arrI and cat_prod_2_is_arrI'[cat_prod_cs_intros] = smc_prod_2_is_arrI' and cat_prod_2_is_arrE = smc_prod_2_is_arrE and cat_prod_2_Dom_vsv = smc_prod_2_Dom_vsv and cat_prod_2_Dom_vdomain[cat_cs_simps] = smc_prod_2_Dom_vdomain and cat_prod_2_Dom_app[cat_prod_cs_simps] = smc_prod_2_Dom_app and cat_prod_2_Dom_vrange = smc_prod_2_Dom_vrange and cat_prod_2_Cod_vsv = smc_prod_2_Cod_vsv and cat_prod_2_Cod_vdomain[cat_cs_simps] = smc_prod_2_Cod_vdomain and cat_prod_2_Cod_app[cat_prod_cs_simps] = smc_prod_2_Cod_app and cat_prod_2_Cod_vrange = smc_prod_2_Cod_vrange and cat_prod_2_op_cat_cat_Obj[cat_op_simps] = smc_prod_2_op_smc_smc_Obj and cat_prod_2_cat_op_cat_Obj[cat_op_simps] = smc_prod_2_smc_op_smc_Obj and cat_prod_2_op_cat_cat_Arr[cat_op_simps] = smc_prod_2_op_smc_smc_Arr and cat_prod_2_cat_op_cat_Arr[cat_op_simps] = smc_prod_2_smc_op_smc_Arr lemmas_with [ where \=\cat_smc \\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute, OF \.cat_semicategory \.cat_semicategory ]: cat_prod_2_Comp_app[cat_prod_cs_simps] = smc_prod_2_Comp_app end subsubsection\Product of two categories is a category\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \ \ by (rule categoryD[OF \]) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemma finite_pcategory_cat_prod_2: "finite_pcategory \ (2\<^sub>\) (if2 \ \)" proof(intro finite_pcategoryI pcategory_baseI) from Axiom_of_Infinity show z1_in_Vset: "2\<^sub>\ \\<^sub>\ Vset \" by blast show "category \ (i = 0 ? \ : \)" if "i \\<^sub>\ 2\<^sub>\" for i by (auto simp: cat_cs_intros) qed auto interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemma category_cat_prod_2[cat_cs_intros]: "category \ (\ \\<^sub>C \)" unfolding cat_prod_2_def by (rule pcat_category_cat_prod) end subsubsection\Identity\ lemma cat_prod_2_CId_vsv[cat_cs_intros]: "vsv ((\ \\<^sub>C \)\CId\)" unfolding cat_prod_2_def cat_prod_components by simp lemma cat_prod_2_CId_vdomain[cat_cs_simps]: "\\<^sub>\ ((\ \\<^sub>C \)\CId\) = (\ \\<^sub>C \)\Obj\" unfolding cat_prod_2_def cat_prod_components by simp context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \(\i. if i = 0 then \ else \)\ by (intro finite_pcategory_cat_prod_2 \ \) lemma cat_prod_2_CId_app[cat_prod_cs_simps]: assumes "[a, b]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "(\ \\<^sub>C \)\CId\\a, b\\<^sub>\ = [\\CId\\a\, \\CId\\b\]\<^sub>\" proof- - have "(\ \\<^sub>C \)\CId\ \a, b\\<^sub>\ = + have "(\ \\<^sub>C \)\CId\\a, b\\<^sub>\ = (\i\\<^sub>\2\<^sub>\. (if i = 0 then \ else \)\CId\\[a, b]\<^sub>\\i\\)" by ( rule cat_prod_CId_app[ OF assms[unfolded cat_prod_2_def], folded cat_prod_2_def ] ) also have "(\i\\<^sub>\2\<^sub>\. (if i = 0 then \ else \)\CId\\[a, b]\<^sub>\\i\\) = [\\CId\\a\, \\CId\\b\]\<^sub>\" proof(rule vsv_eqI, unfold vdomain_VLambda) fix i assume "i \\<^sub>\ 2\<^sub>\" then consider \i = 0\ | \i = 1\<^sub>\\ unfolding two by auto then show "(\i\\<^sub>\2\<^sub>\. (if i = 0 then \ else \)\CId\\[a, b]\<^sub>\\i\\)\i\ = [\\CId\\a\, \\CId\\b\]\<^sub>\\i\" by cases (simp_all add: two nat_omega_simps) qed (auto simp: two nat_omega_simps) finally show ?thesis by simp qed lemma cat_prod_2_CId_vrange: "\\<^sub>\ ((\ \\<^sub>C \)\CId\) \\<^sub>\ (\ \\<^sub>C \)\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_cs_simps) show "vsv ((\ \\<^sub>C \)\CId\)" by (rule cat_prod_2_CId_vsv) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF \ \]) from \ \ a b show "(\ \\<^sub>C \)\CId\\ab\ \\<^sub>\ (\ \\<^sub>C \)\Arr\" unfolding ab_def by (cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros) qed end subsubsection\Opposite product category\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemma op_smc_smc_prod_2[smc_op_simps]: "op_cat (\ \\<^sub>C \) = op_cat \ \\<^sub>C op_cat \" proof(rule cat_smc_eqI [of \]) from \ \ show cat_lhs: "category \ (op_cat (\ \\<^sub>C \))" by ( cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) interpret cat_lhs: category \ \op_cat (\ \\<^sub>C \)\ by (rule cat_lhs) from \ \ show cat_rhs: "category \ (op_cat \ \\<^sub>C op_cat \)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) interpret cat_rhs: category \ \op_cat \ \\<^sub>C op_cat \\ by (rule cat_rhs) show "op_cat (\ \\<^sub>C \)\CId\ = (op_cat \ \\<^sub>C op_cat \)\CId\" unfolding cat_op_simps proof(rule vsv_eqI, unfold cat_cs_simps) show "vsv ((\ \\<^sub>C \)\CId\)" by (rule cat_prod_2_CId_vsv) show "vsv ((op_cat \ \\<^sub>C op_cat \)\CId\)" by (rule cat_prod_2_CId_vsv) from \ \ show "(\ \\<^sub>C \)\Obj\ = (op_cat \ \\<^sub>C op_cat \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_op_intros) show "(\ \\<^sub>C \)\CId\\ab\ = (op_cat \ \\<^sub>C op_cat \)\CId\\ab\" if "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" for ab using that unfolding cat_cs_simps proof- from that obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF \ \]) from \ \ a b show "(\ \\<^sub>C \)\CId\\ab\ = (op_cat \ \\<^sub>C op_cat \)\CId\\ab\" unfolding ab_def by ( cs_concl cs_simp: cat_op_simps cat_prod_cs_simps cs_intro: cat_op_intros cat_prod_cs_intros ) qed qed from \ \ show "cat_smc (op_cat (\ \\<^sub>C \)) = cat_smc (op_cat \ \\<^sub>C op_cat \)" unfolding slicing_commute[symmetric] by (cs_concl cs_simp: smc_op_simps cs_intro: slicing_intros) qed end subsubsection\Flip\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemma cat_prod_2_Obj_fconverse[cat_cs_simps]: "((\ \\<^sub>C \)\Obj\)\\<^sub>\ = (\ \\<^sub>C \)\Obj\" proof- interpret fbrelation \((\ \\<^sub>C \)\Obj\)\ by (auto elim: cat_prod_2_ObjE[OF \ \]) show ?thesis proof(intro vsubset_antisym vsubsetI) fix ba assume prems: "ba \\<^sub>\ ((\ \\<^sub>C \)\Obj\)\\<^sub>\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" by clarsimp from prems[unfolded ba_def] have "[a, b]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" by auto then have "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" by (auto elim: cat_prod_2_ObjE[OF \ \]) with \ \ show "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" unfolding ba_def by (cs_concl cs_intro: cat_prod_cs_intros) next fix ba assume "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF \ \]) from b a show "ba \\<^sub>\ ((\ \\<^sub>C \)\Obj\)\\<^sub>\" unfolding ba_def by (auto simp: cat_prod_2_ObjI[OF \ \ a b]) qed qed lemma cat_prod_2_Arr_fconverse[cat_cs_simps]: "((\ \\<^sub>C \)\Arr\)\\<^sub>\ = (\ \\<^sub>C \)\Arr\" proof- interpret fbrelation \((\ \\<^sub>C \)\Arr\)\ by (auto elim: cat_prod_2_ArrE[OF \ \]) show ?thesis proof(intro vsubset_antisym vsubsetI) fix ba assume prems: "ba \\<^sub>\ ((\ \\<^sub>C \)\Arr\)\\<^sub>\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" by clarsimp from prems[unfolded ba_def] have "[a, b]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Arr\" by auto then have "a \\<^sub>\ \\Arr\" and "b \\<^sub>\ \\Arr\" by (auto elim: cat_prod_2_ArrE[OF \ \]) with \ \ show "ba \\<^sub>\ (\ \\<^sub>C \)\Arr\" unfolding ba_def by ( cs_concl cs_simp: cat_prod_cs_simps cs_intro: cat_prod_cs_intros cat_cs_intros ) next fix ba assume "ba \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Arr\" and a: "a \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF \ \]) from b a show "ba \\<^sub>\ ((\ \\<^sub>C \)\Arr\)\\<^sub>\" unfolding ba_def by (auto simp: cat_prod_2_ArrI[OF \ \ a b]) qed qed end subsection\Projections for the product of two categories\ subsubsection\Definition and elementary properties\ text\See Chapter II-3 in \cite{mac_lane_categories_2010}.\ definition cf_proj_fst :: "V \ V \ V" (\\\<^sub>C\<^sub>.\<^sub>1\) where "\\<^sub>C\<^sub>.\<^sub>1 \ \ = cf_proj (2\<^sub>\) (\i. if i = 0 then \ else \) 0" definition cf_proj_snd :: "V \ V \ V" (\\\<^sub>C\<^sub>.\<^sub>2\) where "\\<^sub>C\<^sub>.\<^sub>2 \ \ = cf_proj (2\<^sub>\) (\i. if i = 0 then \ else \) (1\<^sub>\)" text\Slicing\ lemma cf_smcf_cf_proj_fst[slicing_commute]: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>1 (cat_smc \) (cat_smc \) = cf_smcf (\\<^sub>C\<^sub>.\<^sub>1 \ \)" unfolding cf_proj_fst_def smcf_proj_fst_def slicing_commute[symmetric] if_distrib .. lemma cf_smcf_cf_proj_snd[slicing_commute]: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>2 (cat_smc \) (cat_smc \) = cf_smcf (\\<^sub>C\<^sub>.\<^sub>2 \ \)" unfolding cf_proj_snd_def smcf_proj_snd_def slicing_commute[symmetric] if_distrib .. context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemmas_with [ where \=\cat_smc \\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute, OF \.cat_semicategory \.cat_semicategory ]: cf_proj_fst_ObjMap_app = smcf_proj_fst_ObjMap_app and cf_proj_snd_ObjMap_app = smcf_proj_snd_ObjMap_app and cf_proj_fst_ArrMap_app = smcf_proj_fst_ArrMap_app and cf_proj_snd_ArrMap_app = smcf_proj_snd_ArrMap_app end subsubsection\ Domain and codomain of a projection of a product of two categories \ lemma cf_proj_fst_HomDom: "\\<^sub>C\<^sub>.\<^sub>1 \ \\HomDom\ = \ \\<^sub>C \" unfolding cf_proj_fst_def cf_proj_components cat_prod_2_def .. lemma cf_proj_fst_HomCod: "\\<^sub>C\<^sub>.\<^sub>1 \ \\HomCod\ = \" unfolding cf_proj_fst_def cf_proj_components cat_prod_2_def by simp lemma cf_proj_snd_HomDom: "\\<^sub>C\<^sub>.\<^sub>2 \ \\HomDom\ = \ \\<^sub>C \" unfolding cf_proj_snd_def cf_proj_components cat_prod_2_def .. lemma cf_proj_snd_HomCod: "\\<^sub>C\<^sub>.\<^sub>2 \ \\HomCod\ = \" unfolding cf_proj_snd_def cf_proj_components cat_prod_2_def by simp subsubsection\Projection of a product of two categories is a functor\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \ \ by (rule categoryD[OF \]) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemma cf_proj_fst_is_functor: assumes "i \\<^sub>\ I" shows "\\<^sub>C\<^sub>.\<^sub>1 \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( rule pcat_cf_proj_is_functor[ where i=0, simplified, folded cf_proj_fst_def cat_prod_2_def ] ) lemma cf_proj_fst_is_functor'[cat_cs_intros]: assumes "i \\<^sub>\ I" and "\ = \ \\<^sub>C \" and "\
= \" shows "\\<^sub>C\<^sub>.\<^sub>1 \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1) unfolding assms(2,3) by (rule cf_proj_fst_is_functor) lemma cf_proj_snd_is_functor: assumes "i \\<^sub>\ I" shows "\\<^sub>C\<^sub>.\<^sub>2 \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by ( rule pcat_cf_proj_is_functor[ where i=\1\<^sub>\\, simplified, folded cf_proj_snd_def cat_prod_2_def ] ) lemma cf_proj_snd_is_functor'[cat_cs_intros]: assumes "i \\<^sub>\ I" and "\ = \ \\<^sub>C \" and "\
= \" shows "\\<^sub>C\<^sub>.\<^sub>2 \ \ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1) unfolding assms(2,3) by (rule cf_proj_snd_is_functor) end subsection\Product of three categories\ subsubsection\Definition and elementary properties.\ definition cat_prod_3 :: "V \ V \ V \ V" ("(_ \\<^sub>C\<^sub>3 _ \\<^sub>C\<^sub>3 _)" [81, 81, 81] 80) where "\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \ = (\\<^sub>Ci\\<^sub>\3\<^sub>\. if3 \ \ \ i)" abbreviation cat_pow_3 :: "V \ V" (\_^\<^sub>C\<^sub>3\ [81] 80) where "\^\<^sub>C\<^sub>3 \ \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" text\Slicing.\ lemma cat_smc_cat_prod_3[slicing_commute]: "cat_smc \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>3 cat_smc \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>3 cat_smc \ = cat_smc (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)" unfolding cat_prod_3_def smc_prod_3_def slicing_commute[symmetric] if_distrib by (simp add: if_distrib[symmetric]) context fixes \ \ \ \ assumes \: "category \ \" and \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemmas_with [ where \=\cat_smc \\ and \=\cat_smc \\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute, OF \.cat_semicategory \.cat_semicategory \.cat_semicategory ]: cat_prod_3_ObjI = smc_prod_3_ObjI and cat_prod_3_ObjI'[cat_prod_cs_intros] = smc_prod_3_ObjI' and cat_prod_3_ObjE = smc_prod_3_ObjE and cat_prod_3_ArrI = smc_prod_3_ArrI and cat_prod_3_ArrI'[cat_prod_cs_intros] = smc_prod_3_ArrI' and cat_prod_3_ArrE = smc_prod_3_ArrE and cat_prod_3_is_arrI = smc_prod_3_is_arrI and cat_prod_3_is_arrI'[cat_prod_cs_intros] = smc_prod_3_is_arrI' and cat_prod_3_is_arrE = smc_prod_3_is_arrE and cat_prod_3_Dom_vsv = smc_prod_3_Dom_vsv and cat_prod_3_Dom_vdomain[cat_cs_simps] = smc_prod_3_Dom_vdomain and cat_prod_3_Dom_app[cat_prod_cs_simps] = smc_prod_3_Dom_app and cat_prod_3_Dom_vrange = smc_prod_3_Dom_vrange and cat_prod_3_Cod_vsv = smc_prod_3_Cod_vsv and cat_prod_3_Cod_vdomain[cat_cs_simps] = smc_prod_3_Cod_vdomain and cat_prod_3_Cod_app[cat_prod_cs_simps] = smc_prod_3_Cod_app and cat_prod_3_Cod_vrange = smc_prod_3_Cod_vrange lemmas_with [ where \=\cat_smc \\ and \=\cat_smc \\ and \=\cat_smc \\, unfolded slicing_simps slicing_commute, OF \.cat_semicategory \.cat_semicategory \.cat_semicategory ]: cat_prod_3_Comp_app[cat_prod_cs_simps] = smc_prod_3_Comp_app end subsubsection\Product of three categories is a category\ context fixes \ \ \ \ assumes \: "category \ \" and \: "category \ \" and \: "category \ \" begin interpretation \ \ by (rule categoryD[OF \]) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) lemma finite_pcategory_cat_prod_3: "finite_pcategory \ (3\<^sub>\) (if3 \ \ \)" proof(intro finite_pcategoryI pcategory_baseI) from Axiom_of_Infinity show z1_in_Vset: "3\<^sub>\ \\<^sub>\ Vset \" by blast show "category \ (if3 \ \ \ i)" if "i \\<^sub>\ 3\<^sub>\" for i by (auto simp: cat_cs_intros) qed auto interpretation finite_pcategory \ \3\<^sub>\\ \if3 \ \ \\ by (intro finite_pcategory_cat_prod_3 \ \ \) lemma category_cat_prod_3[cat_cs_intros]: "category \ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)" unfolding cat_prod_3_def by (rule pcat_category_cat_prod) end subsubsection\Identity\ lemma cat_prod_3_CId_vsv[cat_cs_intros]: "vsv ((\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\)" unfolding cat_prod_3_def cat_prod_components by simp lemma cat_prod_3_CId_vdomain[cat_cs_simps]: "\\<^sub>\ ((\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\) = (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" unfolding cat_prod_3_def cat_prod_components by simp context fixes \ \ \ \ assumes \: "category \ \" and \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \3\<^sub>\\ \if3 \ \ \\ by (intro finite_pcategory_cat_prod_3 \ \ \) lemma cat_prod_3_CId_app[cat_prod_cs_simps]: assumes "[a, b, c]\<^sub>\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" shows "(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\\a, b, c\\<^sub>\ = [\\CId\\a\, \\CId\\b\, \\CId\\c\]\<^sub>\" proof- have "(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\\a, b, c\\<^sub>\ = (\i\\<^sub>\3\<^sub>\. if3 \ \ \ i\CId\\[a, b, c]\<^sub>\\i\\)" by ( rule cat_prod_CId_app[ OF assms[unfolded cat_prod_3_def], folded cat_prod_3_def ] ) also have "(\i\\<^sub>\3\<^sub>\. if3 \ \ \ i\CId\\[a, b, c]\<^sub>\\i\\) = [\\CId\\a\, \\CId\\b\, \\CId\\c\]\<^sub>\" proof(rule vsv_eqI, unfold vdomain_VLambda) fix i assume "i \\<^sub>\ 3\<^sub>\" then consider \i = 0\ | \i = 1\<^sub>\\ | \i = 2\<^sub>\\ unfolding three by auto then show "(\i\\<^sub>\3\<^sub>\. (if3 \ \ \ i)\CId\\[a, b, c]\<^sub>\\i\\)\i\ = [\\CId\\a\, \\CId\\b\, \\CId\\c\]\<^sub>\\i\" by cases (simp_all add: three nat_omega_simps) qed (auto simp: three nat_omega_simps) finally show ?thesis by simp qed lemma cat_prod_3_CId_vrange: "\\<^sub>\ ((\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\) \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cat_cs_simps) show "vsv ((\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\)" by (rule cat_prod_3_CId_vsv) fix abc assume "abc \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" then obtain a b c where abc_def: "abc = [a, b, c]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" by (elim cat_prod_3_ObjE[OF \ \ \]) from \ \ \ a b c show "(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\\abc\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" unfolding abc_def by (cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros) qed end subsection\ Conversion of a product of three categories to products of two categories \ definition cf_cat_prod_21_of_3 :: "V \ V \ V \ V" where "cf_cat_prod_21_of_3 \ \ \ = [ (\A\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\. [[A\0\, A\1\<^sub>\\]\<^sub>\, A\2\<^sub>\\]\<^sub>\), (\F\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\. [[F\0\, F\1\<^sub>\\]\<^sub>\, F\2\<^sub>\\]\<^sub>\), \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \, (\ \\<^sub>C \) \\<^sub>C \ ]\<^sub>\" definition cf_cat_prod_12_of_3 :: "V \ V \ V \ V" where "cf_cat_prod_12_of_3 \ \ \ = [ (\A\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\. [A\0\, [A\1\<^sub>\\, A\2\<^sub>\\]\<^sub>\]\<^sub>\), (\F\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\. [F\0\, [F\1\<^sub>\\, F\2\<^sub>\\]\<^sub>\]\<^sub>\), \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \, \ \\<^sub>C (\ \\<^sub>C \) ]\<^sub>\" text\Components.\ lemma cf_cat_prod_21_of_3_components: shows "cf_cat_prod_21_of_3 \ \ \\ObjMap\ = (\A\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\. [[A\0\, A\1\<^sub>\\]\<^sub>\, A\2\<^sub>\\]\<^sub>\)" and "cf_cat_prod_21_of_3 \ \ \\ArrMap\ = (\F\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\. [[F\0\, F\1\<^sub>\\]\<^sub>\, F\2\<^sub>\\]\<^sub>\)" and [cat_cs_simps]: "cf_cat_prod_21_of_3 \ \ \\HomDom\ = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" and [cat_cs_simps]: "cf_cat_prod_21_of_3 \ \ \\HomCod\ = (\ \\<^sub>C \) \\<^sub>C \" unfolding cf_cat_prod_21_of_3_def dghm_field_simps by (simp_all add: nat_omega_simps) lemma cf_cat_prod_12_of_3_components: shows "cf_cat_prod_12_of_3 \ \ \\ObjMap\ = (\A\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\. [A\0\, [A\1\<^sub>\\, A\2\<^sub>\\]\<^sub>\]\<^sub>\)" and "cf_cat_prod_12_of_3 \ \ \\ArrMap\ = (\F\\<^sub>\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\. [F\0\, [F\1\<^sub>\\, F\2\<^sub>\\]\<^sub>\]\<^sub>\)" and [cat_cs_simps]: "cf_cat_prod_12_of_3 \ \ \\HomDom\ = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" and [cat_cs_simps]: "cf_cat_prod_12_of_3 \ \ \\HomCod\ = \ \\<^sub>C (\ \\<^sub>C \)" unfolding cf_cat_prod_12_of_3_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object\ mk_VLambda cf_cat_prod_21_of_3_components(1) |vsv cf_cat_prod_21_of_3_ObjMap_vsv[cat_cs_intros]| |vdomain cf_cat_prod_21_of_3_ObjMap_vdomain[cat_cs_simps]| |app cf_cat_prod_21_of_3_ObjMap_app'| mk_VLambda cf_cat_prod_12_of_3_components(1) |vsv cf_cat_prod_12_of_3_ObjMap_vsv[cat_cs_intros]| |vdomain cf_cat_prod_12_of_3_ObjMap_vdomain[cat_cs_simps]| |app cf_cat_prod_12_of_3_ObjMap_app'| lemma cf_cat_prod_21_of_3_ObjMap_app[cat_cs_simps]: assumes "A = [a, b, c]\<^sub>\" and "[a, b, c]\<^sub>\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" shows "cf_cat_prod_21_of_3 \ \ \\ObjMap\\A\ = [[a, b]\<^sub>\, c]\<^sub>\" using assms(2) unfolding assms(1) by (simp add: cf_cat_prod_21_of_3_ObjMap_app' nat_omega_simps) lemma cf_cat_prod_12_of_3_ObjMap_app[cat_cs_simps]: assumes "A = [a, b, c]\<^sub>\" and "[a, b, c]\<^sub>\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" shows "cf_cat_prod_12_of_3 \ \ \\ObjMap\\A\ = [a, [b, c]\<^sub>\]\<^sub>\" using assms(2) unfolding assms(1) by (simp add: cf_cat_prod_12_of_3_ObjMap_app' nat_omega_simps) lemma cf_cat_prod_21_of_3_ObjMap_vrange: assumes "category \ \" and "category \ \" and "category \ \" shows "\\<^sub>\ (cf_cat_prod_21_of_3 \ \ \\ObjMap\) \\<^sub>\ ((\ \\<^sub>C \) \\<^sub>C \)\Obj\" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: category \ \ by (rule assms(3)) show ?thesis proof(rule vsv.vsv_vrange_vsubset, unfold cf_cat_prod_21_of_3_ObjMap_vdomain) fix A assume prems: "A \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" then show "cf_cat_prod_21_of_3 \ \ \\ObjMap\\A\ \\<^sub>\ ((\ \\<^sub>C \) \\<^sub>C \)\Obj\" by (elim cat_prod_3_ObjE[OF assms], insert prems, simp only:) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) qed lemma cf_cat_prod_12_of_3_ObjMap_vrange: assumes "category \ \" and "category \ \" and "category \ \" shows "\\<^sub>\ (cf_cat_prod_12_of_3 \ \ \\ObjMap\) \\<^sub>\ (\ \\<^sub>C (\ \\<^sub>C \))\Obj\" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: category \ \ by (rule assms(3)) show ?thesis proof(rule vsv.vsv_vrange_vsubset, unfold cf_cat_prod_12_of_3_ObjMap_vdomain) fix A assume prems: "A \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" then show "cf_cat_prod_12_of_3 \ \ \\ObjMap\\A\ \\<^sub>\ (\ \\<^sub>C (\ \\<^sub>C \))\Obj\" by (elim cat_prod_3_ObjE[OF assms], insert prems, simp only:) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) qed subsubsection\Arrow\ mk_VLambda cf_cat_prod_21_of_3_components(2) |vsv cf_cat_prod_21_of_3_ArrMap_vsv[cat_cs_intros]| |vdomain cf_cat_prod_21_of_3_ArrMap_vdomain[cat_cs_simps]| |app cf_cat_prod_21_of_3_ArrMap_app'| mk_VLambda cf_cat_prod_12_of_3_components(2) |vsv cf_cat_prod_12_of_3_ArrMap_vsv[cat_cs_intros]| |vdomain cf_cat_prod_12_of_3_ArrMap_vdomain[cat_cs_simps]| |app cf_cat_prod_12_of_3_ArrMap_app'| lemma cf_cat_prod_21_of_3_ArrMap_app[cat_cs_simps]: assumes "F = [h, g, f]\<^sub>\" and "[h, g, f]\<^sub>\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" shows "cf_cat_prod_21_of_3 \ \ \\ArrMap\\F\ = [[h, g]\<^sub>\, f]\<^sub>\" using assms(2) unfolding assms(1) by (simp add: cf_cat_prod_21_of_3_ArrMap_app' nat_omega_simps) lemma cf_cat_prod_12_of_3_ArrMap_app[cat_cs_simps]: assumes "F = [h, g, f]\<^sub>\" and "[h, g, f]\<^sub>\ \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" shows "cf_cat_prod_12_of_3 \ \ \\ArrMap\\F\ = [h, [g, f]\<^sub>\]\<^sub>\" using assms(2) unfolding assms(1) by (simp add: cf_cat_prod_12_of_3_ArrMap_app' nat_omega_simps) subsubsection\ Conversion of a product of three categories to products of two categories is a functor \ lemma cf_cat_prod_21_of_3_is_functor: assumes "category \ \" and "category \ \" and "category \ \" shows "cf_cat_prod_21_of_3 \ \ \ : \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \ \\\<^sub>C\<^bsub>\\<^esub> (\ \\<^sub>C \) \\<^sub>C \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: category \ \ by (rule assms(3)) show ?thesis proof(rule is_functorI') show "vfsequence (cf_cat_prod_21_of_3 \ \ \)" unfolding cf_cat_prod_21_of_3_def by auto show "vcard (cf_cat_prod_21_of_3 \ \ \) = 4\<^sub>\" unfolding cf_cat_prod_21_of_3_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_cat_prod_21_of_3 \ \ \\ObjMap\) \\<^sub>\ ((\ \\<^sub>C \) \\<^sub>C \)\Obj\" by (rule cf_cat_prod_21_of_3_ObjMap_vrange[OF assms]) show "cf_cat_prod_21_of_3 \ \ \\ArrMap\\F\ : cf_cat_prod_21_of_3 \ \ \\ObjMap\\A\ \\<^bsub>(\ \\<^sub>C \) \\<^sub>C \\<^esub> cf_cat_prod_21_of_3 \ \ \\ObjMap\\B\" if "F : A \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> B" for A B F using that by (elim cat_prod_3_is_arrE[OF assms], insert that, simp only:) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) show "cf_cat_prod_21_of_3 \ \ \\ArrMap\\G \\<^sub>A\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> F\ = cf_cat_prod_21_of_3 \ \ \\ArrMap\\G\ \\<^sub>A\<^bsub>(\ \\<^sub>C \) \\<^sub>C \\<^esub> cf_cat_prod_21_of_3 \ \ \\ArrMap\\F\" if "G : B \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> C" and "F : A \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> B" for B C G A F proof- from that(2) obtain f f' f'' a a' a'' b b' b'' where F_def: "F = [f, f', f'']\<^sub>\" and A_def: "A = [a, a', a'']\<^sub>\" and B_def: "B = [b, b', b'']\<^sub>\" and f: "f : a \\<^bsub>\\<^esub> b" and f': "f' : a' \\<^bsub>\\<^esub> b'" and f'': "f'' : a'' \\<^bsub>\\<^esub> b''" by (elim cat_prod_3_is_arrE[OF assms]) with that(1) obtain g g' g'' c c' c'' where G_def: "G = [g, g', g'']\<^sub>\" and C_def: "C = [c, c', c'']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> c" and g': "g' : b' \\<^bsub>\\<^esub> c'" and g'': "g'' : b'' \\<^bsub>\\<^esub> c''" by (auto elim: cat_prod_3_is_arrE[OF assms]) from that f f' f'' g g' g'' show ?thesis unfolding F_def A_def B_def G_def C_def by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_cat_prod_21_of_3 \ \ \\ArrMap\\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\\C\\ = ((\ \\<^sub>C \) \\<^sub>C \)\CId\\cf_cat_prod_21_of_3 \ \ \\ObjMap\\C\\" if "C \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" for C using that by (elim cat_prod_3_ObjE[OF assms], insert that, simp only: ) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ qed lemma cf_cat_prod_21_of_3_is_functor'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "category \ \" and "\' = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" and "\' = (\ \\<^sub>C \) \\<^sub>C \" shows "cf_cat_prod_21_of_3 \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1-3) unfolding assms(4,5) by (rule cf_cat_prod_21_of_3_is_functor) lemma cf_cat_prod_12_of_3_is_functor: assumes "category \ \" and "category \ \" and "category \ \" shows "cf_cat_prod_12_of_3 \ \ \ : \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \ \\\<^sub>C\<^bsub>\\<^esub> \ \\<^sub>C (\ \\<^sub>C \)" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: category \ \ by (rule assms(3)) show ?thesis proof(rule is_functorI') show "vfsequence (cf_cat_prod_12_of_3 \ \ \)" unfolding cf_cat_prod_12_of_3_def by auto show "vcard (cf_cat_prod_12_of_3 \ \ \) = 4\<^sub>\" unfolding cf_cat_prod_12_of_3_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_cat_prod_12_of_3 \ \ \\ObjMap\) \\<^sub>\ (\ \\<^sub>C (\ \\<^sub>C \))\Obj\" by (rule cf_cat_prod_12_of_3_ObjMap_vrange[OF assms]) show "cf_cat_prod_12_of_3 \ \ \\ArrMap\\F\ : cf_cat_prod_12_of_3 \ \ \\ObjMap\\A\ \\<^bsub>\ \\<^sub>C (\ \\<^sub>C \)\<^esub> cf_cat_prod_12_of_3 \ \ \\ObjMap\\B\" if "F : A \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> B" for A B F using that by (elim cat_prod_3_is_arrE[OF assms], insert that, simp only:) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) show "cf_cat_prod_12_of_3 \ \ \\ArrMap\\G \\<^sub>A\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> F\ = cf_cat_prod_12_of_3 \ \ \\ArrMap\\G\ \\<^sub>A\<^bsub>\ \\<^sub>C (\ \\<^sub>C \)\<^esub> cf_cat_prod_12_of_3 \ \ \\ArrMap\\F\" if "G : B \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> C" and "F : A \\<^bsub>\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\<^esub> B" for B C G A F proof- from that(2) obtain f f' f'' a a' a'' b b' b'' where F_def: "F = [f, f', f'']\<^sub>\" and A_def: "A = [a, a', a'']\<^sub>\" and B_def: "B = [b, b', b'']\<^sub>\" and f: "f : a \\<^bsub>\\<^esub> b" and f': "f' : a' \\<^bsub>\\<^esub> b'" and f'': "f'' : a'' \\<^bsub>\\<^esub> b''" by (elim cat_prod_3_is_arrE[OF assms]) with that(1) obtain g g' g'' c c' c'' where G_def: "G = [g, g', g'']\<^sub>\" and C_def: "C = [c, c', c'']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> c" and g': "g' : b' \\<^bsub>\\<^esub> c'" and g'': "g'' : b'' \\<^bsub>\\<^esub> c''" by (auto elim: cat_prod_3_is_arrE[OF assms]) from that f f' f'' g g' g'' show ?thesis unfolding F_def A_def B_def G_def C_def by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_cat_prod_12_of_3 \ \ \\ArrMap\\(\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\CId\\C\\ = (\ \\<^sub>C (\ \\<^sub>C \))\CId\\cf_cat_prod_12_of_3 \ \ \\ObjMap\\C\\" if "C \\<^sub>\ (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" for C using that by (elim cat_prod_3_ObjE[OF assms], insert that, simp only: ) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ qed lemma cf_cat_prod_12_of_3_is_functor'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "category \ \" and "\' = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" and "\' = \ \\<^sub>C (\ \\<^sub>C \)" shows "cf_cat_prod_12_of_3 \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1-3) unfolding assms(4,5) by (rule cf_cat_prod_12_of_3_is_functor) subsection\Bifunctors\ text\ A bifunctor is defined as a functor from a product of two categories to a category (see Chapter II-3 in \cite{mac_lane_categories_2010}). This subsection exposes the elementary properties of the projections of the bifunctors established by fixing an argument in a functor (see Chapter II-3 in \cite{mac_lane_categories_2010} for further information). \ subsubsection\Definitions and elementary properties\ definition bifunctor_proj_fst :: "V \ V \ V \ V \ V" (\(_\<^bsub>_,_\<^esub>/'(/-,_/')/\<^sub>C\<^sub>F)\ [51, 51, 51, 51] 51) where "\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F = (\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\ -\<^sub>\ set {1\<^sub>\}. (i = 0 ? \ : \),\\HomCod\\<^esub>(-,set {\1\<^sub>\, b\})) \\<^sub>C\<^sub>F cf_singleton 0 \" definition bifunctor_proj_snd :: "V \ V \ V \ V \ V" (\(_\<^bsub>_,_\<^esub>/'(/_,-/')/\<^sub>C\<^sub>F)\ [51, 51, 51, 51] 51) where "\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F = (\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\ -\<^sub>\ set {0}. (i = 0 ? \ : \),\\HomCod\\<^esub>(-,set {\0, a\})) \\<^sub>C\<^sub>F cf_singleton (1\<^sub>\) \" abbreviation bcf_ObjMap_app :: "V \ V \ V \ V" (infixl "\\<^sub>H\<^sub>M\<^sub>.\<^sub>O\" 55) where "a \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\<^esub> b \ \\ObjMap\\a, b\\<^sub>\" abbreviation bcf_ArrMap_app :: "V \ V \ V \ V" (infixl "\\<^sub>H\<^sub>M\<^sub>.\<^sub>A\" 55) where "g \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\<^esub> f \ \\ArrMap\\g, f\\<^sub>\" text\Elementary properties.\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemma cat_singleton_qm_fst_def[simp]: "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \)) = (\\<^sub>Ci\\<^sub>\set {0}. \)" proof(rule cat_eqI[of \]) show "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Obj\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Obj\" unfolding cat_prod_components by (subst vproduct_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Arr\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Arr\" unfolding cat_prod_components by (subst vproduct_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Dom\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Dom\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Cod\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Cod\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp have [simp]: "f : a \\<^bsub>\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \)\<^esub> b \ f : a \\<^bsub>\\<^sub>Ci\\<^sub>\set {0}. \\<^esub> b" for f a b unfolding is_arr_def by simp show "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Comp\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Comp\" proof(rule vsv_eqI) show "vsv ((\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Comp\)" unfolding cat_prod_components by simp show "vsv ((\\<^sub>Ci\\<^sub>\set {0}. \)\Comp\)" unfolding cat_prod_components by simp show "\\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Comp\) = \\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {0}. \)\Comp\)" by (simp add: composable_arrs_def cat_cs_simps) show "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Comp\\gf\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\Comp\\gf\" if "gf \\<^sub>\ \\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\Comp\)" for gf proof- from that have "gf \\<^sub>\ composable_arrs (\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))" by (simp add: cat_cs_simps) then obtain g f a b c where gf_def: "gf = [g, f]\<^sub>\" and g: "g : b \\<^bsub>(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\<^esub> c" and f: "f : a \\<^bsub>(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\<^esub> b" by clarsimp then have g': "g : b \\<^bsub>(\\<^sub>Ci\\<^sub>\set {0}. \)\<^esub> c" and f': "f : a \\<^bsub>(\\<^sub>Ci\\<^sub>\set {0}. \)\<^esub> b" by simp_all show ?thesis unfolding gf_def unfolding cat_prod_Comp_app[OF g f] cat_prod_Comp_app[OF g' f'] by (subst (1 2) VLambda_vsingleton_def) simp qed qed show "(\\<^sub>Ci\\<^sub>\set {0}. (i = 0 ? \ : \))\CId\ = (\\<^sub>Ci\\<^sub>\set {0}. \)\CId\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp qed ( simp_all add: \.cat_category_cat_singleton pcategory.pcat_category_cat_prod pcat_vsubset_index_pcategory vsubset_vsingleton_leftI ) lemma cat_singleton_qm_snd_def[simp]: "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \)) = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)" proof(rule cat_eqI[of \]) show "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Obj\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Obj\" unfolding cat_prod_components by (subst vproduct_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Arr\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Arr\" unfolding cat_prod_components by (subst vproduct_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Dom\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Dom\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp show [simp]: "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Cod\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Cod\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp have [simp]: "f : a \\<^bsub>\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \)\<^esub> b \ f : a \\<^bsub>\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \\<^esub> b" for f a b unfolding is_arr_def by simp show "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Comp\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Comp\" proof(rule vsv_eqI) show "vsv ((\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Comp\)" unfolding cat_prod_components by simp show "vsv ((\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Comp\)" unfolding cat_prod_components by simp show "\\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Comp\) = \\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Comp\)" by (simp add: composable_arrs_def cat_cs_simps) show "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Comp\\gf\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Comp\\gf\" if "gf \\<^sub>\ \\<^sub>\ ((\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\Comp\)" for gf proof- from that have "gf \\<^sub>\ composable_arrs (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))" by (simp add: cat_cs_simps) then obtain g f a b c where gf_def: "gf = [g, f]\<^sub>\" and g: "g : b \\<^bsub>(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\<^esub> c" and f: "f : a \\<^bsub>(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\<^esub> b" by clarsimp then have g': "g : b \\<^bsub>(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\<^esub> c" and f': "f : a \\<^bsub>(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\<^esub> b" by simp_all show ?thesis unfolding gf_def unfolding cat_prod_Comp_app[OF g f] cat_prod_Comp_app[OF g' f'] by (subst (1 2) VLambda_vsingleton_def) simp qed qed show "(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. (i = 0 ? \ : \))\CId\ = (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\CId\" unfolding cat_prod_components by (subst vproduct_vsingleton_def, subst (1 2) VLambda_vsingleton_def) simp qed ( simp_all add: \.cat_category_cat_singleton pcategory.pcat_category_cat_prod pcat_vsubset_index_pcategory vsubset_vsingleton_leftI ) end subsubsection\Object map\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemmas_with [OF \.category_axioms \.category_axioms, simp]: cat_singleton_qm_fst_def and cat_singleton_qm_snd_def lemma bifunctor_proj_fst_ObjMap_app[cat_cs_simps]: assumes "[a, b]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\ = \\ObjMap\\a, b\\<^sub>\" proof- let ?\
= \\\HomCod\\ let ?\ = \\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {1\<^sub>\}.(i = 0 ? \ : \),?\
\<^esub>(-,set {\1\<^sub>\, b\})\ let ?cfs = \cf_singleton 0 \\ from assms have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (all\elim cat_prod_2_ObjE[OF \ \]\) auto from a have za: "set {\0, a\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {0}. \)\Obj\" by (intro cat_singleton_ObjI[where a=a]) simp have [simp]: "vinsert \0, a\ (set {\1\<^sub>\, b\}) = [a, b]\<^sub>\" using ord_of_nat_succ_vempty unfolding vcons_def by auto have "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\ = (?\\ObjMap\ \\<^sub>\ ?cfs\ObjMap\)\a\" unfolding bifunctor_proj_fst_def dghm_comp_components by simp also have "\ = ?\\ObjMap\\?cfs\ObjMap\\a\\" by (rule vsv_vcomp_at) ( simp_all add: two a za cf_singleton_components prodfunctor_proj_components cf_singleton_ObjMap_app ) also from za have "\ = \\ObjMap\\a, b\\<^sub>\" unfolding two cf_singleton_ObjMap_app[OF a] prodfunctor_proj_components by simp finally show ?thesis by simp qed lemma bifunctor_proj_snd_ObjMap_app[cat_cs_simps]: assumes "[a, b]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\ = \\ObjMap\\a, b\\<^sub>\" proof- let ?\
= \\\HomCod\\ let ?\ = \\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {0}.(i = 0 ? \ : \),?\
\<^esub>(-,set {\0, a\})\ let ?cfs = \cf_singleton (1\<^sub>\) \\ from assms have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (all\elim cat_prod_2_ObjE[OF \ \]\) auto from a have za: "set {\0, a\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {0}. \)\Obj\" by (intro cat_singleton_ObjI[where a=a]) simp from b have ob: "set {\1\<^sub>\, b\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Obj\" by (intro cat_singleton_ObjI[where a=b]) simp have[simp]: "vinsert \1\<^sub>\, b\ (set {\0, a\}) = [a, b]\<^sub>\" using ord_of_nat_succ_vempty unfolding vcons_def by auto have "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\ = (?\\ObjMap\ \\<^sub>\ ?cfs\ObjMap\)\b\" unfolding bifunctor_proj_snd_def dghm_comp_components by simp also have "\ = ?\\ObjMap\\?cfs\ObjMap\\b\\" by (rule vsv_vcomp_at) ( simp_all add: two cf_singleton_components prodfunctor_proj_components cf_singleton_ObjMap_app ob b ) also from ob have "\ = \\ObjMap\\a, b\\<^sub>\" unfolding two cf_singleton_ObjMap_app[OF b] prodfunctor_proj_components by simp finally show ?thesis by simp qed end subsubsection\Arrow map\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemmas_with [OF \.category_axioms \.category_axioms, simp]: cat_singleton_qm_fst_def and cat_singleton_qm_snd_def lemma bifunctor_proj_fst_ArrMap_app[cat_cs_simps]: assumes "b \\<^sub>\ \\Obj\" and "f \\<^sub>\ \\Arr\" shows "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\\f\ = \\ArrMap\\f, \\CId\\b\\\<^sub>\" proof- let ?\
= \\\HomCod\\ let ?\ = \\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {1\<^sub>\}.(i = 0 ? \ : \),?\
\<^esub>(-,set {\1\<^sub>\, b\})\ let ?cfs = \cf_singleton 0 \\ from assms(1) have "\\CId\\b\ : b \\<^bsub>\\<^esub> b" by (auto intro: cat_cs_intros) then have CId_b: "\\CId\\b\ \\<^sub>\ \\Arr\" by auto from assms(2) have zf: "set {\0, f\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {0}. \)\Arr\" by (intro cat_singleton_ArrI[where a=f]) simp from assms(1) have ob: "set {\1\<^sub>\, b\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Obj\" by (intro cat_singleton_ObjI[where a=b]) simp have [simp]: "vinsert \0, f\ (set {\1\<^sub>\, \\CId\\b\\}) = [f, \\CId\\b\]\<^sub>\" using ord_of_nat_succ_vempty unfolding vcons_def by auto have "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\\f\ = (?\\ArrMap\ \\<^sub>\ ?cfs\ArrMap\)\f\" unfolding bifunctor_proj_fst_def dghm_comp_components by simp also have "\ = ?\\ArrMap\\?cfs\ArrMap\\f\\" by (rule vsv_vcomp_at) ( simp_all add: two assms(2) cf_singleton_components prodfunctor_proj_components cf_singleton_ArrMap_app zf ) also from assms(1) zf have "\ = \\ArrMap\\f, \\CId\\b\\\<^sub>\" unfolding cf_singleton_ArrMap_app[OF assms(2)] prodfunctor_proj_components by (simp add: two cat_singleton_CId_app[OF ob]) finally show ?thesis by simp qed lemma bifunctor_proj_snd_ArrMap_app[cat_cs_simps]: assumes "a \\<^sub>\ \\Obj\" and "g \\<^sub>\ \\Arr\" shows "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\\g\ = \\ArrMap\\\\CId\\a\, g\\<^sub>\" proof- let ?\
= \\\HomCod\\ let ?\ = \\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {0}.(i = 0 ? \ : \),?\
\<^esub>(-,set {\0, a\})\ let ?cfs = \cf_singleton (1\<^sub>\) \\ from assms(1) have "\\CId\\a\ : a \\<^bsub>\\<^esub> a" by (auto intro: cat_cs_intros) then have CId_a: "\\CId\\a\ \\<^sub>\ \\Arr\" by auto from assms(2) have og: "set {\1\<^sub>\, g\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\Arr\" by (intro cat_singleton_ArrI[where a=g]) simp from assms(1) have ob: "set {\0, a\} \\<^sub>\ (\\<^sub>Ci\\<^sub>\set {0}. \)\Obj\" by (intro cat_singleton_ObjI[where a=a]) simp have [simp]: "vinsert \1\<^sub>\, g\ (set {\0, \\CId\\a\\}) = [\\CId\\a\, g]\<^sub>\" using ord_of_nat_succ_vempty unfolding vcons_def by auto have "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\\g\ = (?\\ArrMap\ \\<^sub>\ ?cfs\ArrMap\)\g\" unfolding two bifunctor_proj_snd_def dghm_comp_components by simp also have "\ = ?\\ArrMap\\?cfs\ArrMap\\g\\" by (rule vsv_vcomp_at) ( simp_all add: two assms(2) cf_singleton_components prodfunctor_proj_components cf_singleton_ArrMap_app og ) also from assms(1) og have "\ = \\ArrMap\\\\CId\\a\, g\\<^sub>\" unfolding cf_singleton_ArrMap_app[OF assms(2)] prodfunctor_proj_components by (simp add: two cat_singleton_CId_app[OF ob]) finally show ?thesis by simp qed end subsubsection\Bifunctor projections are functors\ context fixes \ \ \ assumes \: "category \ \" and \: "category \ \" begin interpretation \: category \ \ by (rule \) interpretation \: category \ \ by (rule \) interpretation finite_pcategory \ \2\<^sub>\\ \if2 \ \\ by (intro finite_pcategory_cat_prod_2 \ \) lemmas_with [OF \.category_axioms \.category_axioms, simp]: cat_singleton_qm_fst_def and cat_singleton_qm_snd_def lemma bifunctor_proj_fst_is_functor: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \
\ by (rule assms(1)) show ?thesis unfolding bifunctor_proj_fst_def proof ( intro cf_comp_is_functorI[where \=\(\\<^sub>Ci\\<^sub>\set {0}. \)\], unfold \.cf_HomCod ) from assms(2) have zb: "set {\1\<^sub>\, b\} \\<^sub>\ (\\<^sub>Cj\\<^sub>\set {1\<^sub>\}. if j = 0 then \ else \)\Obj\" unfolding cat_prod_components by (intro vproduct_vsingletonI) simp_all have o_zo: "set {1\<^sub>\} \\<^sub>\ 2\<^sub>\" by clarsimp from pcat_prodfunctor_proj_is_functor[ folded cat_prod_2_def, where J=\set {1\<^sub>\}\, OF assms(1) zb o_zo ] show "\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {1\<^sub>\}.(i = 0 ? \ : \),\
\<^esub>(-,set {\1\<^sub>\, b\}) : (\\<^sub>Ci\\<^sub>\set {0}. \) \\\<^sub>C\<^bsub>\\<^esub> \
" unfolding two by simp from category.cat_cf_singleton_is_functor[OF \.category_axioms, of 0] show "cf_singleton 0 \ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\set {0}. \)" by force qed qed lemma bifunctor_proj_fst_is_functor'[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" and "\' = \" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1,2) unfolding assms(3) by (rule bifunctor_proj_fst_is_functor) lemma bifunctor_proj_fst_ObjMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "vsv ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\)" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_fst_is_functor[OF assms]) show ?thesis by (rule \.cf_ObjMap_vsv) qed lemma bifunctor_proj_fst_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\) = \\Obj\" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_fst_is_functor[OF assms]) show ?thesis by (rule \.cf_ObjMap_vdomain) qed lemma bifunctor_proj_fst_ArrMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "vsv ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\)" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_fst_is_functor[OF assms]) show ?thesis by (rule \.cf_ArrMap_vsv) qed lemma bifunctor_proj_fst_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\) = \\Arr\" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_fst_is_functor[OF assms]) show ?thesis by (rule \.cf_ArrMap_vdomain) qed lemma bifunctor_proj_snd_is_functor: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \
\ by (rule assms(1)) show ?thesis unfolding bifunctor_proj_snd_def proof ( intro cf_comp_is_functorI[where \=\(\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)\], unfold \.cf_HomCod ) from assms(2) have zb: "set {\0, a\} \\<^sub>\ (\\<^sub>Cj\\<^sub>\set {0}. if j = 0 then \ else \)\Obj\" unfolding cat_prod_components by (intro vproduct_vsingletonI) simp_all have o_zo: "set {0} \\<^sub>\ 2\<^sub>\" by clarsimp from pcat_prodfunctor_proj_is_functor[ folded cat_prod_2_def, where J=\set {0}\, OF assms(1) zb o_zo ] show "\\<^bsub>\\<^sub>Ci\\<^sub>\2\<^sub>\-\<^sub>\set {0}.(i = 0 ? \ : \),\
\<^esub>(-,set {\0, a\}) : (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \) \\\<^sub>C\<^bsub>\\<^esub> \
" unfolding two by simp from category.cat_cf_singleton_is_functor[OF \.category_axioms, of \1\<^sub>\\] show "cf_singleton (1\<^sub>\) \ : \ \\\<^sub>C\<^bsub>\\<^esub> (\\<^sub>Ci\\<^sub>\set {1\<^sub>\}. \)" by force qed qed lemma bifunctor_proj_snd_is_functor'[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" and "\' = \" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1,2) unfolding assms(3) by (rule bifunctor_proj_snd_is_functor) lemma bifunctor_proj_snd_ObjMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" shows "vsv ((\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\)" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_snd_is_functor[OF assms]) show ?thesis by (rule \.cf_ObjMap_vsv) qed lemma bifunctor_proj_snd_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\) = \\Obj\" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_snd_is_functor[OF assms]) show ?thesis by (rule \.cf_ObjMap_vdomain) qed lemma bifunctor_proj_snd_ArrMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" shows "vsv ((\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\)" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_snd_is_functor[OF assms]) show ?thesis by (rule \.cf_ArrMap_vsv) qed lemma bifunctor_proj_snd_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "a \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\) = \\Arr\" proof- interpret \: is_functor \ \ \
\\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ by (rule bifunctor_proj_snd_is_functor[OF assms]) show ?thesis by (rule \.cf_ArrMap_vdomain) qed end subsection\Bifunctor flip\ subsubsection\Definition and elementary properties\ definition bifunctor_flip :: "V \ V \ V \ V" where "bifunctor_flip \ \ \ = [fflip (\\ObjMap\), fflip (\\ArrMap\), \ \\<^sub>C \, \\HomCod\]\<^sub>\" text\Components\ lemma bifunctor_flip_components: shows "bifunctor_flip \ \ \\ObjMap\ = fflip (\\ObjMap\)" and "bifunctor_flip \ \ \\ArrMap\ = fflip (\\ArrMap\)" and "bifunctor_flip \ \ \\HomDom\ = \ \\<^sub>C \" and "bifunctor_flip \ \ \\HomCod\ = \\HomCod\" unfolding bifunctor_flip_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Bifunctor flip object map\ lemma bifunctor_flip_ObjMap_vsv[cat_cs_intros]: "vsv (bifunctor_flip \ \ \\ObjMap\)" unfolding bifunctor_flip_components by (rule fflip_vsv) lemma bifunctor_flip_ObjMap_app: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "bifunctor_flip \ \ \\ObjMap\\b, a\\<^sub>\ = \\ObjMap\\a, b\\<^sub>\" using assms unfolding bifunctor_flip_components assms(4,5) by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: cat_prod_cs_intros) lemma bifunctor_flip_ObjMap_app'[cat_cs_simps]: assumes "ba = [b, a]\<^sub>\" and "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "bifunctor_flip \ \ \\ObjMap\\ba\ = \\ObjMap\\a, b\\<^sub>\" using assms(2-6) unfolding assms(1) by (rule bifunctor_flip_ObjMap_app) lemma bifunctor_flip_ObjMap_vdomain[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" using assms unfolding bifunctor_flip_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps) lemma bifunctor_flip_ObjMap_vrange[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) = \\<^sub>\ (\\ObjMap\)" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms(3)) show ?thesis proof(intro vsubset_antisym) show "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) \\<^sub>\ \\<^sub>\ (\\ObjMap\)" proof ( intro vsv.vsv_vrange_vsubset, unfold bifunctor_flip_ObjMap_vdomain[OF assms] ) fix ba assume "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(2,1)]) from assms a b show "bifunctor_flip \ \ \\ObjMap\\ba\ \\<^sub>\ \\<^sub>\ (\\ObjMap\)" unfolding ba_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_prod_cs_intros ) qed (auto intro: cat_cs_intros) show "\\<^sub>\ (\\ObjMap\) \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ObjMap\)" proof(intro vsv.vsv_vrange_vsubset, unfold \.cf_ObjMap_vdomain) fix ab assume prems: "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) from assms a b have ba: "[b, a]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" by (cs_concl cs_intro: cat_prod_cs_intros) from assms bifunctor_flip_ObjMap_vsv prems a b ba show "\\ObjMap\\ab\ \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ObjMap\)" by (cs_concl cs_simp: ab_def cat_cs_simps cs_intro: V_cs_intros) qed auto qed qed subsubsection\Bifunctor flip arrow map\ lemma bifunctor_flip_ArrMap_vsv[cat_cs_intros]: "vsv (bifunctor_flip \ \ \\ArrMap\)" unfolding bifunctor_flip_components by (rule fflip_vsv) lemma bifunctor_flip_ArrMap_app: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g \\<^sub>\ \\Arr\" and "f \\<^sub>\ \\Arr\" shows "bifunctor_flip \ \ \\ArrMap\\f, g\\<^sub>\ = \\ArrMap\\g, f\\<^sub>\" using assms unfolding bifunctor_flip_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: cat_prod_cs_intros) lemma bifunctor_flip_ArrMap_app'[cat_cs_simps]: assumes "fg = [f, g]\<^sub>\" and "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g \\<^sub>\ \\Arr\" and "f \\<^sub>\ \\Arr\" shows "bifunctor_flip \ \ \\ArrMap\\fg\ = \\ArrMap\\g, f\\<^sub>\" using assms(2-6) unfolding assms(1) by (rule bifunctor_flip_ArrMap_app) lemma bifunctor_flip_ArrMap_vdomain[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" using assms unfolding bifunctor_flip_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps) lemma bifunctor_flip_ArrMap_vrange[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) = \\<^sub>\ (\\ArrMap\)" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms(3)) show ?thesis proof(intro vsubset_antisym) show "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) \\<^sub>\ \\<^sub>\ (\\ArrMap\)" proof ( intro vsv.vsv_vrange_vsubset, unfold bifunctor_flip_ArrMap_vdomain[OF assms] ) fix fg assume "fg \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain f g where fg_def: "fg = [f, g]\<^sub>\" and f: "f \\<^sub>\ \\Arr\" and g: "g \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF assms(2,1)]) from f obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by (auto intro: is_arrI) from g obtain a' b' where g: "g : a' \\<^bsub>\\<^esub> b'" by (auto intro: is_arrI) from \.cf_ArrMap_vsv assms f g show "bifunctor_flip \ \ \\ArrMap\\fg\ \\<^sub>\ \\<^sub>\ (\\ArrMap\)" unfolding fg_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_prod_cs_intros ) qed (auto intro: cat_cs_intros) show "\\<^sub>\ (\\ArrMap\) \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ArrMap\)" proof(intro vsv.vsv_vrange_vsubset, unfold \.cf_ArrMap_vdomain) fix gf assume prems: "gf \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain g f where gf_def: "gf = [g, f]\<^sub>\" and g: "g \\<^sub>\ \\Arr\" and f: "f \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF assms(1,2)]) from assms g f have fg: "[f, g]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Arr\" by (cs_concl cs_intro: cat_prod_cs_intros) from assms bifunctor_flip_ArrMap_vsv prems g f fg show "\\ArrMap\\gf\ \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ArrMap\)" unfolding gf_def by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) qed auto qed qed subsubsection\Bifunctor flip is a bifunctor\ lemma bifunctor_flip_is_functor: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bifunctor_flip \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \ " proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) show ?thesis proof(intro is_functorI') show "vfsequence (bifunctor_flip \ \ \)" unfolding bifunctor_flip_def by simp from assms(1,2) show "category \ (\ \\<^sub>C \)" by (cs_concl cs_intro: cat_cs_intros) show "vcard (bifunctor_flip \ \ \) = 4\<^sub>\" unfolding bifunctor_flip_def by (simp add: nat_omega_simps) show "vsv (bifunctor_flip \ \ \\ObjMap\)" by (auto intro: cat_cs_intros) show "vsv (bifunctor_flip \ \ \\ArrMap\)" by (auto intro: cat_cs_intros) from assms show "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms \.cf_ObjMap_vrange show "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) \\<^sub>\ \\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms show "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" by (cs_concl cs_simp: cat_cs_simps) show "bifunctor_flip \ \ \\ArrMap\\gf\ : bifunctor_flip \ \ \\ObjMap\\ba\ \\<^bsub>\\<^esub> bifunctor_flip \ \ \\ObjMap\\b'a'\" if "gf : ba \\<^bsub>\ \\<^sub>C \\<^esub> b'a'" for ba b'a' gf proof- from that obtain g f a b a' b' where gf_def: "gf = [g, f]\<^sub>\" and ba_def: "ba = [b, a]\<^sub>\" and b'a'_def: "b'a' = [b', a']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> b'" and f: "f : a \\<^bsub>\\<^esub> a'" by (elim cat_prod_2_is_arrE[OF assms(2,1)]) from assms g f show ?thesis unfolding gf_def ba_def b'a'_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "bifunctor_flip \ \ \\ArrMap\\gg' \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> ff'\ = bifunctor_flip \ \ \\ArrMap\\gg'\ \\<^sub>A\<^bsub>\\<^esub> bifunctor_flip \ \ \\ArrMap\\ff'\" if gg': "gg' : bb' \\<^bsub>\ \\<^sub>C \\<^esub> cc'" and ff': "ff' : aa' \\<^bsub>\ \\<^sub>C \\<^esub> bb'" for bb' cc' gg' aa' ff' proof- obtain g g' b b' c c' where gg'_def: "gg' = [g, g']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and cc'_def: "cc' = [c, c']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> c" and g': "g' : b' \\<^bsub>\\<^esub> c'" by (elim cat_prod_2_is_arrE[OF assms(2,1) gg']) moreover obtain f f' a a' b'' b''' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and "bb' = [b'', b''']\<^sub>\" and "f : a \\<^bsub>\\<^esub> b''" and "f' : a' \\<^bsub>\\<^esub> b'''" by (elim cat_prod_2_is_arrE[OF assms(2,1) ff']) ultimately have f: "f : a \\<^bsub>\\<^esub> b" and f': "f' : a' \\<^bsub>\\<^esub> b'" by (auto simp: cat_op_simps) from assms g g' f f' have [cat_cs_simps]: "\\ArrMap\\g' \\<^sub>A\<^bsub>\\<^esub> f', g \\<^sub>A\<^bsub>\\<^esub> f\\<^sub>\ = \\ArrMap\\[g', g]\<^sub>\ \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> [f', f]\<^sub>\\" by (cs_concl cs_simp: cat_prod_2_Comp_app cs_intro: cat_prod_cs_intros) from assms g g' f f' show "bifunctor_flip \ \ \\ArrMap\\gg' \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> ff'\ = bifunctor_flip \ \ \\ArrMap\\gg'\ \\<^sub>A\<^bsub>\\<^esub> bifunctor_flip \ \ \\ArrMap\\ff'\" unfolding gg'_def ff'_def (*slow*) by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cs_intro: cat_prod_cs_intros cat_cs_intros ) qed show "bifunctor_flip \ \ \\ArrMap\\(\ \\<^sub>C \)\CId\\ba\\ = \\CId\\bifunctor_flip \ \ \\ObjMap\\ba\\" if "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" for ba proof- from that obtain b a where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) (auto intro: cat_cs_intros) from assms b a have [cat_cs_simps]: "\\ArrMap\\\\CId\\a\, \\CId\\b\\\<^sub>\ = \\ArrMap\\(\ \\<^sub>C \)\CId\\a, b\\<^sub>\\" by (cs_concl cs_simp: cat_prod_2_CId_app cs_intro: cat_prod_cs_intros) from assms b a show ?thesis unfolding ba_def by ( cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros cs_simp: cat_prod_cs_simps cat_cs_simps ) qed qed (auto simp: bifunctor_flip_components cat_cs_simps cat_cs_intros) qed lemma bifunctor_flip_is_functor'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\
= \ \\<^sub>C \" shows "bifunctor_flip \ \ \ : \
\\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4) by (intro bifunctor_flip_is_functor) subsubsection\Double-flip of a bifunctor\ lemma bifunctor_flip_flip[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bifunctor_flip \ \ (bifunctor_flip \ \ \) = \" proof(rule cf_eqI) interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms(3)) from assms show "bifunctor_flip \ \ (bifunctor_flip \ \ \) : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms have ObjMap_dom_lhs: "\\<^sub>\ (bifunctor_flip \ \ (bifunctor_flip \ \ \)\ObjMap\) = (\ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have ObjMap_dom_rhs: "\\<^sub>\ (\\ObjMap\) = (\ \\<^sub>C \)\Obj\" by (simp add: cat_cs_simps) from assms have ArrMap_dom_lhs: "\\<^sub>\ (bifunctor_flip \ \ (bifunctor_flip \ \ \)\ArrMap\) = (\ \\<^sub>C \)\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) have ArrMap_dom_rhs: "\\<^sub>\ (\\ArrMap\) = (\ \\<^sub>C \)\Arr\" by (simp add: cat_cs_simps) show "bifunctor_flip \ \ (bifunctor_flip \ \ \)\ObjMap\ = \\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (rule cat_prod_2_ObjE[OF assms(1,2)]) from assms a b show "bifunctor_flip \ \ (bifunctor_flip \ \ \)\ObjMap\\ab\ = \\ObjMap\\ab\" unfolding ab_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: cat_cs_intros) show "bifunctor_flip \ \ (bifunctor_flip \ \ \)\ArrMap\ = \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Arr\" and b: "b \\<^sub>\ \\Arr\" by (rule cat_prod_2_ArrE[OF assms(1,2)]) from assms a b show "bifunctor_flip \ \ (bifunctor_flip \ \ \)\ArrMap\\ab\ = \\ArrMap\\ab\" unfolding ab_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto simp: cat_cs_intros) qed (simp_all add: assms(3)) subsubsection\A projection of a bifunctor flip\ lemma bifunctor_flip_proj_snd[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F = \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" proof(rule cf_eqI) from assms show f_\b: "bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show \b: "\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms have ObjMap_dom_lhs: "\\<^sub>\ ((bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have ObjMap_dom_rhs: "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms have ArrMap_dom_lhs: "\\<^sub>\ ((bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\) = \\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have ArrMap_dom_rhs: "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\) = \\Arr\" by (cs_concl cs_simp: cat_cs_simps) show "(bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\ = (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) from assms show "vsv ((bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\)" by (intro bifunctor_proj_snd_ObjMap_vsv) (cs_concl cs_intro: cat_cs_intros) from assms show "vsv ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\)" by (intro bifunctor_proj_fst_ObjMap_vsv) (cs_concl cs_intro: cat_cs_intros) fix a assume "a \\<^sub>\ \\Obj\" with assms show "(bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\\a\ = (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_prod_cs_intros) qed simp show "(bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\ = (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) from assms show "vsv ((bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\)" by (intro bifunctor_proj_snd_ArrMap_vsv) (cs_concl cs_intro: cat_cs_intros) from assms show "vsv ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\)" by (intro bifunctor_proj_fst_ArrMap_vsv) (cs_concl cs_intro: cat_cs_intros) fix f assume "f \\<^sub>\ \\Arr\" with assms show "(bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\\f\ = (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\\f\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed simp qed simp_all lemma bifunctor_flip_proj_fst[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(-,a)\<^sub>C\<^sub>F = \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" proof- from assms have f_\: "bifunctor_flip \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show ?thesis by ( rule bifunctor_flip_proj_snd [ OF assms(2,1) f_\ assms(4), unfolded bifunctor_flip_flip[OF assms(1,2,3)], symmetric ] ) qed subsubsection\A flip of a bifunctor isomorphism\ lemma bifunctor_flip_is_iso_functor: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "bifunctor_flip \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \ " proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_iso_functor \ \\ \\<^sub>C \\ \ \ by (rule assms(3)) from assms have f_\: "bifunctor_flip \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \ " by (cs_concl cs_intro: cat_cs_intros) from f_\ have ObjMap_dom: "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps) from f_\ have ArrMap_dom: "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(intro is_iso_functorI' vsv.vsv_valeq_v11I, unfold ObjMap_dom ArrMap_dom) from assms show "bifunctor_flip \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) fix ba b'a' assume prems: "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" "b'a' \\<^sub>\ (\ \\<^sub>C \)\Obj\" "bifunctor_flip \ \ \\ObjMap\\ba\ = bifunctor_flip \ \ \\ObjMap\\b'a'\" from prems(1) obtain b a where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(2,1)]) from prems(2) obtain a' b' where b'a'_def: "b'a' = [b', a']\<^sub>\" and b': "b' \\<^sub>\ \\Obj\" and a': "a' \\<^sub>\ \\Obj\" by (rule cat_prod_2_ObjE[OF assms(2,1)]) from prems(3) assms a b b' a' have \ab_\a'b': "\\ObjMap\\a, b\\<^sub>\ = \\ObjMap\\a', b'\\<^sub>\" unfolding ba_def b'a'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: cf_cs_intros) from assms a b a' b' have "[a, b]\<^sub>\ = [a', b']\<^sub>\" by ( cs_concl cs_intro: \.ObjMap.v11_eq_iff[THEN iffD1, OF _ _ \ab_\a'b'] cat_prod_cs_intros ) then show "ba = b'a'" unfolding ba_def b'a'_def by simp next fix fg f'g' assume prems: "fg \\<^sub>\ (\ \\<^sub>C \)\Arr\" "f'g' \\<^sub>\ (\ \\<^sub>C \)\Arr\" "bifunctor_flip \ \ \\ArrMap\\fg\ = bifunctor_flip \ \ \\ArrMap\\f'g'\" from prems(1) obtain f g where fg_def: "fg = [f, g]\<^sub>\" and f: "f \\<^sub>\ \\Arr\" and g: "g \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF assms(2,1)]) from prems(2) obtain f' g' where f'g'_def: "f'g' = [f', g']\<^sub>\" and f': "f' \\<^sub>\ \\Arr\" and g': "g' \\<^sub>\ \\Arr\" by (rule cat_prod_2_ArrE[OF assms(2,1)]) from prems(3) assms f g f' g' have \gf_\g'f': "\\ArrMap\\g, f\\<^sub>\ = \\ArrMap\\g', f'\\<^sub>\" unfolding fg_def f'g'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: cf_cs_intros) from assms g f g' f' have "[g, f]\<^sub>\ = [g', f']\<^sub>\" by ( cs_concl cs_simp: cs_intro: \.ArrMap.v11_eq_iff[THEN iffD1, OF _ _ \gf_\g'f'] cat_prod_cs_intros ) then show "fg = f'g'" unfolding fg_def f'g'_def by simp next show "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) = \\Obj\" proof(rule vsubset_antisym) show "\\<^sub>\ (bifunctor_flip \ \ \\ObjMap\) \\<^sub>\ \\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold ObjMap_dom) fix ba assume "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain b a where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(2,1)]) from assms b a show "bifunctor_flip \ \ \\ObjMap\\ba\ \\<^sub>\ \\Obj\" unfolding ba_def by (cs_concl cs_intro: cat_cs_intros cf_cs_intros cat_prod_cs_intros) qed (auto simp: cat_cs_intros) show "\\Obj\ \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ObjMap\)" proof(intro vsubsetI) fix c assume prems: "c \\<^sub>\ \\Obj\" from prems obtain ab where ab: "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" and \ab: "\\ObjMap\\ab\ = c" by blast from ab obtain b a where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) show "c \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ObjMap\)" proof(intro vsv.vsv_vimageI2', unfold ObjMap_dom) from assms a b show "[b, a]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" by (cs_concl cs_intro: cat_prod_cs_intros) from assms b a prems show "c = bifunctor_flip \ \ \\ObjMap\\b, a\\<^sub>\" by ( cs_concl cs_simp: \ab[unfolded ab_def] cat_cs_simps cs_intro: cf_cs_intros ) qed (auto intro: cat_cs_intros) qed qed show "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) = \\Arr\" proof(rule vsubset_antisym) show "\\<^sub>\ (bifunctor_flip \ \ \\ArrMap\) \\<^sub>\ \\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold ArrMap_dom) show "vsv (bifunctor_flip \ \ \\ArrMap\)" by (auto intro: cat_cs_intros) fix fg assume "fg \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain f g where fg_def: "fg = [f, g]\<^sub>\" and f: "f \\<^sub>\ \\Arr\" and g: "g \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF assms(2,1)]) from g f obtain a b a' b' where f: "f : a \\<^bsub>\\<^esub> b" and g: "g : a' \\<^bsub>\\<^esub> b'" by (auto intro!: is_arrI) from assms f g show "bifunctor_flip \ \ \\ArrMap\\fg\ \\<^sub>\ \\Arr\" by (cs_concl cs_simp: fg_def cs_intro: cat_cs_intros cat_prod_cs_intros) qed show "\\Arr\ \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ArrMap\)" proof(intro vsubsetI) fix c assume prems: "c \\<^sub>\ \\Arr\" from prems obtain ab where ab: "ab \\<^sub>\ (\ \\<^sub>C \)\Arr\" and \ab: "\\ArrMap\\ab\ = c" by blast from ab obtain b a where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Arr\" and b: "b \\<^sub>\ \\Arr\" by (elim cat_prod_2_ArrE[OF assms(1,2)]) show "c \\<^sub>\ \\<^sub>\ (bifunctor_flip \ \ \\ArrMap\)" proof(intro vsv.vsv_vimageI2', unfold ArrMap_dom) from assms a b show "[b, a]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Arr\" by (cs_concl cs_intro: cat_prod_cs_intros) from assms b a prems show "c = bifunctor_flip \ \ \\ArrMap\\b, a\\<^sub>\" by ( cs_concl cs_simp: \ab[unfolded ab_def] cat_cs_simps cs_intro: cat_cs_intros ) qed (auto intro: cat_cs_intros) qed qed qed (auto intro: cat_cs_intros) qed subsection\Array bifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter II-3 in \cite{mac_lane_categories_2010}.\ definition cf_array :: "V \ V \ V \ (V \ V) \ (V \ V) \ V" where "cf_array \ \ \
\ \ = [ (\a\\<^sub>\(\ \\<^sub>C \)\Obj\. \ (vpfst a)\ObjMap\\vpsnd a\), ( \f\\<^sub>\(\ \\<^sub>C \)\Arr\. \ (\\Cod\\vpfst f\)\ArrMap\\vpsnd f\ \\<^sub>A\<^bsub>\
\<^esub> \ (\\Dom\\vpsnd f\)\ArrMap\\vpfst f\ ), \ \\<^sub>C \, \
]\<^sub>\" text\Components.\ lemma cf_array_components: shows "cf_array \ \ \
\ \\ObjMap\ = (\a\\<^sub>\(\ \\<^sub>C \)\Obj\. \ (vpfst a)\ObjMap\\vpsnd a\)" and "cf_array \ \ \
\ \\ArrMap\ = ( \f\\<^sub>\(\ \\<^sub>C \)\Arr\. \ (\\Cod\\vpfst f\)\ArrMap\\vpsnd f\ \\<^sub>A\<^bsub>\
\<^esub> \ (\\Dom\\vpsnd f\)\ArrMap\\vpfst f\ )" and "cf_array \ \ \
\ \\HomDom\ = \ \\<^sub>C \" and "cf_array \ \ \
\ \\HomCod\ = \
" unfolding cf_array_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma cf_array_ObjMap_vsv: "vsv (cf_array \ \ \
\ \\ObjMap\)" unfolding cf_array_components by simp lemma cf_array_ObjMap_vdomain[cat_cs_simps]: "\\<^sub>\ (cf_array \ \ \
\ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" unfolding cf_array_components by simp lemma cf_array_ObjMap_app[cat_cs_simps]: assumes "[b, c]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "cf_array \ \ \
\ \\ObjMap\\b, c\\<^sub>\ = \ b\ObjMap\\c\" using assms unfolding cf_array_components by (simp add: nat_omega_simps) lemma cf_array_ObjMap_vrange: assumes "category \ \" and "category \ \" and "\b. b \\<^sub>\ \\Obj\ \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_array \ \ \
\ \\ObjMap\) \\<^sub>\ \
\Obj\" proof(rule vsv.vsv_vrange_vsubset, unfold cf_array_ObjMap_vdomain) show "vsv (cf_array \ \ \
\ \\ObjMap\)" by (rule cf_array_ObjMap_vsv) fix x assume prems: "x \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain b c where x_def: "x = [b, c]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) interpret \b: is_functor \ \ \
\\ b\ by (rule assms(3)[OF b]) from prems c show "cf_array \ \ \
\ \\ObjMap\\x\ \\<^sub>\ \
\Obj\" unfolding x_def cf_array_components by (auto simp: nat_omega_simps cat_cs_intros) qed subsubsection\Arrow map\ lemma cf_array_ArrMap_vsv: "vsv (cf_array \ \ \
\ \\ArrMap\)" unfolding cf_array_components by simp lemma cf_array_ArrMap_vdomain[cat_cs_simps]: "\\<^sub>\ (cf_array \ \ \
\ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" unfolding cf_array_components by simp lemma cf_array_ArrMap_app[cat_cs_simps]: assumes "category \ \" and "category \ \" and "g : a \\<^bsub>\\<^esub> b" and "f : a' \\<^bsub>\\<^esub> b'" shows "cf_array \ \ \
\ \\ArrMap\\g, f\\<^sub>\ = \ b\ArrMap\\f\ \\<^sub>A\<^bsub>\
\<^esub> \ a'\ArrMap\\g\" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) from cat_prod_2_is_arrI[OF assms] have "[g, f]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Arr\" by auto with assms show ?thesis unfolding cf_array_components by (simp add: nat_omega_simps cat_cs_simps) qed lemma cf_array_ArrMap_vrange: assumes "category \ \" and "category \ \" and "\c. c \\<^sub>\ \\Obj\ \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\b. b \\<^sub>\ \\Obj\ \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and [cat_cs_simps]: "\b c. b \\<^sub>\ \\Obj\ \ c \\<^sub>\ \\Obj\ \ \ b\ObjMap\\c\ = \ c\ObjMap\\b\" shows "\\<^sub>\ (cf_array \ \ \
\ \\ArrMap\) \\<^sub>\ \
\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cf_array_ArrMap_vdomain) interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \\: category \ \\ \\<^sub>C \\ by (simp add: \.category_axioms \.category_axioms category_cat_prod_2) fix gf assume prems: "gf \\<^sub>\ (\ \\<^sub>C \)\Arr\" then obtain bc b'c' where gf: "gf : bc \\<^bsub>\ \\<^sub>C \\<^esub> b'c'" by auto then obtain g f b c b' c' where gf_def: "gf = [g, f]\<^sub>\" and "bc = [b, c]\<^sub>\" and "b'c' = [b', c']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> b'" and f: "f : c \\<^bsub>\\<^esub> c'" by (elim cat_prod_2_is_arrE[OF assms(1,2)]) then have b: "b \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" and c': "c' \\<^sub>\ \\Obj\" by auto interpret \b: is_functor \ \ \
\\ b\ by (rule assms(4)[OF b]) interpret \c: is_functor \ \ \
\\ c\ by (rule assms(3)[OF c]) interpret \b': is_functor \ \ \
\\ b'\ by (rule assms(4)[OF b']) interpret \c': is_functor \ \ \
\\ c'\ by (rule assms(3)[OF c']) from \b.is_functor_axioms \c.is_functor_axioms \b'.is_functor_axioms \c'.is_functor_axioms \b.HomCod.category_axioms g f have "\ b'\ArrMap\\f\ \\<^sub>A\<^bsub>\
\<^esub> \ c\ArrMap\\g\ \\<^sub>\ \
\Arr\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) with g f prems show "cf_array \ \ \
\ \\ArrMap\\gf\ \\<^sub>\ \
\Arr\" unfolding gf_def cf_array_components by (simp add: nat_omega_simps cat_cs_simps) qed (simp add: cf_array_ArrMap_vsv) subsubsection\Array bifunctor is a bifunctor\ lemma cf_array_specification: \\See Proposition 1 from Chapter II-3 in \cite{mac_lane_categories_2010}.\ assumes "category \ \" and "category \ \" and "category \ \
" and "\c. c \\<^sub>\ \\Obj\ \ \ c : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\b. b \\<^sub>\ \\Obj\ \ \ b : \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\b c. b \\<^sub>\ \\Obj\ \ c \\<^sub>\ \\Obj\ \ \ b\ObjMap\\c\ = \ c\ObjMap\\b\" and "\b c b' c' f g. \ f : b \\<^bsub>\\<^esub> b'; g : c \\<^bsub>\\<^esub> c' \ \ \ b'\ArrMap\\g\ \\<^sub>A\<^bsub>\
\<^esub> \ c\ArrMap\\f\ = \ c'\ArrMap\\f\ \\<^sub>A\<^bsub>\
\<^esub> \ b\ArrMap\\g\" shows cf_array_is_functor: "cf_array \ \ \
\ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and cf_array_ObjMap_app_fst: "\b c. \ b \\<^sub>\ \\Obj\; c \\<^sub>\ \\Obj\ \ \ cf_array \ \ \
\ \\ObjMap\\b, c\\<^sub>\ = \ c\ObjMap\\b\" and cf_array_ObjMap_app_snd: "\b c. \ b \\<^sub>\ \\Obj\; c \\<^sub>\ \\Obj\ \ \ cf_array \ \ \
\ \\ObjMap\\b, c\\<^sub>\ = \ b\ObjMap\\c\" and cf_array_ArrMap_app_fst: "\a b f c. \ f : a \\<^bsub>\\<^esub> b; c \\<^sub>\ \\Obj\\ \ cf_array \ \ \
\ \\ArrMap\\f, \\CId\\c\\\<^sub>\ = \ c\ArrMap\\f\" and cf_array_ArrMap_app_snd: "\a b g c. \ g : a \\<^bsub>\\<^esub> b; c \\<^sub>\ \\Obj\ \ \ cf_array \ \ \
\ \\ArrMap\\\\CId\\c\, g\\<^sub>\ = \ c\ArrMap\\g\" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \
: category \ \
by (rule assms(3)) from assms(4) have [cat_cs_intros]: "\ c : \' \\\<^sub>C\<^bsub>\'\<^esub> \
'" if "c \\<^sub>\ \\Obj\" "\' = \" "\
' = \
" "\' = \" for \' c \' \
' using that(1) unfolding that(2-4) by (intro assms(4)) from assms(4) have [cat_cs_intros]: "\ c : \' \\\<^sub>C\<^bsub>\'\<^esub> \
'" if "c \\<^sub>\ \\Obj\" "\' = \" "\
' = \
" "\' = \" for \' c \' \
' using that(1) unfolding that(2-4) by (intro assms(5)) show "cf_array \ \ \
\ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" proof(intro is_functorI') show "vfsequence (cf_array \ \ \
\ \)" unfolding cf_array_def by auto from assms(1,2) show "category \ (\ \\<^sub>C \)" by (simp add: category_cat_prod_2) show "vcard (cf_array \ \ \
\ \) = 4\<^sub>\" unfolding cf_array_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_array \ \ \
\ \\ObjMap\) \\<^sub>\ \
\Obj\" by (rule cf_array_ObjMap_vrange) (auto simp: assms intro: cat_cs_intros) show cf_array_is_arrI: "cf_array \ \ \
\ \\ArrMap\\ff'\ : cf_array \ \ \
\ \\ObjMap\\aa'\ \\<^bsub>\
\<^esub> cf_array \ \ \
\ \\ObjMap\\bb'\" if ff': "ff' : aa' \\<^bsub>\ \\<^sub>C \\<^esub> bb'" for aa' bb' ff' proof- obtain f f' a a' b b' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and f: "f : a \\<^bsub>\\<^esub> b" and f': "f' : a' \\<^bsub>\\<^esub> b'" by (elim cat_prod_2_is_arrE[OF \.category_axioms \.category_axioms ff']) then have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" and a': "a' \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ \\Obj\" by auto from f' assms(5)[OF a] a have "\ a\ArrMap\\f'\ : \ a'\ObjMap\\a\ \\<^bsub>\
\<^esub> \ b'\ObjMap\\a\" by (cs_concl cs_simp: assms(6)[symmetric] cs_intro: cat_cs_intros) with assms(1-3) f f' assms(4)[OF b'] show ?thesis unfolding ff'_def aa'_def bb'_def by ( cs_concl cs_simp: cat_cs_simps assms(6) cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_array \ \ \
\ \\ArrMap\\gg' \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> ff'\ = cf_array \ \ \
\ \\ArrMap\\gg'\ \\<^sub>A\<^bsub>\
\<^esub> cf_array \ \ \
\ \\ArrMap\\ff'\" if gg': "gg' : bb' \\<^bsub>\ \\<^sub>C \\<^esub> cc'" and ff': "ff' : aa' \\<^bsub>\ \\<^sub>C \\<^esub> bb'" for bb' cc' gg' aa' ff' proof- obtain g g' b b' c c' where gg'_def: "gg' = [g, g']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and cc'_def: "cc' = [c, c']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> c" and g': "g' : b' \\<^bsub>\\<^esub> c'" by (elim cat_prod_2_is_arrE[OF \.category_axioms \.category_axioms gg']) moreover obtain f f' a a' b'' b''' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and "bb' = [b'', b''']\<^sub>\" and f: "f : a \\<^bsub>\\<^esub> b''" and f': "f' : a' \\<^bsub>\\<^esub> b'''" by (elim cat_prod_2_is_arrE[OF \.category_axioms \.category_axioms ff']) ultimately have f: "f : a \\<^bsub>\\<^esub> b" and f': "f' : a' \\<^bsub>\\<^esub> b'" by auto with g have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" and c: "c \\<^sub>\ \\Obj\" and a': "a' \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ \\Obj\" and c': "b' \\<^sub>\ \\Obj\" by auto from f' assms(5)[OF a] a have \a_f': "\ a\ArrMap\\f'\ : \ a'\ObjMap\\a\ \\<^bsub>\
\<^esub> \ b'\ObjMap\\a\" by (cs_concl cs_simp: assms(6)[symmetric] cs_intro: cat_cs_intros) from f' b assms(5)[OF b] have \b_f': "\ b\ArrMap\\f'\ : \ a'\ObjMap\\b\ \\<^bsub>\
\<^esub> \ b'\ObjMap\\b\" by (cs_concl cs_simp: assms(6)[symmetric] cs_intro: cat_cs_intros) from f' c assms(5)[OF c] have \c_f': "\ c\ArrMap\\f'\ : \ a'\ObjMap\\c\ \\<^bsub>\
\<^esub> \ b'\ObjMap\\c\" by (cs_concl cs_simp: assms(6)[symmetric] cs_intro: cat_cs_intros) have "\ b'\ArrMap\\g\ \\<^sub>A\<^bsub>\
\<^esub> (\ b'\ArrMap\\f\ \\<^sub>A\<^bsub>\
\<^esub> \ a\ArrMap\\f'\) = (\ c\ArrMap\\f'\ \\<^sub>A\<^bsub>\
\<^esub> \ a'\ArrMap\\g\) \\<^sub>A\<^bsub>\
\<^esub> \ a'\ArrMap\\f\" using f' f g \b_f' assms(4)[OF a'] assms(4)[OF b'] by (cs_concl cs_simp: cat_cs_simps assms(7) cs_intro: cat_cs_intros) also have "\ = \ c\ArrMap\\f'\ \\<^sub>A\<^bsub>\
\<^esub> (\ a'\ArrMap\\g\ \\<^sub>A\<^bsub>\
\<^esub> \ a'\ArrMap\\f\)" using assms(2) f f' g g' assms(4)[OF a'] assms(5)[OF c] by (cs_concl cs_simp: assms(6) cat_cs_simps cs_intro: cat_cs_intros) finally have [cat_cs_simps]: "\ b'\ArrMap\\g\ \\<^sub>A\<^bsub>\
\<^esub> (\ b'\ArrMap\\f\ \\<^sub>A\<^bsub>\
\<^esub> \ a\ArrMap\\f'\) = \ c\ArrMap\\f'\ \\<^sub>A\<^bsub>\
\<^esub> (\ a'\ArrMap\\g\ \\<^sub>A\<^bsub>\
\<^esub> \ a'\ArrMap\\f\)" by simp show ?thesis using \a_f' \c_f' f f' g g' assms(1,2) assms(4)[OF a'] assms(4)[OF c'] assms(5)[OF c] unfolding gg'_def ff'_def aa'_def bb'_def cc'_def (*slow*) by ( cs_concl cs_simp: assms(6,7) cat_prod_cs_simps cat_cs_simps cs_intro: cat_prod_cs_intros cat_cs_intros ) qed show "cf_array \ \ \
\ \\ArrMap\\(\ \\<^sub>C \)\CId\\cc'\\ = \
\CId\\cf_array \ \ \
\ \\ObjMap\\cc'\\" if "cc' \\<^sub>\ (\ \\<^sub>C \)\Obj\" for cc' proof- from that obtain c c' where cc'_def: "cc' = [c, c']\<^sub>\" and c: "c \\<^sub>\ \\Obj\" and c': "c' \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) (auto intro: cat_cs_intros) from assms(1,2,3) c c' assms(4)[OF c'] assms(5)[OF c] show ?thesis unfolding cc'_def (*very slow*) by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps assms(6) cs_intro: cat_cs_intros cat_prod_cs_intros ) qed qed (auto simp: cf_array_components cat_cs_intros) - show "cf_array \ \ \
\ \\ObjMap\ \b, c\\<^sub>\ = \ c\ObjMap\\b\" + show "cf_array \ \ \
\ \\ObjMap\\b, c\\<^sub>\ = \ c\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" for b c using that assms(1,2,3) by (cs_concl cs_simp: cat_cs_simps assms(6) cs_intro: cat_prod_cs_intros) - show "cf_array \ \ \
\ \\ObjMap\ \b, c\\<^sub>\ = \ b\ObjMap\\c\" + show "cf_array \ \ \
\ \\ObjMap\\b, c\\<^sub>\ = \ b\ObjMap\\c\" if "b \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" for b c using that assms(1,2,3) by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_prod_cs_intros) - show "cf_array \ \ \
\ \\ArrMap\ \f, \\CId\\c\\\<^sub>\ = \ c\ArrMap\\f\" + show "cf_array \ \ \
\ \\ArrMap\\f, \\CId\\c\\\<^sub>\ = \ c\ArrMap\\f\" if f: "f : a \\<^bsub>\\<^esub> b" and c: "c \\<^sub>\ \\Obj\" for a b f c proof- from f have "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" by auto from assms(5)[OF this(1)] assms(5)[OF this(2)] assms(4)[OF c] show ?thesis using assms(1,2,3) f c by (cs_concl cs_simp: cat_cs_simps assms(6) cs_intro: cat_cs_intros) qed - show "cf_array \ \ \
\ \\ArrMap\ \\\CId\\c\, g\\<^sub>\ = \ c\ArrMap\\g\" + show "cf_array \ \ \
\ \\ArrMap\\\\CId\\c\, g\\<^sub>\ = \ c\ArrMap\\g\" if g: "g : a \\<^bsub>\\<^esub> b" and c: "c \\<^sub>\ \\Obj\" for a b g c proof- from g have "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" by auto from assms(4)[OF this(1)] assms(4)[OF this(2)] assms(5)[OF c] show ?thesis using assms(1,2,3) g c by ( cs_concl cs_simp: cat_cs_simps assms(6)[symmetric] cs_intro: cat_cs_intros ) qed qed subsection\Composition of a covariant bifunctor and covariant functors\ subsubsection\Definition and elementary properties.\ definition cf_bcomp :: "V \ V \ V \ V" where "cf_bcomp \ \ \ = [ ( \a\\<^sub>\(\\HomDom\ \\<^sub>C \\HomDom\)\Obj\. \\ObjMap\\\\ObjMap\\vpfst a\, \\ObjMap\\vpsnd a\\\<^sub>\ ), ( \f\\<^sub>\(\\HomDom\ \\<^sub>C \\HomDom\)\Arr\. \\ArrMap\\\\ArrMap\\vpfst f\, \\ArrMap\\vpsnd f\\\<^sub>\ ), \\HomDom\ \\<^sub>C \\HomDom\, \\HomCod\ ]\<^sub>\" text\Components.\ lemma cf_bcomp_components: shows "cf_bcomp \ \ \\ObjMap\ = ( \a\\<^sub>\(\\HomDom\ \\<^sub>C \\HomDom\)\Obj\. \\ObjMap\\\\ObjMap\\vpfst a\, \\ObjMap\\vpsnd a\\\<^sub>\ )" and "cf_bcomp \ \ \\ArrMap\ = ( \f\\<^sub>\(\\HomDom\ \\<^sub>C \\HomDom\)\Arr\. \\ArrMap\\\\ArrMap\\vpfst f\, \\ArrMap\\vpsnd f\\\<^sub>\ )" and "cf_bcomp \ \ \\HomDom\ = \\HomDom\ \\<^sub>C \\HomDom\" and "cf_bcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_bcomp_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma cf_bcomp_ObjMap_vsv: "vsv (cf_bcomp \ \ \\ObjMap\)" unfolding cf_bcomp_components by simp lemma cf_bcomp_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_bcomp \ \ \\ObjMap\) = (\' \\<^sub>C \')\Obj\" proof- interpret \: is_functor \ \' \ \ by (rule assms) interpret \: is_functor \ \' \ \ by (rule assms) show ?thesis unfolding cf_bcomp_components by (simp add: cat_cs_simps) qed lemma cf_bcomp_ObjMap_app[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "[a, b]\<^sub>\ \\<^sub>\ (\' \\<^sub>C \')\Obj\" shows "cf_bcomp \ \ \\ObjMap\\a, b\\<^sub>\ = \\ObjMap\\\\ObjMap\\a\, \\ObjMap\\b\\\<^sub>\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) from assms show ?thesis unfolding cf_bcomp_components by (simp_all add: cat_cs_simps nat_omega_simps) qed lemma cf_bcomp_ObjMap_vrange: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_bcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" proof ( rule vsv.vsv_vrange_vsubset, unfold cf_bcomp_ObjMap_vdomain[OF assms(1,2)] ) interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) show "vsv (cf_bcomp \ \ \\ObjMap\)" by (rule cf_bcomp_ObjMap_vsv) fix bc assume "bc \\<^sub>\ (\' \\<^sub>C \')\Obj\" with \.HomDom.category_axioms \.HomDom.category_axioms obtain b c where bc_def: "bc = [b, c]\<^sub>\" and b: "b \\<^sub>\ \'\Obj\" and c: "c \\<^sub>\ \'\Obj\" by (elim cat_prod_2_ObjE[rotated -1]) from assms b c show "cf_bcomp \ \ \\ObjMap\\bc\ \\<^sub>\ \
\Obj\" unfolding bc_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed subsubsection\Arrow map\ lemma cf_bcomp_ArrMap_vsv: "vsv (cf_bcomp \ \ \\ArrMap\)" unfolding cf_bcomp_components by simp lemma cf_bcomp_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_bcomp \ \ \\ArrMap\) = (\' \\<^sub>C \')\Arr\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) show ?thesis unfolding cf_bcomp_components by (simp add: cat_cs_simps) qed lemma cf_bcomp_ArrMap_app[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "[g, f]\<^sub>\ \\<^sub>\ (\' \\<^sub>C \')\Arr\" shows "cf_bcomp \ \ \\ArrMap\\g, f\\<^sub>\ = \\ArrMap\\\\ArrMap\\g\, \\ArrMap\\f\\\<^sub>\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) from assms show ?thesis unfolding cf_bcomp_components by (simp_all add: nat_omega_simps cat_cs_simps) qed lemma cf_bcomp_ArrMap_vrange: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_bcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cf_bcomp_ArrMap_vdomain[OF assms(1,2)]) interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) fix gf assume "gf \\<^sub>\ (\' \\<^sub>C \')\Arr\" with \.HomDom.category_axioms \.HomDom.category_axioms obtain g f where gf_def: "gf = [g, f]\<^sub>\" and g: "g \\<^sub>\ \'\Arr\" and f: "f \\<^sub>\ \'\Arr\" by (elim cat_prod_2_ArrE[rotated -1]) from g obtain a b where g: "g : a \\<^bsub>\'\<^esub> b" by auto from f obtain a' b' where f: "f : a' \\<^bsub>\'\<^esub> b'" by auto from assms g f show "cf_bcomp \ \ \\ArrMap\\gf\ \\<^sub>\ \
\Arr\" unfolding gf_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (simp add: cf_bcomp_ArrMap_vsv) subsubsection\ Composition of a covariant bifunctor and covariant functors is a functor \ lemma cf_bcomp_is_functor: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_bcomp \ \ \ : \' \\<^sub>C \' \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) interpret \: is_functor \ \\ \\<^sub>C \\ \
\ by (rule assms(3)) show ?thesis proof(intro is_functorI') show "vfsequence (cf_bcomp \ \ \)" unfolding cf_bcomp_def by simp show "category \ (\' \\<^sub>C \')" by ( simp add: \.HomDom.category_axioms \.HomDom.category_axioms category_cat_prod_2 ) show "vcard (cf_bcomp \ \ \) = 4\<^sub>\" unfolding cf_bcomp_def by (simp add: nat_omega_simps) from assms show "\\<^sub>\ (cf_bcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" by (rule cf_bcomp_ObjMap_vrange) show "cf_bcomp \ \ \\ArrMap\\ff'\ : cf_bcomp \ \ \\ObjMap\\aa'\ \\<^bsub>\
\<^esub> cf_bcomp \ \ \\ObjMap\\bb'\" if ff': "ff' : aa' \\<^bsub>\' \\<^sub>C \'\<^esub> bb'" for aa' bb' ff' proof- obtain f f' a a' b b' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and f: "f : a \\<^bsub>\'\<^esub> b" and f': "f' : a' \\<^bsub>\'\<^esub> b'" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_axioms \.HomDom.category_axioms ff' ] ) from assms f f' show ?thesis unfolding ff'_def aa'_def bb'_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_bcomp \ \ \\ArrMap\\gg' \\<^sub>A\<^bsub>\' \\<^sub>C \'\<^esub> ff'\ = cf_bcomp \ \ \\ArrMap\\gg'\ \\<^sub>A\<^bsub>\
\<^esub> cf_bcomp \ \ \\ArrMap\\ff'\" if gg': "gg' : bb' \\<^bsub>\' \\<^sub>C \'\<^esub> cc'" and ff': "ff' : aa' \\<^bsub>\' \\<^sub>C \'\<^esub> bb'" for bb' cc' gg' aa' ff' proof- obtain g g' b b' c c' where gg'_def: "gg' = [g, g']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and cc'_def: "cc' = [c, c']\<^sub>\" and g: "g : b \\<^bsub>\'\<^esub> c" and g': "g' : b' \\<^bsub>\'\<^esub> c'" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_axioms \.HomDom.category_axioms gg' ] ) moreover obtain f f' a a' b'' b''' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and "bb' = [b'', b''']\<^sub>\" and f: "f : a \\<^bsub>\'\<^esub> b''" and f': "f' : a' \\<^bsub>\'\<^esub> b'''" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_axioms \.HomDom.category_axioms ff' ] ) ultimately have f: "f : a \\<^bsub>\'\<^esub> b" and f': "f' : a' \\<^bsub>\'\<^esub> b'" by auto from assms f f' g g' have [cat_cs_simps]: "[\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\, \\ArrMap\\g'\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\]\<^sub>\ = [\\ArrMap\\g\, \\ArrMap\\g'\]\<^sub>\ \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> [\\ArrMap\\f\, \\ArrMap\\f'\]\<^sub>\" by ( cs_concl cs_simp: cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) from assms f f' g g' show ?thesis unfolding gg'_def ff'_def aa'_def bb'_def cc'_def by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_bcomp \ \ \\ArrMap\\(\' \\<^sub>C \')\CId\\cc'\\ = \
\CId\\cf_bcomp \ \ \\ObjMap\\cc'\\" if "cc' \\<^sub>\ (\' \\<^sub>C \')\Obj\" for cc' proof- from that obtain c c' where cc'_def: "cc' = [c, c']\<^sub>\" and c: "c \\<^sub>\ \'\Obj\" and c': "c' \\<^sub>\ \'\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) (auto intro: cat_cs_intros) from assms c c' have [cat_cs_simps]: "[\\CId\\\\ObjMap\\c\\, \\CId\\\\ObjMap\\c'\\]\<^sub>\ = (\ \\<^sub>C \)\CId\\\\ObjMap\\c\, \\ObjMap\\c'\\\<^sub>\" by ( cs_concl cs_simp: cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) from assms c c' show ?thesis unfolding cc'_def by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed qed (auto simp: cf_bcomp_components cat_cs_intros cat_cs_simps) qed lemma cf_bcomp_is_functor'[cat_cs_intros]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\' = \' \\<^sub>C \'" shows "cf_bcomp \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_bcomp_is_functor) subsection\Composition of a contracovariant bifunctor and covariant functors\ text\ The term \contracovariant bifunctor\ is used to refer to a bifunctor that is contravariant in the first argument and covariant in the second argument. \ definition cf_cn_cov_bcomp :: "V \ V \ V \ V" where "cf_cn_cov_bcomp \ \ \ = [ ( \a\\<^sub>\(op_cat (\\HomDom\) \\<^sub>C \\HomDom\)\Obj\. \\ObjMap\\\\ObjMap\\vpfst a\, \\ObjMap\\vpsnd a\\\<^sub>\ ), ( \f\\<^sub>\(op_cat (\\HomDom\) \\<^sub>C \\HomDom\)\Arr\. \\ArrMap\\\\ArrMap\\vpfst f\, \\ArrMap\\vpsnd f\\\<^sub>\ ), op_cat (\\HomDom\) \\<^sub>C \\HomDom\, \\HomCod\ ]\<^sub>\" text\Components.\ lemma cf_cn_cov_bcomp_components: shows "cf_cn_cov_bcomp \ \ \\ObjMap\ = ( \a\\<^sub>\(op_cat (\\HomDom\) \\<^sub>C \\HomDom\)\Obj\. \\ObjMap\\\\ObjMap\\vpfst a\, \\ObjMap\\vpsnd a\\\<^sub>\ )" and "cf_cn_cov_bcomp \ \ \\ArrMap\ = ( \f\\<^sub>\(op_cat (\\HomDom\) \\<^sub>C \\HomDom\)\Arr\. \\ArrMap\\\\ArrMap\\vpfst f\, \\ArrMap\\vpsnd f\\\<^sub>\ )" and "cf_cn_cov_bcomp \ \ \\HomDom\ = op_cat (\\HomDom\) \\<^sub>C \\HomDom\" and "cf_cn_cov_bcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_cn_cov_bcomp_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma cf_cn_cov_bcomp_ObjMap_vsv: "vsv (cf_cn_cov_bcomp \ \ \\ObjMap\)" unfolding cf_cn_cov_bcomp_components by simp lemma cf_cn_cov_bcomp_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_bcomp \ \ \\ObjMap\) = (op_cat \' \\<^sub>C \')\Obj\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) show ?thesis unfolding cf_cn_cov_bcomp_components by (simp add: nat_omega_simps cat_cs_simps) qed lemma cf_cn_cov_bcomp_ObjMap_app[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "[a, b]\<^sub>\ \\<^sub>\ (op_cat \' \\<^sub>C \')\Obj\" shows "cf_cn_cov_bcomp \ \ \\ObjMap\\a, b\\<^sub>\ = \\ObjMap\\\\ObjMap\\a\, \\ObjMap\\b\\\<^sub>\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) from assms show ?thesis unfolding cf_cn_cov_bcomp_components by (simp_all add: cat_cs_simps nat_omega_simps) qed lemma cf_cn_cov_bcomp_ObjMap_vrange: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_bcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" proof ( rule vsv.vsv_vrange_vsubset, unfold cf_cn_cov_bcomp_ObjMap_vdomain[OF assms(1,2)] ) interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) show "vsv (cf_cn_cov_bcomp \ \ \\ObjMap\)" by (rule cf_cn_cov_bcomp_ObjMap_vsv) fix bc assume "bc \\<^sub>\ (op_cat \' \\<^sub>C \')\Obj\" with \.HomDom.category_op \.HomDom.category_axioms obtain b c where bc_def: "bc = [b, c]\<^sub>\" and b: "b \\<^sub>\ op_cat \'\Obj\" and c: "c \\<^sub>\ \'\Obj\" by (elim cat_prod_2_ObjE[rotated -1]) from assms b c show "cf_cn_cov_bcomp \ \ \\ObjMap\\bc\ \\<^sub>\ \
\Obj\" unfolding bc_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed subsubsection\Arrow map\ lemma cf_cn_cov_bcomp_ArrMap_vsv: "vsv (cf_cn_cov_bcomp \ \ \\ArrMap\)" unfolding cf_cn_cov_bcomp_components by simp lemma cf_cn_cov_bcomp_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_bcomp \ \ \\ArrMap\) = (op_cat \' \\<^sub>C \')\Arr\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) show ?thesis unfolding cf_cn_cov_bcomp_components by (simp add: cat_cs_simps) qed lemma cf_cn_cov_bcomp_ArrMap_app[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "[g, f]\<^sub>\ \\<^sub>\ (op_cat \' \\<^sub>C \')\Arr\" shows "cf_cn_cov_bcomp \ \ \\ArrMap\\g, f\\<^sub>\ = \\ArrMap\\\\ArrMap\\g\, \\ArrMap\\f\\\<^sub>\" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) from assms show ?thesis unfolding cf_cn_cov_bcomp_components by (simp_all add: nat_omega_simps cat_cs_simps) qed lemma cf_cn_cov_bcomp_ArrMap_vrange: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_bcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold cf_cn_cov_bcomp_ArrMap_vdomain[OF assms(1,2)]) interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) fix gf assume "gf \\<^sub>\ (op_cat \' \\<^sub>C \')\Arr\" with \.HomDom.category_op \.HomDom.category_axioms obtain g f where gf_def: "gf = [g, f]\<^sub>\" and g: "g \\<^sub>\ op_cat \'\Arr\" and f: "f \\<^sub>\ \'\Arr\" by (elim cat_prod_2_ArrE[rotated -1]) from g obtain a b where g: "g : a \\<^bsub>\'\<^esub> b" unfolding cat_op_simps by auto from f obtain a' b' where f: "f : a' \\<^bsub>\'\<^esub> b'" by auto from assms g f show "cf_cn_cov_bcomp \ \ \\ArrMap\\gf\ \\<^sub>\ \
\Arr\" unfolding gf_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (rule cf_cn_cov_bcomp_ArrMap_vsv) subsubsection\ Composition of a contracovariant bifunctor and functors is a functor \ lemma cf_cn_cov_bcomp_is_functor: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_cn_cov_bcomp \ \ \ : op_cat \' \\<^sub>C \' \\\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: is_functor \ \' \ \ by (rule assms(1)) interpret \: is_functor \ \' \ \ by (rule assms(2)) interpret \: is_functor \ \op_cat \ \\<^sub>C \\ \
\ by (rule assms(3)) show ?thesis proof(intro is_functorI') show "vfsequence (cf_cn_cov_bcomp \ \ \)" unfolding cf_cn_cov_bcomp_def by simp show "category \ (op_cat \' \\<^sub>C \')" by ( simp add: \.HomDom.category_op \.HomDom.category_axioms category_cat_prod_2 ) show "vcard (cf_cn_cov_bcomp \ \ \) = 4\<^sub>\" unfolding cf_cn_cov_bcomp_def by (simp add: nat_omega_simps) from assms show "\\<^sub>\ (cf_cn_cov_bcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" by (rule cf_cn_cov_bcomp_ObjMap_vrange) show "cf_cn_cov_bcomp \ \ \\ArrMap\\ff'\ : cf_cn_cov_bcomp \ \ \\ObjMap\\aa'\ \\<^bsub>\
\<^esub> cf_cn_cov_bcomp \ \ \\ObjMap\\bb'\" if ff': "ff' : aa' \\<^bsub>op_cat \' \\<^sub>C \'\<^esub> bb'" for aa' bb' ff' proof- obtain f f' a a' b b' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and f: "f : a \\<^bsub>op_cat \'\<^esub> b" and f': "f' : a' \\<^bsub>\'\<^esub> b'" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_op \.HomDom.category_axioms ff' ] ) from assms f f' show ?thesis unfolding ff'_def aa'_def bb'_def cat_op_simps 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 show "cf_cn_cov_bcomp \ \ \\ArrMap\\gg' \\<^sub>A\<^bsub>op_cat \' \\<^sub>C \'\<^esub> ff'\ = cf_cn_cov_bcomp \ \ \\ArrMap\\gg'\ \\<^sub>A\<^bsub>\
\<^esub> cf_cn_cov_bcomp \ \ \\ArrMap\\ff'\" if gg': "gg' : bb' \\<^bsub>op_cat \' \\<^sub>C \'\<^esub> cc'" and ff': "ff' : aa' \\<^bsub>op_cat \' \\<^sub>C \'\<^esub> bb'" for bb' cc' gg' aa' ff' proof- obtain g g' b b' c c' where gg'_def: "gg' = [g, g']\<^sub>\" and bb'_def: "bb' = [b, b']\<^sub>\" and cc'_def: "cc' = [c, c']\<^sub>\" and g: "g : b \\<^bsub>op_cat \'\<^esub> c" and g': "g' : b' \\<^bsub>\'\<^esub> c'" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_op \.HomDom.category_axioms gg' ] ) moreover obtain f f' a a' b'' b''' where ff'_def: "ff' = [f, f']\<^sub>\" and aa'_def: "aa' = [a, a']\<^sub>\" and "bb' = [b'', b''']\<^sub>\" and f: "f : a \\<^bsub>op_cat \'\<^esub> b''" and "f' : a' \\<^bsub>\'\<^esub> b'''" by ( elim cat_prod_2_is_arrE[ OF \.HomDom.category_op \.HomDom.category_axioms ff' ] ) ultimately have f: "f : a \\<^bsub>op_cat \'\<^esub> b" and f': "f' : a' \\<^bsub>\'\<^esub> b'" by auto from assms f f' g g' have [cat_cs_simps]: "[ \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\, \\ArrMap\\g'\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f'\ ]\<^sub>\ = [\\ArrMap\\g\, \\ArrMap\\g'\]\<^sub>\ \\<^sub>A\<^bsub>op_cat \ \\<^sub>C \\<^esub> [\\ArrMap\\f\, \\ArrMap\\f'\]\<^sub>\" unfolding cat_op_simps by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from assms f f' g g' show ?thesis unfolding gg'_def ff'_def aa'_def bb'_def cc'_def cat_op_simps by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed show "cf_cn_cov_bcomp \ \ \\ArrMap\\(op_cat \' \\<^sub>C \')\CId\\cc'\\ = \
\CId\\cf_cn_cov_bcomp \ \ \\ObjMap\\cc'\\" if "cc' \\<^sub>\ (op_cat \' \\<^sub>C \')\Obj\" for cc' proof- from that obtain c c' where cc'_def: "cc' = [c, c']\<^sub>\" and c: "c \\<^sub>\ op_cat \'\Obj\" and c': "c' \\<^sub>\ \'\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) (auto intro: cat_cs_intros) from assms c c' have [cat_cs_simps]: "[\\CId\\\\ObjMap\\c\\, \\CId\\\\ObjMap\\c'\\]\<^sub>\ = (op_cat \ \\<^sub>C \)\CId\\\\ObjMap\\c\, \\ObjMap\\c'\\\<^sub>\" unfolding cat_op_simps by ( cs_concl cs_simp: cat_prod_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from assms c c' show ?thesis unfolding cc'_def cat_op_simps by (*slow*) ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed qed (auto simp: cf_cn_cov_bcomp_components cat_cs_simps intro: cat_cs_intros) qed lemma cf_cn_cov_bcomp_is_functor'[cat_cs_intros]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\' = op_cat \' \\<^sub>C \'" shows "cf_cn_cov_bcomp \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_cn_cov_bcomp_is_functor) subsubsection\Projection of a contracovariant bifunctor and functors\ lemma cf_cn_cov_bcomp_bifunctor_proj_snd[cat_cs_simps]: assumes "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \'\Obj\" shows "cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F = (\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \" proof(rule cf_eqI) from assms show [intro]: "cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F : \' \\\<^sub>C\<^bsub>\\<^esub> \
" "(\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \
" by (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from assms have ObjMap_dom_lhs: "\\<^sub>\ ((cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\) = \'\Obj\" and ObjMap_dom_rhs: "\\<^sub>\ (((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ObjMap\) = \'\Obj\" and ArrMap_dom_lhs: "\\<^sub>\ ((cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\) = \'\Arr\" and ArrMap_dom_rhs: "\\<^sub>\ (((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ArrMap\) = \'\Arr\" by (cs_concl cs_intro: cat_cs_intros cat_op_intros cs_simp: cat_cs_simps)+ show "(cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\ = ((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume "a \\<^sub>\ \'\Obj\" with assms show "(cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ObjMap\\a\ = ((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ObjMap\\a\" by ( cs_concl cs_simp: cat_prod_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (auto intro: is_functor.cf_ObjMap_vsv) show "(cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\ = ((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume "f \\<^sub>\ \'\Arr\" then obtain a' b' where "f : a' \\<^bsub>\'\<^esub> b'" by (auto intro: is_arrI) with assms show "(cf_cn_cov_bcomp \ \ \\<^bsub>op_cat \',\'\<^esub>(b,-)\<^sub>C\<^sub>F)\ArrMap\\f\ = ((\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \)\ArrMap\\f\" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (auto intro: is_functor.cf_ArrMap_vsv) qed simp_all subsection\Composition of a covariant bifunctor and a covariant functor\ subsubsection\Definition and elementary properties\ definition cf_lcomp :: "V \ V \ V \ V" where "cf_lcomp \ \ \ = cf_bcomp \ \ (cf_id \)" definition cf_rcomp :: "V \ V \ V \ V" where "cf_rcomp \ \ \ = cf_bcomp \ (cf_id \) \" text\Components.\ lemma cf_lcomp_components: shows "cf_lcomp \ \ \\HomDom\ = \\HomDom\ \\<^sub>C \" and "cf_lcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_lcomp_def cf_bcomp_components dghm_id_components by simp_all lemma cf_rcomp_components: shows "cf_rcomp \ \ \\HomDom\ = \ \\<^sub>C \\HomDom\" and "cf_rcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_rcomp_def cf_bcomp_components dghm_id_components by simp_all subsubsection\Object map\ lemma cf_lcomp_ObjMap_vsv: "vsv (cf_lcomp \ \ \\ObjMap\)" unfolding cf_lcomp_def by (rule cf_bcomp_ObjMap_vsv) lemma cf_rcomp_ObjMap_vsv: "vsv (cf_rcomp \ \ \\ObjMap\)" unfolding cf_rcomp_def by (rule cf_bcomp_ObjMap_vsv) lemma cf_lcomp_ObjMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_lcomp \ \ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" using assms unfolding cf_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_rcomp_ObjMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_rcomp \ \ \\ObjMap\) = (\ \\<^sub>C \)\Obj\" using assms unfolding cf_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_lcomp_ObjMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" shows "cf_lcomp \ \ \\ObjMap\\a, c\\<^sub>\ = \\ObjMap\\\\ObjMap\\a\, c\\<^sub>\" using assms unfolding cf_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) lemma cf_rcomp_ObjMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" and "a \\<^sub>\ \\Obj\" shows "cf_rcomp \ \ \\ObjMap\\b, a\\<^sub>\ = \\ObjMap\\b, \\ObjMap\\a\\\<^sub>\" using assms unfolding cf_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) lemma cf_lcomp_ObjMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_lcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" using assms unfolding cf_lcomp_def by (intro cf_bcomp_ObjMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ lemma cf_rcomp_ObjMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_rcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" using assms unfolding cf_rcomp_def by (intro cf_bcomp_ObjMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ subsubsection\Arrow map\ lemma cf_lcomp_ArrMap_vsv: "vsv (cf_lcomp \ \ \\ArrMap\)" unfolding cf_lcomp_def by (rule cf_bcomp_ArrMap_vsv) lemma cf_rcomp_ArrMap_vsv: "vsv (cf_rcomp \ \ \\ArrMap\)" unfolding cf_rcomp_def by (rule cf_bcomp_ArrMap_vsv) lemma cf_lcomp_ArrMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_lcomp \ \ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" using assms unfolding cf_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_rcomp_ArrMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_rcomp \ \ \\ArrMap\) = (\ \\<^sub>C \)\Arr\" using assms unfolding cf_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_lcomp_ArrMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f \\<^sub>\ \\Arr\" and "g \\<^sub>\ \\Arr\" shows "cf_lcomp \ \ \\ArrMap\\f, g\\<^sub>\ = \\ArrMap\\\\ArrMap\\f\, g\\<^sub>\" using assms unfolding cf_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) lemma cf_rcomp_ArrMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f \\<^sub>\ \\Arr\" and "g \\<^sub>\ \\Arr\" shows "cf_rcomp \ \ \\ArrMap\\f, g\\<^sub>\ = \\ArrMap\\f, \\ArrMap\\g\\\<^sub>\" using assms unfolding cf_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros) lemma cf_lcomp_ArrMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_lcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" using assms unfolding cf_lcomp_def by (intro cf_bcomp_ArrMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ lemma cf_rcomp_ArrMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_rcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" using assms unfolding cf_rcomp_def by (intro cf_bcomp_ArrMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ subsubsection\ Composition of a covariant bifunctor and a covariant functor is a functor \ lemma cf_lcomp_is_functor: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_lcomp \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms unfolding cf_lcomp_def by (cs_concl cs_intro: cat_cs_intros)+ lemma cf_lcomp_is_functor'[cat_cs_intros]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\' = \ \\<^sub>C \" shows "cf_lcomp \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_lcomp_is_functor) lemma cf_rcomp_is_functor: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_rcomp \ \ \ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms unfolding cf_rcomp_def by (cs_concl cs_intro: cat_cs_intros)+ lemma cf_rcomp_is_functor'[cat_cs_intros]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\' = \ \\<^sub>C \" shows "cf_rcomp \ \ \ : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_rcomp_is_functor) subsection\Composition of a contracovariant bifunctor and a covariant functor\ definition cf_cn_cov_lcomp :: "V \ V \ V \ V" where "cf_cn_cov_lcomp \ \ \ = cf_cn_cov_bcomp \ \ (cf_id \)" definition cf_cn_cov_rcomp :: "V \ V \ V \ V" where "cf_cn_cov_rcomp \ \ \ = cf_cn_cov_bcomp \ (cf_id \) \" text\Components.\ lemma cf_cn_cov_lcomp_components: shows "cf_cn_cov_lcomp \ \ \\HomDom\ = op_cat (\\HomDom\) \\<^sub>C \" and "cf_cn_cov_lcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_cn_cov_lcomp_def cf_cn_cov_bcomp_components dghm_id_components by simp_all lemma cf_cn_cov_rcomp_components: shows "cf_cn_cov_rcomp \ \ \\HomDom\ = op_cat \ \\<^sub>C \\HomDom\" and "cf_cn_cov_rcomp \ \ \\HomCod\ = \\HomCod\" unfolding cf_cn_cov_rcomp_def cf_cn_cov_bcomp_components dghm_id_components by simp_all subsubsection\Object map\ lemma cf_cn_cov_lcomp_ObjMap_vsv: "vsv (cf_cn_cov_lcomp \ \ \\ObjMap\)" unfolding cf_cn_cov_lcomp_def by (rule cf_cn_cov_bcomp_ObjMap_vsv) lemma cf_cn_cov_rcomp_ObjMap_vsv: "vsv (cf_cn_cov_rcomp \ \ \\ObjMap\)" unfolding cf_cn_cov_rcomp_def by (rule cf_cn_cov_bcomp_ObjMap_vsv) lemma cf_cn_cov_lcomp_ObjMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_lcomp \ \ \\ObjMap\) = (op_cat \ \\<^sub>C \)\Obj\" using assms unfolding cf_cn_cov_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_cn_cov_rcomp_ObjMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_rcomp \ \ \\ObjMap\) = (op_cat \ \\<^sub>C \)\Obj\" using assms unfolding cf_cn_cov_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_cn_cov_lcomp_ObjMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "c \\<^sub>\ \\Obj\" shows "cf_cn_cov_lcomp \ \ \\ObjMap\\a, c\\<^sub>\ = \\ObjMap\\\\ObjMap\\a\, c\\<^sub>\" using assms unfolding cf_cn_cov_lcomp_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) lemma cf_cn_cov_rcomp_ObjMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ op_cat \\Obj\" and "a \\<^sub>\ \\Obj\" shows "cf_cn_cov_rcomp \ \ \\ObjMap\\b, a\\<^sub>\ = \\ObjMap\\b, \\ObjMap\\a\\\<^sub>\" using assms unfolding cf_cn_cov_rcomp_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) lemma cf_cn_cov_lcomp_ObjMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_lcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" using assms unfolding cf_cn_cov_lcomp_def by (intro cf_cn_cov_bcomp_ObjMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ lemma cf_cn_cov_rcomp_ObjMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_rcomp \ \ \\ObjMap\) \\<^sub>\ \
\Obj\" using assms unfolding cf_cn_cov_rcomp_def by (intro cf_cn_cov_bcomp_ObjMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ subsubsection\Arrow map\ lemma cf_cn_cov_lcomp_ArrMap_vsv: "vsv (cf_cn_cov_lcomp \ \ \\ArrMap\)" unfolding cf_cn_cov_lcomp_def by (rule cf_cn_cov_bcomp_ArrMap_vsv) lemma cf_cn_cov_rcomp_ArrMap_vsv: "vsv (cf_cn_cov_rcomp \ \ \\ArrMap\)" unfolding cf_cn_cov_rcomp_def by (rule cf_cn_cov_bcomp_ArrMap_vsv) lemma cf_cn_cov_lcomp_ArrMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_lcomp \ \ \\ArrMap\) = (op_cat \ \\<^sub>C \)\Arr\" using assms unfolding cf_cn_cov_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_cn_cov_rcomp_ArrMap_vdomain[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_cn_cov_rcomp \ \ \\ArrMap\) = (op_cat \ \\<^sub>C \)\Arr\" using assms unfolding cf_cn_cov_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_cn_cov_lcomp_ArrMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f \\<^sub>\ op_cat \\Arr\" and "g \\<^sub>\ \\Arr\" shows "cf_cn_cov_lcomp \ \ \\ArrMap\\f, g\\<^sub>\ = \\ArrMap\\\\ArrMap\\f\, g\\<^sub>\" using assms unfolding cf_cn_cov_lcomp_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) lemma cf_cn_cov_rcomp_ArrMap_app[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "f \\<^sub>\ op_cat \\Arr\" and "g \\<^sub>\ \\Arr\" shows "cf_cn_cov_rcomp \ \ \\ArrMap\\f, g\\<^sub>\ = \\ArrMap\\f, \\ArrMap\\g\\\<^sub>\" using assms unfolding cf_cn_cov_rcomp_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) lemma cf_cn_cov_lcomp_ArrMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_lcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" using assms unfolding cf_cn_cov_lcomp_def by (intro cf_cn_cov_bcomp_ArrMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ lemma cf_cn_cov_rcomp_ArrMap_vrange: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "\\<^sub>\ (cf_cn_cov_rcomp \ \ \\ArrMap\) \\<^sub>\ \
\Arr\" using assms unfolding cf_cn_cov_rcomp_def cat_op_simps by (intro cf_cn_cov_bcomp_ArrMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ subsubsection\ Composition of a contracovariant bifunctor and a covariant functor is a functor \ lemma cf_cn_cov_lcomp_is_functor: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_cn_cov_lcomp \ \ \ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms unfolding cf_cn_cov_lcomp_def cat_op_simps by (cs_concl cs_intro: cat_cs_intros)+ lemma cf_cn_cov_lcomp_is_functor'[cat_cs_intros]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\\ = op_cat \ \\<^sub>C \" shows "cf_cn_cov_lcomp \ \ \ : \\ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_cn_cov_lcomp_is_functor) lemma cf_cn_cov_rcomp_is_functor: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" shows "cf_cn_cov_rcomp \ \ \ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms unfolding cf_cn_cov_rcomp_def cat_op_simps by (cs_concl cs_intro: cat_cs_intros)+ lemma cf_cn_cov_rcomp_is_functor'[cat_cs_intros]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "\\ = op_cat \ \\<^sub>C \" shows "cf_cn_cov_rcomp \ \ \ : \\ \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1-3) unfolding assms(4) by (rule cf_cn_cov_rcomp_is_functor) subsubsection\ Projection of a composition of a contracovariant bifunctor and a covariant functor \ lemma cf_cn_cov_rcomp_bifunctor_proj_snd[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "cf_cn_cov_rcomp \ \ \\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>C\<^sub>F = (\\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>C\<^sub>F) \\<^sub>C\<^sub>F \" using assms unfolding cf_cn_cov_rcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) lemma cf_cn_cov_lcomp_bifunctor_proj_snd[cat_cs_simps]: assumes "category \ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \
" and "b \\<^sub>\ \\Obj\" shows "cf_cn_cov_lcomp \ \ \\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>C\<^sub>F = (\\<^bsub>op_cat \,\\<^esub>(\\ObjMap\\b\,-)\<^sub>C\<^sub>F)" using assms unfolding cf_cn_cov_lcomp_def by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) subsection\Composition of bifunctors\ subsubsection\Definitions and elementary properties\ definition cf_blcomp :: "V \ V" where "cf_blcomp \ = cf_lcomp (\\HomCod\) \ \ \\<^sub>C\<^sub>F cf_cat_prod_21_of_3 (\\HomCod\) (\\HomCod\) (\\HomCod\)" definition cf_brcomp :: "V \ V" where "cf_brcomp \ = cf_rcomp (\\HomCod\) \ \ \\<^sub>C\<^sub>F cf_cat_prod_12_of_3 (\\HomCod\) (\\HomCod\) (\\HomCod\)" text\Alternative forms of the definitions.\ lemma cf_blcomp_def': assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "cf_blcomp \ = cf_lcomp \ \ \ \\<^sub>C\<^sub>F cf_cat_prod_21_of_3 \ \ \" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) show ?thesis by (cs_concl cs_simp: cat_cs_simps cf_blcomp_def cs_intro: cat_cs_intros) qed lemma cf_brcomp_def': assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "cf_brcomp \ = cf_rcomp \ \ \ \\<^sub>C\<^sub>F cf_cat_prod_12_of_3 \ \ \" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) show ?thesis by (cs_concl cs_simp: cat_cs_simps cf_brcomp_def cs_intro: cat_cs_intros) qed subsubsection\Compositions of bifunctors are functors\ lemma cf_blcomp_is_functor: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "cf_blcomp \ : \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) show ?thesis by (cs_concl cs_simp: cat_cs_simps cf_blcomp_def' cs_intro: cat_cs_intros) qed lemma cf_blcomp_is_functor'[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" shows "cf_blcomp \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_blcomp_is_functor) lemma cf_brcomp_is_functor: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "cf_brcomp \ : \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) show ?thesis by (cs_concl cs_simp: cat_cs_simps cf_brcomp_def' cs_intro: cat_cs_intros) qed lemma cf_brcomp_is_functor'[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' = \ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \" shows "cf_brcomp \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1) unfolding assms(2) by (rule cf_brcomp_is_functor) subsubsection\Object map\ lemma cf_blcomp_ObjMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv (cf_blcomp \\ObjMap\)" proof- interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms]) show ?thesis by auto qed lemma cf_brcomp_ObjMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv (cf_brcomp \\ObjMap\)" proof- interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms]) show ?thesis by auto qed lemma cf_blcomp_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_blcomp \\ObjMap\) = (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms]) show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma cf_brcomp_ObjMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_brcomp \\ObjMap\) = (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Obj\" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms]) show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma cf_blcomp_ObjMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "A = [a, b, c]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" shows "cf_blcomp \\ObjMap\\A\ = (a \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\<^esub> b) \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\<^esub> c" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms(1)]) from assms(3-5) show ?thesis unfolding assms(2) by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cf_blcomp_def' cs_intro: cat_cs_intros cat_prod_cs_intros ) qed lemma cf_brcomp_ObjMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "A = [a, b, c]\<^sub>\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" shows "cf_brcomp \\ObjMap\\A\ = a \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\<^esub> (b \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\<^esub> c)" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms(1)]) from assms(3-5) show ?thesis unfolding assms(2) by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cf_brcomp_def' cs_intro: cat_cs_intros cat_prod_cs_intros ) qed subsubsection\Arrow map\ lemma cf_blcomp_ArrMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv (cf_blcomp \\ArrMap\)" proof- interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms]) show ?thesis by auto qed lemma cf_brcomp_ArrMap_vsv[cat_cs_intros]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "vsv (cf_brcomp \\ArrMap\)" proof- interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms]) show ?thesis by auto qed lemma cf_blcomp_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_blcomp \\ArrMap\) = (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms]) show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma cf_brcomp_ArrMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (cf_brcomp \\ArrMap\) = (\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \)\Arr\" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms]) show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma cf_blcomp_ArrMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "F = [h, g, f]\<^sub>\" and "h \\<^sub>\ \\Arr\" and "g \\<^sub>\ \\Arr\" and "f \\<^sub>\ \\Arr\" shows "cf_blcomp \\ArrMap\\F\ = (h \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\<^esub> g) \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\<^esub> f" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_blcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_blcomp \\ by (rule cf_blcomp_is_functor[OF assms(1)]) from assms(3-5) show ?thesis unfolding assms(2) by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cf_blcomp_def' cs_intro: cat_cs_intros cat_prod_cs_intros ) qed lemma cf_brcomp_ArrMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "F = [h, g, f]\<^sub>\" and "h \\<^sub>\ \\Arr\" and "g \\<^sub>\ \\Arr\" and "f \\<^sub>\ \\Arr\" shows "cf_brcomp \\ArrMap\\F\ = h \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\<^esub> (g \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\<^esub> f)" proof- interpret \: is_functor \ \\ \\<^sub>C \\ \ \ by (rule assms) interpret cf_brcomp: is_functor \ \\ \\<^sub>C\<^sub>3 \ \\<^sub>C\<^sub>3 \\ \ \cf_brcomp \\ by (rule cf_brcomp_is_functor[OF assms(1)]) from assms(3-5) show ?thesis unfolding assms(2) by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cf_brcomp_def' cs_intro: cat_cs_intros cat_prod_cs_intros ) qed subsection\Binatural transformation\ subsubsection\Definitions and elementary properties\ text\ In this work, a \binatural transformation\ is used to denote a natural transformation of bifunctors. \ definition bnt_proj_fst :: "V \ V \ V \ V \ V" (\(_\<^bsub>_,_\<^esub>/'(/-,_/')/\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\ [51, 51, 51, 51] 51) where "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F = [ (\a\\<^sub>\\\Obj\. \\NTMap\\a, b\\<^sub>\), \\NTDom\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F, \\NTCod\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F, \, \\NTDGCod\ ]\<^sub>\" definition bnt_proj_snd :: "V \ V \ V \ V \ V" (\(_\<^bsub>_,_\<^esub>/'(/_,-/')/\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\ [51, 51, 51, 51] 51) where "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F = [ (\b\\<^sub>\\\Obj\. \\NTMap\\a, b\\<^sub>\), \\NTDom\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F, \\NTCod\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F, \, \\NTDGCod\ ]\<^sub>\" text\Components\ lemma bnt_proj_fst_components: shows "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\ = (\a\\<^sub>\\\Obj\. \\NTMap\\a, b\\<^sub>\)" and "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDom\ = \\NTDom\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTCod\ = \\NTCod\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDGDom\ = \" and "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDGCod\ = \\NTDGCod\" unfolding bnt_proj_fst_def nt_field_simps by (simp_all add: nat_omega_simps) lemma bnt_proj_snd_components: shows "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\ = (\b\\<^sub>\\\Obj\. \\NTMap\\a, b\\<^sub>\)" and "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDom\ = \\NTDom\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" and "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTCod\ = \\NTCod\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" and "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDGDom\ = \" and "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTDGCod\ = \\NTDGCod\" unfolding bnt_proj_snd_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation maps\ mk_VLambda bnt_proj_fst_components(1)[folded VLambda_vconst_on] |vsv bnt_proj_fst_NTMap_vsv[cat_cs_intros]| |vdomain bnt_proj_fst_NTMap_vdomain[cat_cs_simps]| |app bnt_proj_fst_NTMap_app[cat_cs_simps]| lemma bnt_proj_fst_vrange: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis unfolding bnt_proj_fst_components proof(rule vrange_VLambda_vsubset) fix a assume "a \\<^sub>\ \\Obj\" with assms show "\\NTMap\\a, b\\<^sub>\ \\<^sub>\ \\Arr\" by (cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros) qed qed mk_VLambda bnt_proj_snd_components(1)[folded VLambda_vconst_on] |vsv bnt_proj_snd_NTMap_vsv[intro]| |vdomain bnt_proj_snd_NTMap_vdomain[cat_cs_simps]| |app bnt_proj_snd_NTMap_app[cat_cs_simps]| lemma bnt_proj_snd_vrange: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\) \\<^sub>\ \\Arr\" proof- interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis unfolding bnt_proj_snd_components proof(rule vrange_VLambda_vsubset) fix b assume "b \\<^sub>\ \\Obj\" with assms show "\\NTMap\\a, b\\<^sub>\ \\<^sub>\ \\Arr\" by (cs_concl cs_intro: cat_cs_intros cat_prod_cs_intros) qed qed subsubsection\Binatural transformation projection is a natural transformation\ lemma bnt_proj_snd_is_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis proof(intro is_ntcfI') show "vfsequence (\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)" unfolding bnt_proj_snd_def by simp show "vcard (\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F) = 5\<^sub>\" unfolding bnt_proj_snd_def by (simp add: nat_omega_simps) from assms show "\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms show "\'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b\ : (\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\ \\<^bsub>\\<^esub> (\'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" for b using that assms by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) show "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> (\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\\f\ = (\'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a'\" if "f : a' \\<^bsub>\\<^esub> b" for a' b f using that assms by ( cs_concl cs_simp: is_ntcf.ntcf_Comp_commute cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (auto simp: bnt_proj_snd_components cat_cs_simps) qed lemma bnt_proj_snd_is_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "\ = \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" and "\ = \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms by (auto intro: bnt_proj_snd_is_ntcf) lemma bnt_proj_fst_is_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis proof(intro is_ntcfI') show "vfsequence (\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)" unfolding bnt_proj_fst_def by simp show "vcard (\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F) = 5\<^sub>\" unfolding bnt_proj_fst_def by (simp add: nat_omega_simps) from assms show "\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show "\'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\ : (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\ \\<^bsub>\\<^esub> (\'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using that assms by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) show "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\\f\ = (\'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\" if "f : a \\<^bsub>\\<^esub> b'" for a b' f using that assms by ( cs_concl cs_simp: is_ntcf.ntcf_Comp_commute cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed (auto simp: bnt_proj_fst_components cat_cs_simps) qed lemma bnt_proj_fst_is_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" and "\ = \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "\ = \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "\' = \" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \ \\<^sub>C\<^sub>F \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-4) unfolding assms(5-7) by (rule bnt_proj_fst_is_ntcf) subsubsection\Array binatural transformation is a natural transformation\ lemma ntcf_array_is_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "vfsequence \" and "vcard \ = 5\<^sub>\" and "\\NTDom\ = \" and "\\NTCod\ = \'" and "\\NTDGDom\ = \ \\<^sub>C \" and "\\NTDGCod\ = \" and "vsv (\\NTMap\)" and "\\<^sub>\ (\\NTMap\) = (\ \\<^sub>C \)\Obj\" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ \\NTMap\\a, b\\<^sub>\ : \\ObjMap\\a, b\\<^sub>\ \\<^bsub>\\<^esub> \'\ObjMap\\a, b\\<^sub>\" and "\a. a \\<^sub>\ \\Obj\ \ \\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\b. b \\<^sub>\ \\Obj\ \ \\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: vsv \\\NTMap\\ by (rule assms(11)) have [cat_cs_intros]: "\ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\; A = \\ObjMap\\a, b\\<^sub>\; B = \'\ObjMap\\a, b\\<^sub>\ \ \ \\NTMap\\a, b\\<^sub>\ : A \\<^bsub>\\<^esub> B" for a b A B by (auto intro: assms(13)) show ?thesis proof(intro is_ntcfI') show "\\NTMap\\ab\ : \\ObjMap\\ab\ \\<^bsub>\\<^esub> \'\ObjMap\\ab\" if "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" for ab proof- from that obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) from a b show ?thesis unfolding ab_def by (rule assms(13)) qed show "\\NTMap\\a'b'\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\gf\ = \'\ArrMap\\gf\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\ab\" if "gf : ab \\<^bsub>\ \\<^sub>C \\<^esub> a'b'" for ab a'b' gf proof- from that obtain g f a b a' b' where gf_def: "gf = [g, f]\<^sub>\" and ab_def: "ab = [a, b]\<^sub>\" and a'b'_def: "a'b' = [a', b']\<^sub>\" and g: "g : a \\<^bsub>\\<^esub> a'" and f: "f : b \\<^bsub>\\<^esub> b'" by (elim cat_prod_2_is_arrE[OF assms(1,2)]) then have a: "a \\<^sub>\ \\Obj\" and a': "a' \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" and b': "b' \\<^sub>\ \\Obj\" by auto show ?thesis unfolding gf_def ab_def a'b'_def proof- from is_ntcfD'(13)[OF assms(15)[OF b] g] g f assms(1,2,3,4) have [cat_cs_simps]: "(\'\ArrMap\\g, \\CId\\b\\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a, b\\<^sub>\) = (\\NTMap\\a', b\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\)" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) auto from is_ntcfD'(13)[OF assms(14)[OF a'] f] g f assms(1,2) have \'\: "\'\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a', b\\<^sub>\ = \\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\\\CId\\a'\,f\\<^sub>\" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) auto from g f assms(1-4) have [cat_cs_simps]: "\'\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a', b\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> q) = \\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> (\\ArrMap\\\\CId\\a'\,f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> q)" if "q : r \\<^bsub>\\<^esub> \\ObjMap\\a', b\\<^sub>\" for q r using that by ( cs_concl cs_simp: \'\ category.cat_Comp_assoc[symmetric] cs_intro: cat_cs_intros cat_prod_cs_intros ) from assms(1-4) g f have "\'\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \'\ArrMap\\g, \\CId\\b\\\<^sub>\ = \'\ArrMap\\[\\CId\\a'\, f]\<^sub>\ \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> [g, \\CId\\b\]\<^sub>\\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) - also from assms(1-4) g f have "\ = \'\ArrMap\ \g, f\\<^sub>\" + also from assms(1-4) g f have "\ = \'\ArrMap\\g, f\\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) - finally have \'_gf: "\'\ArrMap\ \g, f\\<^sub>\ = + finally have \'_gf: "\'\ArrMap\\g, f\\<^sub>\ = \'\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \'\ArrMap\\g, \\CId\\b\\\<^sub>\" by simp from assms(1-4) g f have "\\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\ = \\ArrMap\\[\\CId\\a'\, f]\<^sub>\ \\<^sub>A\<^bsub>\ \\<^sub>C \\<^esub> [g, \\CId\\b\]\<^sub>\\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) - also from assms(1-4) g f have "\ = \\ArrMap\ \g, f\\<^sub>\" + also from assms(1-4) g f have "\ = \\ArrMap\\g, f\\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) finally have \_gf: "\\ArrMap\\g, f\\<^sub>\ = \\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\" by simp from assms(1-4) g f assms(13)[OF a b] assms(13)[OF a' b] have "\'\ArrMap\\g, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a, b\\<^sub>\ = (\'\ArrMap\\\\CId\\a'\, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a', b\\<^sub>\) \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\" unfolding \'_gf by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) also from assms(1-4) g f have "\ = (\\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\\\CId\\a'\,f\\<^sub>\) \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) also from assms(1-4) g f assms(13)[OF a' b'] have "\ = \\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> (\\ArrMap\\\\CId\\a'\,f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, \\CId\\b\\\<^sub>\)" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) also from assms(1-4) g f assms(13)[OF a' b'] have "\ = \\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, f\\<^sub>\" unfolding \_gf[symmetric] by simp finally show "\\NTMap\\a', b'\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g, f\\<^sub>\ = \'\ArrMap\\g, f\\<^sub>\ \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a, b\\<^sub>\" by simp qed qed qed (auto simp: assms) qed subsubsection\Binatural transformation projections and isomorphisms\ lemma is_iso_ntcf_if_bnt_proj_snd_is_iso_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\a. a \\<^sub>\ \\Obj\ \ \\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) show ?thesis proof(intro is_iso_ntcfI) show "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms(3)) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) interpret \a: is_iso_ntcf \ \ \ \\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ \\'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F\ \\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F\ by (rule assms(4)[OF a]) from b have \ab: "\\NTMap\\a, b\\<^sub>\ = (\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b\" by (cs_concl cs_simp: cat_cs_simps) from \a.iso_ntcf_is_arr_isomorphism[OF b] assms(1,2,3) a b show "\\NTMap\\ab\ : \\ObjMap\\ab\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'\ObjMap\\ab\" by (cs_prems cs_simp: cat_cs_simps ab_def cs_intro: cat_prod_cs_intros) qed qed lemma is_iso_ntcf_if_bnt_proj_fst_is_iso_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\b. b \\<^sub>\ \\Obj\ \ \\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) show ?thesis proof(intro is_iso_ntcfI) show "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms(3)) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) interpret \a: is_iso_ntcf \ \ \ \\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ \\'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F\ \\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F\ by (rule assms(4)[OF b]) from b have \ab: "\\NTMap\\a, b\\<^sub>\ = (\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \a.iso_ntcf_is_arr_isomorphism[OF a] assms(1,2,3) a b show "\\NTMap\\ab\ : \\ObjMap\\ab\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'\ObjMap\\ab\" unfolding ab_def by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_prod_cs_intros) qed qed lemma bnt_proj_snd_is_iso_ntcf_if_is_iso_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_iso_ntcfI) from assms show "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) show "(\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\b\ : (\\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> (\'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F)\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" for b using assms that by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_prod_cs_intros cat_arrow_cs_intros ) qed lemma bnt_proj_snd_is_iso_ntcf_if_is_iso_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" and "\ = \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F" and "\' = \" and "a \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" unfolding assms(4-6) by (rule bnt_proj_snd_is_iso_ntcf_if_is_iso_ntcf[OF assms(1-3,7)]) lemma bnt_proj_fst_is_iso_ntcf_if_is_iso_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(intro is_iso_ntcfI) from assms show "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) show "(\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\ : (\\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> (\'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F)\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using assms that by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_prod_cs_intros cat_arrow_cs_intros ) qed lemma bnt_proj_fst_is_iso_ntcf_if_is_iso_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "\ = \'\<^bsub>\,\\<^esub>(-,b)\<^sub>C\<^sub>F" and "\' = \" and "b \\<^sub>\ \\Obj\" shows "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \' \\\<^sub>C\<^bsub>\\<^esub> \" unfolding assms(4-6) by (rule bnt_proj_fst_is_iso_ntcf_if_is_iso_ntcf[OF assms(1-3,7)]) subsection\Binatural transformation flip\ subsubsection\Definition and elementary properties\ definition bnt_flip :: "V \ V \ V \ V" where "bnt_flip \ \ \ = [ fflip (\\NTMap\), bifunctor_flip \ \ (\\NTDom\), bifunctor_flip \ \ (\\NTCod\), \ \\<^sub>C \, \\NTDGCod\ ]\<^sub>\" text\Components.\ lemma bnt_flip_components: shows "bnt_flip \ \ \\NTMap\ = fflip (\\NTMap\)" and "bnt_flip \ \ \\NTDom\ = bifunctor_flip \ \ (\\NTDom\)" and "bnt_flip \ \ \\NTCod\ = bifunctor_flip \ \ (\\NTCod\)" and "bnt_flip \ \ \\NTDGDom\ = \ \\<^sub>C \" and "bnt_flip \ \ \\NTDGCod\ = \\NTDGCod\" unfolding bnt_flip_def nt_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ \' \ assumes \: "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule \) lemmas bnt_flip_components' = bnt_flip_components[where \=\ and \=\ and \=\, unfolded cat_cs_simps] lemmas [cat_cs_simps] = bnt_flip_components'(2-5) end subsubsection\Natural transformation map\ lemma bnt_flip_NTMap_vsv[cat_cs_intros]: "vsv (bnt_flip \ \ \\NTMap\)" unfolding bnt_flip_components by (rule fflip_vsv) lemma bnt_flip_NTMap_app: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "bnt_flip \ \ \\NTMap\\b, a\\<^sub>\ = \\NTMap\\a, b\\<^sub>\" using assms unfolding bnt_flip_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: cat_prod_cs_intros) lemma bnt_flip_NTMap_app'[cat_cs_simps]: assumes "ba = [b, a]\<^sub>\" and "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "bnt_flip \ \ \\NTMap\\ba\ = \\NTMap\\a, b\\<^sub>\" using assms(2-6) unfolding assms(1) by (rule bnt_flip_NTMap_app) lemma bnt_flip_NTMap_vdomain[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bnt_flip \ \ \\NTMap\) = (\ \\<^sub>C \)\Obj\" using assms unfolding bnt_flip_components by (cs_concl cs_simp: V_cs_simps cat_cs_simps) lemma bnt_flip_NTMap_vrange[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (bnt_flip \ \ \\NTMap\) = \\<^sub>\ (\\NTMap\)" proof- interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis proof(intro vsubset_antisym) show "\\<^sub>\ (bnt_flip \ \ \\NTMap\) \\<^sub>\ \\<^sub>\ (\\NTMap\)" proof ( intro vsv.vsv_vrange_vsubset, unfold bnt_flip_NTMap_vdomain[OF assms] ) fix ba assume "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(2,1)]) from \.ntcf_NTMap_vsv assms a b show "bnt_flip \ \ \\NTMap\\ba\ \\<^sub>\ \\<^sub>\ (\\NTMap\)" unfolding ba_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_prod_cs_intros ) qed (cs_concl cs_intro: cat_cs_intros) show "\\<^sub>\ (\\NTMap\) \\<^sub>\ \\<^sub>\ (bnt_flip \ \ \\NTMap\)" proof(intro vsv.vsv_vrange_vsubset, unfold \.ntcf_NTMap_vdomain) fix ab assume prems: "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[OF assms(1,2)]) from assms a b have ba: "[b, a]\<^sub>\ \\<^sub>\ (\ \\<^sub>C \)\Obj\" by (cs_concl cs_intro: cat_prod_cs_intros) from assms bnt_flip_NTMap_vsv prems a b ba show "\\NTMap\\ab\ \\<^sub>\ \\<^sub>\ (bnt_flip \ \ \\NTMap\)" unfolding ab_def by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) qed auto qed qed subsubsection\Binatural transformation flip natural transformation map\ lemma bnt_flip_NTMap_is_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bnt_flip \ \ \ : bifunctor_flip \ \ \ \\<^sub>C\<^sub>F bifunctor_flip \ \ \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) show ?thesis proof(intro is_ntcfI') show "vfsequence (bnt_flip \ \ \)" unfolding bnt_flip_def by simp show "vcard (bnt_flip \ \ \) = 5\<^sub>\" unfolding bnt_flip_def by (simp add: nat_omega_simps) show "bnt_flip \ \ \\NTMap\\ba\ : bifunctor_flip \ \ \\ObjMap\\ba\ \\<^bsub>\\<^esub> bifunctor_flip \ \ \'\ObjMap\\ba\" if "ba \\<^sub>\ (\ \\<^sub>C \)\Obj\" for ba proof- from that obtain b a where ba_def: "ba = [b, a]\<^sub>\" and b: "b \\<^sub>\ \\Obj\" and a: "a \\<^sub>\ \\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) (auto intro: cat_cs_intros) from assms a b show ?thesis by ( cs_concl cs_simp: cat_cs_simps ba_def cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "bnt_flip \ \ \\NTMap\\b'a'\ \\<^sub>A\<^bsub>\\<^esub> bifunctor_flip \ \ \\ArrMap\\gf\ = bifunctor_flip \ \ \'\ArrMap\\gf\ \\<^sub>A\<^bsub>\\<^esub> bnt_flip \ \ \\NTMap\\ba\" if "gf : ba \\<^bsub>\ \\<^sub>C \\<^esub> b'a'" for ba b'a' gf proof- from that obtain g f a b a' b' where gf_def: "gf = [g, f]\<^sub>\" and ba_def: "ba = [b, a]\<^sub>\" and b'a'_def: "b'a' = [b', a']\<^sub>\" and g: "g : b \\<^bsub>\\<^esub> b'" and f: "f : a \\<^bsub>\\<^esub> a'" by (elim cat_prod_2_is_arrE[OF assms(2,1)]) from assms g f show ?thesis unfolding gf_def ba_def b'a'_def by ( cs_concl cs_simp: cat_cs_simps cat_cs_simps \.ntcf_Comp_commute cs_intro: cat_cs_intros cat_prod_cs_intros ) qed qed (use assms in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\)+ qed lemma bnt_flip_NTMap_is_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = bifunctor_flip \ \ \" and "\' = bifunctor_flip \ \ \'" and "\
= \ \\<^sub>C \" shows "bnt_flip \ \ \ : \ \\<^sub>C\<^sub>F \' : \
\\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4-6) by (intro bnt_flip_NTMap_is_ntcf) subsubsection\Double-flip of a binatural transformation\ lemma bnt_flip_flip[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bnt_flip \ \ (bnt_flip \ \ \) = \" proof(rule ntcf_eqI) interpret \: category \ \ by (rule assms(1)) interpret \: category \ \ by (rule assms(2)) interpret \: is_ntcf \ \\ \\<^sub>C \\ \ \ \' \ by (rule assms(3)) from assms show "bnt_flip \ \ (bnt_flip \ \ \) : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) then have dom_lhs: "\\<^sub>\ (bnt_flip \ \ (bnt_flip \ \ \)\NTMap\) = (\ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule assms(3)) then have dom_rhs: "\\<^sub>\ (\\NTMap\) = (\ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "bnt_flip \ \ (bnt_flip \ \ \)\NTMap\ = \\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix ab assume "ab \\<^sub>\ (\ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by (rule cat_prod_2_ObjE[OF assms(1,2)]) from assms a b show "bnt_flip \ \ (bnt_flip \ \ \)\NTMap\\ab\ = \\NTMap\\ab\" by (cs_concl cs_simp: cat_cs_simps ab_def cs_intro: cat_cs_intros) qed (cs_concl cs_intro: V_cs_intros cat_cs_intros)+ qed simp_all subsubsection\A projection of a flip of a binatural transformation\ lemma bnt_flip_proj_snd[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "bnt_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F = \\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F" proof(rule ntcf_eqI) from assms show "bnt_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F bifunctor_flip \ \ \'\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) from assms show "\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F bifunctor_flip \ \ \'\<^bsub>\,\\<^esub>(b,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have dom_lhs: "\\<^sub>\ ((bnt_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) from assms have dom_rhs: "\\<^sub>\ ((\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps) show "(bnt_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\ = (\\<^bsub>\,\\<^esub>(-,b)\<^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 assms show "(bnt_flip \ \ \\<^bsub>\,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\ = (\\<^bsub>\,\\<^esub>(-,b)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps) qed (auto simp: cat_cs_intros) qed simp_all lemma bnt_flip_proj_fst[cat_cs_simps]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "bnt_flip \ \ \\<^bsub>\,\\<^esub>(-,a)\<^sub>N\<^sub>T\<^sub>C\<^sub>F = \\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F" proof- from assms have f_\: "bnt_flip \ \ \ : bifunctor_flip \ \ \ \\<^sub>C\<^sub>F bifunctor_flip \ \ \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) show ?thesis by ( rule bnt_flip_proj_snd [ OF assms(2,1) f_\ assms(4), unfolded bnt_flip_flip[OF assms(1,2,3)], symmetric ] ) qed subsubsection\A flip of a binatural isomorphism\ lemma bnt_flip_is_iso_ntcf: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "bnt_flip \ \ \ : bifunctor_flip \ \ \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o bifunctor_flip \ \ \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_iso_ntcf_if_bnt_proj_snd_is_iso_ntcf) from assms show f_\: "bnt_flip \ \ \ : bifunctor_flip \ \ \ \\<^sub>C\<^sub>F bifunctor_flip \ \ \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros ntcf_cs_intros) fix a assume "a \\<^sub>\ \\Obj\" with assms f_\ show "bnt_flip \ \ \\<^bsub>\,\\<^esub>(a,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : bifunctor_flip \ \ \\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o bifunctor_flip \ \ \'\<^bsub>\,\\<^esub>(a,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros ntcf_cs_intros) qed (simp_all add: assms) lemma bnt_flip_is_iso_ntcf'[cat_cs_intros]: assumes "category \ \" and "category \ \" and "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = bifunctor_flip \ \ \" and "\ = bifunctor_flip \ \ \'" and "\
= \ \\<^sub>C \" shows "bnt_flip \ \ \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \
\\\<^sub>C\<^bsub>\\<^esub> \" using bnt_flip_is_iso_ntcf[OF assms(1-3)] unfolding assms(4-6) by simp text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Simple.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Simple.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Simple.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Simple.thy @@ -1,336 +1,339 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Simple categories\ theory CZH_ECAT_Simple imports CZH_Foundations.CZH_SMC_Simple CZH_ECAT_Functor CZH_ECAT_Small_Functor begin subsection\Background\ text\ The section presents a variety of simple categories, (such as the empty category \0\ and the singleton category \1\) and functors between them (see \cite{mac_lane_categories_2010} for further information). \ subsection\Empty category \0\\ subsubsection\Definition and elementary properties\ text\See Chapter I-2 in \cite{mac_lane_categories_2010}.\ definition cat_0 :: "V" where "cat_0 = [0, 0, 0, 0, 0, 0]\<^sub>\" text\Components.\ lemma cat_0_components: shows "cat_0\Obj\ = 0" and "cat_0\Arr\ = 0" and "cat_0\Dom\ = 0" and "cat_0\Cod\ = 0" and "cat_0\Comp\ = 0" and "cat_0\CId\ = 0" unfolding cat_0_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smc_cat_0: "cat_smc cat_0 = smc_0" unfolding cat_smc_def cat_0_def smc_0_def dg_field_simps by (simp add: nat_omega_simps) lemmas_with (in \) [folded smc_cat_0, unfolded slicing_simps]: cat_0_is_arr_iff = smc_0_is_arr_iff subsubsection\\0\ is a category\ lemma (in \) category_cat_0: "category \ cat_0" proof(intro categoryI) show "vfsequence cat_0" "vcard cat_0 = 6\<^sub>\" by (simp_all add: cat_0_def nat_omega_simps) qed ( auto simp: cat_0_components \_axioms cat_0_is_arr_iff smc_cat_0 \.semicategory_smc_0 ) subsection\Empty functors\ subsubsection\Definition and elementary properties\ definition cf_0 :: "V \ V" where "cf_0 \ = [0, 0, cat_0, \]\<^sub>\" text\Components.\ lemma cf_0_components: shows "cf_0 \\ObjMap\ = 0" and "cf_0 \\ArrMap\ = 0" and "cf_0 \\HomDom\ = cat_0" and "cf_0 \\HomCod\ = \" unfolding cf_0_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cf_smcf_cf_0: "cf_smcf (cf_0 \) = smcf_0 (cat_smc \)" unfolding dg_field_simps dghm_field_simps cf_smcf_def cf_0_def smc_0_def cat_0_def smcf_0_def cat_smc_def by (simp add: nat_omega_simps) subsubsection\Empty functor is a faithful functor\ lemma (in \) cf_0_is_functor: assumes "category \ \" shows "cf_0 \ : cat_0 \\\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" proof(rule is_ft_functorI) show "cf_0 \ : cat_0 \\\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_functorI, unfold smc_cat_0 cf_smcf_cf_0) show "vfsequence (cf_0 \)" unfolding cf_0_def by simp show "vcard (cf_0 \) = 4\<^sub>\" unfolding cf_0_def by (simp add: nat_omega_simps) from \.smcf_0_is_semifunctor assms show "smcf_0 (cat_smc \) : smc_0 \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> cat_smc \" by auto qed (auto simp: assms category_cat_0 cat_0_components cf_0_components) show "cf_smcf (cf_0 \) : cat_smc cat_0 \\\<^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 ( auto simp: assms \_axioms \.smcf_0_is_semifunctor category.cat_semicategory cf_smcf_cf_0 smc_cat_0 ) qed subsection\\1\: category with one object and one arrow\ subsubsection\Definition and elementary properties\ text\See Chapter I-2 in \cite{mac_lane_categories_2010}.\ definition cat_1 :: "V \ V \ V" where "cat_1 \ \ = [ set {\}, set {\}, set {\\, \\}, set {\\, \\}, set {\[\, \]\<^sub>\, \\}, set {\\, \\} ]\<^sub>\" text\Components.\ lemma cat_1_components: shows "cat_1 \ \\Obj\ = set {\}" and "cat_1 \ \\Arr\ = set {\}" and "cat_1 \ \\Dom\ = set {\\, \\}" and "cat_1 \ \\Cod\ = set {\\, \\}" and "cat_1 \ \\Comp\ = set {\[\, \]\<^sub>\, \\}" and "cat_1 \ \\CId\ = set {\\, \\}" unfolding cat_1_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smc_cat_1: "cat_smc (cat_1 \ \) = smc_1 \ \" unfolding cat_smc_def cat_1_def smc_1_def dg_field_simps by (simp add: nat_omega_simps) lemmas_with (in \) [folded smc_cat_1, unfolded slicing_simps]: cat_1_is_arrI = smc_1_is_arrI and cat_1_is_arrD = smc_1_is_arrD and cat_1_is_arrE = smc_1_is_arrE and cat_1_is_arr_iff = smc_1_is_arr_iff and cat_1_Comp_app[cat_cs_simps] = smc_1_Comp_app subsubsection\Object\ lemma cat_1_ObjI[cat_cs_intros]: assumes "a = \" shows "a \\<^sub>\ cat_1 \ \ \Obj\" unfolding cat_1_components(1) assms by simp subsubsection\Identity\ lemma cat_1_CId_app: "cat_1 \ \\CId\\\\ = \" unfolding cat_1_components by simp subsubsection\Arrow with a domain and a codomain\ lemma cat_1_is_arrI: assumes "f = \" and "a = \" and "b = \" shows "f : a \\<^bsub>cat_1 \ \\<^esub> b" by (rule is_arrI, unfold assms cat_1_components) auto subsubsection\\1\ is a category\ lemma (in \) category_cat_1: assumes "\ \\<^sub>\ Vset \" and "\ \\<^sub>\ Vset \" shows "category \ (cat_1 \ \)" proof(intro categoryI, unfold smc_cat_1) show "vfsequence (cat_1 \ \)" unfolding cat_1_def by (simp add: nat_omega_simps) show "vcard (cat_1 \ \) = 6\<^sub>\" unfolding cat_1_def by (simp add: nat_omega_simps) qed (auto simp: assms semicategory_smc_1 cat_1_is_arr_iff cat_1_components) lemmas [cat_cs_intros] = \.category_cat_1 lemma (in \) finite_category_cat_1: assumes "\ \\<^sub>\ Vset \" and "\ \\<^sub>\ Vset \" shows "finite_category \ (cat_1 \ \)" by (intro finite_categoryI') (auto simp: cat_1_components intro: category_cat_1[OF assms]) lemmas [cat_small_cs_intros] = \.finite_category_cat_1 subsubsection\Opposite of the category \1\\ lemma (in \) cat_1_op[cat_op_simps]: assumes "\ \\<^sub>\ Vset \" and "\ \\<^sub>\ Vset \" shows "op_cat (cat_1 \ \) = cat_1 \ \" proof(rule cat_eqI, unfold cat_op_simps) from assms show "category \ (op_cat (cat_1 \ \))" by (cs_concl cs_intro: cat_cs_intros cat_op_intros) from assms show "category \ (cat_1 \ \)" by (cs_concl cs_intro: cat_cs_intros) show "op_cat (cat_1 \ \)\Comp\ = cat_1 \ \\Comp\" unfolding cat_1_components op_cat_components fflip_vsingleton .. qed (simp_all add: cat_1_components) +lemma (in \) cat_1_op_0[cat_op_simps]: "op_cat (cat_1 0 0) = cat_1 0 0" + by (cs_concl cs_simp: cat_op_simps cs_intro: V_cs_intros cat_cs_intros) + subsubsection\Further properties\ lemma cf_const_if_HomCod_is_cat_1: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_1 \ \" shows "\ = cf_const \ (cat_1 \ \) \" proof(rule cf_eqI) interpret \: is_functor \ \ \cat_1 \ \\ \ by (rule assms(1)) show "cf_const \ (cat_1 \ \) \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_1 \ \" by (cs_concl cs_intro: cat_cs_intros) have ObjMap_dom_lhs: "\\<^sub>\ (\\ObjMap\) = \\Obj\" by (simp add: cat_cs_simps) have ObjMap_dom_rhs: "\\<^sub>\ (cf_const \ (cat_1 \ \) \\ObjMap\) = \\Obj\" by (simp add: cat_cs_simps) have ArrMap_dom_lhs: "\\<^sub>\ (\\ArrMap\) = \\Arr\" by (simp add: cat_cs_simps) have ArrMap_dom_rhs: "\\<^sub>\ (cf_const \ (cat_1 \ \) \\ArrMap\) = \\Arr\" by (simp add: cat_cs_simps) show "\\ObjMap\ = cf_const \ (cat_1 \ \) \\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume prems: "a \\<^sub>\ \\Obj\" then have "\\ObjMap\\a\ \\<^sub>\ cat_1 \ \\Obj\" by (auto intro: \.cf_ObjMap_app_in_HomCod_Obj) then have "\\ObjMap\\a\ = \" by (auto simp: cat_1_components) with prems show "\\ObjMap\\a\ = cf_const \ (cat_1 \ \) \\ObjMap\\a\" by (auto simp: cat_cs_simps) qed (auto intro: cat_cs_intros) show "\\ArrMap\ = cf_const \ (cat_1 \ \) \\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix a assume prems: "a \\<^sub>\ \\Arr\" then have "\\ArrMap\\a\ \\<^sub>\ cat_1 \ \\Arr\" by (auto intro: \.cf_ArrMap_app_in_HomCod_Arr) then have "\\ArrMap\\a\ = \" by (auto simp: cat_1_components) with prems show "\\ArrMap\\a\ = cf_const \ (cat_1 \ \) \\ArrMap\\a\" by (auto simp: cat_1_CId_app cat_cs_simps) qed (auto intro: cat_cs_intros) qed (simp_all add: assms) lemma cf_const_if_HomDom_is_cat_1: assumes "\ : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)" proof- interpret \: is_functor \ \cat_1 \ \\ \ \ by (rule assms(1)) from cat_1_components(1) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.HomDom.cat_in_Obj_in_Vset) from cat_1_components(2) have \: "\ \\<^sub>\ Vset \" by (auto simp: \.HomDom.cat_in_Arr_in_Vset) from \ \ interpret cf_1: is_tm_functor \ \cat_1 \ \\ \ \cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ by (cs_concl cs_intro: cat_small_cs_intros cat_cs_intros) show ?thesis proof(rule cf_eqI) show "cf_const (cat_1 \ \) \ (\\ObjMap\\\\) : cat_1 \ \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) have ObjMap_dom_lhs: "\\<^sub>\ (\\ObjMap\) = set {\}" by (simp add: cat_cs_simps cat_1_components) have ObjMap_dom_rhs: "\\<^sub>\ (cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ObjMap\) = set {\}" by (simp add: cat_cs_simps cat_1_components) show "\\ObjMap\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ObjMap\" proof(rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) fix a assume "a \\<^sub>\ set {\}" then have a_def: "a = \" by simp show "\\ObjMap\\a\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ObjMap\\a\" by ( cs_concl cs_simp: cat_1_components(1) cat_cs_simps a_def cs_intro: V_cs_intros ) qed auto have ArrMap_dom_lhs: "\\<^sub>\ (\\ArrMap\) = set {\}" by (simp add: cat_cs_simps cat_1_components) have ArrMap_dom_rhs: "\\<^sub>\ (cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ArrMap\) = set {\}" by (simp add: cat_cs_simps cat_1_components) show "\\ArrMap\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ArrMap\" proof(rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) fix f assume "f \\<^sub>\ set {\}" then have f_def: "f = \" by simp show "\\ArrMap\\f\ = cf_const (cat_1 \ \) \ (\\ObjMap\\\\)\ArrMap\\f\" unfolding f_def by (subst cat_1_CId_app[symmetric, of \ \]) ( cs_concl cs_simp: cat_1_components(1,2) cat_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed auto qed (simp_all add: assms) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Yoneda.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Yoneda.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Yoneda.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Yoneda.thy @@ -1,5624 +1,5624 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Yoneda Lemma\ theory CZH_ECAT_Yoneda imports CZH_ECAT_FUNCT CZH_ECAT_Hom begin subsection\Yoneda map\ text\ The Yoneda map is the bijection that is used in the statement of the Yoneda Lemma, as presented, for example, in Chapter III-2 in \cite{mac_lane_categories_2010} or in subsection 1.15 in \cite{bodo_categories_1970}. \ definition Yoneda_map :: "V \ V \ V \ V" where "Yoneda_map \ \ r = ( \\\\<^sub>\these_ntcfs \ (\\HomDom\) (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomDom\(r,-) \. \\NTMap\\r\\ArrVal\\\\HomDom\\CId\\r\\ )" text\Elementary properties.\ mk_VLambda Yoneda_map_def |vsv Yoneda_map_vsv[cat_cs_intros]| mk_VLambda (in is_functor) Yoneda_map_def[where \=\ and \=\, unfolded cf_HomDom] |vdomain Yoneda_map_vdomain| |app Yoneda_map_app[unfolded these_ntcfs_iff]| lemmas [cat_cs_simps] = is_functor.Yoneda_map_vdomain lemmas Yoneda_map_app[cat_cs_simps] = is_functor.Yoneda_map_app[unfolded these_ntcfs_iff] subsection\Yoneda component\ subsubsection\Definition and elementary properties\ text\ The Yoneda components are the components of the natural transformations that appear in the statement of the Yoneda Lemma (e.g., see Chapter III-2 in \cite{mac_lane_categories_2010} or subsection 1.15 in \cite{bodo_categories_1970}). \ definition Yoneda_component :: "V \ V \ V \ V \ V" where "Yoneda_component \ r u d = [ (\f\\<^sub>\Hom (\\HomDom\) r d. \\ArrMap\\f\\ArrVal\\u\), Hom (\\HomDom\) r d, \\ObjMap\\d\ ]\<^sub>\" text\Components.\ lemma (in is_functor) Yoneda_component_components: shows "Yoneda_component \ r u d\ArrVal\ = (\f\\<^sub>\Hom \ r d. \\ArrMap\\f\\ArrVal\\u\)" and "Yoneda_component \ r u d\ArrDom\ = Hom \ r d" and "Yoneda_component \ r u d\ArrCod\ = \\ObjMap\\d\" unfolding Yoneda_component_def arr_field_simps by (simp_all add: nat_omega_simps cat_cs_simps) subsubsection\Arrow value\ mk_VLambda (in is_functor) Yoneda_component_components(1) |vsv Yoneda_component_ArrVal_vsv| |vdomain Yoneda_component_ArrVal_vdomain| |app Yoneda_component_ArrVal_app[unfolded in_Hom_iff]| lemmas [cat_cs_simps] = is_functor.Yoneda_component_ArrVal_vdomain lemmas Yoneda_component_ArrVal_app[cat_cs_simps] = is_functor.Yoneda_component_ArrVal_app[unfolded in_Hom_iff] subsubsection\Yoneda component is an arrow in the category \Set\\ lemma (in category) cat_Yoneda_component_is_arr: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" and "u \\<^sub>\ \\ObjMap\\r\" and "d \\<^sub>\ \\Obj\" shows "Yoneda_component \ r u d : Hom \ r d \\<^bsub>cat_Set \\<^esub> \\ObjMap\\d\" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(1)) show ?thesis proof(intro cat_Set_is_arrI arr_SetI, unfold \.Yoneda_component_components) show "vfsequence (Yoneda_component \ r u d)" unfolding Yoneda_component_def by simp show "vcard (Yoneda_component \ r u d) = 3\<^sub>\" unfolding Yoneda_component_def by (simp add: nat_omega_simps) show "\\<^sub>\ (\f\\<^sub>\Hom \ r d. \\ArrMap\\f\\ArrVal\\u\) \\<^sub>\ \\ObjMap\\d\" proof(rule vrange_VLambda_vsubset) fix f assume "f \\<^sub>\ Hom \ r d" then have \f: "\\ArrMap\\f\ : \\ObjMap\\r\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\d\" by (auto simp: cat_cs_intros) note \f_simps = cat_Set_is_arrD[OF \f] interpret \f: arr_Set \ \\\ArrMap\\f\\ by (rule \f_simps(1)) have "u \\<^sub>\ \\<^sub>\ (\\ArrMap\\f\\ArrVal\)" by (simp add: \f_simps assms cat_Set_cs_simps) with \f.arr_Set_ArrVal_vrange[unfolded \f_simps] show "\\ArrMap\\f\\ArrVal\\u\ \\<^sub>\ \\ObjMap\\d\" by (blast elim: \f.ArrVal.vsv_value) qed from assms \.HomCod.cat_Obj_vsubset_Vset show "\\ObjMap\\d\ \\<^sub>\ Vset \" by (auto dest: \.cf_ObjMap_app_in_HomCod_Obj) qed (auto simp: assms cat_cs_intros) qed lemma (in category) cat_Yoneda_component_is_arr': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" and "u \\<^sub>\ \\ObjMap\\r\" and "d \\<^sub>\ \\Obj\" and "s = Hom \ r d" and "t = \\ObjMap\\d\" and "\
= cat_Set \" shows "Yoneda_component \ r u d : s \\<^bsub>\
\<^esub> t" unfolding assms(5-7) using assms(1-4) by (rule cat_Yoneda_component_is_arr) lemmas [cat_cs_intros] = category.cat_Yoneda_component_is_arr'[rotated 1] subsection\Yoneda arrow\ subsubsection\Definition and elementary properties\ text\ The Yoneda arrows are the natural transformations that appear in the statement of the Yoneda Lemma in Chapter III-2 in \cite{mac_lane_categories_2010} and subsection 1.15 in \cite{bodo_categories_1970}. \ definition Yoneda_arrow :: "V \ V \ V \ V \ V" where "Yoneda_arrow \ \ r u = [ (\d\\<^sub>\\\HomDom\\Obj\. Yoneda_component \ r u d), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomDom\(r,-), \, \\HomDom\, cat_Set \ ]\<^sub>\" text\Components.\ lemma (in is_functor) Yoneda_arrow_components: shows "Yoneda_arrow \ \ r u\NTMap\ = (\d\\<^sub>\\\Obj\. Yoneda_component \ r u d)" and "Yoneda_arrow \ \ r u\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" and "Yoneda_arrow \ \ r u\NTCod\ = \" and "Yoneda_arrow \ \ r u\NTDGDom\ = \" and "Yoneda_arrow \ \ r u\NTDGCod\ = cat_Set \" unfolding Yoneda_arrow_def nt_field_simps by (simp_all add: nat_omega_simps cat_cs_simps) subsubsection\Natural transformation map\ mk_VLambda (in is_functor) Yoneda_arrow_components(1) |vsv Yoneda_arrow_NTMap_vsv| |vdomain Yoneda_arrow_NTMap_vdomain| |app Yoneda_arrow_NTMap_app| lemmas [cat_cs_simps] = is_functor.Yoneda_arrow_NTMap_vdomain lemmas Yoneda_arrow_NTMap_app[cat_cs_simps] = is_functor.Yoneda_arrow_NTMap_app subsubsection\Yoneda arrow is a natural transformation\ lemma (in category) cat_Yoneda_arrow_is_ntcf: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" and "u \\<^sub>\ \\ObjMap\\r\" shows "Yoneda_arrow \ \ r u : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(1)) note \ru = cat_Yoneda_component_is_arr[OF assms] let ?\ru = \Yoneda_component \ r u\ show ?thesis proof(intro is_ntcfI', unfold \.Yoneda_arrow_components) show "vfsequence (Yoneda_arrow \ \ r u)" unfolding Yoneda_arrow_def by simp show "vcard (Yoneda_arrow \ \ r u) = 5\<^sub>\" unfolding Yoneda_arrow_def by (simp add: nat_omega_simps) show "(\d\\<^sub>\\\Obj\. ?\ru d)\a\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\a\" if "a \\<^sub>\ \\Obj\" for a using that assms category_axioms by ( cs_concl cs_simp: cat_cs_simps cat_op_simps V_cs_simps cs_intro: cat_cs_intros ) show "(\d\\<^sub>\\\Obj\. ?\ru d)\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> (\d\\<^sub>\\\Obj\. ?\ru d)\a\" if "f : a \\<^bsub>\\<^esub> b" for a b f proof- note \a = \ru[OF cat_is_arrD(2)[OF that]] note \b = \ru[OF cat_is_arrD(3)[OF that]] from category_axioms assms that \b have b_f: "?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\ : Hom \ r a \\<^bsub>cat_Set \\<^esub> \\ObjMap\\b\" by ( cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) then have dom_lhs: "\\<^sub>\ ((?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\)\ArrVal\) = Hom \ r a" by (cs_concl cs_simp: cat_cs_simps) from assms that \a have f_a: "\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a : Hom \ r a \\<^bsub>cat_Set \\<^esub> \\ObjMap\\b\" by (cs_concl cs_intro: cat_cs_intros) then have dom_rhs: "\\<^sub>\ ((\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a)\ArrVal\) = Hom \ r a" by (cs_concl cs_simp: cat_cs_simps) have [cat_cs_simps]: "?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a" proof(rule arr_Set_eqI[of \]) from b_f show arr_Set_b_f: "arr_Set \ (?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\)" by (auto simp: cat_Set_is_arrD(1)) interpret b_f: arr_Set \ \?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\\ by (rule arr_Set_b_f) from f_a show arr_Set_f_a: "arr_Set \ (\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a)" by (auto simp: cat_Set_is_arrD(1)) interpret f_a: arr_Set \ \\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a\ by (rule arr_Set_f_a) show "(?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\)\ArrVal\ = (\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix q assume "q : r \\<^bsub>\\<^esub> a" from category_axioms assms that this \a \b show "(?\ru b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\CId\\r\, f]\<^sub>\)\ArrVal\\q\ = (\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ru a)\ArrVal\\q\" by ( 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_f arr_Set_f_a in auto) qed (use b_f f_a in \cs_concl cs_simp: cat_cs_simps\)+ from that category_axioms assms \a \b show ?thesis by ( cs_concl cs_simp: V_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_cs_intros ) qed qed (auto simp: assms(2) cat_cs_intros) qed subsection\Yoneda Lemma\ text\ The following lemma is approximately equivalent to the Yoneda Lemma stated in subsection 1.15 in \cite{bodo_categories_1970} (the first two conclusions correspond to the statement of the Yoneda lemma in Chapter III-2 in \cite{mac_lane_categories_2010}). \ lemma (in category) cat_Yoneda_Lemma: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "v11 (Yoneda_map \ \ r)" and "\\<^sub>\ (Yoneda_map \ \ r) = \\ObjMap\\r\" and "(Yoneda_map \ \ r)\\<^sub>\ = (\u\\<^sub>\\\ObjMap\\r\. Yoneda_arrow \ \ r u)" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(1)) from assms(2) \.HomCod.cat_Obj_vsubset_Vset \.cf_ObjMap_app_in_HomCod_Obj have \r_in_Vset: "\\ObjMap\\r\ \\<^sub>\ Vset \" by auto show Ym: "v11 (Yoneda_map \ \ r)" proof(intro vsv.vsv_valeq_v11I, unfold \.Yoneda_map_vdomain these_ntcfs_iff) fix \ \ assume prems: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" "Yoneda_map \ \ r\\\ = Yoneda_map \ \ r\\\" from prems(3) have \r_\r: "\\NTMap\\r\\ArrVal\\\\CId\\r\\ = \\NTMap\\r\\ArrVal\\\\CId\\r\\" unfolding Yoneda_map_app[OF assms(1) prems(1)] Yoneda_map_app[OF assms(1) prems(2)] by simp interpret \: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \ \ by (rule prems(1)) interpret \: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \ \ by (rule prems(2)) show "\ = \" proof ( rule ntcf_eqI[OF prems(1,2)]; (rule refl)?; rule vsv_eqI, unfold \.ntcf_NTMap_vdomain \.ntcf_NTMap_vdomain ) fix d assume prems': "d \\<^sub>\ \\Obj\" note \d_simps = cat_Set_is_arrD[OF \.ntcf_NTMap_is_arr[OF prems']] interpret \d: arr_Set \ \\\NTMap\\d\\ by (rule \d_simps(1)) note \d_simps = cat_Set_is_arrD[OF \.ntcf_NTMap_is_arr[OF prems']] interpret \d: arr_Set \ \\\NTMap\\d\\ by (rule \d_simps(1)) show "\\NTMap\\d\ = \\NTMap\\d\" proof(rule arr_Set_eqI[of \]) show "\\NTMap\\d\\ArrVal\ = \\NTMap\\d\\ArrVal\" proof ( rule vsv_eqI, unfold \d.arr_Set_ArrVal_vdomain \d.arr_Set_ArrVal_vdomain \d_simps \d_simps ) fix f assume prems'': "f \\<^sub>\ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ObjMap\\d\" from prems'' prems' category_axioms assms(2) have f: "f : r \\<^bsub>\\<^esub> d" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_op_intros) from \.ntcf_Comp_commute[OF f] have "( \\NTMap\\d\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f\ )\ArrVal\\\\CId\\r\\ = (\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\)\ArrVal\\\\CId\\r\\" by simp from this category_axioms assms(2) f prems prems' have \df: "\\NTMap\\d\\ArrVal\\f\ = \\ArrMap\\f\\ArrVal\\\\NTMap\\r\\ArrVal\\\\CId\\r\\\" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from \.ntcf_Comp_commute[OF f] have "( \\NTMap\\d\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ArrMap\\f\ )\ArrVal\\\\CId\\r\\ = (\\ArrMap\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\r\)\ArrVal\\\\CId\\r\\" by simp from this category_axioms assms(2) f prems prems' have \df: "\\NTMap\\d\\ArrVal\\f\ = \\ArrMap\\f\\ArrVal\\\\NTMap\\r\\ArrVal\\\\CId\\r\\\" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) show "\\NTMap\\d\\ArrVal\\f\ = \\NTMap\\d\\ArrVal\\f\" unfolding \df \df \r_\r by simp qed auto qed (simp_all add: \d_simps \d_simps) qed auto qed (auto simp: Yoneda_map_vsv) interpret Ym: v11 \Yoneda_map \ \ r\ by (rule Ym) have YY: "Yoneda_map \ \ r\Yoneda_arrow \ \ r a\ = a" if "a \\<^sub>\ \\ObjMap\\r\" for a proof- note cat_Yoneda_arrow_is_ntcf[OF assms that] moreover with assms have Ya: "Yoneda_arrow \ \ r a \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" by (cs_concl cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros) ultimately show "Yoneda_map \ \ r\Yoneda_arrow \ \ r a\ = a" using assms that \r_in_Vset by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed show [simp]: "\\<^sub>\ (Yoneda_map \ \ r) = \\ObjMap\\r\" proof(intro vsubset_antisym) show "\\<^sub>\ (Yoneda_map \ \ r) \\<^sub>\ \\ObjMap\\r\" unfolding Yoneda_map_def proof(intro vrange_VLambda_vsubset, unfold these_ntcfs_iff \.cf_HomDom) fix \ assume prems: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" then interpret \: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \ \ . note \r_simps = cat_Set_is_arrD[OF \.ntcf_NTMap_is_arr[OF assms(2)]] interpret \r: arr_Set \ \\\NTMap\\r\\ by (rule \r_simps(1)) from prems category_axioms assms(2) have "\\CId\\r\ \\<^sub>\ \\<^sub>\ (\\NTMap\\r\\ArrVal\)" unfolding \r.arr_Set_ArrVal_vdomain \r_simps by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) then have "\\NTMap\\r\\ArrVal\\\\CId\\r\\ \\<^sub>\ \\<^sub>\ (\\NTMap\\r\\ArrVal\)" by (blast elim: \r.ArrVal.vsv_value) then show "\\NTMap\\r\\ArrVal\\\\CId\\r\\ \\<^sub>\ \\ObjMap\\r\" by (auto simp: \r_simps dest!: vsubsetD[OF \r.arr_Set_ArrVal_vrange]) qed show "\\ObjMap\\r\ \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" proof(intro vsubsetI) fix u assume prems: "u \\<^sub>\ \\ObjMap\\r\" from cat_Yoneda_arrow_is_ntcf[OF assms prems] have "Yoneda_arrow \ \ r u \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" by (cs_concl cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros) with YY[OF prems] show "u \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" by (force dest!: vdomain_atD) qed qed show "(Yoneda_map \ \ r)\\<^sub>\ = (\u\\<^sub>\\\ObjMap\\r\. Yoneda_arrow \ \ r u)" proof(rule vsv_eqI, unfold vdomain_vconverse vdomain_VLambda) from Ym show "vsv ((Yoneda_map \ \ r)\\<^sub>\)" by auto show "(Yoneda_map \ \ r)\\<^sub>\\a\ = (\u\\<^sub>\\\ObjMap\\r\. Yoneda_arrow \ \ r u)\a\" if "a \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" for a proof- from that have a: "a \\<^sub>\ \\ObjMap\\r\" by simp note Ya = cat_Yoneda_arrow_is_ntcf[OF assms a] then have "Yoneda_arrow \ \ r a \\<^sub>\ \\<^sub>\ (Yoneda_map \ \ r)" by ( cs_concl cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros ) with Ya YY[OF a] a show ?thesis by ( intro Ym.v11_vconverse_app[ unfolded \.Yoneda_map_vdomain these_ntcfs_iff ] ) (simp_all add: these_ntcfs_iff cat_cs_simps) qed qed auto qed subsection\Inverse of the Yoneda map\ lemma (in category) inv_Yoneda_map_v11: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "v11 ((Yoneda_map \ \ r)\\<^sub>\)" using cat_Yoneda_Lemma(1)[OF assms] by (simp add: v11.v11_vconverse) lemma (in category) inv_Yoneda_map_vdomain: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((Yoneda_map \ \ r)\\<^sub>\) = \\ObjMap\\r\" unfolding cat_Yoneda_Lemma(3)[OF assms] by simp lemmas [cat_cs_simps] = category.inv_Yoneda_map_vdomain lemma (in category) inv_Yoneda_map_app: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" and "u \\<^sub>\ \\ObjMap\\r\" shows "(Yoneda_map \ \ r)\\<^sub>\\u\ = Yoneda_arrow \ \ r u" using assms(3) unfolding cat_Yoneda_Lemma(3)[OF assms(1,2)] by simp lemmas [cat_cs_simps] = category.inv_Yoneda_map_app lemma (in category) inv_Yoneda_map_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "\\<^sub>\ ((Yoneda_map \ \ r)\\<^sub>\) = these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(1)) show ?thesis unfolding Yoneda_map_def by (simp add: cat_cs_simps) qed subsection\ Component of a composition of a \Hom\-natural transformation with natural transformations \ subsubsection\Definition and elementary properties\ text\ The following definition is merely a technical generalization that is used in the context of the description of the composition of a \Hom\-natural transformation with a natural transformation later in this section (also see subsection 1.15 in \cite{bodo_categories_1970}). \ definition ntcf_Hom_component :: "V \ V \ V \ V \ V" where "ntcf_Hom_component \ \ a b = [ ( \f\\<^sub>\Hom (\\NTDGCod\) (\\NTCod\\ObjMap\\a\) (\\NTDom\\ObjMap\\b\). \\NTMap\\b\ \\<^sub>A\<^bsub>\\NTDGCod\\<^esub> f \\<^sub>A\<^bsub>\\NTDGCod\\<^esub> \\NTMap\\a\ ), Hom (\\NTDGCod\) (\\NTCod\\ObjMap\\a\) (\\NTDom\\ObjMap\\b\), Hom (\\NTDGCod\) (\\NTDom\\ObjMap\\a\) (\\NTCod\\ObjMap\\b\) ]\<^sub>\" text\Components.\ lemma ntcf_Hom_component_components: shows "ntcf_Hom_component \ \ a b\ArrVal\ = ( \f\\<^sub>\Hom (\\NTDGCod\) (\\NTCod\\ObjMap\\a\) (\\NTDom\\ObjMap\\b\). \\NTMap\\b\ \\<^sub>A\<^bsub>\\NTDGCod\\<^esub> f \\<^sub>A\<^bsub>\\NTDGCod\\<^esub> \\NTMap\\a\ )" and "ntcf_Hom_component \ \ a b\ArrDom\ = Hom (\\NTDGCod\) (\\NTCod\\ObjMap\\a\) (\\NTDom\\ObjMap\\b\)" and "ntcf_Hom_component \ \ a b\ArrCod\ = Hom (\\NTDGCod\) (\\NTDom\\ObjMap\\a\) (\\NTCod\\ObjMap\\b\)" unfolding ntcf_Hom_component_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow value\ mk_VLambda ntcf_Hom_component_components(1) |vsv ntcf_Hom_component_ArrVal_vsv[intro]| context fixes \ \ \ \ \ \' \' \ \ \ assumes \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_ntcf \ \ \ \ \ \ by (rule \) interpretation \: is_ntcf \ \ \ \' \' \ by (rule \) mk_VLambda ntcf_Hom_component_components(1) [ of \ \, unfolded \.ntcf_NTDom \.ntcf_NTDom \.ntcf_NTCod \.ntcf_NTCod \.ntcf_NTDGDom \.ntcf_NTDGDom \.ntcf_NTDGCod \.ntcf_NTDGCod ] |vdomain ntcf_Hom_component_ArrVal_vdomain| |app ntcf_Hom_component_ArrVal_app[unfolded in_Hom_iff]| lemmas [cat_cs_simps] = ntcf_Hom_component_ArrVal_vdomain ntcf_Hom_component_ArrVal_app lemma ntcf_Hom_component_ArrVal_vrange: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (ntcf_Hom_component \ \ a b\ArrVal\) \\<^sub>\ Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" proof ( rule vsv.vsv_vrange_vsubset, unfold ntcf_Hom_component_ArrVal_vdomain in_Hom_iff ) fix f assume "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\b\" with assms \ \ show "ntcf_Hom_component \ \ a b\ArrVal\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\b\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (rule ntcf_Hom_component_ArrVal_vsv) end subsubsection\Arrow domain and codomain\ context fixes \ \ \ \ \ \' \' \ \ \ assumes \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_ntcf \ \ \ \ \ \ by (rule \) interpretation \: is_ntcf \ \ \ \' \' \ by (rule \) lemma ntcf_Hom_component_ArrDom[cat_cs_simps]: "ntcf_Hom_component \ \ a b\ArrDom\ = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" unfolding ntcf_Hom_component_components by (simp add: cat_cs_simps) lemma ntcf_Hom_component_ArrCod[cat_cs_simps]: "ntcf_Hom_component \ \ a b\ArrCod\ = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" unfolding ntcf_Hom_component_components by (simp add: cat_cs_simps) end subsubsection\ Component of a composition of a \Hom\-natural transformation with natural transformations is an arrow in the category \Set\ \ lemma (in category) cat_ntcf_Hom_component_is_arr: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "ntcf_Hom_component \ \ a b : Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(2)) from assms have a: "a \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp show ?thesis proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (ntcf_Hom_component \ \ a b)" unfolding ntcf_Hom_component_def by (simp add: nat_omega_simps) show "vcard (ntcf_Hom_component \ \ a b) = 3\<^sub>\" unfolding ntcf_Hom_component_def by (simp add: nat_omega_simps) from assms ntcf_Hom_component_ArrVal_vrange[OF assms(1,2) a assms(4)] show "\\<^sub>\ (ntcf_Hom_component \ \ a b\ArrVal\) \\<^sub>\ ntcf_Hom_component \ \ a b\ArrCod\" by (cs_concl cs_simp: cat_cs_simps) from assms(1,2,4) a show "ntcf_Hom_component \ \ a b\ArrDom\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(1,2,4) a show "ntcf_Hom_component \ \ a b\ArrCod\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (use assms in \auto simp: ntcf_Hom_component_components cat_cs_simps\) qed lemma (in category) cat_ntcf_Hom_component_is_arr': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" and "\' = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" and "\' = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" and "\' = cat_Set \" shows "ntcf_Hom_component \ \ a b : \' \\<^bsub>\'\<^esub> \'" using assms(1-4) unfolding assms(5-7) by (rule cat_ntcf_Hom_component_is_arr) lemmas [cat_cs_intros] = category.cat_ntcf_Hom_component_is_arr' subsubsection\ Naturality of the components of a composition of a \Hom\-natural transformation with natural transformations \ lemma (in category) cat_ntcf_Hom_component_nat: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "g : a \\<^bsub>op_cat \\<^esub> a'" and "f : b \\<^bsub>\\<^esub> b'" shows "ntcf_Hom_component \ \ a' b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [\\ArrMap\\g\, \'\ArrMap\\f\]\<^sub>\ = cf_hom \ [\\ArrMap\\g\, \'\ArrMap\\f\]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Hom_component \ \ a b" proof- let ?Y_ab = \ntcf_Hom_component \ \ a b\ and ?Y_a'b' = \ntcf_Hom_component \ \ a' b'\ and ?\g = \\\ArrMap\\g\\ and ?\'f = \\'\ArrMap\\f\\ and ?\g = \\\ArrMap\\g\\ and ?\'f = \\'\ArrMap\\f\\ and ?\a = \\\ObjMap\\a\\ and ?\'b = \\'\ObjMap\\b\\ and ?\a' = \\\ObjMap\\a'\\ and ?\'b' = \\'\ObjMap\\b'\\ interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(2)) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) from assms(3) have g: "g : a' \\<^bsub>\\<^esub> a" unfolding cat_op_simps by simp from Set.category_axioms category_axioms assms g have a'b_Gg\'f: "?Y_a'b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [?\g, ?\'f]\<^sub>\ : Hom \ ?\a ?\'b \\<^bsub>cat_Set \\<^esub> Hom \ ?\a' ?\'b'" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) then have dom_lhs: "\\<^sub>\ ((?Y_a'b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [?\g, ?\'f]\<^sub>\)\ArrVal\) = Hom \ ?\a ?\'b" by (cs_concl cs_simp: cat_cs_simps) from Set.category_axioms category_axioms assms g have \g\'f_ab: "cf_hom \ [?\g, ?\'f]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Y_ab : Hom \ ?\a ?\'b \\<^bsub>cat_Set \\<^esub> Hom \ ?\a' ?\'b'" by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) then have dom_rhs: "\\<^sub>\ ((cf_hom \ [?\g, ?\'f]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Y_ab)\ArrVal\) = Hom \ ?\a ?\'b" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from a'b_Gg\'f show arr_Set_a'b_Gg\'f: "arr_Set \ (?Y_a'b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [?\g, ?\'f]\<^sub>\)" by (auto dest: cat_Set_is_arrD(1)) from \g\'f_ab show arr_Set_\g\'f_ab: "arr_Set \ (cf_hom \ [?\g, ?\'f]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Y_ab)" by (auto dest: cat_Set_is_arrD(1)) show "(?Y_a'b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [?\g, ?\'f]\<^sub>\)\ArrVal\ = (cf_hom \ [?\g, ?\'f]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Y_ab)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix h assume prems: "h : \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\b\" from assms(1,2) g have [cat_cs_simps]: "\\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> (?\'f \\<^sub>A\<^bsub>\\<^esub> (h \\<^sub>A\<^bsub>\\<^esub> (?\g \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a'\))) = \\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> (?\'f \\<^sub>A\<^bsub>\\<^esub> (h \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> ?\g)))" by (cs_concl cs_simp: is_ntcf.ntcf_Comp_commute cs_intro: cat_cs_intros) also from assms(1,2,4) prems g have "\ = (((\\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> ?\'f) \\<^sub>A\<^bsub>\\<^esub> h) \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\) \\<^sub>A\<^bsub>\\<^esub> ?\g" by (cs_concl cs_simp: cat_Comp_assoc cs_intro: cat_cs_intros) (*slow*) also from assms(1,2,4) have "\ = (((?\'f \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\b\) \\<^sub>A\<^bsub>\\<^esub> h) \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\) \\<^sub>A\<^bsub>\\<^esub> ?\g" by (cs_concl cs_simp: is_ntcf.ntcf_Comp_commute cs_intro: cat_cs_intros) also from assms(1,2,4) prems g have "\ = ?\'f \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> (h \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> ?\g)))" by (cs_concl cs_simp: cat_Comp_assoc cs_intro: cat_cs_intros) (*slow*) finally have nat: "\\NTMap\\b'\ \\<^sub>A\<^bsub>\\<^esub> (?\'f \\<^sub>A\<^bsub>\\<^esub> (h \\<^sub>A\<^bsub>\\<^esub> (?\g \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a'\))) = ?\'f \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> (h \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a\ \\<^sub>A\<^bsub>\\<^esub> ?\g)))". from prems Set.category_axioms category_axioms assms(1,2,4) g show "(?Y_a'b' \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom \ [?\g, ?\'f]\<^sub>\)\ArrVal\\h\ = (cf_hom \ [?\g, ?\'f]\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Y_ab)\ArrVal\\h\" by (*slow*) ( cs_concl cs_simp: nat cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (use arr_Set_a'b_Gg\'f arr_Set_\g\'f_ab in auto) qed (use a'b_Gg\'f \g\'f_ab in \cs_concl cs_simp: cat_cs_simps\)+ qed subsubsection\ Composition of the components of a composition of a \Hom\-natural transformation with natural transformations \ lemma (in category) cat_ntcf_Hom_component_Comp: assumes "\' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "ntcf_Hom_component \ \' a b \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Hom_component \' \ a b = ntcf_Hom_component (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) (\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \) a b" (is \?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\ = ?\'\\'\\) proof- interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) from assms Set.category_axioms category_axioms have \\'_\'\: "?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\ : Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_intro: cat_cs_intros cat_op_intros) then have dom_lhs: "\\<^sub>\ ((?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\)\ArrVal\) = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) from assms Set.category_axioms category_axioms have \'\\'\: "?\'\\'\ : Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_intro: cat_cs_intros cat_op_intros) then have dom_rhs: "\\<^sub>\ (?\'\\'\\ArrVal\) = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from \\'_\'\ show arr_Set_\\'_\'\: "arr_Set \ (?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\)" by (auto dest: cat_Set_is_arrD(1)) from \'\\'\ show arr_Set_\'\\'\: "arr_Set \ ?\'\\'\" by (auto dest: cat_Set_is_arrD(1)) show "(?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\)\ArrVal\ = ?\'\\'\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix f assume "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\b\" with category_axioms assms Set.category_axioms show "(?\\' \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\'\)\ArrVal\\f\ = ?\'\\'\\ArrVal\\f\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (use arr_Set_\'\\'\ arr_Set_\\'_\'\ in auto) qed (use \\'_\'\ \'\\'\ in \cs_concl cs_simp: cat_cs_simps\)+ qed lemmas [cat_cs_simps] = category.cat_ntcf_Hom_component_Comp subsubsection\ Component of a composition of \Hom\-natural transformation with the identity natural transformations \ lemma (in category) cat_ntcf_Hom_component_ntcf_id: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\': \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "ntcf_Hom_component (ntcf_id \) (ntcf_id \') a b = cat_Set \\CId\\Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)\" (is \?\\' = cat_Set \\CId\\?\a\'b\\) proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \': is_functor \ \ \ \' by (rule assms(2)) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) from assms Set.category_axioms category_axioms have \\': "?\\' : Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_intro: cat_cs_intros cat_op_intros) then have dom_lhs: "\\<^sub>\ (?\\'\ArrVal\) = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) from category_axioms assms Set.category_axioms have \a\'b: "cat_Set \\CId\\?\a\'b\ : Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by ( cs_concl cs_full cs_simp: cat_Set_cs_simps cat_Set_components(1) cs_intro: cat_cs_intros ) then have dom_rhs: "\\<^sub>\ (cat_Set \\CId\\?\a\'b\\ArrVal\) = Hom \ (\\ObjMap\\a\) (\'\ObjMap\\b\)" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from \\' show arr_Set_\\: "arr_Set \ ?\\'" by (auto dest: cat_Set_is_arrD(1)) from \a\'b show arr_Set_\a\'b: "arr_Set \ (cat_Set \\CId\\?\a\'b\)" by (auto dest: cat_Set_is_arrD(1)) show "?\\'\ArrVal\ = cat_Set \\CId\\?\a\'b\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix f assume "f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \'\ObjMap\\b\" with category_axioms Set.category_axioms assms show "?\\'\ArrVal\\f\ = cat_Set \\CId\\?\a\'b\\ArrVal\\f\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (use arr_Set_\a\'b in auto) qed (use \\' \a\'b in \cs_concl cs_simp: cat_cs_simps\)+ qed lemmas [cat_cs_simps] = category.cat_ntcf_Hom_component_ntcf_id subsection\ Component of a composition of a \Hom\-natural transformation with a natural transformation \ subsubsection\Definition and elementary properties\ definition ntcf_lcomp_Hom_component :: "V \ V \ V \ V" where "ntcf_lcomp_Hom_component \ a b = ntcf_Hom_component \ (ntcf_id (cf_id (\\NTDGCod\))) a b" definition ntcf_rcomp_Hom_component :: "V \ V \ V \ V" where "ntcf_rcomp_Hom_component \ a b = ntcf_Hom_component (ntcf_id (cf_id (\\NTDGCod\))) \ a b" subsubsection\Arrow value\ lemma ntcf_lcomp_Hom_component_ArrVal_vsv: "vsv (ntcf_lcomp_Hom_component \ a b\ArrVal\)" unfolding ntcf_lcomp_Hom_component_def by (rule ntcf_Hom_component_ArrVal_vsv) lemma ntcf_rcomp_Hom_component_ArrVal_vsv: "vsv (ntcf_rcomp_Hom_component \ a b\ArrVal\)" unfolding ntcf_rcomp_Hom_component_def by (rule ntcf_Hom_component_ArrVal_vsv) lemma ntcf_lcomp_Hom_component_ArrVal_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (ntcf_lcomp_Hom_component \ a b\ArrVal\) = Hom \ (\\ObjMap\\a\) b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis using assms unfolding ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_rcomp_Hom_component_ArrVal_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" shows "\\<^sub>\ (ntcf_rcomp_Hom_component \ a b\ArrVal\) = Hom \ a (\\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis using assms unfolding cat_op_simps ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_lcomp_Hom_component_ArrVal_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" and "h : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" shows "ntcf_lcomp_Hom_component \ a b\ArrVal\\h\ = h \\<^sub>A\<^bsub>\\<^esub> \\NTMap\\a\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis using assms unfolding cat_op_simps ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_rcomp_Hom_component_ArrVal_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" and "h : a \\<^bsub>\\<^esub> \\ObjMap\\b\" shows "ntcf_rcomp_Hom_component \ a b\ArrVal\\h\ = \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> h" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis using assms unfolding cat_op_simps ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_lcomp_Hom_component_ArrVal_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (ntcf_lcomp_Hom_component \ a b\ArrVal\) \\<^sub>\ Hom \ (\\ObjMap\\a\) b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms(2) have a: "a \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp from assms(1,3) a have "\\<^sub>\ (ntcf_lcomp_Hom_component \ a b\ArrVal\) \\<^sub>\ Hom \ (\\ObjMap\\a\) (cf_id \\ObjMap\\b\)" by ( unfold cat_op_simps ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod, intro ntcf_Hom_component_ArrVal_vrange ) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ from this assms(3) show ?thesis by (cs_prems cs_simp: cat_cs_simps) qed lemma ntcf_rcomp_Hom_component_ArrVal_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ (ntcf_rcomp_Hom_component \ a b\ArrVal\) \\<^sub>\ Hom \ a (\\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms(2) have a: "a \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp from assms(1,3) a have "\\<^sub>\ (ntcf_rcomp_Hom_component \ a b\ArrVal\) \\<^sub>\ Hom \ (cf_id \\ObjMap\\a\) (\\ObjMap\\b\)" by ( unfold ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod, intro ntcf_Hom_component_ArrVal_vrange ) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from this a show ?thesis by (cs_prems cs_simp: cat_cs_simps) qed subsubsection\Arrow domain and codomain\ lemma ntcf_lcomp_Hom_component_ArrDom[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "ntcf_lcomp_Hom_component \ a b\ArrDom\ = Hom \ (\\ObjMap\\a\) b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_rcomp_Hom_component_ArrDom[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" shows "ntcf_rcomp_Hom_component \ a b\ArrDom\ = Hom \ a (\\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding cat_op_simps ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_lcomp_Hom_component_ArrCod[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "ntcf_lcomp_Hom_component \ a b\ArrCod\ = Hom \ (\\ObjMap\\a\) b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_rcomp_Hom_component_ArrCod[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" shows "ntcf_rcomp_Hom_component \ a b\ArrCod\ = Hom \ a (\\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding cat_op_simps ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed subsubsection\ Component of a composition of a \Hom\-natural transformation with a natural transformation is an arrow in the category \Set\ \ lemma (in category) cat_ntcf_lcomp_Hom_component_is_arr: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "ntcf_lcomp_Hom_component \ a b : Hom \ (\\ObjMap\\a\) b \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms have a: "a \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp from assms(1,3) a have "ntcf_lcomp_Hom_component \ a b : Hom \ (\\ObjMap\\a\) (cf_id \\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) (cf_id \\ObjMap\\b\)" unfolding ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (intro cat_ntcf_Hom_component_is_arr) (cs_concl cs_intro: cat_cs_intros cat_op_intros)+ from this assms(1,3) a show ?thesis by (cs_prems cs_simp: cat_cs_simps) qed lemma (in category) cat_ntcf_lcomp_Hom_component_is_arr': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" and "\' = Hom \ (\\ObjMap\\a\) b" and "\' = Hom \ (\\ObjMap\\a\) b" and "\' = cat_Set \" shows "ntcf_lcomp_Hom_component \ a b : \' \\<^bsub>\'\<^esub> \'" using assms(1-3) unfolding assms(4-6) by (rule cat_ntcf_lcomp_Hom_component_is_arr) lemmas [cat_cs_intros] = category.cat_ntcf_lcomp_Hom_component_is_arr' lemma (in category) cat_ntcf_rcomp_Hom_component_is_arr: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "ntcf_rcomp_Hom_component \ a b : Hom \ a (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ a (\\ObjMap\\b\)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms have a: "a \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp from assms(1,3) a have "ntcf_rcomp_Hom_component \ a b : Hom \ (cf_id \\ObjMap\\a\) (\\ObjMap\\b\) \\<^bsub>cat_Set \\<^esub> Hom \ (cf_id \\ObjMap\\a\) (\\ObjMap\\b\)" unfolding ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (intro cat_ntcf_Hom_component_is_arr) (cs_concl cs_intro: cat_cs_intros cat_op_intros) from this assms(1,3) a show ?thesis by (cs_prems cs_simp: cat_cs_simps) qed lemma (in category) cat_ntcf_rcomp_Hom_component_is_arr': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" and "\' = Hom \ a (\\ObjMap\\b\)" and "\' = Hom \ a (\\ObjMap\\b\)" and "\' = cat_Set \" shows "ntcf_rcomp_Hom_component \ a b : \' \\<^bsub>\'\<^esub> \'" using assms(1-3) unfolding assms(4-6) by (rule cat_ntcf_rcomp_Hom_component_is_arr) lemmas [cat_cs_intros] = category.cat_ntcf_rcomp_Hom_component_is_arr' subsection\ Composition of a \Hom\-natural transformation with two natural transformations \ subsubsection\Definition and elementary properties\ text\See subsection 1.15 in \cite{bodo_categories_1970}.\ definition ntcf_Hom :: "V \ V \ V \ V" (\Hom\<^sub>A\<^sub>.\<^sub>C\'(/_-,_-/')\) where "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-) = [ ( \ab\\<^sub>\(op_cat (\\NTDGDom\) \\<^sub>C \\NTDGDom\)\Obj\. ntcf_Hom_component \ \ (vpfst ab) (vpsnd ab) ), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\NTDGCod\(\\NTCod\-,\\NTDom\-), Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\NTDGCod\(\\NTDom\-,\\NTCod\-), op_cat (\\NTDGDom\) \\<^sub>C \\NTDGDom\, cat_Set \ ]\<^sub>\" text\Components.\ lemma ntcf_Hom_components: shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\ = ( \ab\\<^sub>\(op_cat (\\NTDGDom\) \\<^sub>C \\NTDGDom\)\Obj\. ntcf_Hom_component \ \ (vpfst ab) (vpsnd ab) )" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\NTDGCod\(\\NTCod\-,\\NTDom\-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\NTDGCod\(\\NTDom\-,\\NTCod\-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTDGDom\ = op_cat (\\NTDGDom\) \\<^sub>C \\NTDGDom\" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTDGCod\ = cat_Set \" unfolding ntcf_Hom_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda ntcf_Hom_components(1) |vsv ntcf_Hom_NTMap_vsv| context fixes \ \ \ \ \ \' \' \ \ \ assumes \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and \: "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_ntcf \ \ \ \ \ \ by (rule \) interpretation \: is_ntcf \ \ \ \' \' \ by (rule \) mk_VLambda ntcf_Hom_components(1)[of _ \ \, simplified] |vdomain ntcf_Hom_NTMap_vdomain[unfolded in_Hom_iff]| lemmas [cat_cs_simps] = ntcf_Hom_NTMap_vdomain lemma ntcf_Hom_NTMap_app[cat_cs_simps]: assumes "[a, b]\<^sub>\ \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\a, b\\<^sub>\ = ntcf_Hom_component \ \ a b" using assms unfolding ntcf_Hom_components by (simp add: nat_omega_simps cat_cs_simps) end lemma (in category) ntcf_Hom_NTMap_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\) \\<^sub>\ cat_Set \\Arr\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(2)) show ?thesis proof ( rule vsv.vsv_vrange_vsubset, unfold ntcf_Hom_NTMap_vdomain[OF assms] cat_cs_simps ) fix ab assume "ab \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ op_cat \\Obj\" and b: "b \\<^sub>\ \\Obj\" by ( rule cat_prod_2_ObjE[ OF \.NTDom.HomDom.category_op \.NTDom.HomDom.category_axioms ] ) from assms a b category_cat_Set category_axioms show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\ab\ \\<^sub>\ cat_Set \\Arr\" unfolding ab_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (simp add: ntcf_Hom_NTMap_vsv) qed subsubsection\ Composition of a \Hom\-natural transformation with two natural transformations is a natural transformation \ lemma (in category) cat_ntcf_Hom_is_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(2)) show ?thesis proof(intro is_ntcfI') show "vfsequence (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-))" unfolding ntcf_Hom_def by simp show "vcard (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)) = 5\<^sub>\" unfolding ntcf_Hom_def by (simp add: nat_omega_simps) from assms category_axioms show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) from assms category_axioms show "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_intro: cat_cs_intros) from assms show "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\ab\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ObjMap\\ab\ \\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ObjMap\\ab\" if "ab \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" for ab proof- from that obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ op_cat \\Obj\" and b: "b \\<^sub>\ \\Obj\" by ( rule cat_prod_2_ObjE[ OF \.NTDom.HomDom.category_op \.NTDom.HomDom.category_axioms ] ) from category_axioms assms a b show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\ab\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ObjMap\\ab\ \\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ObjMap\\ab\" unfolding ab_def cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\a'b'\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ArrMap\\gf\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\ArrMap\\gf\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-)\NTMap\\ab\" if "gf : ab \\<^bsub>op_cat \ \\<^sub>C \\<^esub> a'b'" for ab a'b' gf proof- from that obtain g f a b a' b' where gf_def: "gf = [g, f]\<^sub>\" and ab_def: "ab = [a, b]\<^sub>\" and a'b'_def: "a'b' = [a', b']\<^sub>\" and g: "g : a \\<^bsub>op_cat \\<^esub> a'" and f: "f : b \\<^bsub>\\<^esub> b'" by ( elim cat_prod_2_is_arrE[ OF \.NTDom.HomDom.category_op \.NTDom.HomDom.category_axioms ] ) from assms category_axioms that g f show ?thesis unfolding gf_def ab_def a'b'_def cat_op_simps by (*slow*) ( cs_concl cs_simp: cat_ntcf_Hom_component_nat cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed qed (auto simp: ntcf_Hom_components cat_cs_simps) qed lemma (in category) cat_ntcf_Hom_is_ntcf': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)" and "\' = op_cat \ \\<^sub>C \" and "\
' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\-) : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \
'" using assms(1-2) unfolding assms(3-7) by (rule cat_ntcf_Hom_is_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_Hom_is_ntcf' subsubsection\ Composition of a \Hom\-natural transformation with two vertical compositions of natural transformations \ lemma (in category) cat_ntcf_Hom_vcomp: assumes "\' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\' : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,\-)" proof(rule ntcf_eqI[of \]) interpret \': is_ntcf \ \ \ \ \ \' by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) interpret \': is_ntcf \ \ \ \' \' \' by (rule assms(3)) interpret \: is_ntcf \ \ \ \' \' \ by (rule assms(4)) from category_axioms assms show H_vcomp: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms show vcomp_H: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,\-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms H_vcomp have dom_H_vcomp: "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms H_vcomp have dom_vcomp_H: "\\<^sub>\ ((Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,\-))\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-)\NTMap\ = (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,\-))\NTMap\" proof(rule vsv_eqI, unfold dom_H_vcomp dom_vcomp_H) fix ab assume prems: "ab \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by ( auto elim: cat_prod_2_ObjE[ OF \'.NTDom.HomDom.category_op \'.NTDom.HomDom.category_axioms ] simp: cat_op_simps ) from assms a b category_axioms \'.NTDom.HomDom.category_axioms \'.NTDom.HomDom.category_axioms show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-)\NTMap\\ab\ = (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,\-))\NTMap\\ab\" by ( cs_concl cs_simp: cat_cs_simps ab_def cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed (auto simp: ntcf_Hom_NTMap_vsv cat_cs_intros) qed simp_all lemmas [cat_cs_simps] = category.cat_ntcf_Hom_vcomp lemma (in category) cat_ntcf_Hom_ntcf_id: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\': \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)" proof(rule ntcf_eqI[of \]) interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \': is_functor \ \ \ \' by (rule assms(2)) from category_axioms assms show H_id: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms show id_H: "ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms H_id have dom_H_id: "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms H_id have dom_id_H: "\\<^sub>\ (ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-)\NTMap\ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\NTMap\" proof(rule vsv_eqI, unfold dom_H_id dom_id_H) show "vsv (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-)\NTMap\)" by (rule ntcf_Hom_NTMap_vsv) from id_H show "vsv (ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\NTMap\)" by (intro is_functor.ntcf_id_NTMap_vsv) (cs_concl cs_simp: cs_intro: cat_cs_intros) fix ab assume "ab \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" then obtain a b where ab_def: "ab = [a, b]\<^sub>\" and a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by ( auto elim: cat_prod_2_ObjE[OF \.HomDom.category_op \'.HomDom.category_axioms] simp: cat_op_simps ) from category_axioms assms a b H_id id_H show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,ntcf_id \'-)\NTMap\\ab\ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,\'-)\NTMap\\ab\" unfolding ab_def by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed simp qed simp_all lemmas [cat_cs_simps] = category.cat_ntcf_Hom_ntcf_id subsection\ Composition of a \Hom\-natural transformation with a natural transformation \ subsubsection\Definition and elementary properties\ text\See subsection 1.15 in \cite{bodo_categories_1970}.\ definition ntcf_lcomp_Hom :: "V \ V \ V" (\Hom\<^sub>A\<^sub>.\<^sub>C\'(/_-,-/')\) where "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,ntcf_id (cf_id (\\NTDGCod\))-)" definition ntcf_rcomp_Hom :: "V \ V \ V" (\Hom\<^sub>A\<^sub>.\<^sub>C\'(/-,_-/')\) where "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id (cf_id (\\NTDGCod\))-,\-)" subsubsection\Natural transformation map\ lemma ntcf_lcomp_Hom_NTMap_vsv: "vsv (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\NTMap\)" unfolding ntcf_lcomp_Hom_def by (rule ntcf_Hom_NTMap_vsv) lemma ntcf_rcomp_Hom_NTMap_vsv: "vsv (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-)\NTMap\)" unfolding ntcf_rcomp_Hom_def by (rule ntcf_Hom_NTMap_vsv) lemma ntcf_lcomp_Hom_NTMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_lcomp_Hom_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_rcomp_Hom_NTMap_vdomain[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-)\NTMap\) = (op_cat \ \\<^sub>C \)\Obj\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_rcomp_Hom_def \.ntcf_NTDGCod by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemma ntcf_lcomp_Hom_NTMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\NTMap\\a, b\\<^sub>\ = ntcf_lcomp_Hom_component \ a b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis unfolding ntcf_lcomp_Hom_def ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod using assms unfolding cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed lemma ntcf_rcomp_Hom_NTMap_app[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ op_cat \\Obj\" and "b \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-)\NTMap\\a, b\\<^sub>\ = ntcf_rcomp_Hom_component \ a b" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis unfolding ntcf_rcomp_Hom_def ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod using assms unfolding cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) qed lemma (in category) ntcf_lcomp_Hom_NTMap_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\NTMap\) \\<^sub>\ cat_Set \\Arr\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_lcomp_Hom_def ntcf_lcomp_Hom_component_def \.ntcf_NTDGCod by (intro ntcf_Hom_NTMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ qed lemma (in category) ntcf_rcomp_Hom_NTMap_vrange: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-)\NTMap\) \\<^sub>\ cat_Set \\Arr\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms show ?thesis unfolding ntcf_rcomp_Hom_def ntcf_rcomp_Hom_component_def \.ntcf_NTDGCod by (intro ntcf_Hom_NTMap_vrange) (cs_concl cs_intro: cat_cs_intros)+ qed subsubsection\ Composition of a \Hom\-natural transformation with a natural transformation is a natural transformation \ lemma (in category) cat_ntcf_lcomp_Hom_is_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms category_axioms show ?thesis unfolding ntcf_lcomp_Hom_def cf_bcomp_Hom_cf_lcomp_Hom[symmetric] \.ntcf_NTDGCod by (intro category.cat_ntcf_Hom_is_ntcf) (cs_concl cs_intro: cat_cs_intros)+ qed lemma (in category) cat_ntcf_lcomp_Hom_is_ntcf': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" and "\' = op_cat \ \\<^sub>C \" and "\
' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \
'" using assms(1) unfolding assms(2-6) by (rule cat_ntcf_lcomp_Hom_is_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_lcomp_Hom_is_ntcf' lemma (in category) cat_ntcf_rcomp_Hom_is_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms category_axioms show ?thesis unfolding ntcf_rcomp_Hom_def cf_bcomp_Hom_cf_rcomp_Hom[symmetric] \.ntcf_NTDGCod by (intro category.cat_ntcf_Hom_is_ntcf) (cs_concl cs_intro: cat_cs_intros)+ qed lemma (in category) cat_ntcf_rcomp_Hom_is_ntcf': assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\-)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\-)" and "\' = op_cat \ \\<^sub>C \" and "\
' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-) : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \
'" using assms(1) unfolding assms(2-6) by (rule cat_ntcf_rcomp_Hom_is_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_rcomp_Hom_is_ntcf' subsubsection\ Component of a composition of a \Hom\-natural transformation with a natural transformation and the Yoneda component \ lemma (in category) cat_ntcf_lcomp_Hom_component_is_Yoneda_component: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ op_cat \\Obj\" and "c \\<^sub>\ \\Obj\" shows "ntcf_lcomp_Hom_component \ b c = Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) (\\ObjMap\\b\) (\\NTMap\\b\) c" (is \?lcomp = ?Yc\) proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms(2) have b: "b \\<^sub>\ \\Obj\" unfolding cat_op_simps by clarsimp from b have \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" and \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" by (auto intro: cat_cs_intros) from assms(1,3) b category_axioms have \b: "\\NTMap\\b\ \\<^sub>\ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-)\ObjMap\\\\ObjMap\\b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) have lcomp: "?lcomp : Hom \ (\\ObjMap\\b\) c \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\b\) c" by (rule cat_ntcf_lcomp_Hom_component_is_arr[OF assms]) then have dom_lhs: "\\<^sub>\ (?lcomp\ArrVal\) = Hom \ (\\ObjMap\\b\) c" by (cs_concl cs_simp: cat_cs_simps) have Yc: "?Yc : Hom \ (\\ObjMap\\b\) c \\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-)\ObjMap\\c\" by ( rule cat_Yoneda_component_is_arr[ OF cat_cf_Hom_snd_is_functor[OF \b] \b \b assms(3) ] ) then have dom_rhs: "\\<^sub>\ (?Yc\ArrVal\) = Hom \ (\\ObjMap\\b\) c" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from lcomp show "arr_Set \ ?lcomp" by (auto dest: cat_Set_is_arrD(1)) from Yc show "arr_Set \ ?Yc" by (auto dest: cat_Set_is_arrD(1)) show "?lcomp\ArrVal\ = ?Yc\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) from assms(1) b category_axioms show "vsv (?Yc\ArrVal\)" by (intro is_functor.Yoneda_component_ArrVal_vsv) (cs_concl cs_intro: cat_cs_intros) show "?lcomp\ArrVal\\f\ = ?Yc\ArrVal\\f\" if "f \\<^sub>\ Hom \ (\\ObjMap\\b\) c" for f proof- from that have "f : \\ObjMap\\b\ \\<^bsub>\\<^esub> c" by simp with category_axioms assms(1,3) b show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_op_intros ) qed qed (simp_all add: ntcf_lcomp_Hom_component_ArrVal_vsv) from Yc category_axioms assms(1,3) b have "?Yc : Hom \ (\\ObjMap\\b\) c \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\b\) c" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) with lcomp show "?lcomp\ArrCod\ = ?Yc\ArrCod\" by (cs_concl cs_simp: cat_cs_simps) qed (use lcomp Yc in \cs_concl cs_simp: cat_cs_simps\) qed subsubsection\ Composition of a \Hom\-natural transformation with a vertical composition of natural transformations \ lemma (in category) cat_ntcf_lcomp_Hom_vcomp: assumes "\' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\'-,-)" proof- interpret \': is_ntcf \ \ \ \ \ \' by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) from category_axioms have ntcf_id_cf_id: "ntcf_id (cf_id \) = ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms show ?thesis unfolding ntcf_lcomp_Hom_def ntsmcf_vcomp_components dghm_id_components \'.ntcf_NTDGCod \.ntcf_NTDGCod by (subst ntcf_id_cf_id) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_vcomp lemma (in category) cat_ntcf_rcomp_Hom_vcomp: assumes "\' : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\'-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,\-)" proof- interpret \': is_ntcf \ \ \ \ \ \' by (rule assms(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(2)) from category_axioms have ntcf_id_cf_id: "ntcf_id (cf_id \) = ntcf_id (cf_id \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_id (cf_id \)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms show ?thesis unfolding ntcf_rcomp_Hom_def ntsmcf_vcomp_components dghm_id_components \'.ntcf_NTDGCod \.ntcf_NTDGCod by (subst ntcf_id_cf_id) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed lemmas [cat_cs_simps] = category.cat_ntcf_rcomp_Hom_vcomp subsubsection\ Composition of a \Hom\-natural transformation with an identity natural transformation \ lemma (in category) cat_ntcf_lcomp_Hom_ntcf_id: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(ntcf_id \-,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) from category_axioms assms show ?thesis unfolding ntcf_lcomp_Hom_def ntcf_id_components \.cf_HomCod by ( cs_concl cs_simp: ntcf_lcomp_Hom_def cat_cs_simps cs_intro: cat_cs_intros ) qed lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_ntcf_id lemma (in category) cat_ntcf_rcomp_Hom_ntcf_id: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(-,ntcf_id \-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\-)" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) from category_axioms assms show ?thesis unfolding ntcf_rcomp_Hom_def ntcf_id_components \.cf_HomCod by (cs_concl cs_simp: ntcf_rcomp_Hom_def cat_cs_simps cs_intro: cat_cs_intros) qed lemmas [cat_cs_simps] = category.cat_ntcf_rcomp_Hom_ntcf_id subsection\Projections of a \Hom\-natural transformation\ text\ The concept of a projection of a \Hom\-natural transformation appears in the corollary to the Yoneda Lemma in Chapter III-2 in \cite{mac_lane_categories_2010} (although the concept has not been given any specific name in the aforementioned reference). \ subsubsection\Definition and elementary properties\ definition ntcf_Hom_snd :: "V \ V \ V \ V" (\Hom\<^sub>A\<^sub>.\<^sub>C\_'(/_,-/')\) where "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) = Yoneda_arrow \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\Dom\\f\,-)) (\\Cod\\f\) f" definition ntcf_Hom_fst :: "V \ V \ V \ V" (\Hom\<^sub>A\<^sub>.\<^sub>C\_'(/-,_/')\) where "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat \(f,-)" text\Components.\ lemma (in category) cat_ntcf_Hom_snd_components: assumes "f : s \\<^bsub>\\<^esub> r" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTMap\ = (\d\\<^sub>\\\Obj\. Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r f d)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDGDom\ = \" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDGCod\ = cat_Set \" proof- interpret is_functor \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)\ using assms category_axioms by (cs_concl cs_simp: cs_intro: cat_cs_intros) show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTMap\ = (\d\\<^sub>\\\Obj\. Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r f d)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDGDom\ = \" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\NTDGCod\ = cat_Set \" unfolding ntcf_Hom_snd_def cat_is_arrD[OF assms] Yoneda_arrow_components by simp_all qed lemma (in category) cat_ntcf_Hom_fst_components: assumes "f : r \\<^bsub>\\<^esub> s" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\NTMap\ = (\d\\<^sub>\op_cat \\Obj\. Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r f d)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\NTDom\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\NTCod\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s)" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\NTDGDom\ = op_cat \" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\NTDGCod\ = cat_Set \" using category_axioms assms unfolding ntcf_Hom_fst_def category.cat_ntcf_Hom_snd_components[ OF category_op, unfolded cat_op_simps, OF assms ] cat_op_simps by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros)+ text\Alternative definition.\ lemma (in category) ntcf_Hom_snd_def': assumes "f : r \\<^bsub>\\<^esub> s" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) = Yoneda_arrow \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)) s f" using assms unfolding ntcf_Hom_snd_def by (simp add: cat_cs_simps) lemma (in category) ntcf_Hom_fst_def': assumes "f : r \\<^bsub>\\<^esub> s" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) = Yoneda_arrow \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r f" proof- from assms category_axioms show ?thesis unfolding ntcf_Hom_fst_def ntcf_Hom_snd_def cat_op_simps by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) qed subsubsection\Natural transformation map\ context category begin context fixes s r f assumes f: "f : s \\<^bsub>\\<^esub> r" begin mk_VLambda cat_ntcf_Hom_snd_components(1)[OF f] |vsv ntcf_Hom_snd_NTMap_vsv[intro]| |vdomain ntcf_Hom_snd_NTMap_vdomain| |app ntcf_Hom_snd_NTMap_app| end context fixes s r f assumes f: "f : r \\<^bsub>\\<^esub> s" begin mk_VLambda cat_ntcf_Hom_fst_components(1)[OF f] |vsv ntcf_Hom_fst_NTMap_vsv[intro]| |vdomain ntcf_Hom_fst_NTMap_vdomain| |app ntcf_Hom_fst_NTMap_app| end end lemmas [cat_cs_simps] = category.ntcf_Hom_snd_NTMap_vdomain category.ntcf_Hom_fst_NTMap_vdomain lemmas ntcf_Hom_snd_NTMap_app[cat_cs_simps] = category.ntcf_Hom_snd_NTMap_app category.ntcf_Hom_fst_NTMap_app subsubsection\ \Hom\-natural transformation projections are natural transformations \ lemma (in category) cat_ntcf_Hom_snd_is_ntcf: assumes "f : s \\<^bsub>\\<^esub> r" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- note f = cat_is_arrD[OF assms] show ?thesis unfolding ntcf_Hom_snd_def f proof(rule category.cat_Yoneda_arrow_is_ntcf) from assms category_axioms show "f \\<^sub>\ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)\ObjMap\\r\" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) qed (intro category_axioms cat_cf_Hom_snd_is_functor f)+ qed lemma (in category) cat_ntcf_Hom_snd_is_ntcf': assumes "f : s \\<^bsub>\\<^esub> r" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" and "\' = \" and "\
' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \
'" using assms(1) unfolding assms(2-6) by (rule cat_ntcf_Hom_snd_is_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_Hom_snd_is_ntcf' lemma (in category) cat_ntcf_Hom_fst_is_ntcf: assumes "f : r \\<^bsub>\\<^esub> s" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- from assms have r: "r \\<^sub>\ \\Obj\" and s: "s \\<^sub>\ \\Obj\" by auto from category.cat_ntcf_Hom_snd_is_ntcf[ OF category_op, unfolded cat_op_simps, OF assms, unfolded cat_op_cat_cf_Hom_snd[OF r] cat_op_cat_cf_Hom_snd[OF s], folded ntcf_Hom_fst_def ] show ?thesis . qed lemma (in category) cat_ntcf_Hom_fst_is_ntcf': assumes "f : r \\<^bsub>\\<^esub> s" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s)" and "\' = op_cat \" and "\
' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) : \' \\<^sub>C\<^sub>F \' : \' \\\<^sub>C\<^bsub>\\<^esub> \
'" using assms(1) unfolding assms(2-6) by (rule cat_ntcf_Hom_fst_is_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_Hom_fst_is_ntcf' subsubsection\Opposite \Hom\-natural transformation projections\ lemma (in category) cat_op_cat_ntcf_Hom_snd: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat \(f,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)" unfolding ntcf_Hom_fst_def by simp lemmas [cat_op_simps] = category.cat_op_cat_ntcf_Hom_snd lemma (in category) cat_op_cat_ntcf_Hom_fst: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>op_cat \(-,f) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" unfolding ntcf_Hom_fst_def cat_op_simps by simp lemmas [cat_op_simps] = category.cat_op_cat_ntcf_Hom_fst subsubsection\ \Hom\-natural transformation projections and the Yoneda component \ lemma (in category) cat_Yoneda_component_cf_Hom_snd_Comp: assumes "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" and "d \\<^sub>\ \\Obj\" shows "Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) c g d = Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) c (g \\<^sub>A\<^bsub>\\<^esub> f) d" (is \?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d = ?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d\) proof- interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) note gD = cat_is_arrD[OF assms(1)] note fD = cat_is_arrD[OF assms(2)] from assms category_axioms have Y_f: "?Ya b f d : Hom \ b d \\<^bsub>cat_Set \\<^esub> Hom \ a d" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) moreover from assms category_axioms have Y_g: "?Yb c g d : Hom \ c d \\<^bsub>cat_Set \\<^esub> Hom \ b d" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) ultimately have Yf_Yg: "?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d : Hom \ c d \\<^bsub>cat_Set \\<^esub> Hom \ a d" by (auto intro: cat_cs_intros) from assms category_axioms have Y_gf: "?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d : Hom \ c d \\<^bsub>cat_Set \\<^esub> Hom \ a d" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) from Yf_Yg have dom_rhs: "\\<^sub>\ ((?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d)\ArrVal\) = Hom \ c d" by (cs_concl cs_simp: cat_cs_simps) from Y_gf have dom_lhs: "\\<^sub>\ (?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d\ArrVal\) = Hom \ c d" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from Yf_Yg show arr_Set_Yf_Yg: "arr_Set \ (?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d)" by (auto dest: cat_Set_is_arrD(1)) interpret Yf_Yg: arr_Set \ \?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d\ by (rule arr_Set_Yf_Yg) from Y_gf show arr_Set_Y_gf: "arr_Set \ (?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d)" by (auto dest: cat_Set_is_arrD(1)) interpret Yf_Yg: arr_Set \ \?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d\ by (rule arr_Set_Y_gf) show "(?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d)\ArrVal\ = ?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix h assume "h : c \\<^bsub>\\<^esub> d" with Y_gf Y_g Y_f category_axioms assms show "(?Ya b f d \\<^sub>A\<^bsub>cat_Set \\<^esub> ?Yb c g d)\ArrVal\\h\ = ?Ya c (g \\<^sub>A\<^bsub>\\<^esub> f) d\ArrVal\\h\" (*slow*) by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) qed auto qed (use Y_gf Yf_Yg in \cs_concl cs_simp: cat_cs_simps\)+ qed lemmas [cat_cs_simps] = category.cat_Yoneda_component_cf_Hom_snd_Comp[symmetric] lemma (in category) cat_Yoneda_component_cf_Hom_snd_CId: assumes "c \\<^sub>\ \\Obj\" and "d \\<^sub>\ \\Obj\" shows "Yoneda_component Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) c (\\CId\\c\) d = cat_Set \\CId\\Hom \ c d\" (is \?Ycd = cat_Set \\CId\\Hom \ c d\\) proof- interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) from assms category_axioms have Y_CId_c: "?Ycd : Hom \ c d \\<^bsub>cat_Set \\<^esub> Hom \ c d" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) from Y_CId_c Set.category_axioms assms category_axioms have CId_cd: "cat_Set \\CId\\Hom \ c d\ : Hom \ c d \\<^bsub>cat_Set \\<^esub> Hom \ c d" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from Y_CId_c have dom_lhs: "\\<^sub>\ (?Ycd\ArrVal\) = Hom \ c d" by (cs_concl cs_simp: cat_cs_simps) from CId_cd have dom_rhs: "\\<^sub>\ (cat_Set \\CId\\Hom \ c d\\ArrVal\) = Hom \ c d" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI[of \]) from Y_CId_c show arr_Set_Y_CId_c: "arr_Set \ ?Ycd" by (auto dest: cat_Set_is_arrD(1)) interpret Yf_Yg: arr_Set \ ?Ycd by (rule arr_Set_Y_CId_c) from CId_cd show arr_Set_CId_cd: "arr_Set \ (cat_Set \\CId\\Hom \ c d\)" by (auto dest: cat_Set_is_arrD(1)) interpret CId_cd: arr_Set \ \cat_Set \\CId\\Hom \ c d\\ by (rule arr_Set_CId_cd) show "?Ycd\ArrVal\ = cat_Set \\CId\\Hom \ c d\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix h assume "h : c \\<^bsub>\\<^esub> d" with CId_cd Y_CId_c category_axioms assms show "?Ycd\ArrVal\\h\ = cat_Set \\CId\\Hom \ c d\\ArrVal\\h\" by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros) qed auto qed (use Y_CId_c CId_cd in \cs_concl cs_simp: cat_cs_simps\)+ qed lemmas [cat_cs_simps] = category.cat_Yoneda_component_cf_Hom_snd_CId subsubsection\\Hom\-natural transformation projection of a composition\ lemma (in category) cat_ntcf_Hom_snd_Comp: assumes "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g \\<^sub>A\<^bsub>\\<^esub> f,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-)" (is \?H_gf = ?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_g\) proof(rule ntcf_eqI[of \]) from assms category_axioms show "?H_gf : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F 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) from assms category_axioms show "?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_g : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F 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) from assms category_axioms have lhs_dom: "\\<^sub>\ (?H_gf\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms category_axioms have rhs_dom: "\\<^sub>\ ((?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_g)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "?H_gf\NTMap\ = (?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_g)\NTMap\" proof(rule vsv_eqI, unfold lhs_dom rhs_dom) fix d assume "d \\<^sub>\ \\Obj\" with assms category_axioms show "?H_gf\NTMap\\d\ = (?H_f \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ?H_g)\NTMap\\d\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (use assms in \auto intro: cat_cs_intros\) qed auto lemmas [cat_cs_simps] = category.cat_ntcf_Hom_snd_Comp lemma (in category) cat_ntcf_Hom_fst_Comp: assumes "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,g \\<^sub>A\<^bsub>\\<^esub> f) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,g) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)" proof- note category.cat_ntcf_Hom_snd_Comp[ OF category_op, unfolded cat_op_simps, OF assms(2,1) ] from this category_axioms assms show ?thesis by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp qed lemmas [cat_cs_simps] = category.cat_ntcf_Hom_fst_Comp subsubsection\\Hom\-natural transformation projection of an identity\ lemma (in category) cat_ntcf_Hom_snd_CId: assumes "c \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\CId\\c\,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-)" (is \?H_c = ?id_H_c\) proof(rule ntcf_eqI[of \]) from assms have "\\CId\\c\ : c \\<^bsub>\\<^esub> c" by (auto simp: cat_cs_intros) from assms category_axioms show "?H_c : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms category_axioms show "?id_H_c : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms category_axioms have lhs_dom: "\\<^sub>\ (?H_c\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms category_axioms have rhs_dom: "\\<^sub>\ (?id_H_c\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "?H_c\NTMap\ = ?id_H_c\NTMap\" proof(rule vsv_eqI, unfold lhs_dom rhs_dom) from assms category_axioms show "vsv (?id_H_c\NTMap\)" by (intro is_functor.ntcf_id_NTMap_vsv) (cs_concl cs_simp: cs_intro: cat_cs_intros) fix d assume "d \\<^sub>\ \\Obj\" with assms category_axioms show "?H_c\NTMap\\d\ = ?id_H_c\NTMap\\d\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) qed (use assms in \auto intro: cat_cs_intros\) qed auto lemmas [cat_cs_simps] = category.cat_ntcf_Hom_snd_CId lemma (in category) cat_ntcf_Hom_fst_CId: assumes "c \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,\\CId\\c\) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,c)" proof- note category.cat_ntcf_Hom_snd_CId[ OF category_op, unfolded cat_op_simps, OF assms ] from this category_axioms assms show ?thesis by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp qed lemmas [cat_cs_simps] = category.cat_ntcf_Hom_fst_CId subsubsection\\Hom\-natural transformation and the Yoneda map\ lemma (in category) cat_Yoneda_map_of_ntcf_Hom_snd: assumes "f : s \\<^bsub>\\<^esub> r" shows "Yoneda_map \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)) r\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\ = f" using category_axioms assms (*slow*) by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) lemmas [cat_cs_simps] = category.cat_Yoneda_map_of_ntcf_Hom_snd lemma (in category) cat_Yoneda_map_of_ntcf_Hom_fst: assumes "f : r \\<^bsub>\\<^esub> s" shows "Yoneda_map \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s)) r\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\ = f" proof- note category.cat_Yoneda_map_of_ntcf_Hom_snd[ OF category_op, unfolded cat_op_simps, OF assms ] from this category_axioms assms show ?thesis by (cs_prems cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp qed lemmas [cat_cs_simps] = category.cat_Yoneda_map_of_ntcf_Hom_fst subsection\Evaluation arrow\ subsubsection\Definition and elementary properties\ text\ The evaluation arrow is a part of the definition of the evaluation functor. The evaluation functor appears in Chapter III-2 in \cite{mac_lane_categories_2010}. \ definition cf_eval_arrow :: "V \ V \ V \ V" where "cf_eval_arrow \ \ f = [ ( \x\\<^sub>\\\NTDom\\ObjMap\\\\Dom\\f\\. \\NTCod\\ArrMap\\f\\ArrVal\\\\NTMap\\\\Dom\\f\\\ArrVal\\x\\ ), \\NTDom\\ObjMap\\\\Dom\\f\\, \\NTCod\\ObjMap\\\\Cod\\f\\ ]\<^sub>\" text\Components.\ lemma cf_eval_arrow_components: shows "cf_eval_arrow \ \ f\ArrVal\ = ( \x\\<^sub>\\\NTDom\\ObjMap\\\\Dom\\f\\. \\NTCod\\ArrMap\\f\\ArrVal\\\\NTMap\\\\Dom\\f\\\ArrVal\\x\\ )" and "cf_eval_arrow \ \ f\ArrDom\ = \\NTDom\\ObjMap\\\\Dom\\f\\" and "cf_eval_arrow \ \ f\ArrCod\ = \\NTCod\\ObjMap\\\\Cod\\f\\" unfolding cf_eval_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) context fixes \ \ \ \ \ a b f assumes \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and f: "f : a \\<^bsub>\\<^esub> b" begin interpretation \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule \) lemmas cf_eval_arrow_components' = cf_eval_arrow_components[ where \=\ and \=\ntcf_arrow \\ and f=f, unfolded ntcf_arrow_components cf_map_components \.NTDom.HomDom.cat_is_arrD[OF f] cat_cs_simps ] lemmas [cat_cs_simps] = cf_eval_arrow_components'(2,3) end subsubsection\Arrow value\ context fixes \ \ \ \ \ a b f assumes \: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and f: "f : a \\<^bsub>\\<^esub> b" begin mk_VLambda cf_eval_arrow_components'(1)[OF \ f] |vsv cf_eval_arrow_ArrVal_vsv[cat_cs_intros]| |vdomain cf_eval_arrow_ArrVal_vdomain[cat_cs_simps]| |app cf_eval_arrow_ArrVal_app[cat_cs_simps]| end subsubsection\Evaluation arrow is an arrow in the category \Set\\ lemma cf_eval_arrow_is_arr: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "f : a \\<^bsub>\\<^esub> b" shows "cf_eval_arrow \ (ntcf_arrow \) f : \\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\b\" proof- interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms) show ?thesis proof ( intro cat_Set_is_arrI arr_SetI, unfold cf_eval_arrow_components'(2,3)[OF assms] ) show "vfsequence (cf_eval_arrow \ (ntcf_arrow \) f)" unfolding cf_eval_arrow_def by simp show "vcard (cf_eval_arrow \ (ntcf_arrow \) f) = 3\<^sub>\" unfolding cf_eval_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_eval_arrow \ (ntcf_arrow \) f\ArrVal\) \\<^sub>\ \\ObjMap\\b\" by ( unfold cf_eval_arrow_components'[OF assms], intro vrange_VLambda_vsubset ) ( use assms in \cs_concl cs_intro: cat_cs_intros cat_Set_cs_intros\ )+ qed ( use assms(2) in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\ )+ qed lemma cf_eval_arrow_is_arr'[cat_cs_intros]: assumes "\' = ntcf_arrow \" and "\a = \\ObjMap\\a\" and "\b = \\ObjMap\\b\" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "f : a \\<^bsub>\\<^esub> b" shows "cf_eval_arrow \ \' f : \a \\<^bsub>cat_Set \\<^esub> \b" using assms(4,5) unfolding assms(1-3) by (rule cf_eval_arrow_is_arr) lemma (in category) cat_cf_eval_arrow_ntcf_vcomp[cat_cs_simps]: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "g : b \\<^bsub>\\<^esub> c" and "f : a \\<^bsub>\\<^esub> b" shows "cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f) = cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f" proof- interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms(2)) have \\: "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3,4) have gf: "g \\<^sub>A\<^bsub>\\<^esub> f : a \\<^bsub>\\<^esub> c" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \\ gf have cf_eval_gf: "cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f) : \\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\c\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(3,4) have cf_eval_g_cf_eval_f: "cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f : \\ObjMap\\a\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\c\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) note cf_eval_gf = cf_eval_gf cat_Set_is_arrD[OF cf_eval_gf] note cf_eval_g_cf_eval_f = cf_eval_g_cf_eval_f cat_Set_is_arrD[OF cf_eval_g_cf_eval_f] interpret arr_Set_cf_eval_gf: arr_Set \ \cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f)\ by (rule cf_eval_gf(2)) interpret arr_Set_cf_eval_g_cf_eval_f: arr_Set \ \ cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f \ by (rule cf_eval_g_cf_eval_f(2)) show ?thesis proof(rule arr_Set_eqI) from \\ gf have dom_lhs: "\\<^sub>\ (cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f)\ArrVal\) = \\ObjMap\\a\" by (cs_concl cs_simp: cat_cs_simps) from cf_eval_g_cf_eval_f(1) have dom_rhs: "\\<^sub>\ ( ( cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f )\ArrVal\ ) = \\ObjMap\\a\" by (cs_concl cs_simp: cat_cs_simps) show "cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f)\ArrVal\ = ( cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f )\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix \a assume prems: "\a \\<^sub>\ \\ObjMap\\a\" from ArrVal_eq_helper [ OF \.ntcf_Comp_commute[OF assms(4), symmetric], where a=\\\NTMap\\a\\ArrVal\\\a\\ ] prems assms(3,4) have [cat_cs_simps]: "\\ArrMap\\f\\ArrVal\\\\NTMap\\a\\ArrVal\\\\NTMap\\a\\ArrVal\\\a\\\ = \\NTMap\\b\\ArrVal\\\\ArrMap\\f\\ArrVal\\\\NTMap\\a\\ArrVal\\\a\\\" by ( cs_prems cs_simp: cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros ) from prems assms(3,4) show "cf_eval_arrow \ (ntcf_arrow (\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \)) (g \\<^sub>A\<^bsub>\\<^esub> f)\ArrVal\\\a\ = ( cf_eval_arrow \ (ntcf_arrow \) g \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval_arrow \ (ntcf_arrow \) f )\ArrVal\\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros ) qed (cs_concl cs_intro: V_cs_intros) qed ( auto simp: cf_eval_gf cf_eval_g_cf_eval_f intro: cf_eval_gf(2) cf_eval_g_cf_eval_f(2) ) qed lemmas [cat_cs_simps] = category.cat_cf_eval_arrow_ntcf_vcomp lemma (in category) cat_cf_eval_arrow_ntcf_id[cat_cs_simps]: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "c \\<^sub>\ \\Obj\" shows "cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\) = cat_Set \\CId\\\\ObjMap\\c\\" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms) from assms(2) have ntcf_id_CId_c: "cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\) : \\ObjMap\\c\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\c\" by (cs_concl cs_intro: cat_cs_intros) from assms(2) have CId_\c: "cat_Set \\CId\\\\ObjMap\\c\\ : \\ObjMap\\c\ \\<^bsub>cat_Set \\<^esub> \\ObjMap\\c\" by (cs_concl cs_intro: cat_cs_intros) show ?thesis proof(rule arr_Set_eqI[of \]) from ntcf_id_CId_c show arr_Set_ntcf_id_CId_c: "arr_Set \ (cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\))" by (auto dest: cat_Set_is_arrD(1)) from ntcf_id_CId_c have dom_lhs: "\\<^sub>\ (cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\)\ArrVal\) = \\ObjMap\\c\" by (cs_concl cs_simp: cat_cs_simps)+ interpret ntcf_id_CId_c: arr_Set \ \cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\)\ by (rule arr_Set_ntcf_id_CId_c) from CId_\c show arr_Set_CId_\c: "arr_Set \ (cat_Set \\CId\\\\ObjMap\\c\\)" by (auto dest: cat_Set_is_arrD(1)) from CId_\c assms(2) have dom_rhs: "\\<^sub>\ ((cat_Set \\CId\\\\ObjMap\\c\\)\ArrVal\) = \\ObjMap\\c\" by (cs_concl cs_simp: cat_cs_simps) show "cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\)\ArrVal\ = cat_Set \\CId\\\\ObjMap\\c\\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\ObjMap\\c\" with category_axioms assms(2) show "cf_eval_arrow \ (ntcf_arrow (ntcf_id \)) (\\CId\\c\)\ArrVal\\a\ = cat_Set \\CId\\\\ObjMap\\c\\\ArrVal\\a\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (use arr_Set_ntcf_id_CId_c arr_Set_CId_\c in auto) qed (use ntcf_id_CId_c CId_\c in \cs_concl cs_simp: cat_cs_simps\)+ qed lemmas [cat_cs_simps] = category.cat_cf_eval_arrow_ntcf_id subsection\\HOM\-functor\ subsubsection\Definition and elementary properties\ text\ The following definition is a technical generalization that is used later in this section. \ definition cf_HOM_snd :: "V \ V \ V" (\HOM\<^sub>C\'(/,_-/')\) where "HOM\<^sub>C\<^bsub>\\<^esub>(,\-) = [ (\a\\<^sub>\op_cat (\\HomCod\)\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\\HomCod\)(a,-) \\<^sub>C\<^sub>F \)), ( \f\\<^sub>\op_cat (\\HomCod\)\Arr\. 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 \) ), op_cat (\\HomCod\), cat_FUNCT \ (\\HomDom\) (cat_Set \) ]\<^sub>\" definition cf_HOM_fst :: "V \ V \ V" (\HOM\<^sub>C\'(/_-,/')\) where "HOM\<^sub>C\<^bsub>\\<^esub>(\-,) = [ (\a\\<^sub>\(\\HomCod\)\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\\HomCod\)(-,a) \\<^sub>C\<^sub>F op_cf \)), ( \f\\<^sub>\(\\HomCod\)\Arr\. 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 op_cf \) ), \\HomCod\, cat_FUNCT \ (op_cat (\\HomDom\)) (cat_Set \) ]\<^sub>\" text\Components.\ lemma cf_HOM_snd_components: shows "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ObjMap\ = (\a\\<^sub>\op_cat (\\HomCod\)\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\\HomCod\)(a,-) \\<^sub>C\<^sub>F \))" and "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ArrMap\ = ( \f\\<^sub>\op_cat (\\HomCod\)\Arr\. 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 \) )" and [cat_cs_simps]: "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\HomDom\ = op_cat (\\HomCod\)" and [cat_cs_simps]: "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\HomCod\ = cat_FUNCT \ (\\HomDom\) (cat_Set \)" unfolding cf_HOM_snd_def dghm_field_simps by (simp_all add: nat_omega_simps) lemma cf_HOM_fst_components: shows "HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\ObjMap\ = (\a\\<^sub>\(\\HomCod\)\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\\HomCod\)(-,a) \\<^sub>C\<^sub>F op_cf \))" and "HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\ArrMap\ = ( \f\\<^sub>\(\\HomCod\)\Arr\. 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 op_cf \) )" and "HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\HomDom\ = \\HomCod\" and "HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\HomCod\ = cat_FUNCT \ (op_cat (\\HomDom\)) (cat_Set \)" unfolding cf_HOM_fst_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_functor begin lemmas cf_HOM_snd_components' = cf_HOM_snd_components[where \=\, unfolded cf_HomDom cf_HomCod] lemmas [cat_cs_simps] = cf_HOM_snd_components'(3,4) lemmas cf_HOM_fst_components' = cf_HOM_fst_components[where \=\, unfolded cf_HomDom cf_HomCod] lemmas [cat_cs_simps] = cf_HOM_snd_components'(3,4) end subsubsection\Object map\ mk_VLambda cf_HOM_snd_components(1) |vsv cf_HOM_snd_ObjMap_vsv[cat_cs_intros]| mk_VLambda (in is_functor) cf_HOM_snd_components'(1)[unfolded cat_op_simps] |vdomain cf_HOM_snd_ObjMap_vdomain[cat_cs_simps]| |app cf_HOM_snd_ObjMap_app[cat_cs_simps]| mk_VLambda cf_HOM_snd_components(1) |vsv cf_HOM_fst_ObjMap_vsv[cat_cs_intros]| mk_VLambda (in is_functor) cf_HOM_fst_components'(1)[unfolded cat_op_simps] |vdomain cf_HOM_fst_ObjMap_vdomain[cat_cs_simps]| |app cf_HOM_fst_ObjMap_app[cat_cs_simps]| subsubsection\Arrow map\ mk_VLambda cf_HOM_snd_components(2) |vsv cf_HOM_snd_ArrMap_vsv[cat_cs_intros]| mk_VLambda (in is_functor) cf_HOM_snd_components'(2)[unfolded cat_op_simps] |vdomain cf_HOM_snd_ArrMap_vdomain[cat_cs_simps]| |app cf_HOM_snd_ArrMap_app[cat_cs_simps]| mk_VLambda cf_HOM_fst_components(2) |vsv cf_HOM_fst_ArrMap_vsv[cat_cs_intros]| mk_VLambda (in is_functor) cf_HOM_fst_components'(2)[unfolded cat_op_simps] |vdomain cf_HOM_fst_ArrMap_vdomain[cat_cs_simps]| |app cf_HOM_fst_ArrMap_app[cat_cs_simps]| subsubsection\Opposite \HOM\-functor\ lemma (in is_functor) cf_HOM_snd_op[cat_op_simps]: "HOM\<^sub>C\<^bsub>\\<^esub>(,op_cf \-) = HOM\<^sub>C\<^bsub>\\<^esub>(\-,)" proof- have dom_lhs: "\\<^sub>\ HOM\<^sub>C\<^bsub>\\<^esub>(,op_cf \-) = 4\<^sub>\" unfolding cf_HOM_snd_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ HOM\<^sub>C\<^bsub>\\<^esub>(\-,) = 4\<^sub>\" unfolding cf_HOM_fst_def by (simp add: nat_omega_simps) show ?thesis proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ 4\<^sub>\" then show "HOM\<^sub>C\<^bsub>\\<^esub>(,op_cf \-)\a\ = HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\a\" proof ( elim_in_numeral, use nothing in \fold dghm_field_simps, unfold cat_cs_simps\ ) show "HOM\<^sub>C\<^bsub>\\<^esub>(,op_cf \-)\ObjMap\ = HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\ObjMap\" unfolding cf_HOM_fst_components' is_functor.cf_HOM_snd_components'[OF is_functor_op] by (rule VLambda_eqI, unfold cat_op_simps) (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)+ show "HOM\<^sub>C\<^bsub>\\<^esub>(,op_cf \-)\ArrMap\ = HOM\<^sub>C\<^bsub>\\<^esub>(\-,)\ArrMap\" unfolding cf_HOM_fst_components' is_functor.cf_HOM_snd_components'[OF is_functor_op] by (rule VLambda_eqI, unfold cat_op_simps) (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros)+ qed ( auto simp: cf_HOM_fst_components' cat_cs_simps cat_op_simps cat_op_intros ) qed (auto simp: cf_HOM_snd_def cf_HOM_fst_def) qed lemmas [cat_op_simps] = is_functor.cf_HOM_snd_op context is_functor begin lemmas cf_HOM_fst_op[cat_op_simps] = is_functor.cf_HOM_snd_op[OF is_functor_op, unfolded cat_op_simps, symmetric] end lemmas [cat_op_simps] = is_functor.cf_HOM_fst_op subsubsection\\HOM\-functor is a functor\ lemma (in is_functor) cf_HOM_snd_is_functor: assumes "\ \" and "\ \\<^sub>\ \" shows "HOM\<^sub>C\<^bsub>\\<^esub>(,\-) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ \ (cat_Set \)" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ show ?thesis proof(intro is_functorI', unfold cat_op_simps) show "vfsequence HOM\<^sub>C\<^bsub>\\<^esub>(,\-)" unfolding cf_HOM_snd_def by auto show "vcard HOM\<^sub>C\<^bsub>\\<^esub>(,\-) = 4\<^sub>\" unfolding cf_HOM_snd_def by (simp add: nat_omega_simps) show "\\<^sub>\ (HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ObjMap\) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Obj\" unfolding cf_HOM_snd_components' proof(rule vrange_VLambda_vsubset, unfold cat_op_simps) fix b assume prems: "b \\<^sub>\ \\Obj\" with assms(2) show "cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F \) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Obj\" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed show "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ArrMap\\f \\<^sub>A\<^bsub>\\<^esub> g\ = HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ArrMap\\g\ \\<^sub>A\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ArrMap\\f\" if "g : c \\<^bsub>\\<^esub> b" and "f : b \\<^bsub>\\<^esub> a" for b c g a f using that by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) show "HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ArrMap\\\\CId\\c\\ = cat_FUNCT \ \ (cat_Set \)\CId\\HOM\<^sub>C\<^bsub>\\<^esub>(,\-)\ObjMap\\c\\" if "c \\<^sub>\ \\Obj\" for c using that by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed ( use assms(2) in \ cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros \ )+ qed lemma (in is_functor) cf_HOM_snd_is_functor'[cat_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "\' = op_cat \" and "\
= cat_FUNCT \ \ (cat_Set \)" shows "HOM\<^sub>C\<^bsub>\\<^esub>(,\-) : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1,2) unfolding assms(3,4) by (rule cf_HOM_snd_is_functor) lemmas [cat_cs_intros] = is_functor.cf_HOM_snd_is_functor' lemma (in is_functor) cf_HOM_fst_is_functor: assumes "\ \" and "\ \\<^sub>\ \" shows "HOM\<^sub>C\<^bsub>\\<^esub>(\-,) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_FUNCT \ (op_cat \) (cat_Set \)" by ( rule is_functor.cf_HOM_snd_is_functor[ OF is_functor_op assms, unfolded cat_op_simps ] ) lemma (in is_functor) cf_HOM_fst_is_functor'[cat_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "\' = \" and "\
= cat_FUNCT \ (op_cat \) (cat_Set \)" shows "HOM\<^sub>C\<^bsub>\\<^esub>(\-,) : \' \\\<^sub>C\<^bsub>\\<^esub> \
" using assms(1,2) unfolding assms(3,4) by (rule cf_HOM_fst_is_functor) lemmas [cat_cs_intros] = is_functor.cf_HOM_fst_is_functor' subsection\Evaluation functor\ subsubsection\Definition and elementary properties\ text\See Chapter III-2 in \cite{mac_lane_categories_2010}.\ definition cf_eval :: "V \ V \ V \ V" where "cf_eval \ \ \ = [ (\\d\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\. \d\0\\ObjMap\\\d\1\<^sub>\\\), ( \\f\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Arr\. cf_eval_arrow \ (\f\0\) (\f\1\<^sub>\\) ), cat_FUNCT \ \ (cat_Set \) \\<^sub>C \, cat_Set \ ]\<^sub>\" text\Components.\ lemma cf_eval_components: shows "cf_eval \ \ \\ObjMap\ = (\\d\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\. \d\0\\ObjMap\\\d\1\<^sub>\\\)" and "cf_eval \ \ \\ArrMap\ = ( \\f\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Arr\. cf_eval_arrow \ (\f\0\) (\f\1\<^sub>\\) )" and [cat_cs_simps]: "cf_eval \ \ \\HomDom\ = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" and [cat_cs_simps]: "cf_eval \ \ \\HomCod\ = cat_Set \" unfolding cf_eval_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma cf_eval_ObjMap_vsv[cat_cs_intros]: "vsv (cf_eval \ \ \\ObjMap\)" unfolding cf_eval_components by simp lemma cf_eval_ObjMap_vdomain[cat_cs_simps]: "\\<^sub>\ (cf_eval \ \ \\ObjMap\) = (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" unfolding cf_eval_components by simp lemma (in category) cf_eval_ObjMap_app[cat_cs_simps]: assumes "\c = [cf_map \, c]\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" (*the order of premises is important*) and "c \\<^sub>\ \\Obj\" shows "cf_eval \ \ \\ObjMap\\\c\ = \\ObjMap\\c\" proof- interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(2)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) then interpret \: \ \ by simp note [cat_small_cs_intros] = cat_category_if_ge_Limit from assms(2,3) \\ have "\c \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" by ( cs_concl cs_simp: assms(1) cat_FUNCT_components(1) cs_intro: cat_cs_intros cat_small_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) then show ?thesis by (simp add: assms(1) cf_map_components cf_eval_components nat_omega_simps) qed lemmas [cat_cs_simps] = category.cf_eval_ObjMap_app subsubsection\Arrow map\ lemma cf_eval_ArrMap_vsv[cat_cs_intros]: "vsv (cf_eval \ \ \\ArrMap\)" unfolding cf_eval_components by simp lemma cf_eval_ArrMap_vdomain[cat_cs_simps]: "\\<^sub>\ (cf_eval \ \ \\ArrMap\) = (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Arr\" unfolding cf_eval_components by simp lemma (in category) cf_eval_ArrMap_app[cat_cs_simps]: assumes "\f = [ntcf_arrow \, f]\<^sub>\" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "f : a \\<^bsub>\\<^esub> b" shows "cf_eval \ \ \\ArrMap\\\f\ = cf_eval_arrow \ (ntcf_arrow \) f" proof- interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms(2)) define \ where "\ = \ + \" have "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_def \_Limit_\\ \_\_\\ \_def \_\_\\) then interpret \: \ \ by simp note [cat_small_cs_intros] = cat_category_if_ge_Limit from assms(1,3) \\ have "\f \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Arr\" by ( cs_concl cs_simp: assms(1) cat_FUNCT_components(1) cs_intro: cat_cs_intros cat_small_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) then show ?thesis by (simp add: assms(1) cf_map_components cf_eval_components nat_omega_simps) qed lemmas [cat_cs_simps] = category.cf_eval_ArrMap_app subsubsection\Evaluation functor is a functor\ lemma (in category) cat_cf_eval_is_functor: assumes "\ \" and "\ \\<^sub>\ \" shows "cf_eval \ \ \ : cat_FUNCT \ \ (cat_Set \) \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: \ \ by (rule assms(1)) from assms(2) cat_category_if_ge_Limit[OF assms] interpret FUNCT: category \ \(cat_FUNCT \ \ (cat_Set \))\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret cat_Set_\\: subcategory \ \cat_Set \\ \cat_Set \\ by (rule subcategory_cat_Set_cat_Set[OF assms]) show ?thesis proof(intro is_functorI') show "vfsequence (cf_eval \ \ \)" unfolding cf_eval_def by simp from cat_category_if_ge_Limit[OF assms] show "category \ ((cat_FUNCT \ \ (cat_Set \)) \\<^sub>C \)" by (cs_concl cs_simp: cs_intro: cat_small_cs_intros cat_cs_intros) show "vcard (cf_eval \ \ \) = 4\<^sub>\" unfolding cf_eval_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_eval \ \ \\ObjMap\) \\<^sub>\ cat_Set \\Obj\" proof(intro vsv.vsv_vrange_vsubset, unfold cat_cs_simps) fix \c assume prems: "\c \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" then obtain \ c where \c_def: "\c = [\, c]\<^sub>\" and \: "\ \\<^sub>\ cf_maps \ \ (cat_Set \)" and c: "c \\<^sub>\ \\Obj\" by ( auto elim: cat_prod_2_ObjE[rotated 2] intro: FUNCT.category_axioms \\.category_axioms simp: cat_FUNCT_components(1) ) from \ obtain \ where \_def: "\ = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (elim cf_mapsE) interpret \: is_functor \ \ \cat_Set \\ \ by (rule \) from \ c show "cf_eval \ \ \\ObjMap\\\c\ \\<^sub>\ cat_Set \\Obj\" unfolding \c_def \_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Set_\\.subcat_Obj_vsubset ) qed (cs_concl cs_intro: cat_cs_intros) show "cf_eval \ \ \\ArrMap\\\f\ : cf_eval \ \ \\ObjMap\\\a\ \\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ObjMap\\\b\" if \f: "\f : \a \\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \b" for \a \b \f proof- obtain \ f \ a \ b where \f_def: "\f = [\, f]\<^sub>\" and \a_def: "\a = [\, a]\<^sub>\" and \b_def: "\b = [\, b]\<^sub>\" and \: "\ : \ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> \" and f: "f : a \\<^bsub>\\<^esub> b" by ( auto intro: cat_prod_2_is_arrE[rotated 2, OF \f] FUNCT.category_axioms \\.category_axioms ) note \ = cat_FUNCT_is_arrD[OF \] from \(1) f assms(2) show "cf_eval \ \ \\ArrMap\\\f\ : cf_eval \ \ \\ObjMap\\\a\ \\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ObjMap\\\b\" unfolding \f_def \a_def \b_def by ( intro cat_Set_\\.subcat_is_arrD, use nothing in \subst \(2), subst \(3), subst \(4)\ ) ( cs_concl cs_simp: cat_FUNCT_cs_simps cat_cs_simps cs_intro: cat_cs_intros ) (*slow*) qed show "cf_eval \ \ \\ArrMap\\\g \\<^sub>A\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \f\ = cf_eval \ \ \\ArrMap\\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ArrMap\\\f\" if \g: "\g : \b \\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \c" and \f: "\f : \a \\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \b" for \f \g \a \b \c proof- obtain \ f \ a \ b where \f_def: "\f = [\, f]\<^sub>\" and \a_def: "\a = [\, a]\<^sub>\" and \b_def: "\b = [\, b]\<^sub>\" and \: "\ : \ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> \" and f: "f : a \\<^bsub>\\<^esub> b" by ( auto intro: cat_prod_2_is_arrE[rotated 2, OF \f] FUNCT.category_axioms \\.category_axioms ) then obtain \ g \ c where \g_def: "\g = [\, g]\<^sub>\" and \c_def: "\c = [\, c]\<^sub>\" and \: "\ : \ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> \" and g: "g : b \\<^bsub>\\<^esub> c" by ( auto intro: cat_prod_2_is_arrE[rotated 2, OF \g] FUNCT.category_axioms \\.category_axioms ) note \ = cat_FUNCT_is_arrD[OF \] and \ = cat_FUNCT_is_arrD[OF \] from \(1) \(1) f g show "cf_eval \ \ \\ArrMap\\\g \\<^sub>A\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \f\ = cf_eval \ \ \\ArrMap\\\g\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ArrMap\\\f\" unfolding \g_def \f_def \a_def \b_def \c_def by ( subst (1 2) \(2), use nothing in \subst (1 2) \(2)\, cs_concl_step cat_Set_\\.subcat_Comp_simp[symmetric] ) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) qed show "cf_eval \ \ \\ArrMap\\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\CId\\\c\\ = cat_Set \\CId\\cf_eval \ \ \\ObjMap\\\c\\" if "\c \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" for \c proof- from that obtain \ c where \c_def: "\c = [\, c]\<^sub>\" and \: "\ \\<^sub>\ cf_maps \ \ (cat_Set \)" and c: "c \\<^sub>\ \\Obj\" by ( auto elim: cat_prod_2_ObjE[rotated 2] intro: FUNCT.category_axioms \\.category_axioms simp: cat_FUNCT_components(1) ) from \ obtain \ where \_def: "\ = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (elim cf_mapsE) interpret \: is_functor \ \ \cat_Set \\ \ by (rule \) from \ c show "cf_eval \ \ \\ArrMap\\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\CId\\\c\\ = cat_Set \\CId\\cf_eval \ \ \\ObjMap\\\c\\" unfolding \c_def \_def by (cs_concl_step cat_Set_\\.subcat_CId[symmetric]) ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps cat_FUNCT_cs_simps 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)+ qed lemma (in category) cat_cf_eval_is_functor': assumes "\ \" and "\ \\<^sub>\ \" and "\' = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" and "\' = cat_Set \" and "\' = \" shows "cf_eval \ \ \ : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1,2) unfolding assms(3-5) by (rule cat_cf_eval_is_functor) lemmas [cat_cs_intros] = category.cat_cf_eval_is_functor' subsection\\N\-functor\ subsubsection\Definition and elementary properties\ text\See Chapter III-2 in \cite{mac_lane_categories_2010}.\ definition cf_nt :: "V \ V \ V \ V" where "cf_nt \ \ \ = bifunctor_flip (\\HomCod\) (cat_FUNCT \ (\\HomDom\) (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-))" text\Alternative definition.\ lemma (in is_functor) cf_nt_def': "cf_nt \ \ \ = bifunctor_flip \ (cat_FUNCT \ \ (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-))" unfolding cf_nt_def cf_HomDom cf_HomCod by simp text\Components.\ lemma cf_nt_components: shows "cf_nt \ \ \\ObjMap\ = ( bifunctor_flip (\\HomCod\) (cat_FUNCT \ (\\HomDom\) (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ObjMap\" and "cf_nt \ \ \\ArrMap\ = ( bifunctor_flip (\\HomCod\) (cat_FUNCT \ (\\HomDom\) (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ArrMap\" and "cf_nt \ \ \\HomDom\ = ( bifunctor_flip (\\HomCod\) (cat_FUNCT \ (\\HomDom\) (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\HomDom\" and "cf_nt \ \ \\HomCod\ = ( bifunctor_flip (\\HomCod\) (cat_FUNCT \ (\\HomDom\) (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ (\\HomDom\) (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\HomCod\" unfolding cf_nt_def by simp_all lemma (in is_functor) cf_nt_components': assumes "\ \" and "\ \\<^sub>\ \" shows "cf_nt \ \ \\ObjMap\ = ( bifunctor_flip \ (cat_FUNCT \ \ (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ObjMap\" and "cf_nt \ \ \\ArrMap\ = ( bifunctor_flip \ (cat_FUNCT \ \ (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ArrMap\" and [cat_cs_simps]: "cf_nt \ \ \\HomDom\ = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" and [cat_cs_simps]: "cf_nt \ \ \\HomCod\ = cat_Set \" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ show "cf_nt \ \ \\ObjMap\ = ( bifunctor_flip \ (cat_FUNCT \ \ (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ObjMap\" "cf_nt \ \ \\ArrMap\ = ( bifunctor_flip \ (cat_FUNCT \ \ (cat_Set \)) (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>cat_FUNCT \ \ (cat_Set \)(HOM\<^sub>C\<^bsub>\\<^esub>(,\-)-,-)) )\ArrMap\" "cf_nt \ \ \\HomDom\ = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" "cf_nt \ \ \\HomCod\ = cat_Set \" unfolding cf_nt_def using assms(2) by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros )+ qed lemmas [cat_cs_simps] = is_functor.cf_nt_components'(3,4) subsubsection\Object map\ lemma cf_nt_ObjMap_vsv[cat_cs_intros]: "vsv (cf_nt \ \ \\ObjMap\)" unfolding cf_nt_components by (cs_intro_step cat_cs_intros) lemma (in is_functor) cf_nt_ObjMap_vdomain[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" shows "\\<^sub>\ (cf_nt \ \ \\ObjMap\) = (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) show ?thesis unfolding cf_nt_components by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros ) qed lemmas [cat_cs_simps] = is_functor.cf_nt_ObjMap_vdomain lemma (in is_functor) cf_nt_ObjMap_app[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "\b = [cf_map \, b]\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "b \\<^sub>\ \\Obj\" shows "cf_nt \ \ \\ObjMap\\\b\ = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F \)) (cf_map \)" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(4)) from assms(2,5) show ?thesis unfolding assms(3) cf_nt_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) qed lemmas [cat_cs_simps] = is_functor.cf_nt_ObjMap_app subsubsection\Arrow map\ lemma cf_nt_ArrMap_vsv[cat_cs_intros]: "vsv (cf_nt \ \ \\ArrMap\)" unfolding cf_nt_components by (cs_intro_step cat_cs_intros) lemma (in is_functor) cf_nt_ArrMap_vdomain[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" shows "\\<^sub>\ (cf_nt \ \ \\ArrMap\) = (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Arr\" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) show ?thesis unfolding cf_nt_components by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros ) qed lemmas [cat_cs_simps] = is_functor.cf_nt_ArrMap_vdomain lemma (in is_functor) cf_nt_ArrMap_app[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "\f = [ntcf_arrow \, f]\<^sub>\" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "f : a \\<^bsub>\\<^esub> b" shows "cf_nt \ \ \\ArrMap\\\f\ = cf_hom (cat_FUNCT \ \ (cat_Set \)) [ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \), ntcf_arrow \]\<^sub>\" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms(4)) from assms(2,5) show ?thesis unfolding assms(3) cf_nt_def by ( cs_concl cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros ) qed lemmas [cat_cs_simps] = is_functor.cf_nt_ArrMap_app subsubsection\\N\-functor is a functor\ lemma (in is_functor) cf_nt_is_functor: assumes "\ \" and "\ \\<^sub>\ \" shows "cf_nt \ \ \ : cat_FUNCT \ \ (cat_Set \) \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) show ?thesis unfolding cf_nt_def' by ( cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) qed lemma (in is_functor) cf_nt_is_functor': assumes "\ \" and "\ \\<^sub>\ \" and "\' = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" and "\' = cat_Set \" and "\' = \" shows "cf_nt \ \ \ : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1,2) unfolding assms(3-5) by (rule cf_nt_is_functor) lemmas [cat_cs_intros] = is_functor.cf_nt_is_functor' subsection\Yoneda natural transformation arrow\ subsubsection\Definition and elementary properties\ text\ The following subsection is based on the elements of the content of Chapter III-2 in \cite{mac_lane_categories_2010}. \ definition ntcf_Yoneda_arrow :: "V \ V \ V \ V \ V" where "ntcf_Yoneda_arrow \ \ \ r = [ ( \\\\<^sub>\Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) \. Yoneda_map \ (cf_of_cf_map \ (cat_Set \) \) r\ ntcf_of_ntcf_arrow \ (cat_Set \) \ \ ), Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) \, \\ObjMap\\r\ ]\<^sub>\" text\Components\ lemma ntcf_Yoneda_arrow_components: shows "ntcf_Yoneda_arrow \ \ \ r\ArrVal\ = ( \\\\<^sub>\Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) \. Yoneda_map \ (cf_of_cf_map \ (cat_Set \) \) r\ ntcf_of_ntcf_arrow \ (cat_Set \) \ \ )" and [cat_cs_simps]: "ntcf_Yoneda_arrow \ \ \ r\ArrDom\ = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) \" and [cat_cs_simps]: "ntcf_Yoneda_arrow \ \ \ r\ArrCod\ = \\ObjMap\\r\" unfolding ntcf_Yoneda_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow map\ mk_VLambda ntcf_Yoneda_arrow_components(1) |vsv ntcf_Yoneda_arrow_vsv[cat_cs_intros]| |vdomain ntcf_Yoneda_arrow_vdomain[cat_cs_simps]| context category begin context fixes \ :: V begin mk_VLambda ntcf_Yoneda_arrow_components(1)[where \=\ and \=\ and \=\cf_map \\] |app ntcf_Yoneda_arrow_app'| lemmas ntcf_Yoneda_arrow_app = ntcf_Yoneda_arrow_app'[unfolded in_Hom_iff, cat_cs_simps] end end lemmas [cat_cs_simps] = category.ntcf_Yoneda_arrow_app subsubsection\Several technical lemmas\ lemma (in vsv) vsv_vrange_VLambda_app: assumes "g ` elts A = elts (\\<^sub>\ r)" shows "\\<^sub>\ (\x\\<^sub>\A. r\g x\) = \\<^sub>\ r" proof(intro vsubset_antisym vsv.vsv_vrange_vsubset, unfold vdomain_VLambda) show "(\x\\<^sub>\A. r\g x\)\x\ \\<^sub>\ \\<^sub>\ r" if "x \\<^sub>\ A" for x proof- from assms that have "g x \\<^sub>\ \\<^sub>\ r" by auto then have "r\g x\ \\<^sub>\ \\<^sub>\ r" by force with that show ?thesis by simp qed show "r\x\ \\<^sub>\ \\<^sub>\ (\x\\<^sub>\A. r\g x\)" if "x \\<^sub>\ \\<^sub>\ r" for x proof- from that assms have "x \ g ` elts A" by simp then obtain c where c: "c \\<^sub>\ A" and x_def: "x = g c" by clarsimp from c show ?thesis unfolding x_def by auto qed qed auto lemma (in vsv) vsv_vrange_VLambda_app': assumes "g ` elts A = elts (\\<^sub>\ r)" and "R = \\<^sub>\ r" shows "\\<^sub>\ (\x\\<^sub>\A. r\g x\) = R" using assms(1) unfolding assms(2) by (rule vsv_vrange_VLambda_app) lemma (in v11) v11_VLambda_v11_bij_betw_comp: assumes "bij_betw g (elts A) (elts (\\<^sub>\ r))" shows "v11 (\x\\<^sub>\A. r\g x\)" proof(rule vsv.vsv_valeq_v11I, unfold vdomain_VLambda beta) fix x y assume prems: "x \\<^sub>\ A" "y \\<^sub>\ A" "r\g x\ = r\g y\" from assms prems(1,2) have "g x \\<^sub>\ \\<^sub>\ r" and "g y \\<^sub>\ \\<^sub>\ r" by auto from v11_injective[OF this prems(3)] have "g x = g y". with assms prems(1,2) show "x = y" unfolding bij_betw_def inj_on_def by simp qed simp subsubsection\ Yoneda natural transformation arrow is an arrow in the category \Set\ \ lemma (in category) cat_ntcf_Yoneda_arrow_is_arr_isomoprhism: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "ntcf_Yoneda_arrow \ \ (cf_map \) r : Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) (cf_map \) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> \\ObjMap\\r\" proof- interpret \: \ \ by (rule assms(1)) interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms) from assms(2) interpret FUNCT: tiny_category \ \cat_FUNCT \ \ (cat_Set \)\ by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) let ?Hom_r = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ from assms have [cat_cs_simps]: "cf_of_cf_map \ (cat_Set \) (cf_map \) = \" by (cs_concl cs_simp: cat_FUNCT_cs_simps) note Yoneda = cat_Yoneda_Lemma[OF assms(3,4)] show ?thesis proof ( intro cat_Set_is_arr_isomorphismI cat_Set_is_arrI arr_SetI, unfold cat_cs_simps cf_map_components ) show "vfsequence (ntcf_Yoneda_arrow \ \ (cf_map \) r)" unfolding ntcf_Yoneda_arrow_def by simp show "vcard (ntcf_Yoneda_arrow \ \ (cf_map \) r) = 3\<^sub>\" unfolding ntcf_Yoneda_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (ntcf_Yoneda_arrow \ \ (cf_map \) r\ArrVal\) = \\ObjMap\\r\" unfolding cat_cs_simps cf_map_components ntcf_Yoneda_arrow_components by (intro vsv.vsv_vrange_VLambda_app', unfold Yoneda(2)) ( use assms(4) in \ cs_concl cs_simp: cat_cs_simps bij_betwD(2)[OF bij_betw_ntcf_of_ntcf_arrow_Hom] cs_intro: cat_cs_intros \ )+ then show "\\<^sub>\ (ntcf_Yoneda_arrow \ \ (cf_map \) r\ArrVal\) \\<^sub>\ \\ObjMap\\r\" by auto from assms(4) show "v11 (ntcf_Yoneda_arrow \ \ (cf_map \) r\ArrVal\)" unfolding ntcf_Yoneda_arrow_components by ( intro v11.v11_VLambda_v11_bij_betw_comp, unfold cat_cs_simps \.Yoneda_map_vdomain; intro Yoneda bij_betw_ntcf_of_ntcf_arrow_Hom ) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms(4) show "Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map ?Hom_r) (cf_map \) \\<^sub>\ Vset \" by (intro FUNCT.cat_Hom_in_Vset) ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) from assms(4) have "\\ObjMap\\r\ \\<^sub>\ Vset \" by (cs_concl cs_intro: cat_cs_intros) then show "\\ObjMap\\r\ \\<^sub>\ Vset \" by (auto simp: assms(2) Vset_trans Vset_in_mono) qed (auto intro: cat_cs_intros) qed lemma (in category) cat_ntcf_Yoneda_arrow_is_arr_isomoprhism': assumes "\ \" and "\ \\<^sub>\ \" and "\' = cf_map \" and "B = \\ObjMap\\r\" and "A = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) (cf_map \)" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "ntcf_Yoneda_arrow \ \ \' r : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" using assms(1,2,6,7) unfolding assms(3-5) by (rule cat_ntcf_Yoneda_arrow_is_arr_isomoprhism) lemmas [cat_arrow_cs_intros] = category.cat_ntcf_Yoneda_arrow_is_arr_isomoprhism' lemma (in category) cat_ntcf_Yoneda_arrow_is_arr: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "ntcf_Yoneda_arrow \ \ (cf_map \) r : Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) (cf_map \) \\<^bsub>cat_Set \\<^esub> \\ObjMap\\r\" by ( rule cat_Set_is_arr_isomorphismD[ OF cat_ntcf_Yoneda_arrow_is_arr_isomoprhism[OF assms] ] ) lemma (in category) cat_ntcf_Yoneda_arrow_is_arr'[cat_cs_intros]: assumes "\ \" and "\ \\<^sub>\ \" and "\' = cf_map \" and "B = \\ObjMap\\r\" and "A = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-))) (cf_map \)" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "ntcf_Yoneda_arrow \ \ \' r : A \\<^bsub>cat_Set \\<^esub> B" using assms(1,2,6,7) unfolding assms(3-5) by (rule cat_ntcf_Yoneda_arrow_is_arr) lemmas [cat_arrow_cs_intros] = category.cat_ntcf_Yoneda_arrow_is_arr' subsection\Commutativity law for the Yoneda natural transformation arrow\ lemma (in category) cat_ntcf_Yoneda_arrow_commutativity: assumes "\ \" and "\ \\<^sub>\ \" and "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "f : a \\<^bsub>\\<^esub> b" shows "ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_hom (cat_FUNCT \ \ (cat_Set \)) [ntcf_arrow Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-), ntcf_arrow \]\<^sub>\ = cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a" proof- let ?hom = \ cf_hom (cat_FUNCT \ \ (cat_Set \)) [ntcf_arrow Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-), ntcf_arrow \]\<^sub>\ \ interpret \: \ \ by (rule assms(1)) interpret \: is_ntcf \ \ \cat_Set \\ \ \ \ by (rule assms(3)) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ interpret cat_Set_\\: subcategory \ \cat_Set \\ \cat_Set \\ by (rule subcategory_cat_Set_cat_Set[OF assms(1,2)]) from assms(2,4) have \b_\f: "ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom : Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-))) (cf_map \) \\<^bsub>cat_Set \\<^esub> \\ObjMap\\b\" by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_prod_cs_intros cat_op_intros cat_FUNCT_cs_intros ) from assms(2,4) have \f_\a: "cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a : Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-))) (cf_map \) \\<^bsub>cat_Set \\<^esub> \\ObjMap\\b\" by (cs_concl cs_intro: cat_cs_intros cat_Set_\\.subcat_is_arrD) show ?thesis proof(rule arr_Set_eqI[of \]) from \b_\f show arr_Set_\b_\f: "arr_Set \ (ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom)" by (auto dest: cat_Set_is_arrD(1)) from \b_\f have dom_lhs: "\\<^sub>\ ((ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom)\ArrVal\) = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-))) (cf_map \)" by (cs_concl cs_simp: cat_cs_simps)+ interpret \f_\a: arr_Set \ \ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom\ by (rule arr_Set_\b_\f) from \f_\a show arr_Set_\f_\a: "arr_Set \ ( cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a )" by (auto dest: cat_Set_is_arrD(1)) from \f_\a have dom_rhs: "\\<^sub>\ ( ( cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a )\ArrVal\ ) = Hom (cat_FUNCT \ \ (cat_Set \)) (cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-))) (cf_map \)" by (cs_concl cs_simp: cat_cs_simps) show "(ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom)\ArrVal\ = ( cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a )\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff) fix \ assume prems: "\ : cf_map Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> cf_map \" from assms(4) have [cat_cs_simps]: "cf_of_cf_map \ (cat_Set \) (cf_map Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)) = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" "cf_of_cf_map \ (cat_Set \) (cf_map \) = \" 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] interpret \: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)\ \ \ntcf_of_ntcf_arrow \ (cat_Set \) \\ by (rule \(1)) have \\_eq_\\: "\\ArrMap\\f\\ArrVal\\\\NTMap\\a\\ArrVal\\A\\ = \\NTMap\\b\\ArrVal\\\\ArrMap\\f\\ArrVal\\A\\" if "A \\<^sub>\ \\ObjMap\\a\" for A using ArrVal_eq_helper[ OF \.ntcf_Comp_commute[OF assms(4), symmetric], where a=A ] assms(4) that by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from \(1) assms(2,3,4) have \a_CId_a: "\\NTMap\\a\\ArrVal\\\\CId\\a\\ \\<^sub>\ \\ObjMap\\a\" by (subst \) ( cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros ) have \f_\a_eq_\b: "\\ArrMap\\f\\ArrVal\\\\NTMap\\a\\ArrVal\\h\\ = \\NTMap\\b\\ArrVal\\f \\<^sub>A\<^bsub>\\<^esub> h\" if "h : a \\<^bsub>\\<^esub> a" for h using ArrVal_eq_helper[ OF \.ntcf_Comp_commute[OF assms(4), symmetric], where a=h ] that assms(4) category_axioms by ( cs_prems cs_simp: cat_FUNCT_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros ) from \(1) assms(2,3,4) \a_CId_a category_axioms show "(ntcf_Yoneda_arrow \ \ (cf_map \) b \\<^sub>A\<^bsub>cat_Set \\<^esub> ?hom)\ArrVal\\\\ = ( cf_eval_arrow \ (ntcf_arrow \) f \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda_arrow \ \ (cf_map \) a )\ArrVal\\\\" by (subst (1 2) \(2)) (*very slow*) ( cs_concl cs_simp: \f_\a_eq_\b \\_eq_\\ cat_FUNCT_cs_simps cat_cs_simps cat_op_simps cs_intro: cat_Set_\\.subcat_is_arrD cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_prod_cs_intros cat_op_intros )+ qed (use arr_Set_\b_\f arr_Set_\f_\a in auto) qed (use \b_\f \f_\a in \cs_concl cs_simp: cat_cs_simps\)+ qed subsection\Yoneda Lemma: naturality\ subsubsection\ The Yoneda natural transformation: definition and elementary properties \ text\ The main result of this subsection corresponds to the corollary to the Yoneda Lemma on page 61 in \cite{mac_lane_categories_2010}. \ definition ntcf_Yoneda :: "V \ V \ V \ V" where "ntcf_Yoneda \ \ \ = [ ( \\r\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\. ntcf_Yoneda_arrow \ \ (\r\0\) (\r\1\<^sub>\\) ), cf_nt \ \ (cf_id \), cf_eval \ \ \, cat_FUNCT \ \ (cat_Set \) \\<^sub>C \, cat_Set \ ]\<^sub>\" text\Components.\ lemma ntcf_Yoneda_components: shows "ntcf_Yoneda \ \ \\NTMap\ = ( \\r\\<^sub>\(cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\. ntcf_Yoneda_arrow \ \ (\r\0\) (\r\1\<^sub>\\) )" and [cat_cs_simps]: "ntcf_Yoneda \ \ \\NTDom\ = cf_nt \ \ (cf_id \)" and [cat_cs_simps]: "ntcf_Yoneda \ \ \\NTCod\ = cf_eval \ \ \" and [cat_cs_simps]: "ntcf_Yoneda \ \ \\NTDGDom\ = cat_FUNCT \ \ (cat_Set \) \\<^sub>C \" and [cat_cs_simps]: "ntcf_Yoneda \ \ \\NTDGCod\ = cat_Set \" unfolding ntcf_Yoneda_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda ntcf_Yoneda_components(1) |vsv ntcf_Yoneda_NTMap_vsv[cat_cs_intros]| |vdomain ntcf_Yoneda_NTMap_vdomain[cat_cs_intros]| lemma (in category) ntcf_Yoneda_NTMap_app[cat_cs_simps]: assumes "\ \" and "\ \\<^sub>\ \" and "\r = [cf_map \, r]\<^sub>\" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "r \\<^sub>\ \\Obj\" shows "ntcf_Yoneda \ \ \\NTMap\\\r\ = ntcf_Yoneda_arrow \ \ (cf_map \) r" proof- interpret \: \ \ by (rule assms(1)) interpret \: is_functor \ \ \cat_Set \\ \ by (rule assms(4)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) interpret FUNCT: category \ \cat_FUNCT \ \ (cat_Set \)\ by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) from assms(5) have "[cf_map \, r]\<^sub>\ \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" by ( cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros ) then show ?thesis unfolding assms(3) ntcf_Yoneda_components by (simp add: nat_omega_simps) qed lemmas [cat_cs_simps] = category.ntcf_Yoneda_NTMap_app subsubsection\The Yoneda natural transformation is a natural transformation\ lemma (in category) cat_ntcf_Yoneda_is_ntcf: assumes "\ \" and "\ \\<^sub>\ \" shows "ntcf_Yoneda \ \ \ : cf_nt \ \ (cf_id \) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_eval \ \ \ : cat_FUNCT \ \ (cat_Set \) \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: \ \ by (rule assms(1)) interpret \\: category \ \ by (rule category.cat_category_if_ge_Limit) (use assms(2) in \cs_concl cs_intro: cat_cs_intros\)+ from assms(2) interpret FUNCT: category \ \cat_FUNCT \ \ (cat_Set \)\ by ( cs_concl cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros ) show ?thesis proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence (ntcf_Yoneda \ \ \)" unfolding ntcf_Yoneda_def by simp show "vcard (ntcf_Yoneda \ \ \) = 5\<^sub>\" unfolding ntcf_Yoneda_def by (simp add: nat_omega_simps) show ntcf_Yoneda_\r: "ntcf_Yoneda \ \ \\NTMap\\\r\ : cf_nt \ \ (cf_id \)\ObjMap\\\r\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ObjMap\\\r\" if "\r \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" for \r proof- from that obtain \ r where \r_def: "\r = [\, r]\<^sub>\" and \: "\ \\<^sub>\ cf_maps \ \ (cat_Set \)" and r: "r \\<^sub>\ \\Obj\" by ( auto elim: cat_prod_2_ObjE[rotated 2] simp: cat_FUNCT_cs_simps intro: cat_cs_intros ) from \ obtain \ where \_def: "\ = cf_map \" and \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by clarsimp from assms(2) \ r show ?thesis unfolding \r_def \_def by ( cs_concl! cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_arrow_cs_intros ) qed show "ntcf_Yoneda \ \ \\NTMap\\\r\ : cf_nt \ \ (cf_id \)\ObjMap\\\r\ \\<^bsub>cat_Set \\<^esub> cf_eval \ \ \\ObjMap\\\r\" if "\r \\<^sub>\ (cat_FUNCT \ \ (cat_Set \) \\<^sub>C \)\Obj\" for \r by (rule is_arr_isomorphismD[OF ntcf_Yoneda_\r[OF that]]) show "ntcf_Yoneda \ \ \\NTMap\\\b\ \\<^sub>A\<^bsub>cat_Set \\<^esub> cf_nt \ \ (cf_id \)\ArrMap\\\f\ = cf_eval \ \ \\ArrMap\\\f\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ntcf_Yoneda \ \ \\NTMap\\\a\" if \f: "\f : \a \\<^bsub>cat_FUNCT \ \ (cat_Set \) \\<^sub>C \\<^esub> \b" for \a \b \f proof- obtain \ f \ a \ b where \f_def: "\f = [\, f]\<^sub>\" and \a_def: "\a = [\, a]\<^sub>\" and \b_def: "\b = [\, b]\<^sub>\" and \: "\ : \ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> \" and f: "f : a \\<^bsub>\\<^esub> b" by ( auto intro: cat_prod_2_is_arrE[rotated 2, OF \f] FUNCT.category_axioms \\.category_axioms ) note \ = cat_FUNCT_is_arrD[OF \] note [cat_cs_simps] = cat_ntcf_Yoneda_arrow_commutativity[OF assms \(1) f, folded \(2,3,4)] from \(1) assms(2) f show ?thesis unfolding \f_def \a_def \b_def by (subst (1 2) \(2), use nothing in \subst \(3), subst \(4)\) ( cs_concl cs_simp: \(2,3,4)[symmetric] cat_cs_simps cs_intro: cat_cs_intros )+ qed qed (use assms(2) in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\)+ qed subsection\\Hom\-map\ text\ This subsection presents some of the results stated as Corollary 2 in subsection 1.15 in \cite{bodo_categories_1970} and the corollary following the statement of the Yoneda Lemma on page 61 in \cite{mac_lane_categories_2010} in a variety of forms. \ subsubsection\Definition and elementary properties\ text\ The following function makes an explicit appearance in subsection 1.15 in \cite{bodo_categories_1970}. \ definition ntcf_Hom_map :: "V \ V \ V \ V \ V" where "ntcf_Hom_map \ \ a b = (\f\\<^sub>\Hom \ a b. Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-))" text\Elementary properties.\ mk_VLambda ntcf_Hom_map_def |vsv ntcf_Hom_map_vsv| |vdomain ntcf_Hom_map_vdomain[cat_cs_simps]| |app ntcf_Hom_map_app[unfolded in_Hom_iff, cat_cs_simps]| subsubsection\\Hom\-map is a bijection\ lemma (in category) cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique: \\The following lemma approximately corresponds to the corollary on page 61 in \cite{mac_lane_categories_2010}.\ assumes "r \\<^sub>\ \\Obj\" and "s \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\ : s \\<^bsub>\\<^esub> r" and "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\,-)" and "\f. \ f \\<^sub>\ \\Arr\; \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \ \ f = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\" proof- interpret \: is_ntcf \ \ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)\ \ by (rule assms(3)) let ?Y_Hom_s = \Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\ note Yoneda = cat_Yoneda_Lemma[OF cat_cf_Hom_snd_is_functor[OF assms(2)] assms(1)] interpret Y: v11 \?Y_Hom_s\ by (rule Yoneda(1)) from category_axioms assms have \_in_vdomain: "\ \\<^sub>\ \\<^sub>\ (?Y_Hom_s)" by (cs_concl cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros) then have "?Y_Hom_s\\\ \\<^sub>\ \\<^sub>\ (?Y_Hom_s)" by (simp add: Y.vsv_vimageI2) from this category_axioms assms show Ym_\: "?Y_Hom_s\\\ : s \\<^bsub>\\<^esub> r" unfolding Yoneda(2) by (cs_prems_step cs_simp: cat_cs_simps cat_op_simps)+ simp then have "?Y_Hom_s\\\ \\<^sub>\ \\Arr\" by (simp add: cat_cs_intros) have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Y_Hom_s\\\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (intro cat_ntcf_Hom_snd_is_ntcf Ym_\) from assms Ym_\ this category_axioms assms have "(?Y_Hom_s)\\<^sub>\\?Y_Hom_s\\\\ = Yoneda_arrow \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r (?Y_Hom_s\\\)" by (intro category.inv_Yoneda_map_app) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) then have "(?Y_Hom_s)\\<^sub>\\?Y_Hom_s\\\\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Y_Hom_s\\\,-)" by (simp add: ntcf_Hom_snd_def'[OF Ym_\]) with \_in_vdomain show "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Y_Hom_s\\\,-)" by auto fix f assume prems: "f \\<^sub>\ \\Arr\" "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" then obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto have "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (rule cat_ntcf_Hom_snd_is_ntcf[OF f, folded prems(2)]) with f \.ntcf_NTDom \.ntcf_NTCod assms cat_is_arrD(2,3)[OF f] have ba_simps: "b = r" "a = s" by ( simp_all add: prems(2) cat_cf_Hom_snd_inj cat_ntcf_Hom_snd_components(2,3) ) from f have "f : s \\<^bsub>\\<^esub> r" unfolding ba_simps . with category_axioms show "f = ?Y_Hom_s\\\" unfolding prems(2) by (cs_concl cs_simp: cat_cs_simps cat_op_simps) qed lemma (in category) cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique: assumes "r \\<^sub>\ \\Obj\" and "s \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\ : r \\<^bsub>\\<^esub> s" and "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\)" and "\f. \ f \\<^sub>\ \\Arr\; \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) \ \ f = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\" by ( intro category.cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[ OF category_op, unfolded cat_op_simps cat_op_cat_ntcf_Hom_snd, OF assms(1,2), unfolded assms(1,2)[THEN cat_op_cat_cf_Hom_snd], OF assms(3) ] )+ lemma (in category) cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique': assumes "r \\<^sub>\ \\Obj\" and "s \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "\!f. f \\<^sub>\ \\Arr\ \ \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" using cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms] by blast lemma (in category) cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique': assumes "r \\<^sub>\ \\Obj\" and "s \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,r) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "\!f. f \\<^sub>\ \\Arr\ \ \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)" using cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique[OF assms] by blast lemma (in category) cat_ntcf_Hom_snd_inj: assumes "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" and "g : a \\<^bsub>\\<^esub> b" and "f : a \\<^bsub>\\<^esub> b" shows "g = f" proof- from assms have "Yoneda_map \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)) b\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-)\ = Yoneda_map \ (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)) b\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\" by simp from this assms category_axioms show "g = f" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) simp (*slow*) qed lemma (in category) cat_ntcf_Hom_fst_inj: assumes "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,g) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)" and "g : a \\<^bsub>\\<^esub> b" and "f : a \\<^bsub>\\<^esub> b" shows "g = f" proof- from category.cat_ntcf_Hom_snd_inj [ OF category_op, unfolded cat_op_simps, unfolded cat_op_cat_ntcf_Hom_snd, OF assms ] show ?thesis . qed lemma (in category) cat_ntcf_Hom_map: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "v11 (ntcf_Hom_map \ \ a b)" and "\\<^sub>\ (ntcf_Hom_map \ \ a b) = these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" and "(ntcf_Hom_map \ \ a b)\\<^sub>\ = (\\\\<^sub>\these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-). Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\)" proof- show "v11 (ntcf_Hom_map \ \ a b)" proof(rule vsv.vsv_valeq_v11I, unfold ntcf_Hom_map_vdomain in_Hom_iff) show "vsv (ntcf_Hom_map \ \ a b)" unfolding ntcf_Hom_map_def by simp fix g f assume prems: "g : a \\<^bsub>\\<^esub> b" "f : a \\<^bsub>\\<^esub> b" "ntcf_Hom_map \ \ a b\g\ = ntcf_Hom_map \ \ a b\f\" from prems(3,1,2) have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) with prems(1,2) show "g = f" by (intro cat_ntcf_Hom_snd_inj[of g f]) qed then interpret Hm: v11 \ntcf_Hom_map \ \ a b\ . show Hm_vrange: "\\<^sub>\ (ntcf_Hom_map \ \ a b) = these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" proof(intro vsubset_antisym) show "\\<^sub>\ (ntcf_Hom_map \ \ a b) \\<^sub>\ these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" by ( unfold ntcf_Hom_map_def, intro vrange_VLambda_vsubset, unfold these_ntcfs_iff in_Hom_iff, intro cat_ntcf_Hom_snd_is_ntcf ) show "these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>\ \\<^sub>\ (ntcf_Hom_map \ \ a b)" proof(intro vsubsetI, unfold these_ntcfs_iff) fix \ assume prems: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" note unique = cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) prems] from unique(1) have "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\ \\<^sub>\ \\<^sub>\ (ntcf_Hom_map \ \ a b)" by (cs_concl cs_simp: cat_cs_simps) moreover from cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(1,2)[OF assms(2,1) prems] have \_def: "\ = ntcf_Hom_map \ \ a b\Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\\" by (cs_concl cs_simp: cat_cs_simps) ultimately show "\ \\<^sub>\ \\<^sub>\ (ntcf_Hom_map \ \ a b)" by force qed qed show "(ntcf_Hom_map \ \ a b)\\<^sub>\ = ( \\\\<^sub>\these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-). Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\ )" proof ( rule vsv_eqI, unfold vdomain_vconverse vdomain_VLambda Hm_vrange these_ntcfs_iff ) from Hm.v11_axioms show "vsv ((ntcf_Hom_map \ \ a b)\\<^sub>\)" by auto show "vsv ( \\\\<^sub>\these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-). Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\ )" by simp fix \ assume prems: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" then have \: "\ \\<^sub>\ these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" unfolding these_ntcfs_iff by simp show "(ntcf_Hom_map \ \ a b)\\<^sub>\\\\ = ( \\\\<^sub>\these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-). Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\ )\\\" proof ( intro Hm.v11_vconverse_app, unfold ntcf_Hom_map_vdomain in_Hom_iff beta[OF \] ) note unique = cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) prems] show "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\ : a \\<^bsub>\\<^esub> b" by (rule unique(1)) then show "ntcf_Hom_map \ \ a b\Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\\ = \" by (cs_concl cs_simp: unique(2)[symmetric] cat_cs_simps) qed qed simp qed subsubsection\Inverse of a \Hom\-map\ lemma (in category) inv_ntcf_Hom_map_v11: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "v11 ((ntcf_Hom_map \ \ a b)\\<^sub>\)" using cat_ntcf_Hom_map(1)[OF assms] by (simp add: v11.v11_vconverse) lemma (in category) inv_ntcf_Hom_map_vdomain: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "\\<^sub>\ ((ntcf_Hom_map \ \ a b)\\<^sub>\) = these_ntcfs \ \ (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-)" unfolding cat_ntcf_Hom_map(3)[OF assms] by simp lemmas [cat_cs_simps] = category.inv_ntcf_Hom_map_vdomain lemma (in category) inv_ntcf_Hom_map_app: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "(ntcf_Hom_map \ \ a b)\\<^sub>\\\\ = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) b\\\" using assms(3) unfolding cat_ntcf_Hom_map(3)[OF assms(1,2)] by simp lemmas [cat_cs_simps] = category.inv_ntcf_Hom_map_app lemma inv_ntcf_Hom_map_vrange: "\\<^sub>\ ((ntcf_Hom_map \ \ a b)\\<^sub>\) = Hom \ a b" unfolding ntcf_Hom_map_def by simp subsubsection\\Hom\-natural transformation and isomorphisms\ text\ This subsection presents further results that were stated as Corollary 2 in subsection 1.15 in \cite{bodo_categories_1970}. \ lemma (in category) cat_is_arr_isomorphism_ntcf_Hom_snd_is_iso_ntcf: assumes "f : s \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) : 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>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- from assms obtain g where iso_g: "g : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> s" and gf: "g \\<^sub>A\<^bsub>\\<^esub> f = \\CId\\s\" and fg: "f \\<^sub>A\<^bsub>\\<^esub> g = \\CId\\r\" by ( auto intro: cat_the_inverse_Comp_CId_left cat_the_inverse_Comp_CId_right cat_the_inverse_is_arr_isomorphism' ) then have g: "g : r \\<^bsub>\\<^esub> s" by auto show ?thesis proof(intro is_arr_isomorphism_is_iso_ntcf) from assms have f: "f : s \\<^bsub>\\<^esub> r" by auto with category_axioms show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms g show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms f g have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g \\<^sub>A\<^bsub>\\<^esub> f,-)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) also from category_axioms f g have "\ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" by (cs_concl cs_simp: gf cat_cs_simps cs_intro: cat_cs_intros) finally show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" by simp from category_axioms f g have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f \\<^sub>A\<^bsub>\\<^esub> g,-)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) also from category_axioms f g have "\ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" by (cs_concl cs_simp: fg cat_cs_simps cs_intro: cat_cs_intros) finally show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" by simp qed qed lemma (in category) cat_is_arr_isomorphism_ntcf_Hom_fst_is_iso_ntcf: assumes "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> s" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) : 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>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- from assms have r: "r \\<^sub>\ \\Obj\" and s: "s \\<^sub>\ \\Obj\" by auto from category.cat_is_arr_isomorphism_ntcf_Hom_snd_is_iso_ntcf [ OF category_op, unfolded cat_op_simps, OF assms, unfolded category.cat_op_cat_cf_Hom_snd[OF category_axioms r] category.cat_op_cat_cf_Hom_snd[OF category_axioms s] category.cat_op_cat_ntcf_Hom_snd[OF category_axioms] ] show ?thesis. qed lemma (in category) cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique: assumes "r \\<^sub>\ \\Obj\" and "s \\<^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>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\ : s \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" and "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\,-)" and "\f. \ f \\<^sub>\ \\Arr\; \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \ \ f = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\" proof- let ?Ym_\ = \Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\\ and ?Ym_inv_\ = \Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) s\inv_ntcf \\\ from assms(3) have \: "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by auto from iso_ntcf_is_arr_isomorphism[OF assms(3)] have iso_inv_\: "inv_ntcf \ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and [simp]: "\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" and [simp]: "inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" by auto from iso_inv_\ have inv_\: "inv_ntcf \ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by auto note unique = cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(1,2) \] and inv_unique = cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) inv_\] have Ym_\: "?Ym_\ : s \\<^bsub>\\<^esub> r" by (rule unique(1)) show "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\,-)" and "\f. \ f \\<^sub>\ \\Arr\; \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \ \ f = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\" by (intro unique)+ show "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\\\ : s \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" proof(intro is_arr_isomorphismI[OF Ym_\, of \?Ym_inv_\\] is_inverseI) show Ym_inv_\: "?Ym_inv_\ : r \\<^bsub>\\<^esub> s" by (rule inv_unique(1)) have "ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) = \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F inv_ntcf \" by simp also have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_\,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_inv_\,-)" by (subst unique(2), subst inv_unique(2)) simp also from category_axioms Ym_\ inv_unique(1) assms(3) have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_inv_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_\,-)" by (cs_concl cs_simp: cat_cs_simps) finally have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_inv_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_\,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-)" by simp also from category_axioms assms(1,2) have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\CId\\s\,-)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) finally have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_inv_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_\,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\CId\\s\,-)" by simp then show "?Ym_inv_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_\ = \\CId\\s\" by (rule cat_ntcf_Hom_snd_inj) ( all\ use category_axioms Ym_\ Ym_inv_\ assms in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\ \ ) have "ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-) = inv_ntcf \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F \" by simp also have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_inv_\,-) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_\,-)" by (subst unique(2), subst inv_unique(2)) simp also from category_axioms Ym_\ inv_unique(1) have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_inv_\,-)" by (cs_concl cs_simp: cat_cs_simps) finally have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_inv_\,-) = ntcf_id Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(r,-)" by simp also from category_axioms assms(1,2) have "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\CId\\r\,-)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) finally have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?Ym_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_inv_\,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\CId\\r\,-)" by simp then show "?Ym_\ \\<^sub>A\<^bsub>\\<^esub> ?Ym_inv_\ = \\CId\\r\" by (rule cat_ntcf_Hom_snd_inj) ( all\ use category_axioms Ym_\ Ym_inv_\ assms in \cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros\ \ ) qed (intro Ym_\) qed lemma (in category) cat_ntcf_Hom_fst_is_iso_ntcf_Hom_fst_unique: assumes "r \\<^sub>\ \\Obj\" and "s \\<^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>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\ : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> s" and "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\)" and "\f. \ f \\<^sub>\ \\Arr\; \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) \ \ f = Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\\\" by ( intro category.cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique[ OF category_op, unfolded cat_op_simps cat_op_cat_ntcf_Hom_snd, OF assms(1,2), unfolded assms(1,2)[THEN cat_op_cat_cf_Hom_snd], OF assms(3) ] )+ lemma (in category) cat_is_arr_isomorphism_if_ntcf_Hom_snd_is_iso_ntcf: assumes "f : s \\<^bsub>\\<^esub> r" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) : 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>\(s,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "f : s \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> r" proof- from assms(1) have r: "r \\<^sub>\ \\Obj\" and s: "s \\<^sub>\ \\Obj\" by auto note unique = cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique[OF r s assms(2)] from unique(1) have Ym_Hf: "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(s,-) r\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)\ : s \\<^bsub>\\<^esub> r" by auto from unique(1) show ?thesis unfolding cat_ntcf_Hom_snd_inj[OF unique(2) assms(1) Ym_Hf, symmetric] by simp qed lemma (in category) cat_is_arr_isomorphism_if_ntcf_Hom_fst_is_iso_ntcf: assumes "f : r \\<^bsub>\\<^esub> s" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f) : 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>\(-,s) : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "f : r \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> s" proof- from assms(1) have r: "r \\<^sub>\ \\Obj\" and s: "s \\<^sub>\ \\Obj\" by auto note unique = cat_ntcf_Hom_fst_is_iso_ntcf_Hom_fst_unique[OF r s assms(2)] from unique(1) have Ym_Hf: "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,s) r\Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(-,f)\ : r \\<^bsub>\\<^esub> s" by auto from unique(1) show ?thesis unfolding cat_ntcf_Hom_fst_inj[OF unique(2) assms(1) Ym_Hf, symmetric] by simp qed subsubsection\ The relationship between a \Hom\-natural transformation and the compositions of a \Hom\-natural transformation and a natural transformation \ lemma (in category) cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" and "c \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\NTMap\\b, c\\<^sub>\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\\c\" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from assms(2) have b: "b \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp from category_axioms assms(1,3) b show ?thesis by ( cs_concl cs_simp: cat_ntcf_lcomp_Hom_component_is_Yoneda_component cat_cs_simps cs_intro: cat_cs_intros cat_op_intros ) qed lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app lemma (in category) cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "b \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis proof(rule ntcf_eqI[of \]) from category_axioms assms show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) from assms this have dom_lhs: "\\<^sub>\ ((Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms this have dom_rhs: "\\<^sub>\ (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\) = \\Obj\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "(Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ \\Obj\" with category_axioms assms show "(Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F)\NTMap\\a\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\\a\" by (cs_concl cs_simp: cat_cs_simps) qed (use assms(2) in \auto intro: cat_cs_intros\) qed simp_all qed lemmas [cat_cs_simps] = category.cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd subsubsection\ The relationship between the \Hom\-natural isomorphisms and the compositions of a \Hom\-natural isomorphism and a natural transformation \ lemma (in category) cat_ntcf_lcomp_Hom_if_ntcf_Hom_snd_is_iso_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\b. b \\<^sub>\ \\Obj\ \ Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>C\<^sub>F : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" if "b \\<^sub>\ \\Obj\" for b unfolding cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd[OF assms(1) that] cat_cf_lcomp_Hom_cf_Hom_snd[OF \.NTDom.is_functor_axioms that] cat_cf_lcomp_Hom_cf_Hom_snd[OF \.NTCod.is_functor_axioms that] by (intro assms(2) that) from is_iso_ntcf_if_bnt_proj_snd_is_iso_ntcf[ OF \.NTDom.HomDom.category_op category_axioms cat_ntcf_lcomp_Hom_is_ntcf[OF assms(1)], unfolded cat_op_simps, OF this ] show ?thesis . qed lemma (in category) cat_ntcf_Hom_snd_if_ntcf_lcomp_Hom_is_iso_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "b \\<^sub>\ \\Obj\" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_ntcf \ \ \ \ \ \ by (rule assms(1)) from category_axioms assms show ?thesis by ( fold cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd[OF assms(1,3)] cat_cf_lcomp_Hom_cf_Hom_snd[OF \.NTDom.is_functor_axioms assms(3)] cat_cf_lcomp_Hom_cf_Hom_snd[OF \.NTCod.is_functor_axioms assms(3)], intro bnt_proj_snd_is_iso_ntcf_if_is_iso_ntcf ) (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros) qed subsection\Yoneda map for arbitrary functors\ text\ The concept of the Yoneda map for arbitrary functors was developed based on the function that was used in the statement of Lemma 3 in subsection 1.15 in \cite{bodo_categories_1970}. \ definition af_Yoneda_map :: "V \ V \ V \ V" where "af_Yoneda_map \ \ \ = (\\\\<^sub>\these_ntcfs \ (\\HomDom\) (\\HomCod\) \ \. Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-))" text\Elementary properties.\ 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 af_Yoneda_map_def[where \=\ and \=\, unfolded \.cf_HomDom \.cf_HomCod] |vsv af_Yoneda_map_vsv| |vdomain af_Yoneda_map_vdomain[cat_cs_simps]| |app af_Yoneda_map_app[unfolded these_ntcfs_iff, cat_cs_simps]| end subsection\Yoneda arrow for arbitrary functors\ subsubsection\Definition and elementary properties\ text\ The following natural transformation is used in the proof of Lemma 3 in subsection 1.15 in \cite{bodo_categories_1970}. \ definition af_Yoneda_arrow :: "V \ V \ V \ V \ V" where "af_Yoneda_arrow \ \ \ \ = [ ( \b\\<^sub>\(\\HomDom\)\Obj\. Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(\\ObjMap\\b\,-) (\\ObjMap\\b\)\ \\<^bsub>op_cat (\\HomDom\),\\HomCod\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ ), \, \, \\HomDom\, \\HomCod\ ]\<^sub>\" text\Components.\ lemma af_Yoneda_arrow_components: shows "af_Yoneda_arrow \ \ \ \\NTMap\ = ( \b\\<^sub>\\\HomDom\\Obj\. Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\\HomCod\(\\ObjMap\\b\,-) (\\ObjMap\\b\)\ \\<^bsub>op_cat (\\HomDom\),\\HomCod\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ )" and "af_Yoneda_arrow \ \ \ \\NTDom\ = \" and "af_Yoneda_arrow \ \ \ \\NTCod\ = \" and "af_Yoneda_arrow \ \ \ \\NTDGDom\ = \\HomDom\" and "af_Yoneda_arrow \ \ \ \\NTDGCod\ = \\HomCod\" unfolding af_Yoneda_arrow_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda af_Yoneda_arrow_components(1) |vsv af_Yoneda_arrow_NTMap_vsv| context fixes \ \ \ \ assumes \: "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" begin interpretation \: is_functor \ \ \ \ by (rule \) mk_VLambda af_Yoneda_arrow_components(1)[where \=\, unfolded \.cf_HomDom \.cf_HomCod] |vdomain af_Yoneda_arrow_NTMap_vdomain[cat_cs_simps]| |app af_Yoneda_arrow_NTMap_app[cat_cs_simps]| end lemma (in category) cat_af_Yoneda_arrow_is_ntcf: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "af_Yoneda_arrow \ \ \ \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?Set = \cat_Set \\ and ?Ym = \ \b. Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) (\\ObjMap\\b\)\\\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F\ \ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \op_cat \ \\<^sub>C \\ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ \ by (rule assms) have comm[unfolded cat_op_simps]: - "(\\NTMap\ \c, d\\<^sub>\)\ArrVal\\f \\<^sub>A\<^bsub>\\<^esub> (q \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)\ = - f \\<^sub>A\<^bsub>\\<^esub> ((\\NTMap\ \a, b\\<^sub>\)\ArrVal\\q\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)" + "(\\NTMap\\c, d\\<^sub>\)\ArrVal\\f \\<^sub>A\<^bsub>\\<^esub> (q \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)\ = + f \\<^sub>A\<^bsub>\\<^esub> ((\\NTMap\\a, b\\<^sub>\)\ArrVal\\q\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)" if "g : a \\<^bsub>op_cat \\<^esub> c" and "f : b \\<^bsub>\\<^esub> d" and "q : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" for q g f a b c d proof- from that(1) have g: "g : c \\<^bsub>\\<^esub> a" unfolding cat_op_simps by simp from category_axioms assms g that(2) have ab: "[a, b]\<^sub>\ \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" by ( cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from \.ntcf_NTMap_is_arr[OF ab] category_axioms assms g that(2) have \ab: "\\NTMap\\a, b\\<^sub>\ : Hom \ (\\ObjMap\\a\) b \\<^bsub>cat_Set \\<^esub> Hom \ (\\ObjMap\\a\) b" by ( cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) have \_abq: "(\\NTMap\\a, b\\<^sub>\)\ArrVal\\q\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> b" by ( rule cat_Set_ArrVal_app_vrange[ OF \ab, unfolded in_Hom_iff, OF that(3) ] ) have "[g, f]\<^sub>\ : [a, b]\<^sub>\ \\<^bsub>op_cat \ \\<^sub>C \\<^esub> [c, d]\<^sub>\" by ( rule cat_prod_2_is_arrI[ OF \.HomDom.category_op category_axioms that(1,2) ] ) then have "\\NTMap\\c, d\\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ArrMap\\g, f\\<^sub>\ = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ArrMap\\g, f\\<^sub>\ \\<^sub>A\<^bsub>cat_Set \\<^esub> \\NTMap\\a, b\\<^sub>\" by (rule is_ntcf.ntcf_Comp_commute[OF assms(3)]) then have "(\\NTMap\\c, d\\<^sub>\ \\<^sub>A\<^bsub>?Set\<^esub> ?H\\ArrMap\\g, f\\<^sub>\)\ArrVal\\q\ = (?H\\ArrMap\\g, f\\<^sub>\ \\<^sub>A\<^bsub>?Set\<^esub> \\NTMap\\a, b\\<^sub>\)\ArrVal\\q\" by auto from this that(2,3) assms category_axioms \.HomDom.category_axioms \.HomDom.category_op category_op g \ab \_abq show "(\\NTMap\\c, d\\<^sub>\)\ArrVal\\f \\<^sub>A\<^bsub>\\<^esub> (q \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)\ = f \\<^sub>A\<^bsub>\\<^esub> ((\\NTMap\\a, b\\<^sub>\)\ArrVal\\q\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\g\)" by ( cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) qed show ?thesis proof(rule is_ntcfI') show "vfsequence (af_Yoneda_arrow \ \ \ \)" unfolding af_Yoneda_arrow_def by simp show "vcard (af_Yoneda_arrow \ \ \ \) = 5\<^sub>\" unfolding af_Yoneda_arrow_def by (simp add: nat_omega_simps) have \b: "\\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" if "b \\<^sub>\ \\Obj\" for b by ( rule bnt_proj_snd_is_ntcf [ OF \.HomDom.category_op category_axioms assms(3), unfolded cat_op_simps, OF that, unfolded cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(1) that] cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(2) that] ] ) show "af_Yoneda_arrow \ \ \ \\NTMap\\b\ : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" if "b \\<^sub>\ \\Obj\" for b proof- let ?\b = \\\ObjMap\\b\\ and ?\b = \\\ObjMap\\b\\ and ?\\b = \\\CId\\\\ObjMap\\b\\\ from that have \\b: "?\\b : ?\b \\<^bsub>\\<^esub> ?\b" by (auto simp: cat_cs_intros) from assms that have "[b, ?\b]\<^sub>\ \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from \.ntcf_NTMap_is_arr[OF this] category_axioms assms that have \_b\b: "\\NTMap\\b, ?\b\\<^sub>\ : Hom \ ?\b ?\b \\<^bsub>cat_Set \\<^esub> Hom \ ?\b ?\b" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) from \\b have \_b\b_\\b: "(\\NTMap\\b, ?\b\\<^sub>\)\ArrVal\\?\\b\ : ?\b \\<^bsub>\\<^esub> ?\b" by (rule cat_Set_ArrVal_app_vrange[OF \_b\b, unfolded in_Hom_iff]) with category_axioms assms that \b[OF that] show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) qed show "af_Yoneda_arrow \ \ \ \\NTMap\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> af_Yoneda_arrow \ \ \ \\NTMap\\a\" if "f : a \\<^bsub>\\<^esub> b" for a b f proof- from that have a: "a \\<^sub>\ \\Obj\" and b: "b \\<^sub>\ \\Obj\" by auto let ?\a = \\\CId\\a\\ and ?\b = \\\CId\\b\\ and ?\a = \\\ObjMap\\a\\ and ?\b = \\\ObjMap\\b\\ and ?\a = \\\ObjMap\\a\\ and ?\b = \\\ObjMap\\b\\ and ?\\a = \\\CId\\\\ObjMap\\a\\\ and ?\\b = \\\CId\\\\ObjMap\\b\\\ from that have \\a: "?\\a : ?\a \\<^bsub>\\<^esub> ?\a" by (auto intro: cat_cs_intros) from that have \\b: "?\\b : ?\b \\<^bsub>\\<^esub> ?\b" by (auto intro: cat_cs_intros) from that have \a: "?\a : a \\<^bsub>\\<^esub> a" by (auto intro: cat_cs_intros) from assms that have "[b, ?\b]\<^sub>\ \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from \.ntcf_NTMap_is_arr[OF this] category_axioms assms that have \_b\b: "\\NTMap\\b, ?\b\\<^sub>\ : Hom \ ?\b ?\b \\<^bsub>cat_Set \\<^esub> Hom \ ?\b ?\b" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) from \\b have \_b\b_\\b: "(\\NTMap\\b, ?\b\\<^sub>\)\ArrVal\\?\\b\ : ?\b \\<^bsub>\\<^esub> ?\b" by (rule cat_Set_ArrVal_app_vrange[OF \_b\b, unfolded in_Hom_iff]) from assms that have "[a, ?\a]\<^sub>\ \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros ) from \.ntcf_NTMap_is_arr[OF this] category_axioms assms that have \_a\a: "\\NTMap\\a, ?\a\\<^sub>\ : Hom \ ?\a ?\a \\<^bsub>cat_Set \\<^esub> Hom \ ?\a ?\a" by ( cs_prems cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_prod_cs_intros ) from \\a have \_a\a_\\a: "(\\NTMap\\a, ?\a\\<^sub>\)\ArrVal\\?\\a\ : ?\a \\<^bsub>\\<^esub> ?\a" by (rule cat_Set_ArrVal_app_vrange[OF \_a\a, unfolded in_Hom_iff]) from comm[OF \a \.cf_ArrMap_is_arr[OF that] \\a] category_axioms assms that \_a\a_\\a have \_a_\b[symmetric, cat_cs_simps]: "(\\NTMap\\a, ?\b\\<^sub>\)\ArrVal\\\\ArrMap\\f\\ = \\ArrMap\\f\ \\<^sub>A\<^bsub>\\<^esub> (\\NTMap\\a, ?\a\\<^sub>\)\ArrVal\\?\\a\" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from comm[OF that \\b \\b] category_axioms assms that \_b\b_\\b have \_a_\b'[cat_cs_simps]: "(\\NTMap\\a, ?\b\\<^sub>\)\ArrVal\\\\ArrMap\\f\\ = (\\NTMap\\b, ?\b\\<^sub>\)\ArrVal\\?\\b\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\" by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from category_axioms assms that \b[OF a] \b[OF b] show ?thesis by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros) qed qed (auto simp: af_Yoneda_arrow_components cat_cs_simps intro: cat_cs_intros) qed lemma (in category) cat_af_Yoneda_arrow_is_ntcf': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "\ = \" and "\' = \" and "\' = \" shows "af_Yoneda_arrow \ \ \ \ : \' \\<^sub>C\<^sub>F \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4-6) by (rule cat_af_Yoneda_arrow_is_ntcf) lemmas [cat_cs_intros] = category.cat_af_Yoneda_arrow_is_ntcf' subsubsection\Yoneda Lemma for arbitrary functors\ text\ The following lemmas correspond to variants of the elements of Lemma 3 in subsection 1.15 in \cite{bodo_categories_1970}. \ lemma (in category) cat_af_Yoneda_map_af_Yoneda_arrow_app: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(af_Yoneda_arrow \ \ \ \-,-)" proof- let ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?aYa = \\\. af_Yoneda_arrow \ \ \ \\ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_ntcf \ \op_cat \ \\<^sub>C \\ \cat_Set \\ \?H\\ \?H\\ \ by (rule assms(3)) interpret aY\: is_ntcf \ \ \ \ \ \?aYa \\ by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms]) interpret HY\: is_ntcf \ \op_cat \ \\<^sub>C \\ \cat_Set \\ \?H\\ \?H\\ \Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(?aYa \-,-)\ by (rule cat_ntcf_lcomp_Hom_is_ntcf[OF aY\.is_ntcf_axioms]) show [cat_cs_simps]: "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(?aYa \-,-)" proof ( rule sym, rule ntcf_eqI[OF HY\.is_ntcf_axioms assms(3)], rule vsv_eqI; (intro HY\.NTMap.vsv_axioms \.NTMap.vsv_axioms)?; (unfold \.ntcf_NTMap_vdomain HY\.ntcf_NTMap_vdomain)? ) fix bc assume prems': "bc \\<^sub>\ (op_cat \ \\<^sub>C \)\Obj\" then obtain b c where bc_def: "bc = [b, c]\<^sub>\" and op_b: "b \\<^sub>\ op_cat \\Obj\" and c: "c \\<^sub>\ \\Obj\" by (auto intro: cat_prod_2_ObjE cat_cs_intros) from op_b have b: "b \\<^sub>\ \\Obj\" unfolding cat_op_simps by simp then have \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" and \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" by (auto intro: cat_cs_intros) have Ym_\: "Yoneda_map \ Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) (\\ObjMap\\b\)\ \\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F \ = ?aYa \\NTMap\\b\" unfolding af_Yoneda_arrow_NTMap_app[OF assms(1) b] by simp from bnt_proj_snd_is_ntcf [ OF \.HomDom.category_op category_axioms assms(3) op_b, unfolded cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(1) b] cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(2) b] ] have \b: "\\<^bsub>op_cat \,\\<^esub>(b,-)\<^sub>N\<^sub>T\<^sub>C\<^sub>F : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by simp from c show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(?aYa \-,-)\NTMap\\bc\ = \\NTMap\\bc\" unfolding bc_def cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app[OF aY\.is_ntcf_axioms b c] cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(2)[ OF \b \b \b, unfolded Ym_\, symmetric ] by (cs_concl cs_simp: cat_cs_simps) qed simp_all qed lemma (in category) cat_af_Yoneda_Lemma: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "v11 (af_Yoneda_map \ \ \)" and "\\<^sub>\ (af_Yoneda_map \ \ \) = these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" and "(af_Yoneda_map \ \ \)\\<^sub>\ = ( \\\\<^sub>\these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-). af_Yoneda_arrow \ \ \ \ )" proof- let ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?aYm = \af_Yoneda_map \ \ \\ and ?aYa = \\\. af_Yoneda_arrow \ \ \ \\ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) show v11_aY: "v11 ?aYm" proof ( intro vsv.vsv_valeq_v11I, unfold af_Yoneda_map_vdomain[OF assms] these_ntcfs_iff ) show "vsv (af_Yoneda_map \ \ \)" by (rule af_Yoneda_map_vsv[OF assms]) fix \ \ assume prems: "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" "?aYm\\\ = ?aYm\\\" interpret \: is_ntcf \ \ \ \ \ \ by (rule prems(1)) interpret \: is_ntcf \ \ \ \ \ \ by (rule prems(2)) from prems(3) have H\_H\: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-)" unfolding af_Yoneda_map_app[OF assms prems(1)] af_Yoneda_map_app[OF assms prems(2)] by simp show "\ = \" proof ( rule ntcf_eqI[OF prems(1,2)], rule vsv_eqI, unfold \.ntcf_NTMap_vdomain \.ntcf_NTMap_vdomain ) fix b assume prems': "b \\<^sub>\ \\Obj\" from prems' have \b: "\\NTMap\\b\ : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and \b: "\\NTMap\\b\ : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" and \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" and \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" by (auto intro: cat_cs_intros cat_prod_cs_intros) have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)" proof ( rule ntcf_eqI [ OF cat_ntcf_Hom_snd_is_ntcf[OF \b] cat_ntcf_Hom_snd_is_ntcf[OF \b] ] ) show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\" proof ( rule vsv_eqI, unfold ntcf_Hom_snd_NTMap_vdomain[OF \b] ntcf_Hom_snd_NTMap_vdomain[OF \b] ) fix c assume prems'': "c \\<^sub>\ \\Obj\" note H = cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\\c\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-)\NTMap\\c\" unfolding H[OF prems(1) prems' prems'', symmetric] H[OF prems(2) prems' prems'', symmetric] H\_H\ by simp qed ( simp_all add: ntcf_Hom_snd_NTMap_vsv[OF \b] ntcf_Hom_snd_NTMap_vsv[OF \b] ) qed simp_all with \b \b show "\\NTMap\\b\ = \\NTMap\\b\" by (auto intro: cat_ntcf_Hom_snd_inj) qed auto qed interpret aYm: v11 ?aYm by (rule v11_aY) have [cat_cs_simps]: "?aYm\?aYa \\ = \" if "\ : ?H\ \\<^sub>C\<^sub>F ?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" for \ using category_axioms assms that by ( cs_concl cs_simp: cat_af_Yoneda_map_af_Yoneda_arrow_app[symmetric] cat_cs_simps cs_intro: cat_cs_intros ) show aYm_vrange: "\\<^sub>\ ?aYm = these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\" proof(intro vsubset_antisym) show "\\<^sub>\ ?aYm \\<^sub>\ these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\" proof ( rule vsv.vsv_vrange_vsubset, unfold these_ntcfs_iff af_Yoneda_map_vdomain[OF assms] ) fix \ assume "\ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" with category_axioms assms show "?aYm\\\ : ?H\ \\<^sub>C\<^sub>F ?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) qed (auto intro: af_Yoneda_map_vsv) show "these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\ \\<^sub>\ \\<^sub>\ ?aYm" proof(rule vsubsetI, unfold these_ntcfs_iff) fix \ assume prems: "\ : ?H\ \\<^sub>C\<^sub>F ?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" interpret aY\: is_ntcf \ \ \ \ \ \?aYa \\ by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms prems]) from prems have \_def: "\ = ?aYm\?aYa \\" by (cs_concl cs_simp: cat_cs_simps) from assms aY\.is_ntcf_axioms have "?aYa \ \\<^sub>\ \\<^sub>\ ?aYm" by (cs_concl cs_simp: these_ntcfs_iff cat_cs_simps) then show "\ \\<^sub>\ \\<^sub>\ ?aYm" by (subst \_def, intro aYm.vsv_vimageI2) auto qed qed show "?aYm\\<^sub>\ = (\\\\<^sub>\these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\. ?aYa \)" proof ( rule vsv_eqI, unfold vdomain_vconverse vdomain_VLambda aYm_vrange these_ntcfs_iff ) from aYm.v11_axioms show "vsv ((af_Yoneda_map \ \ \)\\<^sub>\)" by auto fix \ assume prems: "\ : ?H\ \\<^sub>C\<^sub>F ?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" then have \: "\ \\<^sub>\ these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\" by simp show "?aYm\\<^sub>\\\\ = (\\\\<^sub>\these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) ?H\ ?H\. ?aYa \)\\\" proof ( intro aYm.v11_vconverse_app, unfold beta[OF \] af_Yoneda_map_vdomain[OF assms] these_ntcfs_iff ) from prems show \_def: "?aYm\?aYa \\ = \" by (cs_concl cs_simp: cat_cs_simps) show "?aYa \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms prems]) qed qed simp_all qed subsubsection\Inverse of the Yoneda map for arbitrary functors\ lemma (in category) inv_af_Yoneda_map_v11: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "v11 ((af_Yoneda_map \ \ \)\\<^sub>\)" using cat_af_Yoneda_Lemma(1)[OF assms] by (simp add: v11.v11_vconverse) lemma (in category) inv_af_Yoneda_map_vdomain: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((af_Yoneda_map \ \ \)\\<^sub>\) = these_ntcfs \ (op_cat \ \\<^sub>C \) (cat_Set \) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" unfolding cat_af_Yoneda_Lemma(3)[OF assms] by simp lemmas [cat_cs_simps] = category.inv_af_Yoneda_map_vdomain lemma (in category) inv_af_Yoneda_map_app: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "(af_Yoneda_map \ \ \)\\<^sub>\\\\ = af_Yoneda_arrow \ \ \ \" using assms(3) unfolding cat_af_Yoneda_Lemma(3)[OF assms(1,2)] by simp lemmas [cat_cs_simps] = category.inv_af_Yoneda_map_app lemma (in category) inv_af_Yoneda_map_vrange: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((af_Yoneda_map \ \ \)\\<^sub>\) = these_ntcfs \ \ \ \ \" proof- interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms show ?thesis unfolding af_Yoneda_map_def by (simp add: cat_cs_simps) qed subsubsection\Yoneda map for arbitrary functors and natural isomorphisms\ text\ The following lemmas correspond to variants of the elements of Lemma 3 in subsection 1.15 in \cite{bodo_categories_1970}. \ lemma (in category) cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf: assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: is_iso_ntcf \ \ \ \ \ \ by (rule assms(1)) show ?thesis proof(intro cat_ntcf_lcomp_Hom_if_ntcf_Hom_snd_is_iso_ntcf) fix b assume "b \\<^sub>\ \\Obj\" then show "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\NTMap\\b\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by ( auto intro!: cat_is_arr_isomorphism_ntcf_Hom_snd_is_iso_ntcf cat_arrow_cs_intros ) qed (auto simp: cat_cs_intros) qed lemma (in category) cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf': assumes "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ = \" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" and "\' = Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" and "\' = op_cat \ \\<^sub>C \" and "\' = cat_Set \" shows "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(\-,-) : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" using assms(1) unfolding assms(2-6) by (rule cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf) lemmas [cat_cs_intros] = category.cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf' lemma (in category) cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" shows "af_Yoneda_arrow \ \ \ \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" proof- let ?aYa = \af_Yoneda_arrow \ \ \ \\ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) interpret \: is_iso_ntcf \ \op_cat \ \\<^sub>C \\ \cat_Set \\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ \ by (rule assms(3)) from assms(1,2) \.is_ntcf_axioms have \_def: "\ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(?aYa-,-)" by (cs_concl cs_simp: cat_af_Yoneda_map_af_Yoneda_arrow_app[symmetric]) from category_axioms assms have aYa: "?aYa : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_intro: cat_cs_intros) have Hom_aYa: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>(?aYa-,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (auto intro: assms(3) simp add: \_def[symmetric]) have Hb: "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(?aYa\NTMap\\b\,-) : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\\ObjMap\\b\,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" if "b \\<^sub>\ \\Obj\" for b by ( rule cat_ntcf_Hom_snd_if_ntcf_lcomp_Hom_is_iso_ntcf[ OF aYa Hom_aYa that ] ) show ?thesis proof(intro is_iso_ntcfI) from category_axioms assms show "af_Yoneda_arrow \ \ \ \ : \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) fix b assume prems: "b \\<^sub>\ \\Obj\" then have \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" and \b: "\\ObjMap\\b\ \\<^sub>\ \\Obj\" by (auto intro: cat_cs_intros) from assms(1,2) aYa prems have aYa_b: "?aYa\NTMap\\b\ : \\ObjMap\\b\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by (cs_concl cs_intro: cat_cs_intros cs_simp: cat_cs_simps) show "af_Yoneda_arrow \ \ \ \\NTMap\\b\ : \\ObjMap\\b\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \\ObjMap\\b\" by ( rule cat_is_arr_isomorphism_if_ntcf_Hom_snd_is_iso_ntcf[ OF aYa_b Hb[OF prems] ] ) qed qed lemma (in category) cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" and "\ = \" and "\' = \" and "\' = \" shows "af_Yoneda_arrow \ \ \ \ : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4-6) by (rule cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf) lemmas [cat_cs_intros] = category.cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf' lemma (in category) cat_iso_functor_if_cf_lcomp_Hom_iso_functor: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^bsub>\\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" shows "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" proof- let ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?aYa = \\\. af_Yoneda_arrow \ \ \ \\ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms(3) obtain \ \ \
where \: "\ : ?H\ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o ?H\ : \ \\\<^sub>C\<^bsub>\\<^esub> \
" by auto interpret \: is_iso_ntcf \ \ \
?H\ ?H\ \ by (rule \) from category_axioms assms have "?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) then have \_def: "\ = op_cat \ \\<^sub>C \" and \
_def: "\
= cat_Set \" by (force simp: cat_cs_simps)+ note \ = \[unfolded \_def \
_def] from \ have "\ : ?H\ \\<^sub>C\<^sub>F ?H\ : op_cat \ \\<^sub>C \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros ntcf_cs_intros) from category_axioms assms \ have "af_Yoneda_arrow \ \ \ \ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) then have "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" by (clarsimp intro!: iso_functorI) then show ?thesis by (rule iso_functor_sym) qed lemma (in category) cat_cf_lcomp_Hom_iso_functor_if_iso_functor: assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" shows "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^bsub>\\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)" proof- let ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?H\ = \Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-)\ and ?aYa = \\\. af_Yoneda_arrow \ \ \ \\ interpret \: is_functor \ \ \ \ by (rule assms(1)) interpret \: is_functor \ \ \ \ by (rule assms(2)) from assms obtain \' \' \ where \: "\ : \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" by auto interpret \: is_iso_ntcf \ \' \' \ \ \ by (rule \) from assms \.NTDom.is_functor_axioms have \'_def: "\' = \" and \'_def: "\' = \" by fast+ note \ = \[unfolded \'_def \'_def] show ?thesis by (rule iso_functor_sym) ( intro iso_functorI[ OF cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf[OF \] ] ) qed lemma (in category) cat_cf_lcomp_Hom_iso_functor_if_iso_functor': assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>C\<^sub>F\<^bsub>\\<^esub> \" and "\' = \" and "\' = \" shows "Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(\-,-) \\<^sub>C\<^sub>F\<^bsub>\\<^esub> Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\'\<^esub>\'(\-,-)" using assms(1-3) unfolding assms(4,5) by (rule cat_cf_lcomp_Hom_iso_functor_if_iso_functor) lemmas [cat_cs_intros] = category.cat_cf_lcomp_Hom_iso_functor_if_iso_functor' subsection\The Yoneda Functor\ subsubsection\Definition and elementary properties\ text\See Chapter III-2 in \cite{mac_lane_categories_2010}.\ definition Yoneda_functor :: "V \ V \ V" where "Yoneda_functor \ \
= [ (\r\\<^sub>\op_cat \
\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\
(r,-))), (\f\\<^sub>\op_cat \
\Arr\. ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\
(f,-))), op_cat \
, cat_FUNCT \ \
(cat_Set \) ]\<^sub>\" text\Components.\ lemma Yoneda_functor_components: shows "Yoneda_functor \ \
\ObjMap\ = (\r\\<^sub>\op_cat \
\Obj\. cf_map (Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\
(r,-)))" and "Yoneda_functor \ \
\ArrMap\ = (\f\\<^sub>\op_cat \
\Arr\. ntcf_arrow (Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\
(f,-)))" and "Yoneda_functor \ \
\HomDom\ = op_cat \
" and "Yoneda_functor \ \
\HomCod\ = cat_FUNCT \ \
(cat_Set \)" unfolding Yoneda_functor_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda Yoneda_functor_components(1) |vsv Yoneda_functor_ObjMap_vsv[cat_cs_intros]| |vdomain Yoneda_functor_ObjMap_vdomain[cat_cs_simps]| |app Yoneda_functor_ObjMap_app[cat_cs_simps]| lemma (in category) Yoneda_functor_ObjMap_vrange: "\\<^sub>\ (Yoneda_functor \ \\ObjMap\) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Obj\" proof ( unfold Yoneda_functor_components, rule vrange_VLambda_vsubset, unfold cat_op_simps ) fix c assume "c \\<^sub>\ \\Obj\" with category_axioms show "cf_map Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(c,-) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Obj\" unfolding cat_op_simps cat_FUNCT_components by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed subsubsection\Arrow map\ mk_VLambda Yoneda_functor_components(2) |vsv Yoneda_functor_ArrMap_vsv[cat_cs_intros]| |vdomain Yoneda_functor_ArrMap_vdomain[cat_cs_simps]| |app Yoneda_functor_ArrMap_app[cat_cs_simps]| lemma (in category) Yoneda_functor_ArrMap_vrange: "\\<^sub>\ (Yoneda_functor \ \\ArrMap\) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Arr\" proof ( unfold Yoneda_functor_components, rule vrange_VLambda_vsubset, unfold cat_op_simps ) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where f: "f : a \\<^bsub>\\<^esub> b" by auto define \ where "\ = \ + \" have \\: "\ \" and \\: "\ \\<^sub>\ \" by (simp_all add: \_\_\\ \.intro \_Limit_\\ \_\_\\ \_def) from tiny_category_cat_FUNCT category_axioms \\ \\ f show "ntcf_arrow Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-) \\<^sub>\ cat_FUNCT \ \ (cat_Set \)\Arr\" unfolding cat_op_simps by (cs_concl cs_intro: cat_cs_intros cat_FUNCT_cs_intros) qed subsubsection\The Yoneda Functor is a fully faithful functor\ lemma (in category) cat_Yoneda_functor_is_functor: assumes "\ \" and "\ \\<^sub>\ \" shows "Yoneda_functor \ \ : op_cat \ \\\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> cat_FUNCT \ \ (cat_Set \)" proof ( intro is_ff_functorI is_ft_functorI' is_fl_functorI' vsubset_antisym vsubsetI, unfold cat_op_simps in_Hom_iff, tactic\distinct_subgoals_tac\ ) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) let ?Yf = \Yoneda_functor \ \\ and ?FUNCT = \cat_FUNCT \ \ (cat_Set \)\ show Yf: "?Yf : op_cat \ \\\<^sub>C\<^bsub>\\<^esub> ?FUNCT" proof(intro is_functorI') show "vfsequence ?Yf" unfolding Yoneda_functor_def by simp from assms have "category \ \" by (intro cat_category_if_ge_Limit) then show "category \ (op_cat \)" by (intro category.category_op) from assms show "category \ ?FUNCT" by (cs_concl cs_intro: cat_small_cs_intros tiny_category_cat_FUNCT) show "vcard ?Yf = 4\<^sub>\" unfolding Yoneda_functor_def by (simp add: nat_omega_simps) show "\\<^sub>\ (?Yf\ObjMap\) \\<^sub>\ ?FUNCT\Obj\" by (rule Yoneda_functor_ObjMap_vrange) show "?Yf\ArrMap\\f\ : ?Yf\ObjMap\\a\ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> ?Yf\ObjMap\\b\" if "f : a \\<^bsub>op_cat \\<^esub> b" for a b f using that category_axioms unfolding cat_op_simps by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros) show "?Yf\ArrMap\\g \\<^sub>A\<^bsub>op_cat \\<^esub> f\ = ?Yf\ArrMap\\g\ \\<^sub>A\<^bsub>?FUNCT\<^esub> ?Yf\ArrMap\\f\" if "g : b \\<^bsub>op_cat \\<^esub> c" and "f : a \\<^bsub>op_cat \\<^esub> b" for b c g a f using that category_axioms unfolding cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) show "?Yf\ArrMap\\op_cat \\CId\\c\\ = ?FUNCT\CId\\?Yf\ObjMap\\c\\" if "c \\<^sub>\ op_cat \\Obj\" for c using that category_axioms unfolding cat_op_simps by ( cs_concl cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros ) qed (auto simp: assms(1) Yoneda_functor_components \.intro \_Limit_\\ \_\_\\) interpret Yf: is_functor \ \op_cat \\ \?FUNCT\ \?Yf\ by (rule Yf) show "v11 (?Yf\ArrMap\ \\<^sup>l\<^sub>\ Hom \ b a)" if "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" for a b proof- from that have dom_Y_ba: "\\<^sub>\ (?Yf\ArrMap\ \\<^sup>l\<^sub>\ Hom \ b a) = Hom \ b a" by ( fastforce simp: cat_op_simps in_Hom_iff vdomain_vlrestriction Yoneda_functor_components ) show "v11 (?Yf\ArrMap\ \\<^sup>l\<^sub>\ Hom \ b a)" proof(intro vsv.vsv_valeq_v11I, unfold dom_Y_ba in_Hom_iff) fix g f assume prems: "g : b \\<^bsub>\\<^esub> a" "f : b \\<^bsub>\\<^esub> a" "(?Yf\ArrMap\ \\<^sup>l\<^sub>\ Hom \ b a)\g\ = (?Yf\ArrMap\ \\<^sup>l\<^sub>\ Hom \ b a)\f\" from prems(3) category_axioms prems(1,2) Yoneda_functor_ArrMap_vsv[of \ \] have "Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(g,-) = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" by ( cs_prems cs_simp: V_cs_simps cat_cs_simps cat_op_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros ) from this prems(1,2) show "g = f" by (rule cat_ntcf_Hom_snd_inj) qed (auto simp: Yoneda_functor_components) qed fix a b assume prems: "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" show "\ : ?Yf\ObjMap\\a\ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> ?Yf\ObjMap\\b\" if "\ \\<^sub>\ ?Yf\ArrMap\ `\<^sub>\ Hom \ b a" for \ proof- from that obtain f where "?Yf\ArrMap\\f\ = \" and f: "f : b \\<^bsub>\\<^esub> a" by (force elim!: Yf.ArrMap.vsv_vimageE) then have \_def: "\ = ntcf_arrow Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" unfolding Yoneda_functor_ArrMap_app[ unfolded cat_op_simps, OF cat_is_arrD(1)[OF f] ] by (simp add: cat_cs_simps cat_op_simps cat_cs_intros) from category_axioms f show ?thesis unfolding \_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_FUNCT_cs_intros ) qed show "\ \\<^sub>\ ?Yf\ArrMap\ `\<^sub>\ Hom \ b a" if "\ : ?Yf\ObjMap\\a\ \\<^bsub>cat_FUNCT \ \ (cat_Set \)\<^esub> ?Yf\ObjMap\\b\" for \ proof- note \ = cat_FUNCT_is_arrD[OF that] from \(1) category_axioms prems have ntcf_\: "ntcf_of_ntcf_arrow \ (cat_Set \) \ : Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(a,-) \\<^sub>C\<^sub>F Hom\<^sub>O\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(b,-) : \ \\\<^sub>C\<^bsub>\\<^esub> cat_Set \" by (subst (asm) \(3), use nothing in \subst (asm) \(4)\) ( cs_prems cs_simp: cat_cs_simps cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_op_intros cat_FUNCT_cs_intros ) from cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(1,2)[OF prems ntcf_\] obtain f where f: "f : b \\<^bsub>\\<^esub> a" and \_def: "ntcf_of_ntcf_arrow \ (cat_Set \) \ = Hom\<^sub>A\<^sub>.\<^sub>C\<^bsub>\\<^esub>\(f,-)" by auto from \(2) f show "\ \\<^sub>\ Yoneda_functor \ \\ArrMap\ `\<^sub>\ Hom \ b a" unfolding \_def by (intro Yf.ArrMap.vsv_vimage_eqI[of f]) (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros)+ qed qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/czh_digraphs/CZH_DG_DGHM.thy b/thys/CZH_Foundations/czh_digraphs/CZH_DG_DGHM.thy --- a/thys/CZH_Foundations/czh_digraphs/CZH_DG_DGHM.thy +++ b/thys/CZH_Foundations/czh_digraphs/CZH_DG_DGHM.thy @@ -1,1814 +1,1844 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Homomorphism of digraphs\ theory CZH_DG_DGHM imports CZH_DG_Digraph begin subsection\Background\ named_theorems dghm_cs_simps named_theorems dghm_cs_intros named_theorems dg_cn_cs_simps named_theorems dg_cn_cs_intros named_theorems dghm_field_simps definition ObjMap :: V where [dghm_field_simps]: "ObjMap = 0" definition ArrMap :: V where [dghm_field_simps]: "ArrMap = 1\<^sub>\" definition HomDom :: V where [dghm_field_simps]: "HomDom = 2\<^sub>\" definition HomCod :: V where [dghm_field_simps]: "HomCod = 3\<^sub>\" subsection\Definition and elementary properties\ text\ A homomorphism of digraphs, as presented in this work, can be seen as a generalization of the concept of a functor between categories, as presented in Chapter I-3 in \cite{mac_lane_categories_2010}, to digraphs. The generalization is performed by removing the axioms (1) from the definition. It is expected that the resulting definition is consistent with the conventional notion of a homomorphism of digraphs in graph theory, but further details are considered to be outside of the scope of this work. The definition of a digraph homomorphism is parameterized by a limit ordinal \\\ such that \\ < \\. Such digraph homomorphisms are referred to either as \\\-digraph homomorphisms or homomorphisms of \\\-digraphs. Following \cite{mac_lane_categories_2010}, all digraph homomorphisms are covariant (see Chapter II-2). However, a special notation is adapted for the digraph homomorphisms from an opposite digraph. Normally, such digraph homomorphisms will be referred to as the contravariant digraph homomorphisms, but this convention will not be enforced. \ locale is_dghm = \ \ + vfsequence \ + HomDom: digraph \ \ + HomCod: digraph \ \ for \ \ \ \ + assumes dghm_length[dg_cs_simps]: "vcard \ = 4\<^sub>\" and dghm_HomDom[dg_cs_simps]: "\\HomDom\ = \" and dghm_HomCod[dg_cs_simps]: "\\HomCod\ = \" and dghm_ObjMap_vsv: "vsv (\\ObjMap\)" and dghm_ArrMap_vsv: "vsv (\\ArrMap\)" and dghm_ObjMap_vdomain[dg_cs_simps]: "\\<^sub>\ (\\ObjMap\) = \\Obj\" and dghm_ObjMap_vrange: "\\<^sub>\ (\\ObjMap\) \\<^sub>\ \\Obj\" and dghm_ArrMap_vdomain[dg_cs_simps]: "\\<^sub>\ (\\ArrMap\) = \\Arr\" and dghm_ArrMap_is_arr: "f : a \\<^bsub>\\<^esub> b \ \\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" syntax "_is_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>D\<^sub>G\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" \ "CONST is_dghm \ \ \ \" abbreviation (input) is_cn_dghm :: "V \ V \ V \ V \ bool" where "is_cn_dghm \ \ \ \ \ \ : op_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" syntax "_is_cn_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \<^sub>D\<^sub>G\\\ _)\ [51, 51, 51] 51) translations "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" \ "CONST is_cn_dghm \ \ \ \" abbreviation all_dghms :: "V \ V" where "all_dghms \ \ set {\. \\ \. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \}" abbreviation dghms :: "V \ V \ V \ V" where "dghms \ \ \ \ set {\. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \}" sublocale is_dghm \ ObjMap: vsv \\\ObjMap\\ rewrites "\\<^sub>\ (\\ObjMap\) = \\Obj\" by (rule dghm_ObjMap_vsv) (simp add: dg_cs_simps) sublocale is_dghm \ ArrMap: vsv \\\ArrMap\\ rewrites "\\<^sub>\ (\\ArrMap\) = \\Arr\" by (rule dghm_ArrMap_vsv) (simp add: dg_cs_simps) lemmas [dg_cs_simps] = is_dghm.dghm_HomDom is_dghm.dghm_HomCod is_dghm.dghm_ObjMap_vdomain is_dghm.dghm_ArrMap_vdomain lemma (in is_dghm) dghm_ArrMap_is_arr''[dg_cs_intros]: assumes "f : a \\<^bsub>\\<^esub> b" and "\f = \\ArrMap\\f\" shows "\f : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" using assms(1) unfolding assms(2) by (rule dghm_ArrMap_is_arr) lemma (in is_dghm) dghm_ArrMap_is_arr'[dg_cs_intros]: assumes "f : a \\<^bsub>\\<^esub> b" and "A = \\ObjMap\\a\" and "B = \\ObjMap\\b\" shows "\\ArrMap\\f\ : A \\<^bsub>\\<^esub> B" using assms(1) unfolding assms(2,3) by (rule dghm_ArrMap_is_arr) lemmas [dg_cs_intros] = is_dghm.dghm_ArrMap_is_arr' text\Rules.\ lemma (in is_dghm) is_dghm_axioms'[dg_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>D\<^sub>G\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_dghm_axioms) mk_ide rf is_dghm_def[unfolded is_dghm_axioms_def] |intro is_dghmI| |dest is_dghmD[dest]| |elim is_dghmE[elim]| lemmas [dg_cs_intros] = is_dghmD(3,4) text\Elementary properties.\ lemma dghm_eqI: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \
" and "\\ObjMap\ = \\ObjMap\" and "\\ArrMap\ = \\ArrMap\" and "\ = \" and "\ = \
" shows "\ = \" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \
\ by (rule assms(2)) show ?thesis proof(rule vsv_eqI) have dom: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: dg_cs_simps V_cs_simps) from assms(5,6) have sup: "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by (simp_all add: dg_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 (cs_concl cs_simp: dg_cs_simps V_cs_simps cs_intro: V_cs_intros)+ qed lemma (in is_dghm) dghm_def: "\ = [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: dg_cs_simps V_cs_simps) have dom_rhs: "\\<^sub>\ [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^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) lemma (in is_dghm) dghm_ObjMap_app_in_HomCod_Obj[dg_cs_intros]: assumes "a \\<^sub>\ \\Obj\" shows "\\ObjMap\\a\ \\<^sub>\ \\Obj\" using assms dghm_ObjMap_vrange by (blast dest: ObjMap.vsv_vimageI2) lemmas [dg_cs_intros] = is_dghm.dghm_ObjMap_app_in_HomCod_Obj lemma (in is_dghm) dghm_ArrMap_vrange: "\\<^sub>\ (\\ArrMap\) \\<^sub>\ \\Arr\" proof(rule vsv.vsv_vrange_vsubset, unfold dg_cs_simps) fix f assume "f \\<^sub>\ \\Arr\" then obtain a b where "f : a \\<^bsub>\\<^esub> b" by auto then have "\\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by (cs_concl cs_intro: dg_cs_intros) then show "\\ArrMap\\f\ \\<^sub>\ \\Arr\" by auto qed auto lemma (in is_dghm) dghm_ArrMap_app_in_HomCod_Arr[dg_cs_intros]: assumes "a \\<^sub>\ \\Arr\" shows "\\ArrMap\\a\ \\<^sub>\ \\Arr\" using assms dghm_ArrMap_vrange by (blast dest: ArrMap.vsv_vimageI2) lemmas [dg_cs_intros] = is_dghm.dghm_ArrMap_app_in_HomCod_Arr text\Size.\ lemma (in is_dghm) dghm_ObjMap_vsubset_Vset: "\\ObjMap\ \\<^sub>\ Vset \" by ( rule ObjMap.vbrelation_Limit_vsubset_VsetI, insert dghm_ObjMap_vrange HomCod.dg_Obj_vsubset_Vset ) (auto intro!: HomDom.dg_Obj_vsubset_Vset) lemma (in is_dghm) dghm_ArrMap_vsubset_Vset: "\\ArrMap\ \\<^sub>\ Vset \" by ( rule ArrMap.vbrelation_Limit_vsubset_VsetI, insert dghm_ArrMap_vrange HomCod.dg_Arr_vsubset_Vset ) (auto intro!: HomDom.dg_Arr_vsubset_Vset) lemma (in is_dghm) dghm_ObjMap_in_Vset: assumes "\ \\<^sub>\ \" shows "\\ObjMap\ \\<^sub>\ Vset \" by (meson assms dghm_ObjMap_vsubset_Vset Vset_in_mono vsubset_in_VsetI) lemma (in is_dghm) dghm_ArrMap_in_Vset: assumes "\ \\<^sub>\ \" shows "\\ArrMap\ \\<^sub>\ Vset \" by (meson assms dghm_ArrMap_vsubset_Vset Vset_in_mono vsubset_in_VsetI) lemma (in is_dghm) dghm_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "\ \\<^sub>\ Vset \" proof- interpret \: \ \ by (rule assms(1)) note [dg_cs_intros] = dghm_ObjMap_in_Vset dghm_ArrMap_in_Vset HomDom.dg_in_Vset HomCod.dg_in_Vset from assms(2) show ?thesis by (subst dghm_def) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros) qed lemma (in is_dghm) dghm_is_dghm_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof(rule is_dghmI) from is_dghm_axioms assms show "digraph \ \" by (cs_concl cs_intro: digraph.dg_digraph_if_ge_Limit dg_cs_intros) from is_dghm_axioms assms show "digraph \ \" by (cs_concl cs_intro: digraph.dg_digraph_if_ge_Limit dg_cs_intros) qed (cs_concl cs_simp: dg_cs_simps cs_intro: assms(1) dg_cs_intros V_cs_intros dghm_ObjMap_vrange)+ lemma small_all_dghms[simp]: "small {\. \\ \. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \}" proof(cases \\ \\) case True from is_dghm.dghm_in_Vset show ?thesis by (intro down[of _ \Vset (\ + \)\] subsetI) (auto simp: True \.\_Limit_\\ \.\_\_\\ \.intro \.\_\_\\) next case False then have "{\. \\ \. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \} = {}" by auto then show ?thesis by simp qed lemma (in is_dghm) dghm_in_Vset_7: "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof- note [folded VPow_iff, folded Vset_succ[OF Ord_\], dg_cs_intros] = dghm_ObjMap_vsubset_Vset dghm_ArrMap_vsubset_Vset from HomDom.dg_digraph_in_Vset_4 have [dg_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.dg_digraph_in_Vset_4 have [dg_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 dghm_def, succ_of_numeral) ( cs_concl cs_simp: plus_V_succ_right V_cs_simps dg_cs_simps cs_intro: dg_cs_intros V_cs_intros ) qed lemma (in \) all_dghms_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "all_dghms \ \\<^sub>\ Vset \" proof(rule vsubset_in_VsetI) interpret \: \ \ by (rule assms(1)) show "all_dghms \ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof(intro vsubsetI) fix \ assume "\ \\<^sub>\ all_dghms \" then obtain \ \ where \: "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" by clarsimp interpret is_dghm \ \ \ \ using \ by simp show "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" by (rule dghm_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_dghms[simp]: "small {\. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \}" by (rule down[of _ \set {\. \\ \. \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \}\]) auto text\Further elementary properties.\ lemma (in is_dghm) dghm_is_arr_HomCod: assumes "f : a \\<^bsub>\\<^esub> b" shows "\\ArrMap\\f\ \\<^sub>\ \\Arr\" "\\ObjMap\\a\ \\<^sub>\ \\Obj\" "\\ObjMap\\b\ \\<^sub>\ \\Obj\" using assms by (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros)+ lemma (in is_dghm) dghm_vimage_dghm_ArrMap_vsubset_Hom: assumes "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" shows "\\ArrMap\ `\<^sub>\ Hom \ a b \\<^sub>\ Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" proof(intro vsubsetI) fix g assume "g \\<^sub>\ \\ArrMap\ `\<^sub>\ Hom \ a b" then obtain f where "f \\<^sub>\ Hom (\\HomDom\) a b" and "g = \\ArrMap\\f\" by (auto simp: dg_cs_simps) then show "g \\<^sub>\ Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" by (simp add: dghm_ArrMap_is_arr dg_cs_simps) qed subsection\Opposite digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter II-2 in \cite{mac_lane_categories_2010}.\ definition op_dghm :: "V \ V" where "op_dghm \ = [\\ObjMap\, \\ArrMap\, op_dg (\\HomDom\), op_dg (\\HomCod\)]\<^sub>\" text\Components.\ lemma op_dghm_components[dg_op_simps]: shows "op_dghm \\ObjMap\ = \\ObjMap\" and "op_dghm \\ArrMap\ = \\ArrMap\" and "op_dghm \\HomDom\ = op_dg (\\HomDom\)" and "op_dghm \\HomCod\ = op_dg (\\HomCod\)" unfolding op_dghm_def dghm_field_simps by (auto simp: nat_omega_simps) subsubsection\Further properties\ lemma (in is_dghm) is_dghm_op: "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> op_dg \" proof(intro is_dghmI, unfold dg_op_simps) show "vfsequence (op_dghm \)" unfolding op_dghm_def by simp show "vcard (op_dghm \) = 4\<^sub>\" unfolding op_dghm_def by (auto simp: nat_omega_simps) qed ( cs_concl cs_intro: dghm_ObjMap_vrange dg_cs_intros dg_op_intros V_cs_intros cs_simp: dg_cs_simps dg_op_simps )+ lemma (in is_dghm) is_dghm_op'[dg_op_intros]: assumes "\' = op_dg \" and "\' = op_dg \" and "\' = \" shows "op_dghm \ : \' \\\<^sub>D\<^sub>G\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_dghm_op) lemmas is_dghm_op[dg_op_intros] = is_dghm.is_dghm_op' lemma (in is_dghm) dghm_op_dghm_op_dghm[dg_op_simps]: "op_dghm (op_dghm \) = \" using is_dghm_axioms by ( cs_concl cs_simp: dg_op_simps cs_intro: dg_op_intros dghm_eqI[where \=\] ) lemmas dghm_op_dghm_op_dghm[dg_op_simps] = is_dghm.dghm_op_dghm_op_dghm lemma eq_op_dghm_iff[dg_op_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \
" shows "op_dghm \ = op_dghm \ \ \ = \" proof interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \
\ by (rule assms(2)) assume prems: "op_dghm \ = op_dghm \" show "\ = \" proof(rule dghm_eqI[OF assms]) from prems R.dghm_op_dghm_op_dghm L.dghm_op_dghm_op_dghm show "\\ObjMap\ = \\ObjMap\" and "\\ArrMap\ = \\ArrMap\" by metis+ from prems R.dghm_op_dghm_op_dghm L.dghm_op_dghm_op_dghm have "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by auto then show "\ = \" "\ = \
" by (auto simp: dg_cs_simps) qed qed auto subsection\Composition of covariant digraph homomorphisms\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ definition dghm_comp :: "V \ V \ V" (infixl \\\<^sub>D\<^sub>G\<^sub>H\<^sub>M\ 55) where "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = [\\ObjMap\ \\<^sub>\ \\ObjMap\, \\ArrMap\ \\<^sub>\ \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\" text\Components.\ lemma dghm_comp_components: shows "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\ = \\ObjMap\ \\<^sub>\ \\ObjMap\" and "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ = \\ArrMap\ \\<^sub>\ \\ArrMap\" and [dg_shared_cs_simps, dg_cs_simps]: "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\HomDom\ = \\HomDom\" and [dg_shared_cs_simps, dg_cs_simps]: "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\HomCod\ = \\HomCod\" unfolding dghm_comp_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma dghm_comp_ObjMap_vsv[dg_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\)" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_comp_components cs_intro: V_cs_intros) qed lemma dghm_comp_ObjMap_vdomain[dg_cs_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\) = \\Obj\" using assms by ( cs_concl cs_simp: dghm_comp_components dg_cs_simps V_cs_simps cs_intro: is_dghm.dghm_ObjMap_vrange ) lemma dghm_comp_ObjMap_vrange: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\) \\<^sub>\ \\Obj\" using assms by ( cs_concl cs_simp: dghm_comp_components cs_intro: is_dghm.dghm_ObjMap_vrange V_cs_intros ) lemma dghm_comp_ObjMap_app[dg_cs_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) from assms(3) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" by ( cs_concl cs_simp: dghm_comp_components dg_cs_simps V_cs_simps cs_intro: V_cs_intros dg_cs_intros ) qed subsubsection\Arrow map\ lemma dghm_comp_ArrMap_vsv[dg_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\)" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_comp_components cs_intro: V_cs_intros) qed lemma dghm_comp_ArrMap_vdomain[dg_cs_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\) = \\Arr\" using assms by ( cs_concl cs_simp: dghm_comp_components dg_cs_simps V_cs_simps cs_intro: is_dghm.dghm_ArrMap_vrange ) lemma dghm_comp_ArrMap_vrange[dg_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\) \\<^sub>\ \\Arr\" using assms by ( cs_concl cs_simp: dghm_comp_components cs_intro: is_dghm.dghm_ArrMap_vrange V_cs_intros ) lemma dghm_comp_ArrMap_app[dg_cs_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "f \\<^sub>\ \\Arr\" shows "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\f\ = \\ArrMap\\\\ArrMap\\f\\" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) from assms(3) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\f\ = \\ArrMap\\\\ArrMap\\f\\" by ( cs_concl cs_simp: dghm_comp_components dg_cs_simps V_cs_simps cs_intro: V_cs_intros dg_cs_intros ) qed subsubsection\Opposite of the composition of covariant digraph homomorphisms\ lemma op_dghm_dghm_comp[dg_op_simps]: "op_dghm (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) = op_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M op_dghm \" unfolding dghm_comp_def op_dghm_def dghm_field_simps by (simp add: nat_omega_simps) subsubsection\Further properties\ lemma dghm_comp_is_dghm[dg_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof- interpret L: is_dghm \ \ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis proof(intro is_dghmI is_dghmI, unfold dg_cs_simps) show "vfsequence (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)" unfolding dghm_comp_def by simp show "vcard (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) = 4\<^sub>\" unfolding dghm_comp_def by (simp add: nat_omega_simps) fix f a b assume "f : a \\<^bsub>\\<^esub> b" with assms show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\f\ : (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ \\<^bsub>\\<^esub> (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\" by (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed ( use assms in \ cs_concl cs_intro: dg_cs_intros dghm_comp_ObjMap_vrange cs_simp: dg_cs_simps \ )+ qed lemma dghm_comp_assoc[dg_cs_simps]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)" proof(rule dghm_eqI [of \ \ \
_ \ \
]) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \))\ObjMap\" proof(rule vsv_eqI) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \))\ObjMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\)" for a using that assms by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (use assms in \cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros\)+ show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \))\ArrMap\" proof(rule vsv_eqI) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\a\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \))\ArrMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\)" for a using that assms by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (use assms in \cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros\)+ qed (use assms in \cs_concl cs_intro: dg_cs_intros\)+ subsection\Composition of contravariant digraph homomorphisms\ subsubsection\Definition and elementary properties\ text\See section 1.2 in \cite{bodo_categories_1970}.\ definition dghm_cn_comp :: "V \ V \ V" (infixl \\<^sub>D\<^sub>G\<^sub>H\<^sub>M\\ 55) where "\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \ = [ \\ObjMap\ \\<^sub>\ \\ObjMap\, \\ArrMap\ \\<^sub>\ \\ArrMap\, op_dg (\\HomDom\), \\HomCod\ ]\<^sub>\" text\Components.\ lemma dghm_cn_comp_components: shows "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\ = \\ObjMap\ \\<^sub>\ \\ObjMap\" and "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\ = \\ArrMap\ \\<^sub>\ \\ArrMap\" and [dg_cn_cs_simps]: "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\HomDom\ = op_dg (\\HomDom\)" and [dg_cn_cs_simps]: "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\HomCod\ = \\HomCod\" unfolding dghm_cn_comp_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map: two contravariant digraph homomorphisms\ lemma dghm_cn_comp_ObjMap_vsv[dg_cn_cs_intros]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\)" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \op_dg \\ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_cn_comp_components cs_intro: V_cs_intros) qed lemma dghm_cn_comp_ObjMap_vdomain[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\) = \\Obj\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: is_dghm.dghm_ObjMap_vrange ) lemma dghm_cn_comp_ObjMap_vrange: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\) \\<^sub>\ \\Obj\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components cs_intro: is_dghm.dghm_ObjMap_vrange V_cs_intros ) lemma dghm_cn_comp_ObjMap_app[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \op_dg \\ \ \ by (rule assms(2)) from assms(3) show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: V_cs_intros dg_cs_intros ) qed subsubsection\Arrow map: two contravariant digraph homomorphisms\ lemma dghm_cn_comp_ArrMap_vsv[dg_cn_cs_intros]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\)" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \op_dg \\ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_cn_comp_components cs_intro: V_cs_intros) qed lemma dghm_cn_comp_ArrMap_vdomain[dg_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\) = \\Arr\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: is_dghm.dghm_ArrMap_vrange ) lemma dghm_cn_comp_ArrMap_vrange: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\) \\<^sub>\ \\Arr\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components cs_intro: is_dghm.dghm_ArrMap_vrange V_cs_intros ) lemma dghm_cn_comp_ArrMap_app[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Arr\" shows "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \op_dg \\ \ \ by (rule assms(2)) from assms(3) show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: V_cs_intros dg_cs_intros ) qed subsubsection\Object map: contravariant and covariant digraph homomorphisms\ lemma dghm_cn_cov_comp_ObjMap_vsv[dg_cn_cs_intros]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\)" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_cn_comp_components cs_intro: V_cs_intros) qed lemma dghm_cn_cov_comp_ObjMap_vdomain[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\) = \\Obj\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: is_dghm.dghm_ObjMap_vrange ) lemma dghm_cn_cov_comp_ObjMap_vrange: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\) \\<^sub>\ \\Obj\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components cs_intro: is_dghm.dghm_ObjMap_vrange V_cs_intros ) lemma dghm_cn_cov_comp_ObjMap_app[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) from assms show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps V_cs_simps cs_intro: V_cs_intros dg_op_intros dg_cs_intros ) qed subsubsection\Arrow map: contravariant and covariant digraph homomorphisms\ lemma dghm_cn_cov_comp_ArrMap_vsv[dg_cn_cs_intros]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\)" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis by (cs_concl cs_simp: dghm_cn_comp_components cs_intro: V_cs_intros) qed lemma dghm_cn_cov_comp_ArrMap_vdomain[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\) = \\Arr\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps dg_op_simps V_cs_simps cs_intro: is_dghm.dghm_ArrMap_vrange ) lemma dghm_cn_cov_comp_ArrMap_vrange: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\) \\<^sub>\ \\Arr\" using assms by ( cs_concl cs_simp: dghm_cn_comp_components cs_intro: is_dghm.dghm_ArrMap_vrange V_cs_intros ) lemma dghm_cn_cov_comp_ArrMap_app[dg_cn_cs_simps]: assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Arr\" shows "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) from assms(3) show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" by ( cs_concl cs_simp: dghm_cn_comp_components dg_cs_simps V_cs_simps cs_intro: V_cs_intros dg_op_intros dg_cs_intros ) qed subsubsection\ Opposite of the contravariant composition of the digraph homomorphisms \ lemma op_dghm_dghm_cn_comp[dg_op_simps]: "op_dghm (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \) = op_dghm \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ op_dghm \" unfolding op_dghm_def dghm_cn_comp_def dghm_field_simps by (auto simp: nat_omega_simps) subsubsection\Further properties\ lemma dghm_cn_comp_is_dghm[dg_cn_cs_intros]: \\See section 1.2 in \cite{bodo_categories_1970}.\ assumes "digraph \ \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof- interpret \: digraph \ \ by (rule assms(1)) interpret L: is_dghm \ \op_dg \\ \ \ using assms(2) by auto interpret R: is_dghm \ \op_dg \\ \ \ using assms(3) by auto show ?thesis proof(intro is_dghmI, unfold dg_op_simps dg_cs_simps dg_cn_cs_simps) show "vfsequence (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)" unfolding dghm_cn_comp_def by auto show "vcard (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \) = 4\<^sub>\" unfolding dghm_cn_comp_def by (simp add: nat_omega_simps) fix f a b assume "f : a \\<^bsub>\\<^esub> b" with assms show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\f\ : (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ \\<^bsub>\\<^esub> (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\b\" by ( cs_concl cs_simp: dg_cn_cs_simps cs_intro: dg_cs_intros dg_op_intros ) qed ( cs_concl cs_simp: dg_cs_simps dg_cn_cs_simps cs_intro: dghm_cn_comp_ObjMap_vrange dg_cs_intros dg_cn_cs_intros )+ qed lemma dghm_cn_cov_comp_is_dghm[dg_cn_cs_intros]: \\See section 1.2 in \cite{bodo_categories_1970}.\ assumes "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" proof- interpret L: is_dghm \ \op_dg \\ \ \ by (rule assms(1)) interpret R: is_dghm \ \ \ \ by (rule assms(2)) show ?thesis proof(intro is_dghmI, unfold dg_op_simps dg_cs_simps) show "vfsequence (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)" unfolding dghm_cn_comp_def by simp show "vcard (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \) = 4\<^sub>\" unfolding dghm_cn_comp_def by (auto simp: nat_omega_simps) fix f b a assume "f : b \\<^bsub>\\<^esub> a" with assms show "(\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ArrMap\\f\ : (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\a\ \\<^bsub>\\<^esub> (\ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \)\ObjMap\\b\" by (cs_concl cs_simp: dg_cn_cs_simps dg_op_simps cs_intro: dg_cs_intros) qed ( cs_concl cs_simp: dg_cs_simps dg_cn_cs_simps cs_intro: dghm_cn_cov_comp_ObjMap_vrange dg_cs_intros dg_cn_cs_intros dg_op_intros )+ qed lemma dghm_cov_cn_comp_is_dghm: \\See section 1.2 in \cite{bodo_categories_1970}\ assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> \" using assms by (rule dghm_comp_is_dghm) subsection\Identity digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ definition dghm_id :: "V \ V" where "dghm_id \ = [vid_on (\\Obj\), vid_on (\\Arr\), \, \]\<^sub>\" text\Components.\ lemma dghm_id_components: shows "dghm_id \\ObjMap\ = vid_on (\\Obj\)" and "dghm_id \\ArrMap\ = vid_on (\\Arr\)" and [dg_shared_cs_simps, dg_cs_simps]: "dghm_id \\HomDom\ = \" and [dg_shared_cs_simps, dg_cs_simps]: "dghm_id \\HomCod\ = \" unfolding dghm_id_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda dghm_id_components(1)[folded VLambda_vid_on] |vsv dghm_id_ObjMap_vsv[dg_shared_cs_intros, dg_cs_intros]| |vdomain dghm_id_ObjMap_vdomain[dg_shared_cs_simps, dg_cs_simps]| |app dghm_id_ObjMap_app[dg_shared_cs_simps, dg_cs_simps]| lemma dghm_id_ObjMap_vrange[dg_shared_cs_simps, dg_cs_simps]: "\\<^sub>\ (dghm_id \\ObjMap\) = \\Obj\" unfolding dghm_id_components by simp subsubsection\Arrow map\ mk_VLambda dghm_id_components(2)[folded VLambda_vid_on] |vsv dghm_id_ArrMap_vsv[dg_shared_cs_intros, dg_cs_intros]| |vdomain dghm_id_ArrMap_vdomain[dg_shared_cs_simps, dg_cs_simps]| |app dghm_id_ArrMap_app[dg_shared_cs_simps, dg_cs_simps]| lemma dghm_id_ArrMap_vrange[dg_shared_cs_simps, dg_cs_simps]: "\\<^sub>\ (dghm_id \\ArrMap\) = \\Arr\" unfolding dghm_id_components by simp subsubsection\Opposite identity digraph homomorphism\ lemma op_dghm_dghm_id[dg_op_simps]: "op_dghm (dghm_id \) = dghm_id (op_dg \)" unfolding dghm_id_def op_dg_def op_dghm_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\An identity digraph homomorphism is a digraph homomorphism\ lemma (in digraph) dg_dghm_id_is_dghm: "dghm_id \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof(intro is_dghmI, unfold dg_cs_simps) show "vfsequence (dghm_id \)" unfolding dghm_id_def by simp show "vcard (dghm_id \) = 4\<^sub>\" unfolding dghm_id_def by (simp add: nat_omega_simps) qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros)+ lemma (in digraph) dg_dghm_id_is_dghm': assumes "\ = \" and "\ = \" shows "dghm_id \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" unfolding assms by (rule dg_dghm_id_is_dghm) lemmas [dg_cs_intros] = digraph.dg_dghm_id_is_dghm' subsubsection\Further properties\ lemma (in is_dghm) dghm_dghm_comp_dghm_id_left: "dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = \" \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ proof(rule dghm_eqI [of \ \ \ _ \ \]) show "(dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\ = \\ObjMap\" proof(rule vsv_eqI) show "(dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = \\ObjMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\)" for a using that by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros)+ show "(dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ = \\ArrMap\" proof(rule vsv_eqI) show "(dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\a\ = \\ArrMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((dghm_id \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\)" for a using that by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros)+ qed (cs_concl cs_simp: cs_intro: dg_cs_intros)+ lemmas [dg_cs_simps] = is_dghm.dghm_dghm_comp_dghm_id_left lemma (in is_dghm) dghm_dghm_comp_dghm_id_right: "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \ = \" \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ proof(rule dghm_eqI [of \ \ \ _ \ \]) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ObjMap\ = \\ObjMap\" proof(rule vsv_eqI) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ObjMap\\a\ = \\ObjMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ObjMap\)" for a using that by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros)+ show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ArrMap\ = \\ArrMap\" proof(rule vsv_eqI) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ArrMap\\a\ = \\ArrMap\\a\" if "a \\<^sub>\ \\<^sub>\ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M dghm_id \)\ArrMap\)" for a using that by (cs_prems cs_simp: dg_cs_simps cs_intro: dg_cs_intros) (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros) qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros)+ qed (cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros)+ lemmas [dg_cs_simps] = is_dghm.dghm_dghm_comp_dghm_id_right subsection\Constant digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter III-3 in \cite{mac_lane_categories_2010}.\ definition dghm_const :: "V \ V \ V \ V \ V" where "dghm_const \ \
a f = [vconst_on (\\Obj\) a, vconst_on (\\Arr\) f, \, \
]\<^sub>\" text\Components.\ lemma dghm_const_components: shows "dghm_const \ \
a f\ObjMap\ = vconst_on (\\Obj\) a" and "dghm_const \ \
a f\ArrMap\ = vconst_on (\\Arr\) f" and [dg_shared_cs_simps, dg_cs_simps]: "dghm_const \ \
a f\HomDom\ = \" and [dg_shared_cs_simps, dg_cs_simps]: "dghm_const \ \
a f\HomCod\ = \
" unfolding dghm_const_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda dghm_const_components(1)[folded VLambda_vconst_on] |vsv dghm_const_ObjMap_vsv[dg_shared_cs_intros, dg_cs_intros]| |vdomain dghm_const_ObjMap_vdomain[dg_shared_cs_simps, dg_cs_simps]| |app dghm_const_ObjMap_app[dg_shared_cs_simps, dg_cs_simps]| subsubsection\Arrow map\ mk_VLambda dghm_const_components(2)[folded VLambda_vconst_on] |vsv dghm_const_ArrMap_vsv[dg_shared_cs_intros, dg_cs_intros]| |vdomain dghm_const_ArrMap_vdomain[dg_shared_cs_simps, dg_cs_simps]| |app dghm_const_ArrMap_app[dg_shared_cs_simps, dg_cs_simps]| subsubsection\Opposite constant digraph homomorphism\ lemma op_dghm_dghm_const[dg_op_simps]: "op_dghm (dghm_const \ \
a f) = dghm_const (op_dg \) (op_dg \
) a f" unfolding dghm_const_def op_dg_def op_dghm_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\A constant digraph homomorphism is a digraph homomorphism\ lemma dghm_const_is_dghm: assumes "digraph \ \" and "digraph \ \
" and "f : a \\<^bsub>\
\<^esub> a" shows "dghm_const \ \
a f : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \
" proof- interpret \
: digraph \ \
by (rule assms(2)) show ?thesis proof(intro is_dghmI) show "vfsequence (dghm_const \ \
a f)" unfolding dghm_const_def by simp show "vcard (dghm_const \ \
a f) = 4\<^sub>\" unfolding dghm_const_def by (simp add: nat_omega_simps) qed ( use assms in \ cs_concl cs_simp: dghm_const_components(1) dg_cs_simps cs_intro: V_cs_intros dg_cs_intros \ )+ qed lemma dghm_const_is_dghm'[dg_cs_intros]: assumes "digraph \ \" and "digraph \ \
" and "f : a \\<^bsub>\
\<^esub> a" and "\ = \" and "\ = \
" shows "dghm_const \ \
a f : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" using assms(1-3) unfolding assms(4,5) by (rule dghm_const_is_dghm) +subsubsection\Further properties\ + +lemma (in is_dghm) dghm_dghm_comp_dghm_const[dg_cs_simps]: + assumes "digraph \ \" and "f : a \\<^bsub>\\<^esub> a" + shows "dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = dghm_const \ \ a f" +proof(rule dghm_eqI) + interpret \: digraph \ \ by (rule assms(1)) + from assms(2) show "dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: dg_cs_intros) + with assms(2) have ObjMap_dom_lhs: + "\\<^sub>\ ((dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\) = \\Obj\" + and ArrMap_dom_lhs: "\\<^sub>\ ((dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\) = \\Arr\" + by (cs_concl cs_simp: dg_cs_simps)+ + from assms(2) show "dghm_const \ \ a f : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" + by (cs_concl cs_intro: dg_cs_intros) + with assms(2) have ObjMap_dom_rhs: + "\\<^sub>\ (dghm_const \ \ a f\ObjMap\) = \\Obj\" + and ArrMap_dom_rhs: "\\<^sub>\ (dghm_const \ \ a f\ArrMap\) = \\Arr\" + by (cs_concl cs_simp: dg_cs_simps)+ + show "(dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\ = dghm_const \ \ a f\ObjMap\" + by (rule vsv_eqI, unfold ObjMap_dom_lhs ObjMap_dom_rhs) + (use assms(2) in \cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros\)+ + show "(dghm_const \ \ a f \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ = dghm_const \ \ a f\ArrMap\" + by (rule vsv_eqI, unfold ArrMap_dom_lhs ArrMap_dom_rhs) + (use assms(2) in \cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros\)+ +qed simp_all + +lemmas [dg_cs_simps] = is_dghm.dghm_dghm_comp_dghm_const + + subsection\Faithful digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_ft_dghm = is_dghm \ \ \ \ for \ \ \ \ + assumes ft_dghm_v11_on_Hom: "\ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" syntax "_is_ft_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_ft_dghm \ \ \ \" text\Rules.\ lemma (in is_ft_dghm) is_ft_dghm_axioms'[dghm_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>D\<^sub>G\<^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_dghm_axioms) mk_ide rf is_ft_dghm_def[unfolded is_ft_dghm_axioms_def] |intro is_ft_dghmI| |dest is_ft_dghmD[dest]| |elim is_ft_dghmE[elim]| lemmas [dghm_cs_intros] = is_ft_dghmD(1) subsubsection\Opposite faithful digraph homomorphism\ lemma (in is_ft_dghm) ft_dghm_op_dghm_is_ft_dghm: "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> op_dg \" by ( rule is_ft_dghmI, unfold dg_op_simps, cs_concl cs_simp: cs_intro: dg_cs_intros dg_op_intros ) (auto simp: ft_dghm_v11_on_Hom) lemma (in is_ft_dghm) ft_dghm_op_dghm_is_ft_dghm'[dg_op_intros]: assumes "\' = op_dg \" and "\' = op_dg \" shows "op_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule ft_dghm_op_dghm_is_ft_dghm) lemmas ft_dghm_op_dghm_is_ft_dghm[dg_op_intros] = is_ft_dghm.ft_dghm_op_dghm_is_ft_dghm' subsubsection\ The composition of faithful digraph homomorphisms is a faithful digraph homomorphism. \ lemma dghm_comp_is_ft_dghm[dghm_cs_intros]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" proof- interpret L: is_ft_dghm \ \ \ \ using assms(1) by auto interpret R: is_ft_dghm \ \ \ \ using assms(2) by auto have inj: "\ a \\<^sub>\ \\Obj\ ; b \\<^sub>\ \\Obj\ \ \ v11 ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" for a b proof- assume prems: "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" then have \_hom_\: "v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\))" by (intro L.ft_dghm_v11_on_Hom) (cs_concl cs_intro: dg_cs_intros)+ have "v11 (\\ArrMap\ \\<^sup>l\<^sub>\ (\\ArrMap\ `\<^sub>\ Hom \ a b))" proof(intro v11_vlrestriction_vsubset[OF \_hom_\] vsubsetI) fix g assume "g \\<^sub>\ \\ArrMap\ `\<^sub>\ Hom \ a b" then obtain f where f: "f : a \\<^bsub>\\<^esub> b" and g_def: "g = \\ArrMap\\f\" by auto from f show "g \\<^sub>\ Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" by (cs_concl cs_simp: g_def cs_intro: dg_cs_intros) qed then show "v11 ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" unfolding dghm_comp_components by (intro v11_vlrestriction_vcomp) (auto simp: R.ft_dghm_v11_on_Hom prems) qed then show "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" by (intro is_ft_dghmI, cs_concl cs_intro: dg_cs_intros) auto qed subsection\Full digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_fl_dghm = is_dghm \ \ \ \ for \ \ \ \ + assumes fl_dghm_surj_on_Hom: "\ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ \\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" syntax "_is_fl_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_fl_dghm \ \ \ \" text\Rules.\ lemma (in is_fl_dghm) is_fl_dghm_axioms'[dghm_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_fl_dghm_axioms) mk_ide rf is_fl_dghm_def[unfolded is_fl_dghm_axioms_def] |intro is_fl_dghmI| |dest is_fl_dghmD[dest]| |elim is_fl_dghmE[elim]| lemmas [dghm_cs_intros] = is_fl_dghmD(1) subsubsection\Opposite full digraph homomorphism\ lemma (in is_fl_dghm) fl_dghm_op_dghm_is_fl_dghm: "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> op_dg \" by ( rule is_fl_dghmI, unfold dg_op_simps, cs_concl cs_intro: dg_cs_intros dg_op_intros ) (auto simp: fl_dghm_surj_on_Hom) lemma (in is_fl_dghm) fl_dghm_op_dghm_is_fl_dghm'[dg_op_intros]: assumes "\' = op_dg \" and "\' = op_dg \" shows "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> op_dg \" unfolding assms by (rule fl_dghm_op_dghm_is_fl_dghm) lemmas fl_dghm_op_dghm_is_fl_dghm[dg_op_intros] = is_fl_dghm.fl_dghm_op_dghm_is_fl_dghm' subsubsection\ The composition of full digraph homomorphisms is a full digraph homomorphism \ lemma dghm_comp_is_fl_dghm[dghm_cs_intros]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" proof- interpret L: is_fl_dghm \ \ \ \ by (rule assms(1)) interpret R: is_fl_dghm \ \ \ \ by (rule assms(2)) have surj: "\ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\) ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\)" for a b proof- assume prems: "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" have surj_\: "\\ArrMap\ `\<^sub>\ Hom \ a b = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" by (rule R.fl_dghm_surj_on_Hom[OF prems]) from prems L.is_dghm_axioms R.is_dghm_axioms have comp_Obj: "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\ = \\ObjMap\\\\ObjMap\\b\\" by (auto simp: dg_cs_simps) have "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ `\<^sub>\ Hom \ a b = \\ArrMap\ `\<^sub>\ \\ArrMap\ `\<^sub>\ Hom \ a b" unfolding dghm_comp_components by (simp add: vcomp_vimage) also from prems have "\ = Hom \ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\) ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\)" unfolding surj_\ comp_Obj by ( simp add: prems(2) L.fl_dghm_surj_on_Hom R.dghm_ObjMap_app_in_HomCod_Obj ) finally show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ `\<^sub>\ (Hom \ a b) = Hom \ ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\) ((\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\)" by simp qed show ?thesis by (rule is_fl_dghmI, cs_concl cs_intro: dg_cs_intros) (auto simp: surj) qed subsection\Fully faithful digraph homomorphism\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_ff_dghm = is_ft_dghm \ \ \ \ + is_fl_dghm \ \ \ \ for \ \ \ \ syntax "_is_ff_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" \ "CONST is_ff_dghm \ \ \ \" text\Rules.\ lemma (in is_ff_dghm) is_ff_dghm_axioms'[dghm_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_ff_dghm_axioms) mk_ide rf is_ff_dghm_def |intro is_ff_dghmI| |dest is_ff_dghmD[dest]| |elim is_ff_dghmE[elim]| lemmas [dghm_cs_intros] = is_ff_dghmD subsubsection\Opposite fully faithful digraph homomorphism.\ lemma (in is_ff_dghm) ff_dghm_op_dghm_is_ff_dghm: "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> op_dg \" by (rule is_ff_dghmI) (cs_concl cs_intro: dg_op_intros)+ lemma (in is_ff_dghm) ff_dghm_op_dghm_is_ff_dghm'[dg_op_intros]: assumes "\' = op_dg \" and "\' = op_dg \" shows "op_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \'" unfolding assms by (rule ff_dghm_op_dghm_is_ff_dghm) lemmas ff_dghm_op_dghm_is_ff_dghm[dg_op_intros] = is_ff_dghm.ff_dghm_op_dghm_is_ff_dghm subsubsection\ The composition of fully faithful digraph homomorphisms is a fully faithful digraph homomorphism. \ lemma dghm_comp_is_ff_dghm[dghm_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" using assms by (intro is_ff_dghmI, elim is_ff_dghmE) (cs_concl cs_intro: dghm_cs_intros) subsection\Isomorphism of digraphs\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_iso_dghm = is_dghm \ \ \ \ for \ \ \ \ + assumes iso_dghm_ObjMap_v11: "v11 (\\ObjMap\)" and iso_dghm_ArrMap_v11: "v11 (\\ArrMap\)" and iso_dghm_ObjMap_vrange[dghm_cs_simps]: "\\<^sub>\ (\\ObjMap\) = \\Obj\" and iso_dghm_ArrMap_vrange[dghm_cs_simps]: "\\<^sub>\ (\\ArrMap\) = \\Arr\" syntax "_is_iso_dghm" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" \ "CONST is_iso_dghm \ \ \ \" sublocale is_iso_dghm \ ObjMap: v11 \\\ObjMap\\ rewrites "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" by (cs_concl cs_simp: dghm_cs_simps dg_cs_simps cs_intro: iso_dghm_ObjMap_v11)+ sublocale is_iso_dghm \ ArrMap: v11 \\\ArrMap\\ rewrites "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" by (cs_concl cs_simp: dghm_cs_simps dg_cs_simps cs_intro: iso_dghm_ArrMap_v11)+ lemmas [dghm_cs_simps] = is_iso_dghm.iso_dghm_ObjMap_vrange is_iso_dghm.iso_dghm_ArrMap_vrange text\Rules.\ lemma (in is_iso_dghm) is_iso_dghm_axioms'[dghm_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_iso_dghm_axioms) mk_ide rf is_iso_dghm_def[unfolded is_iso_dghm_axioms_def] |intro is_iso_dghmI| |dest is_iso_dghmD[dest]| |elim is_iso_dghmE[elim]| text\Elementary properties.\ lemma (in is_iso_dghm) iso_dghm_Obj_HomDom_if_Obj_HomCod[elim]: assumes "b \\<^sub>\ \\Obj\" obtains a where "a \\<^sub>\ \\Obj\" and "b = \\ObjMap\\a\" using assms ObjMap.vrange_atD iso_dghm_ObjMap_vrange by blast lemma (in is_iso_dghm) iso_dghm_Arr_HomDom_if_Arr_HomCod[elim]: assumes "g \\<^sub>\ \\Arr\" obtains f where "f \\<^sub>\ \\Arr\" and "g = \\ArrMap\\f\" using assms ArrMap.vrange_atD iso_dghm_ArrMap_vrange by blast lemma (in is_iso_dghm) iso_dghm_ObjMap_eqE[elim]: assumes "\\ObjMap\\a\ = \\ObjMap\\b\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "a = b \ P" shows P using assms ObjMap.v11_eq_iff by auto lemma (in is_iso_dghm) iso_dghm_ArrMap_eqE[elim]: assumes "\\ArrMap\\f\ = \\ArrMap\\g\" and "f \\<^sub>\ \\Arr\" and "g \\<^sub>\ \\Arr\" and "f = g \ P" shows P using assms ArrMap.v11_eq_iff by auto sublocale is_iso_dghm \ is_ft_dghm by (intro is_ft_dghmI, cs_concl cs_intro: dg_cs_intros) auto sublocale is_iso_dghm \ is_fl_dghm proof fix a b assume [intro]: "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" show "\\ArrMap\ `\<^sub>\ Hom \ a b = Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" proof(intro vsubset_antisym vsubsetI) fix g assume prems: "g \\<^sub>\ Hom \ (\\ObjMap\\a\) (\\ObjMap\\b\)" then have g: "g : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" by auto then have dom_g: "\\Dom\\g\ = \\ObjMap\\a\" and cod_g: "\\Cod\\g\ = \\ObjMap\\b\" by blast+ from prems obtain f where [intro, simp]: "f \\<^sub>\ \\Arr\" and g_def: "g = \\ArrMap\\f\" by auto then obtain a' b' where f: "f : a' \\<^bsub>\\<^esub> b'" by blast then have "g : \\ObjMap\\a'\ \\<^bsub>\\<^esub> \\ObjMap\\b'\" by (cs_concl cs_simp: g_def dg_cs_simps cs_intro: dg_cs_intros) with g have "\\ObjMap\\a\ = \\ObjMap\\a'\" and "\\ObjMap\\b\ = \\ObjMap\\b'\" by (metis HomCod.dg_is_arrE cod_g)+ with f have "a = \\Dom\\f\" "b = \\Cod\\f\" by auto (*slow*) with f show "g \\<^sub>\ \\ArrMap\ `\<^sub>\ Hom \ a b" by (auto simp: g_def HomDom.dg_is_arrD(4,5) ArrMap.vsv_vimageI1) qed (metis ArrMap.vsv_vimageE dghm_ArrMap_is_arr' in_Hom_iff) qed sublocale is_iso_dghm \ is_ff_dghm by unfold_locales lemmas (in is_iso_dghm) iso_dghm_is_ff_dghm = is_ff_dghm_axioms lemmas [dghm_cs_intros] = is_iso_dghm.iso_dghm_is_ff_dghm subsubsection\Opposite digraph isomorphisms\ lemma (in is_iso_dghm) is_iso_dghm_op: "op_dghm \ : op_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> op_dg \" by (intro is_iso_dghmI, unfold dg_op_simps) ( cs_concl cs_simp: dghm_cs_simps cs_intro: V_cs_intros dg_cs_intros dg_op_intros )+ lemma (in is_iso_dghm) is_iso_dghm_op': assumes "\' = op_dg \" and "\' = op_dg \" shows "op_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" unfolding assms by (rule is_iso_dghm_op) lemmas is_iso_dghm_op[dg_op_intros] = is_iso_dghm.is_iso_dghm_op' subsubsection\ The composition of isomorphisms of digraphs is an isomorphism of digraphs \ lemma dghm_comp_is_iso_dghm[dghm_cs_intros]: assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof- interpret \: is_iso_dghm \ \ \ \ using assms by auto interpret \: is_iso_dghm \ \ \ \ using assms by auto show ?thesis by (intro is_iso_dghmI, unfold dghm_comp_components) ( cs_concl cs_simp: V_cs_simps dg_cs_simps dghm_cs_simps cs_intro: dg_cs_intros V_cs_intros )+ qed subsubsection\An identity digraph homomorphism is an isomorphism of digraphs.\ lemma (in digraph) dg_dghm_id_is_iso_dghm: "dghm_id \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (rule is_iso_dghmI) (simp_all add: dg_dghm_id_is_dghm dghm_id_components) lemma (in digraph) dg_dghm_id_is_iso_dghm'[dghm_cs_intros]: assumes "\' = \" and "\' = \" shows "dghm_id \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" unfolding assms by (rule dg_dghm_id_is_iso_dghm) lemmas [dghm_cs_intros] = digraph.dg_dghm_id_is_iso_dghm' subsection\Inverse digraph homomorphism\ subsubsection\Definition and elementary properties\ definition inv_dghm :: "V \ V" where "inv_dghm \ = [(\\ObjMap\)\\<^sub>\, (\\ArrMap\)\\<^sub>\, \\HomCod\, \\HomDom\]\<^sub>\" text\Components.\ lemma inv_dghm_components: shows "inv_dghm \\ObjMap\ = (\\ObjMap\)\\<^sub>\" and "inv_dghm \\ArrMap\ = (\\ArrMap\)\\<^sub>\" and [dghm_cs_simps]: "inv_dghm \\HomDom\ = \\HomCod\" and [dghm_cs_simps]: "inv_dghm \\HomCod\ = \\HomDom\" unfolding inv_dghm_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ lemma (in is_iso_dghm) inv_dghm_ObjMap_v11[dghm_cs_intros]: "v11 (inv_dghm \\ObjMap\)" unfolding inv_dghm_components by (cs_concl cs_intro: V_cs_intros) lemmas [dghm_cs_intros] = is_iso_dghm.inv_dghm_ObjMap_v11 lemma (in is_iso_dghm) inv_dghm_ObjMap_vdomain[dghm_cs_simps]: "\\<^sub>\ (inv_dghm \\ObjMap\) = \\Obj\" unfolding inv_dghm_components by (cs_concl cs_simp: dghm_cs_simps V_cs_simps) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ObjMap_vdomain lemma (in is_iso_dghm) inv_dghm_ObjMap_app[dghm_cs_simps]: assumes "a' = \\ObjMap\\a\" and "a \\<^sub>\ \\Obj\" shows "inv_dghm \\ObjMap\\a'\ = a" unfolding inv_dghm_components by (metis assms ObjMap.vconverse_atI ObjMap.vsv_vconverse vsv.vsv_appI) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ObjMap_app lemma (in is_iso_dghm) inv_dghm_ObjMap_vrange[dghm_cs_simps]: "\\<^sub>\ (inv_dghm \\ObjMap\) = \\Obj\" unfolding inv_dghm_components by (cs_concl cs_simp: dg_cs_simps V_cs_simps) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ObjMap_vrange subsubsection\Arrow map\ lemma (in is_iso_dghm) inv_dghm_ArrMap_v11[dghm_cs_intros]: "v11 (inv_dghm \\ArrMap\)" unfolding inv_dghm_components by (cs_concl cs_intro: V_cs_intros) lemmas [dghm_cs_intros] = is_iso_dghm.inv_dghm_ArrMap_v11 lemma (in is_iso_dghm) inv_dghm_ArrMap_vdomain[dghm_cs_simps]: "\\<^sub>\ (inv_dghm \\ArrMap\) = \\Arr\" unfolding inv_dghm_components by (cs_concl cs_simp: dghm_cs_simps V_cs_simps) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ArrMap_vdomain lemma (in is_iso_dghm) inv_dghm_ArrMap_app[dghm_cs_simps]: assumes "a' = \\ArrMap\\a\" and "a \\<^sub>\ \\Arr\" shows "inv_dghm \\ArrMap\\a'\ = a" unfolding inv_dghm_components by (metis assms ArrMap.vconverse_atI ArrMap.vsv_vconverse vsv.vsv_appI) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ArrMap_app lemma (in is_iso_dghm) inv_dghm_ArrMap_vrange[dghm_cs_simps]: "\\<^sub>\ (inv_dghm \\ArrMap\) = \\Arr\" unfolding inv_dghm_components by (cs_concl cs_simp: dg_cs_simps V_cs_simps) lemmas [dghm_cs_simps] = is_iso_dghm.inv_dghm_ArrMap_vrange subsubsection\Further properties\ lemma (in is_iso_dghm) iso_dghm_ObjMap_inv_dghm_ObjMap_app[dghm_cs_simps]: assumes "a \\<^sub>\ \\Obj\" shows "\\ObjMap\\inv_dghm \\ObjMap\\a\\ = a" using assms by (cs_concl cs_simp: inv_dghm_components V_cs_simps) lemmas [dghm_cs_simps] = is_iso_dghm.iso_dghm_ObjMap_inv_dghm_ObjMap_app lemma (in is_iso_dghm) iso_dghm_ArrMap_inv_dghm_ArrMap_app[dghm_cs_simps]: assumes "f : a \\<^bsub>\\<^esub> b" shows "\\ArrMap\\inv_dghm \\ArrMap\\f\\ = f" using assms by (cs_concl cs_simp: inv_dghm_components V_cs_simps cs_intro: dg_cs_intros) lemmas [dghm_cs_simps] = is_iso_dghm.iso_dghm_ArrMap_inv_dghm_ArrMap_app lemma (in is_iso_dghm) iso_dghm_HomDom_is_arr_conv: assumes "f \\<^sub>\ \\Arr\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "\\ArrMap\\f\ : \\ObjMap\\a\ \\<^bsub>\\<^esub> \\ObjMap\\b\" shows "f : a \\<^bsub>\\<^esub> b" by ( metis assms HomDom.dg_is_arrE is_arr_def dghm_ArrMap_is_arr iso_dghm_ObjMap_eqE ) lemma (in is_iso_dghm) iso_dghm_HomCod_is_arr_conv: assumes "f \\<^sub>\ \\Arr\" and "a \\<^sub>\ \\Obj\" and "b \\<^sub>\ \\Obj\" and "inv_dghm \\ArrMap\\f\ : inv_dghm \\ObjMap\\a\ \\<^bsub>\\<^esub> inv_dghm \\ObjMap\\b\" shows "f : a \\<^bsub>\\<^esub> b" by ( metis assms dghm_ArrMap_is_arr' is_arrI iso_dghm_ArrMap_inv_dghm_ArrMap_app iso_dghm_ObjMap_inv_dghm_ObjMap_app ) subsection\An isomorphism of digraphs is an isomorphism in the category \GRPH\\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ lemma is_arr_isomorphism_is_iso_dghm: assumes "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = dghm_id \" and "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = dghm_id \" shows "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_dghmI) interpret L: is_dghm \ \ \ \ by (rule assms(2)) interpret R: is_dghm \ \ \ \ by (rule assms(1)) show "\ : \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" by (cs_concl cs_intro: dg_cs_intros) show "v11 (\\ObjMap\)" proof(rule R.ObjMap.vsv_valeq_v11I) fix a b assume prems[simp]: "a \\<^sub>\ \\Obj\" "b \\<^sub>\ \\Obj\" "\\ObjMap\\a\ = \\ObjMap\\b\" from assms(1,2) have "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\b\" by (simp add: dg_cs_simps) then show "a = b" by (simp add: assms(3) dghm_id_components) qed show "v11 (\\ArrMap\)" proof(rule R.ArrMap.vsv_valeq_v11I) fix a b assume prems[simp]: "a \\<^sub>\ \\Arr\" "b \\<^sub>\ \\Arr\" "\\ArrMap\\a\ = \\ArrMap\\b\" then have "\\ArrMap\\a\ \\<^sub>\ \\Arr\" by (cs_concl cs_intro: dg_cs_intros) with R.dghm_ArrMap_vsv L.dghm_ArrMap_vsv R.dghm_ArrMap_vrange have "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\a\ = (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\b\" unfolding dghm_comp_components by (simp add: dg_cs_simps) then show "a = b" by (simp add: assms(3) dghm_id_components) qed show "\\<^sub>\ (\\ObjMap\) = \\Obj\" proof(intro vsubset_antisym vsubsetI) from R.dghm_ObjMap_vrange show "a \\<^sub>\ \\<^sub>\ (\\ObjMap\) \ a \\<^sub>\ \\Obj\" for a by auto next fix a assume prems: "a \\<^sub>\ \\Obj\" then have a_def[symmetric]: "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\\a\ = a" unfolding assms(4) dghm_id_components by simp from prems show "a \\<^sub>\ \\<^sub>\ (\\ObjMap\)" by (subst a_def) (cs_concl cs_intro: V_cs_intros dg_cs_intros cs_simp: dg_cs_simps) qed show "\\<^sub>\ (\\ArrMap\) = \\Arr\" proof(intro vsubset_antisym vsubsetI) from R.dghm_ArrMap_vrange show "a \\<^sub>\ \\<^sub>\ (\\ArrMap\) \ a \\<^sub>\ \\Arr\" for a by auto next fix a assume prems: "a \\<^sub>\ \\Arr\" then have a_def[symmetric]: "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\\a\ = a" unfolding assms(4) dghm_id_components by simp with prems show "a \\<^sub>\ \\<^sub>\ (\\ArrMap\)" by (subst a_def) (cs_concl cs_intro: V_cs_intros dg_cs_intros cs_simp: dg_cs_simps) qed qed lemma is_iso_dghm_is_arr_isomorphism: assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows [dghm_cs_intros]: "inv_dghm \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "inv_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = dghm_id \" and "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M inv_dghm \ = dghm_id \" proof- let ?\ = \inv_dghm \\ interpret is_iso_dghm \ \ \ \ by (rule assms(1)) show \: "?\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_dghmI is_dghmI, unfold dghm_cs_simps) show "vfsequence (inv_dghm \)" unfolding inv_dghm_def by auto show "vcard (inv_dghm \) = 4\<^sub>\" unfolding inv_dghm_def by (simp add: nat_omega_simps) show "inv_dghm \\ArrMap\\f\ : inv_dghm \\ObjMap\\a\ \\<^bsub>\\<^esub> inv_dghm \\ObjMap\\b\" if "f : a \\<^bsub>\\<^esub> b" for a b f using that by ( intro iso_dghm_HomDom_is_arr_conv, use nothing in \unfold inv_dghm_components\ ) ( cs_concl cs_simp: V_cs_simps dghm_cs_simps dg_cs_simps cs_intro: dg_cs_intros V_cs_intros )+ qed ( cs_concl cs_simp: dg_cs_simps cs_intro: dg_cs_intros V_cs_intros dghm_cs_intros )+ show "inv_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ = dghm_id \" proof(rule dghm_eqI[of \ \ \ _ \ \]) show "(inv_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ObjMap\ = dghm_id \\ObjMap\" unfolding inv_dghm_components dghm_comp_components dghm_id_components by (rule ObjMap.v11_vcomp_vconverse) show "(inv_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \)\ArrMap\ = dghm_id \\ArrMap\" unfolding inv_dghm_components dghm_comp_components dghm_id_components by (rule ArrMap.v11_vcomp_vconverse) qed (use \ in \cs_concl cs_intro: dghm_cs_intros\) show "\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M inv_dghm \ = dghm_id \" proof(rule dghm_eqI[of \ \ \ _ \ \]) show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M inv_dghm \)\ObjMap\ = dghm_id \\ObjMap\" unfolding inv_dghm_components dghm_comp_components dghm_id_components by (rule ObjMap.v11_vcomp_vconverse') show "(\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M inv_dghm \)\ArrMap\ = dghm_id \\ArrMap\" unfolding inv_dghm_components dghm_comp_components dghm_id_components by (rule ArrMap.v11_vcomp_vconverse') qed (use \ in \cs_concl cs_intro: dghm_cs_intros\) qed subsection\Isomorphic digraphs\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale iso_digraph = fixes \ \ \ :: V assumes iso_digraph_is_iso_dghm: "\\. \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" notation iso_digraph (infixl "\\<^sub>D\<^sub>G\" 50) sublocale iso_digraph \ HomDom: digraph \ \ + HomCod: digraph \ \ using iso_digraph_is_iso_dghm by auto text\Rules.\ lemma iso_digraphI': assumes "\\. \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" using assms iso_digraph.intro by auto lemma iso_digraphI: assumes "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" using assms unfolding iso_digraph_def by auto lemma iso_digraphD[dest]: assumes "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\\. \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms unfolding iso_digraph_def by simp_all lemma iso_digraphE[elim]: assumes "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" obtains \ where "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms by auto subsubsection\A digraph isomorphism is an equivalence relation\ lemma iso_digraph_refl: assumes "digraph \ \" shows "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof(rule iso_digraphI[of _ _ _ \dghm_id \\]) interpret digraph \ \ by (rule assms) show "dghm_id \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (rule dg_dghm_id_is_iso_dghm) qed lemma iso_digraph_sym[sym]: assumes "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof- interpret iso_digraph \ \ \ by (rule assms) from iso_digraph_is_iso_dghm obtain \ where "\ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by clarsimp then have "inv_dghm \ : \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (simp add: is_iso_dghm_is_arr_isomorphism(1)) then show ?thesis by (cs_concl cs_intro: dghm_cs_intros iso_digraphI) qed lemma iso_digraph_trans[trans]: assumes "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" and "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" shows "\ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> \" proof- interpret L: iso_digraph \ \ \ by (rule assms(1)) interpret R: iso_digraph \ \ \ by (rule assms(2)) from L.iso_digraph_is_iso_dghm R.iso_digraph_is_iso_dghm show ?thesis by (meson dghm_comp_is_iso_dghm iso_digraph.intro) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Semifunctor.thy b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Semifunctor.thy --- a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Semifunctor.thy +++ b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Semifunctor.thy @@ -1,2099 +1,2120 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Semifunctor\ theory CZH_SMC_Semifunctor imports CZH_DG_DGHM CZH_SMC_Semicategory begin subsection\Background\ named_theorems smcf_cs_simps named_theorems smcf_cs_intros named_theorems smc_cn_cs_simps named_theorems smc_cn_cs_intros lemmas [smc_cs_simps] = dg_shared_cs_simps lemmas [smc_cs_intros] = dg_shared_cs_intros subsubsection\Slicing\ definition smcf_dghm :: "V \ V" where "smcf_dghm \ = [\\ObjMap\, \\ArrMap\, smc_dg (\\HomDom\), smc_dg (\\HomCod\)]\<^sub>\" text\Components.\ lemma smcf_dghm_components: shows [slicing_simps]: "smcf_dghm \\ObjMap\ = \\ObjMap\" and [slicing_simps]: "smcf_dghm \\ArrMap\ = \\ArrMap\" and [slicing_commute]: "smcf_dghm \\HomDom\ = smc_dg (\\HomDom\)" and [slicing_commute]: "smcf_dghm \\HomCod\ = smc_dg (\\HomCod\)" unfolding smcf_dghm_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} and the description of the concept of a digraph homomorphism in the previous chapter. \ locale is_semifunctor = \ \ + vfsequence \ + HomDom: semicategory \ \ + HomCod: semicategory \ \ for \ \ \ \ + assumes smcf_length[smc_cs_simps]: "vcard \ = 4\<^sub>\" and smcf_is_dghm[slicing_intros]: "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" and smcf_HomDom[smc_cs_simps]: "\\HomDom\ = \" and smcf_HomCod[smc_cs_simps]: "\\HomCod\ = \" and smcf_ArrMap_Comp[smc_cs_simps]: "\ 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\" syntax "_is_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>S\<^sub>M\<^sub>C\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" \ "CONST is_semifunctor \ \ \ \" abbreviation (input) is_cn_semifunctor :: "V \ V \ V \ V \ bool" where "is_cn_semifunctor \ \ \ \ \ \ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" syntax "_is_cn_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \<^sub>S\<^sub>M\<^sub>C\\\ _)\ [51, 51, 51] 51) translations "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" \ "CONST is_cn_semifunctor \ \ \ \" abbreviation all_smcfs :: "V \ V" where "all_smcfs \ \ set {\. \\ \. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \}" abbreviation smcfs :: "V \ V \ V \ V" where "smcfs \ \ \ \ set {\. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \}" lemmas [smc_cs_simps] = is_semifunctor.smcf_HomDom is_semifunctor.smcf_HomCod is_semifunctor.smcf_ArrMap_Comp lemma smcf_is_dghm'[slicing_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\' = smc_dg \" and "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \'" using assms(1) unfolding assms(2,3) by (rule is_semifunctor.smcf_is_dghm) lemma cn_dghm_comp_is_dghm: assumes "\ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "smcf_dghm \ : op_dg (smc_dg \) \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" using assms unfolding slicing_simps slicing_commute by (cs_concl cs_intro: slicing_intros) lemma cn_dghm_comp_is_dghm'[slicing_intros]: assumes "\ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\' = op_dg (smc_dg \)" and "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> \'" using assms(1) unfolding assms(2,3) by (rule cn_dghm_comp_is_dghm) text\Rules.\ lemma (in is_semifunctor) is_semifunctor_axioms'[smc_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_semifunctor_axioms) mk_ide rf is_semifunctor_def[unfolded is_semifunctor_axioms_def] |intro is_semifunctorI| |dest is_semifunctorD[dest]| |elim is_semifunctorE[elim]| lemmas [smc_cs_intros] = is_semifunctorD(3,4) lemma is_semifunctorI': assumes "\ \" and "vfsequence \" and "semicategory \ \" and "semicategory \ \" 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\" shows "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (intro is_semifunctorI is_dghmI, unfold smcf_dghm_components slicing_simps) (simp_all add: assms smcf_dghm_def nat_omega_simps semicategory.smc_digraph) lemma is_semifunctorD': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\ \" and "vfsequence \" and "semicategory \ \" and "semicategory \ \" 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\" by ( simp_all add: is_semifunctorD(2-9)[OF assms] is_dghmD[OF is_semifunctorD(6)[OF assms], unfolded slicing_simps] ) lemma is_semifunctorE': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" obtains "\ \" and "vfsequence \" and "semicategory \ \" and "semicategory \ \" 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\" using assms by (simp add: is_semifunctorD') text\Slicing.\ context is_semifunctor begin interpretation dghm: is_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule smcf_is_dghm) sublocale ObjMap: vsv \\\ObjMap\\ by (rule dghm.ObjMap.vsv_axioms[unfolded slicing_simps]) sublocale ArrMap: vsv \\\ArrMap\\ by (rule dghm.ArrMap.vsv_axioms[unfolded slicing_simps]) lemmas_with [unfolded slicing_simps]: smcf_ObjMap_vsv = dghm.dghm_ObjMap_vsv and smcf_ArrMap_vsv = dghm.dghm_ArrMap_vsv and smcf_ObjMap_vdomain[smc_cs_simps] = dghm.dghm_ObjMap_vdomain and smcf_ObjMap_vrange = dghm.dghm_ObjMap_vrange and smcf_ArrMap_vdomain[smc_cs_simps] = dghm.dghm_ArrMap_vdomain and smcf_ArrMap_is_arr = dghm.dghm_ArrMap_is_arr and smcf_ArrMap_is_arr''[smc_cs_intros] = dghm.dghm_ArrMap_is_arr'' and smcf_ArrMap_is_arr'[smc_cs_intros] = dghm.dghm_ArrMap_is_arr' and smcf_ObjMap_app_in_HomCod_Obj[smc_cs_intros] = dghm.dghm_ObjMap_app_in_HomCod_Obj and smcf_ArrMap_vrange = dghm.dghm_ArrMap_vrange and smcf_ArrMap_app_in_HomCod_Arr[smc_cs_intros] = dghm.dghm_ArrMap_app_in_HomCod_Arr and smcf_ObjMap_vsubset_Vset = dghm.dghm_ObjMap_vsubset_Vset and smcf_ArrMap_vsubset_Vset = dghm.dghm_ArrMap_vsubset_Vset and smcf_ObjMap_in_Vset = dghm.dghm_ObjMap_in_Vset and smcf_ArrMap_in_Vset = dghm.dghm_ArrMap_in_Vset and smcf_is_dghm_if_ge_Limit = dghm.dghm_is_dghm_if_ge_Limit and smcf_is_arr_HomCod = dghm.dghm_is_arr_HomCod and smcf_vimage_dghm_ArrMap_vsubset_Hom = dghm.dghm_vimage_dghm_ArrMap_vsubset_Hom end lemmas [smc_cs_simps] = is_semifunctor.smcf_ObjMap_vdomain is_semifunctor.smcf_ArrMap_vdomain lemmas [smc_cs_intros] = is_semifunctor.smcf_ObjMap_app_in_HomCod_Obj is_semifunctor.smcf_ArrMap_app_in_HomCod_Arr is_semifunctor.smcf_ArrMap_is_arr' text\Elementary properties.\ lemma cn_smcf_ArrMap_Comp[smc_cs_simps]: assumes "semicategory \ \" and "\ : op_smc \ \\\<^sub>S\<^sub>M\<^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- from assms(3,4) have gf: "\\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> \\ArrMap\\f\ = \\ArrMap\\g \\<^sub>A\<^bsub>op_smc \\<^esub> f\" by ( force intro: is_semifunctor.smcf_ArrMap_Comp[OF assms(2), symmetric] simp: slicing_simps smc_op_simps ) from assms show ?thesis unfolding gf by (cs_concl cs_simp: smc_op_simps) qed lemma smcf_eqI: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" and "\\ObjMap\ = \\ObjMap\" and "\\ArrMap\ = \\ArrMap\" and "\ = \" and "\ = \
" shows "\ = \" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \
\ by (rule assms(2)) show ?thesis proof(rule vsv_eqI) have dom: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: smc_cs_simps V_cs_simps) show "\\<^sub>\ \ = \\<^sub>\ \" by (cs_concl cs_simp: smc_cs_simps V_cs_simps) from assms(5,6) have sup: "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by (simp_all add: smc_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 auto qed lemma smcf_dghm_eqI: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" and "\ = \" and "\ = \
" and "smcf_dghm \ = smcf_dghm \" shows "\ = \" proof(rule smcf_eqI) from assms(5) have "smcf_dghm \\ObjMap\ = smcf_dghm \\ObjMap\" "smcf_dghm \\ArrMap\ = smcf_dghm \\ArrMap\" by simp_all then show "\\ObjMap\ = \\ObjMap\" "\\ArrMap\ = \\ArrMap\" unfolding slicing_simps by simp_all qed (auto intro: assms(1,2) simp: assms) lemma (in is_semifunctor) smcf_def: "\ = [\\ObjMap\, \\ArrMap\, \\HomDom\, \\HomCod\]\<^sub>\" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ \ = 4\<^sub>\" by (cs_concl cs_simp: smc_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) lemma (in is_semifunctor) smcf_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "\ \\<^sub>\ Vset \" proof- interpret \: \ \ by (rule assms(1)) note [smc_cs_intros] = smcf_ObjMap_in_Vset smcf_ArrMap_in_Vset HomDom.smc_in_Vset HomCod.smc_in_Vset from assms(2) show ?thesis by (subst smcf_def) (cs_concl cs_simp: smc_cs_simps cs_intro: smc_cs_intros V_cs_intros) qed lemma (in is_semifunctor) smcf_is_semifunctor_if_ge_Limit: assumes "\ \" and "\ \\<^sub>\ \" shows "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (rule is_semifunctorI) ( simp_all add: assms vfsequence_axioms smcf_is_dghm_if_ge_Limit HomDom.smc_semicategory_if_ge_Limit HomCod.smc_semicategory_if_ge_Limit smc_cs_simps ) lemma small_all_smcfs[simp]: "small {\. \\ \. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \}" proof(cases \\ \\) case True from is_semifunctor.smcf_in_Vset show ?thesis by (intro down[of _ \Vset (\ + \)\]) (auto simp: True \.\_Limit_\\ \.\_\_\\ \.intro \.\_\_\\) next case False then have "{\. \\ \. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \} = {}" by auto then show ?thesis by simp qed lemma (in is_semifunctor) smcf_in_Vset_7: "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof- note [folded VPow_iff, folded Vset_succ[OF Ord_\], smc_cs_intros] = smcf_ObjMap_vsubset_Vset smcf_ArrMap_vsubset_Vset from HomDom.smc_semicategory_in_Vset_4 have [smc_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.smc_semicategory_in_Vset_4 have [smc_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 smcf_def, succ_of_numeral) ( cs_concl cs_simp: plus_V_succ_right V_cs_simps smc_cs_simps cs_intro: smc_cs_intros V_cs_intros ) qed lemma (in \) all_smcfs_in_Vset: assumes "\ \" and "\ \\<^sub>\ \" shows "all_smcfs \ \\<^sub>\ Vset \" proof(rule vsubset_in_VsetI) interpret \: \ \ by (rule assms(1)) show "all_smcfs \ \\<^sub>\ Vset (\ + 7\<^sub>\)" proof(intro vsubsetI) fix \ assume "\ \\<^sub>\ all_smcfs \" then obtain \ \ where \: "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by clarsimp then interpret is_semifunctor \ \ \ \ . show "\ \\<^sub>\ Vset (\ + 7\<^sub>\)" by (rule smcf_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_smcfs[simp]: "small {\. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \}" by (rule down[of _ \set {\. \\ \. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \}\]) auto subsection\Opposite semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter II-2 in \cite{mac_lane_categories_2010}.\ definition op_smcf :: "V \ V" where "op_smcf \ = [\\ObjMap\, \\ArrMap\, op_smc (\\HomDom\), op_smc (\\HomCod\)]\<^sub>\" text\Components.\ lemma op_smcf_components[smc_op_simps]: shows "op_smcf \\ObjMap\ = \\ObjMap\" and "op_smcf \\ArrMap\ = \\ArrMap\" and "op_smcf \\HomDom\ = op_smc (\\HomDom\)" and "op_smcf \\HomCod\ = op_smc (\\HomCod\)" unfolding op_smcf_def dghm_field_simps by (auto simp: nat_omega_simps) text\Slicing.\ lemma op_dghm_smcf_dghm[slicing_commute]: "op_dghm (smcf_dghm \) = smcf_dghm (op_smcf \)" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (op_dghm (smcf_dghm \)) = 4\<^sub>\" unfolding op_dghm_def by (auto simp: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smcf_dghm (op_smcf \)) = 4\<^sub>\" unfolding smcf_dghm_def by (auto simp: nat_omega_simps) show "\\<^sub>\ (op_dghm (smcf_dghm \)) = \\<^sub>\ (smcf_dghm (op_smcf \))" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (op_dghm (smcf_dghm \)) \ op_dghm (smcf_dghm \)\a\ = smcf_dghm (op_smcf \)\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold smcf_dghm_def op_smcf_def op_dghm_def dghm_field_simps ) (auto simp: nat_omega_simps slicing_simps slicing_commute) qed (auto simp: smcf_dghm_def op_dghm_def) subsubsection\Further properties\ lemma (in is_semifunctor) is_semifunctor_op: "op_smcf \ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> op_smc \" proof(intro is_semifunctorI) show "vfsequence (op_smcf \)" unfolding op_smcf_def by simp show "vcard (op_smcf \) = 4\<^sub>\" unfolding op_smcf_def by (auto simp: nat_omega_simps) fix g b c f a assume "g : b \\<^bsub>op_smc \\<^esub> c" "f : a \\<^bsub>op_smc \\<^esub> b" then have "g : c \\<^bsub>\\<^esub> b" and "f : b \\<^bsub>\\<^esub> a" unfolding smc_op_simps by simp_all with is_semifunctor_axioms show "op_smcf \\ArrMap\\g \\<^sub>A\<^bsub>op_smc \\<^esub> f\ = op_smcf \\ArrMap\\g\ \\<^sub>A\<^bsub>op_smc \\<^esub> op_smcf \\ArrMap\\f\" by ( cs_concl cs_simp: smc_op_simps smc_cs_simps cs_intro: smc_op_intros smc_cs_intros ) qed ( auto simp: smc_cs_simps smc_op_simps slicing_simps slicing_commute[symmetric] is_dghm.is_dghm_op smcf_is_dghm HomCod.semicategory_op HomDom.semicategory_op ) lemma (in is_semifunctor) is_semifunctor_op': assumes "\' = op_smc \" and "\' = op_smc \" and "\' = \" shows "op_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_semifunctor_op) lemmas is_semifunctor_op'[smc_op_intros] = is_semifunctor.is_semifunctor_op' lemma (in is_semifunctor) smcf_op_smcf_op_smcf[smc_op_simps]: "op_smcf (op_smcf \) = \" proof(rule smcf_eqI, unfold smc_op_simps) show "op_smcf (op_smcf \) : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by ( metis HomCod.smc_op_smc_op_smc HomDom.smc_op_smc_op_smc is_semifunctor.is_semifunctor_op is_semifunctor_op ) qed (simp_all add: is_semifunctor_axioms) lemmas smcf_op_smcf_op_smcf[smc_op_simps] = is_semifunctor.smcf_op_smcf_op_smcf lemma eq_op_smcf_iff[smc_op_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" shows "op_smcf \ = op_smcf \ \ \ = \" proof interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \
\ by (rule assms(2)) assume prems: "op_smcf \ = op_smcf \" show "\ = \" proof(rule smcf_eqI[OF assms]) from prems R.smcf_op_smcf_op_smcf L.smcf_op_smcf_op_smcf show "\\ObjMap\ = \\ObjMap\" and "\\ArrMap\ = \\ArrMap\" by metis+ from prems R.smcf_op_smcf_op_smcf L.smcf_op_smcf_op_smcf have "\\HomDom\ = \\HomDom\" "\\HomCod\ = \\HomCod\" by auto then show "\ = \" "\ = \
" by (simp_all add: smc_cs_simps) qed qed auto subsection\Composition of covariant semifunctors\ subsubsection\Definition and elementary properties\ abbreviation (input) smcf_comp :: "V \ V \ V" (infixl "\\<^sub>S\<^sub>M\<^sub>C\<^sub>F" 55) where "smcf_comp \ dghm_comp" text\Slicing.\ lemma smcf_dghm_smcf_comp[slicing_commute]: "smcf_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M smcf_dghm \ = smcf_dghm (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)" unfolding dghm_comp_def smcf_dghm_def dghm_field_simps by (simp add: nat_omega_simps) subsubsection\Object map\ lemma smcf_comp_ObjMap_vsv[smc_cs_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ObjMap\)" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ObjMap_vsv [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ObjMap_vdomain[smc_cs_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ObjMap\) = \\Obj\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ObjMap_vdomain [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ObjMap_vrange: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ObjMap_vrange [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ObjMap_app[smc_cs_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and [simp]: "a \\<^sub>\ \\Obj\" shows "(\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ObjMap_app [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute, OF assms(3) ] ) qed subsubsection\Arrow map\ lemma smcf_comp_ArrMap_vsv[smc_cs_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\)" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ArrMap_vsv [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ArrMap_vdomain[smc_cs_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\) = \\Arr\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ArrMap_vdomain [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ArrMap_vrange: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ArrMap_vrange [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute ] ) qed lemma smcf_comp_ArrMap_app[smc_cs_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and [simp]: "f \\<^sub>\ \\Arr\" shows "(\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\\f\ = \\ArrMap\\\\ArrMap\\f\\" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_comp_ArrMap_app [ OF L.smcf_is_dghm R.smcf_is_dghm, unfolded slicing_simps slicing_commute, OF assms(3) ] ) qed subsubsection\Further properties\ lemma smcf_comp_is_semifunctor[smc_cs_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: is_semifunctor \ \ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_semifunctorI, unfold dghm_comp_components(3,4)) show "vfsequence (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)" by (simp add: dghm_comp_def) show "vcard (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) = 4\<^sub>\" unfolding dghm_comp_def by (simp add: nat_omega_simps) fix g b c f a assume "g : b \\<^bsub>\\<^esub> c" "f : a \\<^bsub>\\<^esub> b" with assms show "(\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)\ArrMap\\f\" by (cs_concl cs_simp: smc_cs_simps cs_intro: smc_cs_intros) qed ( auto simp: slicing_commute[symmetric] smc_cs_simps smc_cs_intros intro: dg_cs_intros slicing_intros ) qed lemma smcf_comp_assoc[smc_cs_simps]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "(\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \)" proof(rule smcf_eqI[of \ \ \
_ \ \
]) interpret \: is_semifunctor \ \ \
\ by (rule assms(1)) interpret \: is_semifunctor \ \ \ \ by (rule assms(2)) interpret \: is_semifunctor \ \ \ \ by (rule assms(3)) from \.is_semifunctor_axioms \.is_semifunctor_axioms \.is_semifunctor_axioms show "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" and "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" by (auto intro: smc_cs_intros) qed (simp_all add: dghm_comp_components vcomp_assoc) lemma op_smcf_smcf_comp[smc_op_simps]: "op_smcf (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) = op_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F op_smcf \" unfolding dghm_comp_def op_smcf_def dghm_field_simps by (simp add: nat_omega_simps) subsection\Composition of contravariant semifunctors\ subsubsection\Definition and elementary properties\ text\See section 1.2 in \cite{bodo_categories_1970}.\ definition smcf_cn_comp :: "V \ V \ V" (infixl \\<^sub>S\<^sub>M\<^sub>C\<^sub>F\\ 55) where "\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \ = [ \\ObjMap\ \\<^sub>\ \\ObjMap\, \\ArrMap\ \\<^sub>\ \\ArrMap\, op_smc (\\HomDom\), \\HomCod\ ]\<^sub>\" text\Components.\ lemma smcf_cn_comp_components: shows "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\ = \\ObjMap\ \\<^sub>\ \\ObjMap\" and "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\ = \\ArrMap\ \\<^sub>\ \\ArrMap\" and [smc_cn_cs_simps]: "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\HomDom\ = op_smc (\\HomDom\)" and [smc_cn_cs_simps]: "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\HomCod\ = \\HomCod\" unfolding smcf_cn_comp_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smcf_dghm_smcf_cn_comp[slicing_commute]: "smcf_dghm \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ smcf_dghm \ = smcf_dghm (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)" unfolding dghm_cn_comp_def smcf_cn_comp_def smcf_dghm_def by (simp add: nat_omega_simps slicing_commute dghm_field_simps) subsubsection\Object map: two contravariant semifunctors\ lemma smcf_cn_comp_ObjMap_vsv[smc_cn_cs_intros]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\)" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ObjMap_vsv [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ObjMap_vdomain[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\) = \\Obj\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ObjMap_vdomain [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ObjMap_vrange: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ObjMap_vrange [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ObjMap_app[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ObjMap_app [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Arrow map: two contravariant semifunctors\ lemma smcf_cn_comp_ArrMap_vsv[smc_cn_cs_intros]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\)" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ArrMap_vsv [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ArrMap_vdomain[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\) = \\Arr\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ArrMap_vdomain [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ArrMap_vrange: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ArrMap_vrange [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_comp_ArrMap_app[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Arr\" shows "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\a\ = \\ArrMap\\\\ArrMap\\a\\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \op_smc \\ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_comp_ArrMap_app [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Object map: contravariant and covariant semifunctors\ lemma smcf_cn_cov_comp_ObjMap_vsv[smc_cn_cs_intros]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\)" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ObjMap_vsv [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ObjMap_vdomain[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\) = \\Obj\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ObjMap_vdomain [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ObjMap_vrange: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\) \\<^sub>\ \\Obj\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ObjMap_vrange [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ObjMap_app[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "a \\<^sub>\ \\Obj\" shows "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ObjMap\\a\ = \\ObjMap\\\\ObjMap\\a\\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ObjMap_app [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Arrow map: contravariant and covariant semifunctors\ lemma smcf_cn_cov_comp_ArrMap_vsv[smc_cn_cs_intros]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "vsv ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\)" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ArrMap_vsv [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm[unfolded slicing_commute[symmetric]], unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ArrMap_vdomain[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\) = \\Arr\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ArrMap_vdomain [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ArrMap_vrange: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\<^sub>\ ((\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\) \\<^sub>\ \\Arr\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ArrMap_vrange [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps ] ) qed lemma smcf_cn_cov_comp_ArrMap_app[smc_cn_cs_simps]: assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "f \\<^sub>\ \\Arr\" shows "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\f\ = \\ArrMap\\\\ArrMap\\f\\" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ by (rule assms(1)) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis by ( rule dghm_cn_cov_comp_ArrMap_app [ OF L.smcf_is_dghm[unfolded slicing_commute[symmetric]] R.smcf_is_dghm, unfolded slicing_commute slicing_simps, OF assms(3) ] ) qed subsubsection\Opposite of the contravariant composition of semifunctors\ lemma op_smcf_smcf_cn_comp[smc_op_simps]: "op_smcf (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \) = op_smcf \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ op_smcf \" unfolding op_smcf_def smcf_cn_comp_def dghm_field_simps by (auto simp: nat_omega_simps) subsubsection\Further properties\ lemma smcf_cn_comp_is_semifunctor[smc_cn_cs_intros]: \\See section 1.2 in \cite{bodo_categories_1970}.\ assumes "semicategory \ \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a = f : a \\<^bsub>\'\<^esub> b" for \' f b a by (rule assms(2)) (simp_all add: smc_op_simps) interpret R: is_semifunctor \ \op_smc \\ \ \ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a = f : a \\<^bsub>\'\<^esub> b" for \' f b a by (rule assms(3)) (simp_all add: smc_op_simps) interpret \: semicategory \ \ by (rule assms(1)) show ?thesis proof(rule is_semifunctorI, unfold smcf_cn_comp_components(3,4) smc_op_simps) from assms show "smcf_dghm (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \) : smc_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: dg_cn_cs_intros slicing_intros ) fix g b c f a assume "g : b \\<^bsub>\\<^esub> c" "f : a \\<^bsub>\\<^esub> b" with assms show "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\f\" by ( cs_concl cs_simp: smc_cs_simps smc_cn_cs_simps smc_op_simps cs_intro: smc_cs_intros ) qed ( auto simp: smcf_cn_comp_def nat_omega_simps smc_cs_simps smc_op_simps smc_cs_intros ) qed lemma smcf_cn_cov_comp_is_semifunctor[smc_cs_intros]: \\See section 1.2 in \cite{bodo_categories_1970}.\ assumes "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" proof- interpret L: is_semifunctor \ \op_smc \\ \ \ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a = f : a \\<^bsub>\'\<^esub> b" for \' f b a by (rule assms(1)) (simp_all add: smc_op_simps) interpret R: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_semifunctorI, unfold smcf_cn_comp_components(3,4) smc_op_simps) from assms show "smcf_dghm (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \) : smc_dg (op_smc \) \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: dg_cn_cs_intros slicing_intros ) show "vfsequence (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)" unfolding smcf_cn_comp_def by simp show "vcard (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \) = 4\<^sub>\" unfolding smcf_cn_comp_def by (auto simp: nat_omega_simps) show "op_smc (\\HomDom\) = op_smc \" by (simp add: smc_cs_simps) show "\\HomCod\ = \" by (simp add: smc_cs_simps) fix g b c f a assume "g : c \\<^bsub>\\<^esub> b" "f : b \\<^bsub>\\<^esub> a" with assms show "(\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\f \\<^sub>A\<^bsub>\\<^esub> g\ = (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> (\ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \)\ArrMap\\f\" by ( cs_concl cs_simp: smc_cs_simps smc_cn_cs_simps smc_op_simps cs_intro: smc_cs_intros ) qed (auto intro: smc_cs_intros smc_op_intros) qed lemma smcf_cov_cn_comp_is_semifunctor[smc_cn_cs_intros]: \\See section 1.2 in \cite{bodo_categories_1970}.\ assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \<^sub>S\<^sub>M\<^sub>C\\\<^bsub>\\<^esub> \" using assms by (rule smcf_comp_is_semifunctor) subsection\Identity semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ abbreviation (input) smcf_id :: "V \ V" where "smcf_id \ dghm_id" text\Slicing.\ lemma smcf_dghm_smcf_id[slicing_commute]: "dghm_id (smc_dg \) = smcf_dghm (smcf_id \)" unfolding dghm_id_def smc_dg_def smcf_dghm_def dghm_field_simps dg_field_simps by (simp add: nat_omega_simps) context semicategory begin interpretation dg: digraph \ \smc_dg \\ by (rule smc_digraph) lemmas_with [unfolded slicing_simps]: smc_dghm_id_is_dghm = dg.dg_dghm_id_is_dghm end subsubsection\Object map\ lemmas [smc_cs_simps] = dghm_id_ObjMap_app subsubsection\Arrow map\ lemmas [smc_cs_simps] = dghm_id_ArrMap_app subsubsection\Opposite identity semifunctor\ lemma op_smcf_smcf_id[smc_op_simps]: "op_smcf (smcf_id \) = smcf_id (op_smc \)" unfolding dghm_id_def op_smc_def op_smcf_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\An identity semifunctor is a semifunctor\ lemma (in semicategory) smc_smcf_id_is_semifunctor: "smcf_id \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof(rule is_semifunctorI, unfold dghm_id_components) from smc_dghm_id_is_dghm show "smcf_dghm (smcf_id \) : smc_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" by (auto simp: slicing_simps slicing_commute) fix g b c f a assume "g : b \\<^bsub>\\<^esub> c" "f : a \\<^bsub>\\<^esub> b" then show "vid_on (\\Arr\)\g \\<^sub>A\<^bsub>\\<^esub> f\ = vid_on (\\Arr\)\g\ \\<^sub>A\<^bsub>\\<^esub> vid_on (\\Arr\)\f\" by (metis smc_is_arrD(1) smc_Comp_is_arr vid_on_eq_atI) qed (auto simp: semicategory_axioms dghm_id_def nat_omega_simps) lemma (in semicategory) smc_smcf_id_is_semifunctor': assumes "\ = \" and "\ = \" shows "smcf_id \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" unfolding assms by (rule smc_smcf_id_is_semifunctor) lemmas [smc_cs_intros] = semicategory.smc_smcf_id_is_semifunctor' subsubsection\Further properties\ lemma (in is_semifunctor) smcf_smcf_comp_smcf_id_left[smc_cs_simps]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ "smcf_id \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = \" by (rule smcf_eqI, unfold dghm_id_components dghm_comp_components) (auto simp: smcf_ObjMap_vrange smcf_ArrMap_vrange intro: smc_cs_intros) lemmas [smc_cs_simps] = is_semifunctor.smcf_smcf_comp_smcf_id_left lemma (in is_semifunctor) smcf_smcf_comp_smcf_id_right[smc_cs_simps]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F smcf_id \ = \" by (rule smcf_eqI, unfold dghm_id_components dghm_comp_components) ( auto simp: smcf_ObjMap_vrange smcf_ArrMap_vrange smc_cs_simps intro: smc_cs_intros ) lemmas [smc_cs_simps] = is_semifunctor.smcf_smcf_comp_smcf_id_right subsection\Constant semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter III-3 in \cite{mac_lane_categories_2010}.\ abbreviation (input) smcf_const :: "V \ V \ V \ V \ V" where "smcf_const \ dghm_const" text\Slicing.\ lemma smcf_dghm_smcf_const[slicing_commute]: "dghm_const (smc_dg \) (smc_dg \
) a f = smcf_dghm (smcf_const \ \
a f)" unfolding dghm_const_def smc_dg_def smcf_dghm_def dghm_field_simps dg_field_simps by (simp add: nat_omega_simps) subsubsection\Object map\ lemmas [smc_cs_simps] = dghm_const_ObjMap_app subsubsection\Arrow map\ lemmas [smc_cs_simps] = dghm_const_ArrMap_app subsubsection\Opposite constant semifunctor\ lemma op_smcf_smcf_const[smc_op_simps]: "op_smcf (smcf_const \ \
a f) = smcf_const (op_smc \) (op_smc \
) a f" unfolding dghm_const_def op_smc_def op_smcf_def dghm_field_simps dg_field_simps by (auto simp: nat_omega_simps) subsubsection\A constant semifunctor is a semifunctor\ lemma smcf_const_is_semifunctor: assumes "semicategory \ \" and "semicategory \ \
" and "f : a \\<^bsub>\
\<^esub> a" and [simp]: "f \\<^sub>A\<^bsub>\
\<^esub> f = f" shows "smcf_const \ \
a f : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \
" proof- interpret \: semicategory \ \ by (rule assms(1)) interpret \
: semicategory \ \
by (rule assms(2)) show ?thesis proof(intro is_semifunctorI, tactic\distinct_subgoals_tac\) from assms show "smcf_dghm (dghm_const \ \
a f) : smc_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \
" by ( cs_concl cs_simp: slicing_commute[symmetric] cs_intro: dg_cs_intros slicing_intros ) show "vfsequence (smcf_const \ \
a f)" unfolding dghm_const_def by simp show "vcard (smcf_const \ \
a f) = 4\<^sub>\" unfolding dghm_const_def by (simp add: nat_omega_simps) fix g' b c f' a' assume "g' : b \\<^bsub>\\<^esub> c" "f' : a' \\<^bsub>\\<^esub> b" with assms(1-3) show "smcf_const \ \
a f\ArrMap\\g' \\<^sub>A\<^bsub>\\<^esub> f'\ = smcf_const \ \
a f\ArrMap\\g'\ \\<^sub>A\<^bsub>\
\<^esub> smcf_const \ \
a f\ArrMap\\f'\" by (cs_concl cs_simp: assms(4) smc_cs_simps cs_intro: smc_cs_intros) qed (auto simp: assms(1,2) dghm_const_components) qed lemma smcf_const_is_semifunctor'[smc_cs_intros]: assumes "semicategory \ \" and "semicategory \ \
" and "f : a \\<^bsub>\
\<^esub> a" and "f \\<^sub>A\<^bsub>\
\<^esub> f = f" and "\ = \" and "\ = \
" shows "smcf_const \ \
a f : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" using assms(1-4) unfolding assms(5,6) by (rule smcf_const_is_semifunctor) +subsubsection\Further properties\ + +lemma (in is_semifunctor) smcf_smcf_comp_smcf_const[smc_cs_simps]: + assumes "semicategory \ \" and "f : a \\<^bsub>\\<^esub> a" and "f \\<^sub>A\<^bsub>\\<^esub> f = f" + shows "smcf_const \ \ a f \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = smcf_const \ \ a f" +proof(rule smcf_dghm_eqI) + interpret \: semicategory \ \ by (rule assms(1)) + from assms(2) show "smcf_const \ \ a f \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: smc_cs_simps assms(3) cs_intro: smc_cs_intros) + from assms(2) show "smcf_const \ \ a f : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" + by (cs_concl cs_simp: smc_cs_simps assms(3) cs_intro: smc_cs_intros) + from is_dghm.dghm_dghm_comp_dghm_const[ + OF smcf_is_dghm \.smc_digraph, unfolded slicing_simps, OF assms(2) + ] + show "smcf_dghm (smcf_const \ \ a f \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) = smcf_dghm (smcf_const \ \ a f)" + by (cs_prems cs_simp: slicing_simps slicing_commute) +qed simp_all + +lemmas [smc_cs_simps] = is_semifunctor.smcf_smcf_comp_smcf_const + + subsection\Faithful semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_ft_semifunctor = is_semifunctor \ \ \ \ for \ \ \ \ + assumes ft_smcf_is_ft_dghm: "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> smc_dg \" syntax "_is_ft_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>S\<^sub>M\<^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>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_ft_semifunctor \ \ \ \" lemma (in is_ft_semifunctor) ft_smcf_is_ft_dghm'[slicing_intros]: assumes "\' = smc_dg \" and "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule ft_smcf_is_ft_dghm) lemmas [slicing_intros] = is_ft_semifunctor.ft_smcf_is_ft_dghm' text\Rules.\ lemma (in is_ft_semifunctor) is_ft_semifunctor_axioms'[smcf_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^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 is_ft_semifunctor_axioms) mk_ide rf is_ft_semifunctor_def[unfolded is_ft_semifunctor_axioms_def] |intro is_ft_semifunctorI| |dest is_ft_semifunctorD[dest]| |elim is_ft_semifunctorE[elim]| lemmas [smcf_cs_intros] = is_ft_semifunctorD(1) lemma is_ft_semifunctorI': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\a b. \ a \\<^sub>\ \\Obj\; b \\<^sub>\ \\Obj\ \ \ v11 (\\ArrMap\ \\<^sup>l\<^sub>\ Hom \ a b)" shows "\ : \ \\\<^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> \" using assms by (intro is_ft_semifunctorI) ( simp_all add: assms(1) is_ft_dghmI[OF is_semifunctorD(6)[OF assms(1)], unfolded slicing_simps] ) lemma is_ft_semifunctorD': assumes "\ : \ \\\<^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> \" shows "\ : \ \\\<^sub>S\<^sub>M\<^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_semifunctorD[OF assms(1)] is_ft_dghmD(2)[ OF is_ft_semifunctorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_ft_semifunctorE': assumes "\ : \ \\\<^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> \" obtains "\ : \ \\\<^sub>S\<^sub>M\<^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_semifunctorD') text\Elementary properties.\ context is_ft_semifunctor begin interpretation dghm: is_ft_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule ft_smcf_is_ft_dghm) lemmas_with [unfolded slicing_simps]: ft_smcf_v11_on_Hom = dghm.ft_dghm_v11_on_Hom end subsubsection\Opposite faithful semifunctor\ lemma (in is_ft_semifunctor) is_ft_semifunctor_op: "op_smcf \ : op_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> op_smc \" by ( rule is_ft_semifunctorI, unfold smc_op_simps slicing_simps slicing_commute[symmetric] ) ( simp_all add: is_semifunctor_op is_ft_dghm.ft_dghm_op_dghm_is_ft_dghm ft_smcf_is_ft_dghm ) lemma (in is_ft_semifunctor) is_ft_semifunctor_op'[smc_op_intros]: assumes "\' = op_smc \" and "\' = op_smc \" shows "op_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 is_ft_semifunctor_op) lemmas is_ft_semifunctor_op[smc_op_intros] = is_ft_semifunctor.is_ft_semifunctor_op' subsubsection\ The composition of faithful semifunctors is a faithful semifunctor \ lemma smcf_comp_is_ft_semifunctor[smcf_cs_intros]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^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> \" and "\ : \ \\\<^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> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^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> \" proof(intro is_ft_semifunctorI) interpret \: is_ft_semifunctor \ \ \ \ by (simp add: assms(1)) interpret \: is_ft_semifunctor \ \ \ \ by (simp add: assms(2)) from \.is_semifunctor_axioms \.is_semifunctor_axioms show \\: "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: smc_cs_intros) then interpret is_semifunctor \ \ \ \\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \\ . show "smcf_dghm (\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \) : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>a\<^sub>i\<^sub>t\<^sub>h\<^sub>f\<^sub>u\<^sub>l\<^bsub>\\<^esub> smc_dg \" unfolding slicing_simps slicing_commute[symmetric] by (auto intro: dghm_cs_intros slicing_intros) qed subsection\Full semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_fl_semifunctor = is_semifunctor \ \ \ \ for \ \ \ \ + assumes fl_smcf_is_fl_dghm: "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> smc_dg \" syntax "_is_fl_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" \ "CONST is_fl_semifunctor \ \ \ \" lemma (in is_fl_semifunctor) fl_smcf_is_fl_dghm'[slicing_intros]: assumes "\' = smc_dg \" and "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule fl_smcf_is_fl_dghm) lemmas [slicing_intros] = is_fl_semifunctor.fl_smcf_is_fl_dghm' text\Rules.\ mk_ide rf is_fl_semifunctor_def[unfolded is_fl_semifunctor_axioms_def] |intro is_fl_semifunctorI| |dest is_fl_semifunctorD[dest]| |elim is_fl_semifunctorE[elim]| lemmas [smcf_cs_intros] = is_fl_semifunctorD(1) lemma is_fl_semifunctorI': assumes "\ : \ \\\<^sub>S\<^sub>M\<^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>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" using assms by (intro is_fl_semifunctorI) ( simp_all add: assms(1) is_fl_dghmI[OF is_semifunctorD(6)[OF assms(1)], unfolded slicing_simps] ) lemma is_fl_semifunctorD': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" shows "\ : \ \\\<^sub>S\<^sub>M\<^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_semifunctorD[OF assms(1)] is_fl_dghmD(2)[ OF is_fl_semifunctorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_fl_semifunctorE': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" obtains "\ : \ \\\<^sub>S\<^sub>M\<^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_semifunctorD') text\Elementary properties.\ context is_fl_semifunctor begin interpretation dghm: is_fl_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule fl_smcf_is_fl_dghm) lemmas_with [unfolded slicing_simps]: fl_smcf_surj_on_Hom = dghm.fl_dghm_surj_on_Hom end subsubsection\Opposite full semifunctor\ lemma (in is_fl_semifunctor) is_fl_semifunctor_op: "op_smcf \ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> op_smc \" by ( rule is_fl_semifunctorI, unfold smc_op_simps slicing_simps slicing_commute[symmetric] ) ( simp_all add: is_semifunctor_op is_fl_dghm.fl_dghm_op_dghm_is_fl_dghm fl_smcf_is_fl_dghm ) lemma (in is_fl_semifunctor) is_fl_semifunctor_op'[smc_op_intros]: assumes "\' = op_smc \" and "\' = op_smc \" shows "op_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \'" unfolding assms by (rule is_fl_semifunctor_op) lemmas is_fl_semifunctor_op[smc_op_intros] = is_fl_semifunctor.is_fl_semifunctor_op subsubsection\The composition of full semifunctors is a full semifunctor\ lemma smcf_comp_is_fl_semifunctor[smcf_cs_intros]: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> \" proof(intro is_fl_semifunctorI) interpret \: is_fl_semifunctor \ \ \ \ using assms(2) by simp interpret \: is_fl_semifunctor \ \ \ \ using assms(1) by simp from \.is_semifunctor_axioms \.is_semifunctor_axioms show "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: smc_cs_intros) show "smcf_dghm (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>u\<^sub>l\<^sub>l\<^bsub>\\<^esub> smc_dg \" unfolding slicing_commute[symmetric] by (auto intro: dghm_cs_intros slicing_intros) qed subsection\Fully faithful semifunctor\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale is_ff_semifunctor = is_ft_semifunctor \ \ \ \ + is_fl_semifunctor \ \ \ \ for \ \ \ \ syntax "_is_ff_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" \ "CONST is_ff_semifunctor \ \ \ \" text\Rules.\ mk_ide rf is_ff_semifunctor_def |intro is_ff_semifunctorI| |dest is_ff_semifunctorD[dest]| |elim is_ff_semifunctorE[elim]| lemmas [smcf_cs_intros] = is_ff_semifunctorD text\Elementary properties.\ lemma (in is_ff_semifunctor) ff_smcf_is_ff_dghm: "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> smc_dg \" by (rule is_ff_dghmI) (auto intro: slicing_intros) lemma (in is_ff_semifunctor) ff_smcf_is_ff_dghm'[slicing_intros]: assumes "\' = smc_dg \" and "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \'" unfolding assms by (rule ff_smcf_is_ff_dghm) lemmas [slicing_intros] = is_ff_semifunctor.ff_smcf_is_ff_dghm' subsubsection\Opposite fully faithful semifunctor\ lemma (in is_ff_semifunctor) is_ff_semifunctor_op: "op_smcf \ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> op_smc \" by (rule is_ff_semifunctorI) (auto simp: is_fl_semifunctor_op is_ft_semifunctor_op) lemma (in is_ff_semifunctor) is_ff_semifunctor_op'[smc_op_intros]: assumes "\' = op_smc \" and "\' = op_smc \" shows "op_smcf \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \'" unfolding assms by (rule is_ff_semifunctor_op) lemmas is_ff_semifunctor_op[smc_op_intros] = is_ff_semifunctor.is_ff_semifunctor_op' subsubsection\ The composition of fully faithful semifunctors is a fully faithful semifunctor \ lemma smcf_comp_is_ff_semifunctor[smcf_cs_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" using assms by (intro is_ff_semifunctorI, elim is_ff_semifunctorE) (auto intro: smcf_cs_intros) subsection\Isomorphism of semicategories\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ locale is_iso_semifunctor = is_semifunctor \ \ \ \ for \ \ \ \ + assumes iso_smcf_is_iso_dghm: "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_dg \" syntax "_is_iso_semifunctor" :: "V \ V \ V \ V \ bool" (\(_ :/ _ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\ _)\ [51, 51, 51] 51) translations "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" \ "CONST is_iso_semifunctor \ \ \ \" lemma (in is_iso_semifunctor) iso_smcf_is_iso_dghm'[slicing_intros]: assumes "\' = smc_dg \" "\' = smc_dg \" shows "smcf_dghm \ : \' \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" unfolding assms by (rule iso_smcf_is_iso_dghm) lemmas [slicing_intros] = is_iso_semifunctor.iso_smcf_is_iso_dghm' text\Rules.\ lemma (in is_iso_semifunctor) is_iso_semifunctor_axioms'[smcf_cs_intros]: assumes "\' = \" and "\' = \" and "\' = \" shows "\ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> \'" unfolding assms by (rule is_iso_semifunctor_axioms) mk_ide rf is_iso_semifunctor_def[unfolded is_iso_semifunctor_axioms_def] |intro is_iso_semifunctorI| |dest is_iso_semifunctorD[dest]| |elim is_iso_semifunctorE[elim]| lemma is_iso_semifunctorI': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "v11 (\\ObjMap\)" and "v11 (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" shows "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms by (intro is_iso_semifunctorI) ( simp_all add: assms(1) is_iso_dghmI[OF is_semifunctorD(6)[OF assms(1)], unfolded slicing_simps] ) lemma is_iso_semifunctorD': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "v11 (\\ObjMap\)" and "v11 (\\ArrMap\)" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" by ( simp_all add: is_iso_semifunctorD[OF assms(1)] is_iso_dghmD(2-5)[ OF is_iso_semifunctorD(2)[OF assms(1)], unfolded slicing_simps ] ) lemma is_iso_semifunctorE': assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" obtains "\ : \ \\\<^sub>S\<^sub>M\<^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_semifunctorD') text\Elementary properties.\ context is_iso_semifunctor begin interpretation dghm: is_iso_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule iso_smcf_is_iso_dghm) lemmas_with [unfolded slicing_simps]: iso_smcf_ObjMap_vrange[smcf_cs_simps] = dghm.iso_dghm_ObjMap_vrange and iso_smcf_ArrMap_vrange[smcf_cs_simps] = dghm.iso_dghm_ArrMap_vrange sublocale ObjMap: v11 \\\ObjMap\\ rewrites "\\<^sub>\ (\\ObjMap\) = \\Obj\" and "\\<^sub>\ (\\ObjMap\) = \\Obj\" by (rule dghm.iso_dghm_ObjMap_v11[unfolded slicing_simps]) (simp_all add: smc_cs_simps smcf_cs_simps) sublocale ArrMap: v11 \\\ArrMap\\ rewrites "\\<^sub>\ (\\ArrMap\) = \\Arr\" and "\\<^sub>\ (\\ArrMap\) = \\Arr\" by (rule dghm.iso_dghm_ArrMap_v11[unfolded slicing_simps]) (simp_all add: smc_cs_simps smcf_cs_simps) lemmas_with [unfolded slicing_simps]: iso_smcf_Obj_HomDom_if_Obj_HomCod[elim] = dghm.iso_dghm_Obj_HomDom_if_Obj_HomCod and iso_smcf_Arr_HomDom_if_Arr_HomCod[elim] = dghm.iso_dghm_Arr_HomDom_if_Arr_HomCod and iso_smcf_ObjMap_eqE[elim] = dghm.iso_dghm_ObjMap_eqE and iso_smcf_ArrMap_eqE[elim] = dghm.iso_dghm_ArrMap_eqE end sublocale is_iso_semifunctor \ is_ff_semifunctor proof- interpret dghm: is_iso_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule iso_smcf_is_iso_dghm) show "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>f\<^sub>f\<^bsub>\\<^esub> \" by unfold_locales qed lemmas (in is_iso_semifunctor) iso_smcf_is_ff_semifunctor = is_ff_semifunctor_axioms lemmas [smcf_cs_intros] = is_iso_semifunctor.iso_smcf_is_ff_semifunctor subsubsection\Opposite isomorphism of semicategories\ lemma (in is_iso_semifunctor) is_iso_semifunctor_op: "op_smcf \ : op_smc \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> op_smc \" by ( rule is_iso_semifunctorI, unfold smc_op_simps slicing_simps slicing_commute[symmetric] ) ( simp_all add: is_semifunctor_op is_iso_dghm.is_iso_dghm_op iso_smcf_is_iso_dghm ) lemmas is_iso_semifunctor_op[smc_op_intros] = is_iso_semifunctor.is_iso_semifunctor_op subsubsection\ The composition of isomorphisms of semicategories is an isomorphism of semicategories \ lemma smcf_comp_is_iso_semifunctor[smcf_cs_intros]: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof(intro is_iso_semifunctorI) interpret \: is_iso_semifunctor \ \ \ \ using assms by auto interpret \: is_iso_semifunctor \ \ \ \ using assms by auto from \.is_semifunctor_axioms \.is_semifunctor_axioms show "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (auto intro: smcf_cs_intros) show "smcf_dghm (\ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \) : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_dg \" by ( auto intro: dghm_cs_intros slicing_intros simp: slicing_commute[symmetric] ) qed subsection\Inverse semifunctor\ abbreviation (input) inv_smcf :: "V \ V" where "inv_smcf \ inv_dghm" lemmas [smc_cs_simps] = inv_dghm_components(3,4) text\Slicing.\ lemma dghm_inv_smcf[slicing_commute]: "inv_dghm (smcf_dghm \) = smcf_dghm (inv_smcf \)" unfolding smcf_dghm_def inv_dghm_def dghm_field_simps by (simp_all add: nat_omega_simps) context is_iso_semifunctor begin interpretation dghm: is_iso_dghm \ \smc_dg \\ \smc_dg \\ \smcf_dghm \\ by (rule iso_smcf_is_iso_dghm) lemmas_with [unfolded slicing_simps slicing_commute]: inv_smcf_ObjMap_v11 = dghm.inv_dghm_ObjMap_v11 and inv_smcf_ObjMap_vdomain = dghm.inv_dghm_ObjMap_vdomain and inv_smcf_ObjMap_app = dghm.inv_dghm_ObjMap_app and inv_smcf_ObjMap_vrange = dghm.inv_dghm_ObjMap_vrange and inv_smcf_ArrMap_v11 = dghm.inv_dghm_ArrMap_v11 and inv_smcf_ArrMap_vdomain = dghm.inv_dghm_ArrMap_vdomain and inv_smcf_ArrMap_app = dghm.inv_dghm_ArrMap_app and inv_smcf_ArrMap_vrange = dghm.inv_dghm_ArrMap_vrange and iso_smcf_ObjMap_inv_smcf_ObjMap_app = dghm.iso_dghm_ObjMap_inv_dghm_ObjMap_app and iso_smcf_ArrMap_inv_smcf_ArrMap_app = dghm.iso_dghm_ArrMap_inv_dghm_ArrMap_app and iso_smcf_HomDom_is_arr_conv = dghm.iso_dghm_HomDom_is_arr_conv and iso_smcf_HomCod_is_arr_conv = dghm.iso_dghm_HomCod_is_arr_conv end lemmas [smcf_cs_intros] = is_iso_semifunctor.inv_smcf_ObjMap_v11 is_iso_semifunctor.inv_smcf_ArrMap_v11 lemmas [smcf_cs_simps] = is_iso_semifunctor.inv_smcf_ObjMap_vdomain is_iso_semifunctor.inv_smcf_ObjMap_app is_iso_semifunctor.inv_smcf_ObjMap_vrange is_iso_semifunctor.inv_smcf_ArrMap_vdomain is_iso_semifunctor.inv_smcf_ArrMap_app is_iso_semifunctor.inv_smcf_ArrMap_vrange is_iso_semifunctor.iso_smcf_ObjMap_inv_smcf_ObjMap_app is_iso_semifunctor.iso_smcf_ArrMap_inv_smcf_ArrMap_app subsection\ An isomorphism of semicategories is an isomorphism in the category \SemiCAT\ \ lemma is_arr_isomorphism_is_iso_semifunctor: \\See Chapter I-3 in \cite{mac_lane_categories_2010}.\ assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = smcf_id \" and "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = smcf_id \" shows "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof- interpret \: is_semifunctor \ \ \ \ by (rule assms(1)) interpret \: is_semifunctor \ \ \ \ by (rule assms(2)) show ?thesis proof(rule is_iso_semifunctorI) have dg_\\\: "smcf_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M smcf_dghm \ = dghm_id (smc_dg \)" by (simp add: assms(3) smcf_dghm_smcf_id smcf_dghm_smcf_comp) have dg_\\\: "smcf_dghm \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M smcf_dghm \ = dghm_id (smc_dg \)" by (simp add: assms(4) smcf_dghm_smcf_id smcf_dghm_smcf_comp) from \.smcf_is_dghm \.smcf_is_dghm dg_\\\ dg_\\\ show "smcf_dghm \ : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_dg \" by (rule is_arr_isomorphism_is_iso_dghm) qed (simp add: \.is_semifunctor_axioms) qed lemma is_iso_semifunctor_is_arr_isomorphism: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows [smcf_cs_intros]: "inv_smcf \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" and "inv_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = smcf_id \" and "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F inv_smcf \ = smcf_id \" proof- let ?\ = \inv_smcf \\ interpret is_iso_semifunctor \ \ \ \ by (rule assms(1)) note is_iso_dghm = is_iso_dghm_is_arr_isomorphism[OF iso_smcf_is_iso_dghm] show \: "?\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" proof ( intro is_iso_semifunctorI is_semifunctorI; (unfold slicing_commute[symmetric])? ) show "vfsequence (inv_smcf \)" unfolding inv_dghm_def by simp show "vcard (inv_smcf \) = 4\<^sub>\" unfolding inv_dghm_def by (simp add: nat_omega_simps) show inv_iso_dghm_\: "inv_dghm (smcf_dghm \) : smc_dg \ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_dg \" by (rule is_iso_dghm(1)) show inv_dghm_\: "inv_dghm (smcf_dghm \) : smc_dg \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" by (rule is_iso_dghmD(1)[OF inv_iso_dghm_\]) fix b c g a f assume prems: "g : b \\<^bsub>\\<^esub> c" "f : a \\<^bsub>\\<^esub> b" note is_arr_inv = is_dghm.dghm_ArrMap_is_arr[ OF inv_dghm_\, unfolded slicing_simps slicing_commute ] from prems is_arr_inv[OF prems(1)] is_arr_inv[OF prems(2)] show "inv_smcf \\ArrMap\\g \\<^sub>A\<^bsub>\\<^esub> f\ = inv_smcf \\ArrMap\\g\ \\<^sub>A\<^bsub>\\<^esub> inv_smcf \\ArrMap\\f\" unfolding inv_dghm_components by (intro v11.v11_vconverse_app) ( cs_concl cs_intro: smc_cs_intros V_cs_intros cs_simp: V_cs_simps smc_cs_simps )+ qed (auto simp: smc_cs_simps intro: smc_cs_intros) show "?\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ = smcf_id \" proof(rule smcf_eqI, unfold dghm_comp_components inv_dghm_components) from \ is_semifunctor_axioms show "inv_smcf \ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (blast intro: smc_cs_intros) qed ( simp_all add: HomDom.smc_smcf_id_is_semifunctor ObjMap.v11_vcomp_vconverse ArrMap.v11_vcomp_vconverse dghm_id_components ) show "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F inv_smcf \ = smcf_id \" proof(rule smcf_eqI, unfold dghm_comp_components inv_dghm_components) from \ is_semifunctor_axioms show "\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>F inv_smcf \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" by (blast intro: smc_cs_intros) qed ( simp_all add: HomCod.smc_smcf_id_is_semifunctor ObjMap.v11_vcomp_vconverse' ArrMap.v11_vcomp_vconverse' dghm_id_components ) qed subsubsection\An identity semifunctor is an isomorphism of semicategories\ lemma (in semicategory) smc_smcf_id_is_iso_semifunctor: "smcf_id \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (rule is_iso_semifunctorI, unfold slicing_simps slicing_commute[symmetric]) ( simp_all add: smc_smcf_id_is_semifunctor digraph.dg_dghm_id_is_iso_dghm smc_digraph ) lemma (in semicategory) smc_smcf_id_is_iso_semifunctor'[smcf_cs_intros]: assumes "\' = \" and "\' = \" shows "smcf_id \ : \' \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \'" unfolding assms by (rule smc_smcf_id_is_iso_semifunctor) lemmas [smcf_cs_intros] = semicategory.smc_smcf_id_is_iso_semifunctor' subsection\Isomorphic semicategories\ subsubsection\Definition and elementary properties\ text\See Chapter I-3 in \cite{mac_lane_categories_2010}).\ locale iso_semicategory = L: semicategory \ \ + R: semicategory \ \ for \ \ \ + assumes iso_smc_is_iso_semifunctor: "\\. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" notation iso_semicategory (infixl "\\<^sub>S\<^sub>M\<^sub>C\" 50) text\Rules.\ lemma iso_semicategoryI: assumes "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" using assms unfolding iso_semicategory_def iso_semicategory_axioms_def by blast lemma iso_semicategoryD[dest]: assumes "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\\. \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms unfolding iso_semicategory_def iso_semicategory_axioms_def by simp_all lemma iso_semicategoryE[elim]: assumes "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" obtains \ where "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" using assms by auto text\Elementary properties.\ lemma (in iso_semicategory) iso_smc_iso_digraph: "smc_dg \ \\<^sub>D\<^sub>G\<^bsub>\\<^esub> smc_dg \" using iso_smc_is_iso_semifunctor by (auto intro: slicing_intros iso_digraphI) subsubsection\A semicategory isomorphism is an equivalence relation\ lemma iso_semicategory_refl: assumes "semicategory \ \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof(rule iso_semicategoryI[of _ _ _ \smcf_id \\]) interpret semicategory \ \ by (rule assms) show "smcf_id \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (simp add: smc_smcf_id_is_iso_semifunctor) qed lemma iso_semicategory_sym[sym]: assumes "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof- interpret iso_semicategory \ \ \ by (rule assms) from iso_smc_is_iso_semifunctor obtain \ where "\ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by clarsimp then have "inv_smcf \ : \ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" by (simp add: is_iso_semifunctor_is_arr_isomorphism(1)) then show ?thesis by (auto intro: iso_semicategoryI) qed lemma iso_semicategory_trans[trans]: assumes "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" and "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" shows "\ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> \" proof- interpret L: iso_semicategory \ \ \ by (rule assms(1)) interpret R: iso_semicategory \ \ \ by (rule assms(2)) from L.iso_smc_is_iso_semifunctor R.iso_smc_is_iso_semifunctor show ?thesis by (auto intro: iso_semicategoryI smcf_cs_intros) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Universal_Constructions/ROOT b/thys/CZH_Universal_Constructions/ROOT --- a/thys/CZH_Universal_Constructions/ROOT +++ b/thys/CZH_Universal_Constructions/ROOT @@ -1,15 +1,15 @@ chapter AFP session CZH_Universal_Constructions (AFP) = CZH_Elementary_Categories + - options [timeout = 7200] + options [timeout = 14400] directories czh_ucategories theories CZH_UCAT_Conclusions document_files "iman.sty" "extra.sty" "isar.sty" "style.sty" "root.tex" "root.bib" 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,3067 +1,3308 @@ (* 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 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\Opposite universal arrow for Kan extensions\ + + +subsubsection\Definition and elementary properties\ + + +text\ +The following definition is merely a convenience utility for +the exposition of dual results associated with the formula for +the right Kan extension and the pointwise right Kan extension. +\ + +definition op_ua :: "(V \ V) \ V \ V \ V" + where "op_ua lim_Obj \ c = + [ + lim_Obj c\UObj\, + op_ntcf (lim_Obj c\UArr\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf (op_cf_obj_comma \ c) + ]\<^sub>\" + + +text\Components.\ + +lemma op_ua_components: + shows [cat_op_simps]: "op_ua lim_Obj \ c\UObj\ = lim_Obj c\UObj\" + and "op_ua lim_Obj \ c\UArr\ = + op_ntcf (lim_Obj c\UArr\) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf (op_cf_obj_comma \ c)" + unfolding op_ua_def ua_field_simps by (simp_all add: nat_omega_simps) + + +subsubsection\Opposite universal arrow for Kan extensions is a limit\ + +lemma op_ua_UArr_is_cat_limit: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "c \\<^sub>\ \\Obj\" + and "u : \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m r : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" + shows "op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf (op_cf_obj_comma \ c) : + r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) : c \\<^sub>C\<^sub>F (op_cf \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" +proof- + + note [cf_cs_simps] = is_iso_functor_is_arr_isomorphism(2,3) + + let ?op_\ = \\c. op_cf_obj_comma \ c\ + let ?op_\c = \?op_\ c\ + and ?op_ua_UArr = \op_ntcf u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf (op_cf_obj_comma \ c)\ + + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + interpret u: is_cat_colimit \ \\ \<^sub>C\<^sub>F\ c\ \ \\ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c\ r u + by (rule assms(4)) + + from \.op_cf_cf_obj_comma_proj[OF assms(3)] have + "op_cf (\ \<^sub>C\<^sub>F\\<^sub>O c) \\<^sub>C\<^sub>F inv_cf (?op_\ c) = + c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F (?op_\ c) \\<^sub>C\<^sub>F inv_cf (?op_\ c)" + by simp + from this assms(3) have [cat_comma_cs_simps]: + "op_cf (\ \<^sub>C\<^sub>F\\<^sub>O c) \\<^sub>C\<^sub>F inv_cf (?op_\ c) = c \<^sub>O\\<^sub>C\<^sub>F (op_cf \)" + by + ( + cs_prems + cs_simp: cat_cs_simps cat_comma_cs_simps cf_cs_simps cat_op_simps + cs_intro: cf_cs_intros cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + from assms(3) show "?op_ua_UArr : + r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) : c \\<^sub>C\<^sub>F (op_cf \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by + ( + cs_concl + cs_simp: + cf_cs_simps cat_cs_simps cat_comma_cs_simps cat_op_simps + \.op_cf_cf_obj_comma_proj[symmetric] + cs_intro: + cat_cs_intros + cf_cs_intros + cat_lim_cs_intros + cat_comma_cs_intros + cat_op_intros + ) + +qed + +context + fixes lim_Obj :: "V \ V" and c :: V +begin + +lemmas op_ua_UArr_is_cat_limit' = op_ua_UArr_is_cat_limit + [ + unfolded op_ua_components(2)[symmetric], + where u=\lim_Obj c\UArr\\ and r=\lim_Obj c\UObj\\ and c=c, + folded op_ua_components(2)[where lim_Obj=lim_Obj and c=c] + ] + +end + + + 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>\" +definition the_cf_lKe :: "V \ V \ V \ (V \ V) \ V" + where "the_cf_lKe \ \ \ lim_Obj = + op_cf (the_cf_rKe \ (op_cf \) (op_cf \) (op_ua lim_Obj \))" + +definition the_ntcf_lKe :: "V \ V \ V \ (V \ V) \ V" + where "the_ntcf_lKe \ \ \ lim_Obj = + op_ntcf (the_ntcf_rKe \ (op_cf \) (op_cf \) (op_ua lim_Obj \))" + 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\<^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_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 + by ( - cs_concl - cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_comma_cs_intros + cs_concl + cs_simp: cat_cs_simps + cs_intro: cat_lim_cs_intros 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 + cs_simp: cat_cs_simps + cs_intro: cat_lim_cs_intros 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 + cs_intro: cat_lim_cs_intros 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 + cs_intro: cat_lim_cs_intros 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 + cs_intro: cat_lim_cs_intros 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) ( cs_concl cs_intro: cat_cs_intros cat_comma_cs_intros cat_lim_cs_intros cs_simp: const_r_def \c'\ )+ 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\<^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_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\<^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\<^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_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_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\<^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_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_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 + cs_simp: cat_Kan_cs_simps cat_comma_cs_simps cat_cs_simps + cs_intro: + cat_lim_cs_intros + 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 (*very slow*) ( cs_concl cs_intro: cat_cs_intros - cs_simp: \.cf_cf_arr_comma_CId cat_cs_simps + cs_simp: \.cf_arr_cf_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_cf_lKe_is_functor: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" + shows "the_cf_lKe \ \ \ lim_Obj : \ \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + { + fix c assume prems: "c \\<^sub>\ \\Obj\" + from assms(3)[OF this] have lim_Obj_UArr: "lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \". + then interpret lim_Obj_c: is_cat_colimit + \ \\ \<^sub>C\<^sub>F\ c\ \ \\ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c\ \lim_Obj c\UObj\\ \lim_Obj c\UArr\\ + by simp + note op_ua_UArr_is_cat_limit'[ + where lim_Obj=lim_Obj, OF assms(1,2) prems lim_Obj_UArr + ] + } + note the_cf_rKe_is_functor = the_cf_rKe_is_functor + [ + OF \.is_functor_op \.is_functor_op, + unfolded cat_op_simps, + where lim_Obj=\op_ua lim_Obj \\, + unfolded cat_op_simps, + OF this, + simplified, + folded the_cf_lKe_def + ] + show "the_cf_lKe \ \ \ lim_Obj : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by + ( + rule is_functor.is_functor_op + [ + OF the_cf_rKe_is_functor, + folded the_cf_lKe_def, + unfolded cat_op_simps + ] + ) +qed + lemma the_ntcf_rKe_is_ntcf: 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_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_lKe_is_ntcf: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" + shows "the_ntcf_lKe \ \ \ lim_Obj : + \ \\<^sub>C\<^sub>F the_cf_lKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + { + fix c assume prems: "c \\<^sub>\ \\Obj\" + from assms(3)[OF this] have lim_Obj_UArr: "lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \". + then interpret lim_Obj_c: is_cat_colimit + \ \\ \<^sub>C\<^sub>F\ c\ \ \\ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c\ \lim_Obj c\UObj\\ \lim_Obj c\UArr\\ + by simp + note op_ua_UArr_is_cat_limit'[ + where lim_Obj=lim_Obj, OF assms(1,2) prems lim_Obj_UArr + ] + } + note the_ntcf_rKe_is_ntcf = the_ntcf_rKe_is_ntcf + [ + OF \.is_functor_op \.is_functor_op, + unfolded cat_op_simps, + where lim_Obj=\op_ua lim_Obj \\, + unfolded cat_op_simps, + OF this, + simplified + ] + show "the_ntcf_lKe \ \ \ lim_Obj : + \ \\<^sub>C\<^sub>F the_cf_lKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by + ( + rule is_ntcf.is_ntcf_op + [ + OF the_ntcf_rKe_is_ntcf, + unfolded cat_op_simps, + folded the_cf_lKe_def the_ntcf_lKe_def + ] + ) +qed + lemma the_ntcf_rKe_is_cat_rKe: 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_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)+ 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_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_cs_intros\ )+ 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_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_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>\\ + 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 + cat_lim_cs_intros + 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 + cat_lim_cs_intros + 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\ +lemma the_ntcf_lKe_is_cat_lKe: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\c. c \\<^sub>\ \\Obj\ \ lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" + shows "the_ntcf_lKe \ \ \ lim_Obj : + \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> the_cf_lKe \ \ \ lim_Obj \\<^sub>C\<^sub>F \ : \ \\<^sub>C \ \\<^sub>C \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + { + fix c assume prems: "c \\<^sub>\ \\Obj\" + from assms(3)[OF this] have lim_Obj_UArr: "lim_Obj c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m lim_Obj c\UObj\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \". + then interpret lim_Obj_c: is_cat_colimit + \ \\ \<^sub>C\<^sub>F\ c\ \ \\ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c\ \lim_Obj c\UObj\\ \lim_Obj c\UArr\\ + by simp + note op_ua_UArr_is_cat_limit'[ + where lim_Obj=lim_Obj, OF assms(1,2) prems lim_Obj_UArr + ] + } + note the_ntcf_rKe_is_cat_rKe = the_ntcf_rKe_is_cat_rKe + [ + OF \.is_functor_op \.is_functor_op, + unfolded cat_op_simps, + where lim_Obj=\op_ua lim_Obj \\, + unfolded cat_op_simps, + OF this, + simplified, + folded the_cf_lKe_def the_ntcf_lKe_def + ] + show ?thesis + by + ( + rule is_cat_rKe.is_cat_lKe_op + [ + OF the_ntcf_rKe_is_cat_rKe, + unfolded cat_op_simps, + folded the_cf_lKe_def the_ntcf_lKe_def + ] + ) +qed + + + +subsection\Preservation of Kan extensions\ 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\ +subsubsection\Limits and colimits\ 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\<^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_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_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) 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 \)" + moreover with prems'(1) 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\<^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)) 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 \ \]] \.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 \') (cs_concl cs_intro: cat_cs_intros)+ 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 \)" + with f' 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_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 \ \ \ - \' = \" + "\ \' : \ \\<^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,3706 +1,3785 @@ (* 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}. \ 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_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>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 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}. \ 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_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\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 \" 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 \" 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' subsubsection\Universal property\ 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_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_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_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> \" 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\<^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 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> \" 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 +subsubsection\Further properties\ + +lemma ntcf_cf_comp_is_cat_limit_if_is_iso_functor: + assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" + shows "u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" +proof(intro is_cat_limitI) + interpret u: is_cat_limit \ \ \ \ r u by (rule assms(1)) + interpret \: is_iso_functor \ \ \ \ by (rule assms(2)) + note [cf_cs_simps] = is_iso_functor_is_arr_isomorphism(2,3) + show u\: "u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>D\<^sub>G\<^sub>H\<^sub>M \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (intro is_cat_coneI) + ( + cs_concl + cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_lim_cs_intros + ) + fix u' r' assume prems: "u' : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ \\<^sub>C\<^sub>F \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + then interpret u': is_cat_cone \ r' \ \ \\ \\<^sub>C\<^sub>F \\ u' by simp + have "u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \ : r' <\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>n\<^sub>e \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + by (intro is_cat_coneI) + ( + cs_concl + cs_simp: cat_cs_simps cf_cs_simps + cs_intro: cat_cs_intros cat_lim_cs_intros cf_cs_intros + ) + from is_cat_limit.cat_lim_ua_fo[OF assms(1) this] obtain f + where f: "f : r' \\<^bsub>\\<^esub> r" + and u'_\: "u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \ = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" + and f'f: + "\ + f' : r' \\<^bsub>\\<^esub> r; + u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \ = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f' + \ \ f' = f" + for f' + by metis + from u'_\ have u'_inv\_\: + "(u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ = (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f) \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \" + by simp + show "\!f'. f' : r' \\<^bsub>\\<^esub> r \ u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^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) + from u'_inv\_\ f show "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f" + by + ( + cs_prems + cs_simp: + cf_cs_simps cat_cs_simps + ntcf_cf_comp_ntcf_cf_comp_assoc + ntcf_vcomp_ntcf_cf_comp[symmetric] + cs_intro: cat_cs_intros cf_cs_intros + ) + fix f' assume prems: + "f' : r' \\<^bsub>\\<^esub> r" "u' = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" + from prems(2) have + "u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \ = + (u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f') \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \" + by simp + from this f prems(1) have "u' \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F inv_cf \ = u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F ntcf_const \ \ f'" + by + ( + cs_prems + cs_simp: + cat_cs_simps cf_cs_simps + ntcf_vcomp_ntcf_cf_comp[symmetric] + ntcf_cf_comp_ntcf_cf_comp_assoc + cs_intro: cf_cs_intros cat_cs_intros + ) + then show "f' = f" by (intro f'f prems(1)) + qed +qed + +lemma ntcf_cf_comp_is_cat_limit_if_is_iso_functor'[cat_lim_cs_intros]: + assumes "u : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> \" + and "\' = \ \\<^sub>C\<^sub>F \" + shows "u \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F \ : r <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m \' : \ \\\<^sub>C\<^bsub>\\<^esub> \" + using assms(1,2) + unfolding assms(3) + by (rule ntcf_cf_comp_is_cat_limit_if_is_iso_functor) + + 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.\ 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.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.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) 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_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)) 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_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 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 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\<^bsub>\\<^esub> \" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_small_cs_intros cat_parallel_cs_intros cat_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\<^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) 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 ) 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 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 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,3452 +1,3623 @@ (* 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\<^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\<^sub>F \ \ \)" unfolding cf_Cone_def cat_cs_simps by simp end subsubsection\Object map\ 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\<^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_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\<^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_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>\" 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_ArrMap_app subsubsection\The cone functor is a functor\ 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 \ \ \\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from assms interpret op_\: 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_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\<^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_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\<^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 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\<^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_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_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_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_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_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_cs_intros cat_Kan_cs_intros a \ )+ qed lemma L_10_5_\_is_cat_cone'[cat_Kan_cs_intros]: 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> \" 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\<^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_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\: category \ \c \\<^sub>C\<^sub>F \\ by (cs_concl cs_intro: cat_comma_cs_intros) 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_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\<^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\<^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_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_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 \) \) (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(1,2,4) show ?thesis by ( 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_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\<^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 ?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_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\_\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros ) from \.vempty_is_zet assms interpret \c: 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\<^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_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_cs_intros ) 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_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 ( cs_concl 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] 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\<^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\<^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\<^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\<^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 ?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_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(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 (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_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_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_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: "?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_op_simps cs_intro: cat_Kan_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: "?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_op_simps cs_intro: cat_Kan_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] 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: 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_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 ) 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*) 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_op_simps cs_intro: 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_cs_simps cat_op_simps cs_intro: 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 \\\ +The existence of a canonical limit or a canonical colimit for the +pointwise Kan extensions \ 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).\ + \\Based on the elements of Chapter X-5 in \cite{mac_lane_categories_2010}.\ 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 ?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\<^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_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\_\ by ( cs_concl cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_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_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_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_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_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_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\ )+ 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\_\" by ( cs_concl cs_simp: cat_comma_cs_simps 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_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) cs_intro: cat_Kan_cs_intros cat_comma_cs_intros cat_cs_intros cat_FUNCT_cs_intros cat_op_intros ) 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]) have ua: "universal_arrow_fo ?\ (cf_map (?\_c\)) ?\c ?\_\c_CId" by ( 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 "ntcf_of_ntcf_arrow (c \\<^sub>C\<^sub>F \) \ ?\_\c_CId : ?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_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\ +lemma (in is_cat_pw_lKe) cat_pw_lKe_ex_cat_colimit: + \\Based on the elements of Chapter X-5 in \cite{mac_lane_categories_2010}.\ + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + and "c \\<^sub>\ \\Obj\" + obtains UA + where "UA : \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m \\ObjMap\\c\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + from + is_cat_pw_rKe.cat_pw_rKe_ex_cat_limit + [ + OF is_cat_pw_rKe_op AG.is_functor_op ntcf_lKe.NTDom.is_functor_op, + unfolded cat_op_simps, + OF assms(3) + ] + obtain UA where UA: "UA : + \\ObjMap\\c\ <\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>i\<^sub>m op_cf \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) : + c \\<^sub>C\<^sub>F (op_cf \) \\\<^sub>C\<^bsub>\\<^esub> op_cat \" + by auto + from assms(3) have [cat_cs_simps]: + "op_cf \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c = + op_cf \ \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>O c)" + by + ( + cs_concl + cs_simp: cat_cs_simps AG.op_cf_cf_obj_comma_proj[OF assms(3)] + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + from assms(3) have [cat_op_simps]: + "\ \\<^sub>C\<^sub>F op_cf (c \<^sub>O\\<^sub>C\<^sub>F (op_cf \)) \\<^sub>C\<^sub>F op_cf (op_cf_obj_comma \ c) = + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c" + by + ( + cs_concl + cs_simp: + cat_cs_simps cat_op_simps + op_cf_cf_comp[symmetric] AG.op_cf_cf_obj_comma_proj[symmetric] + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + from assms(3) have [cat_op_simps]: "op_cat (op_cat (\ \<^sub>C\<^sub>F\ c)) = \ \<^sub>C\<^sub>F\ c" + by + ( + cs_concl + cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + note ntcf_cf_comp_is_cat_limit_if_is_iso_functor = + ntcf_cf_comp_is_cat_limit_if_is_iso_functor + [ + OF UA AG.op_cf_obj_comma_is_iso_functor[OF assms(3)], + unfolded cat_op_simps + ] + have "op_ntcf UA \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F op_cf (op_cf_obj_comma \ c) : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m \\ObjMap\\c\ : \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" + by + ( + rule is_cat_limit.is_cat_colimit_op + [ + OF ntcf_cf_comp_is_cat_limit_if_is_iso_functor, + unfolded cat_op_simps + ] + ) + then show ?thesis using that by auto +qed + + + +subsection\The limit and the colimit for the pointwise Kan extensions\ 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>\" +definition the_pw_cat_lKe_colimit :: "V \ V \ V \ V \ V \ V" + where "the_pw_cat_lKe_colimit \ \ \ \ c = + [ + \\ObjMap\\c\, + op_ntcf + ( + the_pw_cat_rKe_limit \ (op_cf \) (op_cf \) (op_cf \) c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F + op_cf_obj_comma \ c + ) + ]\<^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) +lemma the_pw_cat_lKe_colimit_components: + shows "the_pw_cat_lKe_colimit \ \ \ \ c\UObj\ = \\ObjMap\\c\" + and "the_pw_cat_lKe_colimit \ \ \ \ c\UArr\ = op_ntcf + ( + the_pw_cat_rKe_limit \ (op_cf \) (op_cf \) (op_cf \) c\UArr\ \\<^sub>N\<^sub>T\<^sub>C\<^sub>F\<^sub>-\<^sub>C\<^sub>F + op_cf_obj_comma \ c + )" + unfolding the_pw_cat_lKe_colimit_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: +subsubsection\ +The limit for the pointwise right Kan extension is a limit, +the colimit for the pointwise left Kan extension is a colimit +\ + +lemma (in is_cat_pw_rKe) cat_pw_rKe_the_pw_cat_rKe_limit_is_cat_limit: 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_lKe) cat_pw_lKe_the_pw_cat_lKe_colimit_is_cat_colimit: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "c \\<^sub>\ \\Obj\" + shows "the_pw_cat_lKe_colimit \ \ \ \ c\UArr\ : + \ \\<^sub>C\<^sub>F \ \<^sub>C\<^sub>F\\<^sub>O c >\<^sub>C\<^sub>F\<^sub>.\<^sub>c\<^sub>o\<^sub>l\<^sub>i\<^sub>m the_pw_cat_lKe_colimit \ \ \ \ c\UObj\ : + \ \<^sub>C\<^sub>F\ c \\\<^sub>C\<^bsub>\\<^esub> \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(1)) + interpret \: is_functor \ \ \ \ by (rule assms(2)) + note cat_pw_rKe_the_pw_cat_rKe_limit_is_cat_limit = + is_cat_pw_rKe.cat_pw_rKe_the_pw_cat_rKe_limit_is_cat_limit + [ + OF is_cat_pw_rKe_op AG.is_functor_op ntcf_lKe.NTDom.is_functor_op, + unfolded cat_op_simps, + OF assms(3) + ] + from assms(3) have + "op_cf \ \\<^sub>C\<^sub>F c \<^sub>O\\<^sub>C\<^sub>F (op_cf \) \\<^sub>C\<^sub>F op_cf_obj_comma \ c = + op_cf \ \\<^sub>C\<^sub>F op_cf (\ \<^sub>C\<^sub>F\\<^sub>O c)" + by + ( + cs_concl + cs_simp: + cat_cs_simps cat_comma_cs_simps cat_op_simps + AG.op_cf_cf_obj_comma_proj[OF assms(3)] + cs_intro: cat_cs_intros cat_comma_cs_intros cat_op_intros + ) + note ntcf_cf_comp_is_cat_limit_if_is_iso_functor = + ntcf_cf_comp_is_cat_limit_if_is_iso_functor + [ + OF + cat_pw_rKe_the_pw_cat_rKe_limit_is_cat_limit + AG.op_cf_obj_comma_is_iso_functor[OF assms(3)], + unfolded this, folded op_cf_cf_comp + ] + from assms(3) have [cat_op_simps]: "op_cat (op_cat (\ \<^sub>C\<^sub>F\ c)) = \ \<^sub>C\<^sub>F\ c" + by + ( + cs_concl + cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + from assms(3) have [cat_op_simps]: "op_cf (op_cf (\ \<^sub>C\<^sub>F\\<^sub>O c)) = \ \<^sub>C\<^sub>F\\<^sub>O c" + by + ( + cs_concl + cs_simp: cat_op_simps cs_intro: cat_cs_intros cat_comma_cs_intros + ) + have [cat_op_simps]: + "the_pw_cat_rKe_limit \ (op_cf \) (op_cf \) (op_cf \) c\UObj\ = + the_pw_cat_lKe_colimit \ \ \ \ c\UObj\" + unfolding + the_pw_cat_lKe_colimit_components + the_pw_cat_rKe_limit_components + cat_op_simps + by simp + show ?thesis + by + ( + rule is_cat_limit.is_cat_colimit_op + [ + OF ntcf_cf_comp_is_cat_limit_if_is_iso_functor, + folded the_pw_cat_lKe_colimit_components, + unfolded cat_op_simps + ] + ) +qed + lemma (in is_cat_pw_rKe) cat_pw_rKe_the_ntcf_rKe_is_cat_rKe: 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_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] + cat_pw_rKe_the_pw_cat_rKe_limit_is_cat_limit[OF assms] ] ) qed +lemma (in is_cat_pw_lKe) cat_pw_lKe_the_ntcf_lKe_is_cat_lKe: + assumes "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" and "\ : \ \\\<^sub>C\<^bsub>\\<^esub> \" + shows "the_ntcf_lKe \ \ \ (the_pw_cat_lKe_colimit \ \ \ \) : + \ \\<^sub>C\<^sub>F\<^sub>.\<^sub>l\<^sub>K\<^sub>e\<^bsub>\\<^esub> the_cf_lKe \ \ \ (the_pw_cat_lKe_colimit \ \ \ \) \\<^sub>C\<^sub>F \ : + \ \\<^sub>C \ \\<^sub>C \" +proof- + interpret \: is_functor \ \ \ \ by (rule assms(2)) + show ?thesis + by + ( + rule the_ntcf_lKe_is_cat_lKe + [ + OF + assms(1,2) + cat_pw_lKe_the_pw_cat_lKe_colimit_is_cat_colimit[OF assms], + simplified + ] + ) +qed + text\\newpage\ end \ No newline at end of file