diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Rel.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Rel.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Rel.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Rel.thy @@ -1,676 +1,687 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\Rel\\ theory CZH_ECAT_Rel imports CZH_Foundations.CZH_SMC_Rel CZH_ECAT_Functor CZH_ECAT_Small_Category begin subsection\Background\ text\ The methodology chosen for the exposition of \Rel\ as a category is analogous to the one used in \cite{milehins_category_2021} for the exposition of \Rel\ as a semicategory. The general references for this section are Chapter I-7 in \cite{mac_lane_categories_2010} and nLab \cite{noauthor_nlab_nodate}\footnote{ \url{https://ncatlab.org/nlab/show/Rel} }. \ named_theorems cat_Rel_cs_simps named_theorems cat_Rel_cs_intros lemmas (in arr_Rel) [cat_Rel_cs_simps] = dg_Rel_shared_cs_simps lemmas [cat_Rel_cs_simps] = dg_Rel_shared_cs_simps arr_Rel.arr_Rel_length arr_Rel_comp_Rel_id_Rel_left arr_Rel_comp_Rel_id_Rel_right arr_Rel.arr_Rel_converse_Rel_converse_Rel arr_Rel_converse_Rel_eq_iff arr_Rel_converse_Rel_comp_Rel arr_Rel_comp_Rel_converse_Rel_left_if_v11 arr_Rel_comp_Rel_converse_Rel_right_if_v11 lemmas [cat_Rel_cs_intros] = dg_Rel_shared_cs_intros arr_Rel_comp_Rel arr_Rel.arr_Rel_converse_Rel subsection\\Rel\ as a category\ subsubsection\Definition and elementary properties\ definition cat_Rel :: "V \ V" where "cat_Rel \ = [ Vset \, set {T. arr_Rel \ T}, (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\), (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\), (\ST\\<^sub>\composable_arrs (dg_Rel \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\), VLambda (Vset \) id_Rel ]\<^sub>\" text\Components.\ lemma cat_Rel_components: shows "cat_Rel \\Obj\ = Vset \" and "cat_Rel \\Arr\ = set {T. arr_Rel \ T}" and "cat_Rel \\Dom\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\)" and "cat_Rel \\Cod\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\)" and "cat_Rel \\Comp\ = (\ST\\<^sub>\composable_arrs (dg_Rel \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\)" and "cat_Rel \\CId\ = VLambda (Vset \) id_Rel" unfolding cat_Rel_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_cat_Rel: "cat_smc (cat_Rel \) = smc_Rel \" proof(rule vsv_eqI) show "vsv (cat_smc (cat_Rel \))" unfolding cat_smc_def by auto show "vsv (smc_Rel \)" unfolding smc_Rel_def by auto have dom_lhs: "\\<^sub>\ (cat_smc (cat_Rel \)) = 5\<^sub>\" unfolding cat_smc_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smc_Rel \) = 5\<^sub>\" unfolding smc_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cat_smc (cat_Rel \)) = \\<^sub>\ (smc_Rel \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (cat_smc (cat_Rel \)) \ cat_smc (cat_Rel \)\a\ = smc_Rel \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cat_smc_def dg_field_simps cat_Rel_def smc_Rel_def ) (auto simp: nat_omega_simps) qed lemmas_with [folded cat_smc_cat_Rel, unfolded slicing_simps]: cat_Rel_Obj_iff = smc_Rel_Obj_iff and cat_Rel_Arr_iff[cat_Rel_cs_simps] = smc_Rel_Arr_iff and cat_Rel_Dom_vsv[cat_Rel_cs_intros] = smc_Rel_Dom_vsv and cat_Rel_Dom_vdomain[cat_Rel_cs_simps] = smc_Rel_Dom_vdomain and cat_Rel_Dom_app[cat_Rel_cs_simps] = smc_Rel_Dom_app and cat_Rel_Dom_vrange = smc_Rel_Dom_vrange and cat_Rel_Cod_vsv[cat_Rel_cs_intros] = smc_Rel_Cod_vsv and cat_Rel_Cod_vdomain[cat_Rel_cs_simps] = smc_Rel_Cod_vdomain and cat_Rel_Cod_app[cat_Rel_cs_simps] = smc_Rel_Cod_app and cat_Rel_Cod_vrange = smc_Rel_Cod_vrange and cat_Rel_is_arrI[cat_Rel_cs_intros] = smc_Rel_is_arrI and cat_Rel_is_arrD = smc_Rel_is_arrD and cat_Rel_is_arrE = smc_Rel_is_arrE lemmas_with [folded cat_smc_cat_Rel, unfolded slicing_simps, unfolded cat_smc_cat_Rel]: cat_Rel_composable_arrs_dg_Rel = smc_Rel_composable_arrs_dg_Rel and cat_Rel_Comp = smc_Rel_Comp and cat_Rel_Comp_app[cat_Rel_cs_simps] = smc_Rel_Comp_app and cat_Rel_Comp_vdomain[simp] = smc_Rel_Comp_vdomain lemmas [cat_cs_simps] = cat_Rel_is_arrD(2,3) lemmas [cat_Rel_cs_intros] = cat_Rel_is_arrI lemmas_with (in \) [folded cat_smc_cat_Rel, unfolded slicing_simps]: cat_Rel_Hom_vifunion_in_Vset = smc_Rel_Hom_vifunion_in_Vset and cat_Rel_incl_Rel_is_arr = smc_Rel_incl_Rel_is_arr and cat_Rel_incl_Rel_is_arr'[cat_Rel_cs_intros] = smc_Rel_incl_Rel_is_arr' + and cat_Rel_is_arr_ArrValE = smc_Rel_is_arr_ArrValE and cat_CAT_Comp_vrange = smc_CAT_Comp_vrange and cat_Rel_is_monic_arrI = smc_Rel_is_monic_arrI and cat_Rel_is_monic_arrD = smc_Rel_is_monic_arrD and cat_Rel_is_monic_arr = smc_Rel_is_monic_arr and cat_Rel_is_monic_arr_is_epic_arr = smc_Rel_is_monic_arr_is_epic_arr and cat_Rel_is_epic_arr_is_monic_arr = smc_Rel_is_epic_arr_is_monic_arr and cat_Rel_is_epic_arrI = smc_Rel_is_epic_arrI and cat_Rel_is_epic_arrD = smc_Rel_is_epic_arrD and cat_Rel_is_epic_arr = smc_Rel_is_epic_arr and cat_Rel_obj_terminal = smc_Rel_obj_terminal and cat_Rel_obj_initial = smc_Rel_obj_initial and cat_Rel_obj_terminal_obj_initial = smc_Rel_obj_terminal_obj_initial and cat_Rel_obj_null = smc_Rel_obj_null and cat_Rel_is_zero_arr = smc_Rel_is_zero_arr lemmas [cat_Rel_cs_intros] = \.cat_Rel_incl_Rel_is_arr' subsubsection\Identity\ lemma (in \) cat_Rel_CId_app[cat_Rel_cs_simps]: assumes "T \\<^sub>\ Vset \" shows "cat_Rel \\CId\\T\ = id_Rel T" using assms unfolding cat_Rel_components by simp lemmas [cat_Rel_cs_simps] = \.cat_Rel_CId_app subsubsection\\Rel\ is a category\ lemma (in \) category_cat_Rel: "category \ (cat_Rel \)" proof(rule categoryI, unfold cat_smc_cat_Rel) interpret Rel: semicategory \ \cat_smc (cat_Rel \)\ unfolding cat_smc_cat_Rel by (simp add: semicategory_smc_Rel) show "vfsequence (cat_Rel \)" unfolding cat_Rel_def by simp show "vcard (cat_Rel \) = 6\<^sub>\" unfolding cat_Rel_def by (simp add: nat_omega_simps) show "cat_Rel \\CId\\A\ : A \\<^bsub>cat_Rel \\<^esub> A" if "A \\<^sub>\ cat_Rel \\Obj\" for A using that unfolding cat_Rel_Obj_iff by ( cs_concl cs_simp: cat_Rel_cs_simps cs_intro: cat_Rel_cs_intros arr_Rel_id_RelI ) show "cat_Rel \\CId\\B\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> F = F" if "F : A \\<^bsub>cat_Rel \\<^esub> B" for F A B proof- from that have "arr_Rel \ F" "A \\<^sub>\ Vset \" "B \\<^sub>\ Vset \" by (auto elim: cat_Rel_is_arrE simp: cat_Rel_cs_simps) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_Rel_cs_intros arr_Rel_id_RelI ) qed show "F \\<^sub>A\<^bsub>cat_Rel \\<^esub> cat_Rel \\CId\\B\ = F" if "F : B \\<^bsub>cat_Rel \\<^esub> C" for F B C proof- from that have "arr_Rel \ F" "B \\<^sub>\ Vset \" "C \\<^sub>\ Vset \" by (auto elim: cat_Rel_is_arrE simp: cat_Rel_cs_simps) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_Rel_cs_intros arr_Rel_id_RelI ) qed qed (auto simp: semicategory_smc_Rel cat_Rel_components) lemma (in \) category_cat_Rel'[cat_Rel_cs_intros]: assumes "\' = \" and "\'' = \" shows "category \' (cat_Rel \'')" unfolding assms by (rule category_cat_Rel) lemmas [cat_Rel_cs_intros] = \.category_cat_Rel' subsection\Canonical dagger for \Rel\\ subsubsection\Definition and elementary properties\ definition cf_dag_Rel :: "V \ V" (\\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l\) where "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = [ vid_on (cat_Rel \\Obj\), VLambda (cat_Rel \\Arr\) converse_Rel, op_cat (cat_Rel \), cat_Rel \ ]\<^sub>\" text\Components.\ lemma cf_dag_Rel_components: shows "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\ = vid_on (cat_Rel \\Obj\)" and "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\ = VLambda (cat_Rel \\Arr\) converse_Rel" and "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomDom\ = op_cat (cat_Rel \)" and "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomCod\ = cat_Rel \" unfolding cf_dag_Rel_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cf_smcf_cf_dag_Rel: "cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) = 4\<^sub>\" unfolding cf_smcf_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = 4\<^sub>\" unfolding smcf_dag_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) = \\<^sub>\ (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding dom_lhs dom_rhs by simp show "A \\<^sub>\ \\<^sub>\ (cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) \ cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\A\ = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\A\" for A by ( unfold dom_lhs, elim_in_numeral, unfold dghm_field_simps[symmetric], unfold cat_smc_cat_Rel slicing_commute[symmetric] cf_smcf_components smcf_dag_Rel_components cf_dag_Rel_components smc_Rel_components cat_Rel_components ) simp_all qed (auto simp: cf_smcf_def smcf_dag_Rel_def) lemmas_with [folded cat_smc_cat_Rel cf_smcf_cf_dag_Rel, unfolded slicing_simps]: cf_dag_Rel_ObjMap_vsv[cat_Rel_cs_intros] = smcf_dag_Rel_ObjMap_vsv and cf_dag_Rel_ObjMap_vdomain[cat_Rel_cs_simps] = smcf_dag_Rel_ObjMap_vdomain and cf_dag_Rel_ObjMap_app[cat_Rel_cs_simps] = smcf_dag_Rel_ObjMap_app and cf_dag_Rel_ObjMap_vrange[cat_Rel_cs_simps] = smcf_dag_Rel_ObjMap_vrange and cf_dag_Rel_ArrMap_vsv[cat_Rel_cs_intros] = smcf_dag_Rel_ArrMap_vsv and cf_dag_Rel_ArrMap_vdomain[cat_Rel_cs_simps] = smcf_dag_Rel_ArrMap_vdomain and cf_dag_Rel_ArrMap_app[cat_Rel_cs_simps] = smcf_dag_Rel_ArrMap_app and cf_dag_Rel_ArrMap_vrange[cat_Rel_cs_simps] = smcf_dag_Rel_ArrMap_vrange lemmas_with (in \) [ folded cat_smc_cat_Rel cf_smcf_cf_dag_Rel, unfolded slicing_simps ]: cf_dag_Rel_app_is_arr[cat_Rel_cs_intros] = smcf_dag_Rel_app_is_arr + and cf_dag_Rel_ArrMap_app_vdomain[cat_cs_simps] = + smcf_dag_Rel_ArrMap_app_vdomain + and cf_dag_Rel_ArrMap_app_vrange[cat_cs_simps] = + smcf_dag_Rel_ArrMap_app_vrange + and cf_dag_Rel_ArrMap_app_iff[cat_cs_simps] = smcf_dag_Rel_ArrMap_app_iff and cf_dag_Rel_ArrMap_smc_Rel_Comp[cat_Rel_cs_simps] = smcf_dag_Rel_ArrMap_smc_Rel_Comp +lemmas [cat_cs_simps] = + \.cf_dag_Rel_ArrMap_app_vdomain + \.cf_dag_Rel_ArrMap_app_vrange + \.cf_dag_Rel_ArrMap_app_iff + subsubsection\Canonical dagger is a contravariant isomorphism of \Rel\\ lemma (in \) cf_dag_Rel_is_iso_functor: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_cat (cat_Rel \) \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> cat_Rel \" proof ( rule is_iso_functorI, unfold cat_smc_cat_Rel cf_smcf_cf_dag_Rel cat_Rel_components cat_op_simps slicing_commute[symmetric] ) interpret is_iso_semifunctor \ \op_smc (smc_Rel \)\ \smc_Rel \\ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (rule smcf_dag_Rel_is_iso_semifunctor) interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) show "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_cat (cat_Rel \) \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof ( rule is_functorI, unfold cat_smc_cat_Rel cf_smcf_cf_dag_Rel cat_op_simps slicing_commute[symmetric] cf_dag_Rel_components(3,4) ) show "vfsequence (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding cf_dag_Rel_def by (simp add: nat_omega_simps) show "vcard (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = 4\<^sub>\" unfolding cf_dag_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\cat_Rel \\CId\\C\\ = cat_Rel \\CId\\\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\C\\" if "C \\<^sub>\ cat_Rel \\Obj\" for C proof- from that have "C \\<^sub>\ Vset \" by (auto elim: cat_Rel_is_arrE simp: cat_Rel_Obj_iff) with that show ?thesis by ( cs_concl cs_simp: cat_Rel_cs_simps cs_intro: cat_cs_intros arr_Rel_id_RelI ) qed qed (auto simp: cat_cs_intros intro: smc_cs_intros) show "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_smc (smc_Rel \) \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_Rel \" by (rule smcf_dag_Rel_is_iso_semifunctor) qed lemma (in \) cf_dag_Rel_is_iso_functor'[cat_cs_intros]: assumes "\' = op_cat (cat_Rel \)" and "\' = cat_Rel \" and "\' = \" shows "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : \' \\\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> \'" unfolding assms by (rule cf_dag_Rel_is_iso_functor) lemmas [cat_cs_intros] = \.cf_dag_Rel_is_iso_functor' subsubsection\Further properties of the canonical dagger\ lemma (in \) cf_cn_comp_cf_dag_Rel_cf_dag_Rel: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>C\<^sub>F\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = cf_id (cat_Rel \)" proof(rule cf_smcf_eqI) interpret category \ \cat_Rel \\ by (rule category_cat_Rel) from cf_dag_Rel_is_iso_functor have dag: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_cat (cat_Rel \) \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" by (simp add: is_iso_functor.axioms(1)) from cf_cn_comp_is_functorI[OF category_axioms dag dag] show "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>C\<^sub>F\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" . show "cf_id (cat_Rel \) : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" by (auto simp: category.cat_cf_id_is_functor category_axioms) show "cf_smcf (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>C\<^sub>F\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = cf_smcf (smcf_id (cat_Rel \))" unfolding slicing_commute[symmetric] cat_smc_cat_Rel cf_smcf_cf_dag_Rel by (simp add: smcf_cn_comp_smcf_dag_Rel_smcf_dag_Rel) qed simp_all subsection\Isomorphism\ context \ begin context begin private lemma cat_Rel_is_arr_isomorphism_right_vsubset: assumes "S : B \\<^bsub>cat_Rel \\<^esub> A" and "T : A \\<^bsub>cat_Rel \\<^esub> B" and "S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A\" and "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S = cat_Rel \\CId\\B\" shows "S\ArrVal\ \\<^sub>\ (T\ArrVal\)\\<^sub>\" proof(rule vsubset_antisym vsubsetI) interpret Rel: category \ \cat_Rel \\ by (simp add: category_cat_Rel) interpret S: arr_Rel \ S rewrites "S\ArrDom\ = B" and "S\ArrCod\ = A" using assms(1) by (all\elim Rel.cat_is_arrE\) (simp_all add: cat_Rel_components) interpret T: arr_Rel \ T rewrites "T\ArrDom\ = A" and "T\ArrCod\ = B" using assms(2) by (all\elim Rel.cat_is_arrE\) (simp_all add: cat_Rel_components) interpret dag: is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (auto simp: cf_dag_Rel_is_iso_functor) from assms(2) have A: "A \\<^sub>\ cat_Rel \\Obj\" by auto from assms(3) have "(S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T)\ArrVal\ = cat_Rel \\CId\\A\\ArrVal\" by simp with A have [simp]: "S\ArrVal\ \\<^sub>\ T\ArrVal\ = vid_on A" unfolding cat_Rel_Comp_app[OF assms(1,2)] by (simp add: id_Rel_components comp_Rel_components cat_Rel_components) from assms(2) have B: "B \\<^sub>\ cat_Rel \\Obj\" by auto from assms(4) have "(T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S)\ArrVal\ = cat_Rel \\CId\\B\\ArrVal\" by simp with B have [simp]: "T\ArrVal\ \\<^sub>\ S\ArrVal\ = vid_on B" unfolding cat_Rel_Comp_app[OF assms(2,1)] by (simp add: id_Rel_components comp_Rel_components cat_Rel_components) fix ab assume ab: "ab \\<^sub>\ S\ArrVal\" with S.vbrelation obtain a b where ab_def: "ab = \a, b\" and "a \\<^sub>\ B" by (metis S.arr_Rel_ArrVal_vdomain S.ArrVal.vbrelation_vinE vsubsetE) then have "\a, a\ \\<^sub>\ T\ArrVal\ \\<^sub>\ S\ArrVal\" by auto then obtain c where "\a, c\ \\<^sub>\ S\ArrVal\" and ca[intro]: "\c, a\ \\<^sub>\ T\ArrVal\" by blast have "\b, a\ \\<^sub>\ T\ArrVal\" proof(rule ccontr) assume "\b, a\ \\<^sub>\ T\ArrVal\" with ca have "c \ b" by clarsimp moreover from ab have "\c, b\ \\<^sub>\ S\ArrVal\ \\<^sub>\ T\ArrVal\" unfolding ab_def by blast ultimately show False by (simp add: vid_on_iff) qed then show "ab \\<^sub>\ (T\ArrVal\)\\<^sub>\" unfolding ab_def by clarsimp qed private lemma cat_Rel_is_arr_isomorphism_left_vsubset: assumes "S : B \\<^bsub>cat_Rel \\<^esub> A" and "T : A \\<^bsub>cat_Rel \\<^esub> B" and "S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A\" and "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S = cat_Rel \\CId\\B\" shows "(T\ArrVal\)\\<^sub>\ \\<^sub>\ S\ArrVal\" using assms(2,3,4) cat_Rel_is_arr_isomorphism_right_vsubset[OF assms(2,1)] by auto private lemma is_arr_isomorphism_dag: assumes "S : B \\<^bsub>cat_Rel \\<^esub> A" and "T : A \\<^bsub>cat_Rel \\<^esub> B" and "S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A\" and "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S = cat_Rel \\CId\\B\" shows "S = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\" proof(rule arr_Rel_eqI[of \]) interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) interpret dag: is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (auto simp: cf_dag_Rel_is_iso_functor) from assms(1) show S: "arr_Rel \ S" by (fastforce simp: cat_Rel_components(2)) from cf_dag_Rel_app_is_arr[OF assms(2)] show "arr_Rel \ (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)" by (auto elim!: cat_Rel_is_arrE) from assms(2) have T: "arr_Rel \ T" by (auto simp: cat_Rel_is_arrD(1)) from S T assms show "S\ArrVal\ = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\" unfolding cf_dag_Rel_ArrMap_app[OF T] converse_Rel_components by (intro vsubset_antisym) ( simp_all add: cat_Rel_is_arr_isomorphism_left_vsubset cat_Rel_is_arr_isomorphism_right_vsubset ) from T assms show "S\ArrDom\ = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrDom\" unfolding cf_dag_Rel_components by (auto simp: cat_cs_simps cat_Rel_cs_simps converse_Rel_components(1)) from S assms show "S\ArrCod\ = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrCod\" by ( cs_concl cs_intro: cat_op_intros cat_cs_intros cs_simp: cat_Rel_cs_simps cat_cs_simps ) qed lemma cat_Rel_is_arr_isomorphismI[intro]: assumes "T : A \\<^bsub>cat_Rel \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" shows "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B" proof(rule is_arr_isomorphismI[where ?g = \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\]) interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) interpret v11: v11 \T\ArrVal\\ by (rule assms(2)) interpret T: arr_Rel \ T rewrites [simp]: "T\ArrDom\ = A" and [simp]: "T\ArrCod\ = B" using assms(1) by (all\elim cat_Rel_is_arrE\) (simp_all add: cat_Rel_components) interpret is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (simp add: cf_dag_Rel_is_iso_functor) show "T : A \\<^bsub>cat_Rel \\<^esub> B" by (rule assms(1)) show "is_inverse (cat_Rel \) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\) T" proof(intro is_inverseI) from assms(1) show dag_T: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^bsub>cat_Rel \\<^esub> A" by ( cs_concl cs_simp: cat_op_simps cat_Rel_cs_simps cs_intro: cat_cs_intros ) show T: "T : A \\<^bsub>cat_Rel \\<^esub> B" by (rule assms(1)) from T T.arr_Rel_axioms v11.v11_axioms assms(3) show "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A\" by ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_cs_intros cat_Rel_cs_intros ) from T T.arr_Rel_axioms v11.v11_axioms assms(4) show "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ = cat_Rel \\CId\\B\" by ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_cs_intros cat_Rel_cs_intros ) qed qed lemma cat_Rel_is_arr_isomorphismD[dest]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B" shows "T : A \\<^bsub>cat_Rel \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" proof- from assms show T: "T : A \\<^bsub>cat_Rel \\<^esub> B" by (simp add: is_arr_isomorphism_def) interpret T: arr_Rel \ T rewrites [simp]: "T\ArrDom\ = A" and [simp]: "T\ArrCod\ = B" using T by (all\elim cat_Rel_is_arrE\) (simp_all add: cat_Rel_components) interpret is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (simp add: cf_dag_Rel_is_iso_functor) from is_arr_isomorphismD[OF assms(1)] obtain S where "is_inverse (cat_Rel \) S T" by clarsimp from is_inverseD[OF this] obtain A' B' where "S : B' \\<^bsub>cat_Rel \\<^esub> A'" and "T : A' \\<^bsub>cat_Rel \\<^esub> B'" and "S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A'\" and "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S = cat_Rel \\CId\\B'\" by auto moreover with T have "A' = A" "B' = B" by auto ultimately have S: "S : B \\<^bsub>cat_Rel \\<^esub> A" and ST: "S \\<^sub>A\<^bsub>cat_Rel \\<^esub> T = cat_Rel \\CId\\A\" and TS: "T \\<^sub>A\<^bsub>cat_Rel \\<^esub> S = cat_Rel \\CId\\B\" by auto from S T ST TS have S_def: "S = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\" by (rule is_arr_isomorphism_dag) interpret S: arr_Rel \ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ rewrites "(\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ArrDom\ = B" and "(\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ArrCod\ = A" by (fold S_def, insert S, all\elim cat_Rel_is_arrE\) (simp_all add: cat_Rel_components) from T.arr_Rel_axioms S_def have S_T: "S\ArrVal\ = (T\ArrVal\)\\<^sub>\" by (simp add: cf_dag_Rel_ArrMap_app converse_Rel_components(1)) from S have A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" by auto from B TS A ST have "(T \\<^sub>R\<^sub>e\<^sub>l S)\ArrVal\ = id_Rel B\ArrVal\" "(S \\<^sub>R\<^sub>e\<^sub>l T)\ArrVal\ = id_Rel A\ArrVal\" unfolding cat_Rel_Comp_app[OF S T] cat_Rel_Comp_app[OF T S] unfolding cat_Rel_components by simp_all then have val_ST: "S\ArrVal\ \\<^sub>\ T\ArrVal\ = vid_on A" and val_TS: "T\ArrVal\ \\<^sub>\ S\ArrVal\ = vid_on B" unfolding comp_Rel_components id_Rel_components by simp_all show "v11 (T\ArrVal\)" proof(rule v11I) show "vsv (T\ArrVal\)" proof(rule vsvI) fix a b c assume prems: "\a, b\ \\<^sub>\ T\ArrVal\" "\a, c\ \\<^sub>\ T\ArrVal\" from prems(1) S_T have "\b, a\ \\<^sub>\ S\ArrVal\" by auto with prems(2) val_TS have "\b, c\ \\<^sub>\ vid_on B" by auto then show "b = c" by clarsimp qed (auto simp: T.ArrVal.vbrelation_axioms) show "vsv ((T\ArrVal\)\\<^sub>\)" proof(rule vsvI) fix a b c assume prems: "\a, b\ \\<^sub>\ (T\ArrVal\)\\<^sub>\" "\a, c\ \\<^sub>\ (T\ArrVal\)\\<^sub>\" with S_T have "\a, b\ \\<^sub>\ S\ArrVal\" and "\a, c\ \\<^sub>\ S\ArrVal\" by auto moreover from prems have "\b, a\ \\<^sub>\ T\ArrVal\" and "\c, a\ \\<^sub>\ T\ArrVal\" by auto ultimately have "\b, c\ \\<^sub>\ vid_on A" using val_ST by auto then show "b = c" by clarsimp qed auto qed show "\\<^sub>\ (T\ArrVal\) = A" proof(intro vsubset_antisym vsubsetI) fix a assume "a \\<^sub>\ A" with val_ST have "\a, a\ \\<^sub>\ S\ArrVal\ \\<^sub>\ T\ArrVal\" by auto then show "a \\<^sub>\ \\<^sub>\ (T\ArrVal\)" by auto qed (use T.arr_Rel_ArrVal_vdomain in auto) show "\\<^sub>\ (T\ArrVal\) = B" proof(intro vsubset_antisym vsubsetI) fix b assume "b \\<^sub>\ B" with val_TS have "\b, b\ \\<^sub>\ T\ArrVal\ \\<^sub>\ S\ArrVal\" by auto then show "b \\<^sub>\ \\<^sub>\ (T\ArrVal\)" by auto qed (use T.arr_Rel_ArrVal_vrange in auto) qed end end lemmas [cat_Rel_cs_simps] = \.cat_Rel_is_arr_isomorphismD(3,4) lemma (in \) cat_Rel_is_arr_isomorphism: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B \ T : A \\<^bsub>cat_Rel \\<^esub> B \ v11 (T\ArrVal\) \ \\<^sub>\ (T\ArrVal\) = A \ \\<^sub>\ (T\ArrVal\) = B" by auto subsection\The inverse arrow\ lemma (in \) cat_Rel_the_inverse: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B" shows "T\\<^sub>C\<^bsub>cat_Rel \\<^esub> = T\\<^sub>R\<^sub>e\<^sub>l" unfolding the_inverse_def proof(rule the_equality) interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) from assms have T: "T : A \\<^bsub>cat_Rel \\<^esub> B" by auto interpret T: arr_Rel \ T rewrites "T\ArrDom\ = A" and "T\ArrCod\ = B" using T by (all\elim cat_Rel_is_arrE\) (simp_all add: cat_Rel_components) from assms T T.arr_Rel_axioms cat_Rel_is_arr_isomorphismD(2)[OF assms] show inv_T_T: "is_inverse (cat_Rel \) (T\\<^sub>R\<^sub>e\<^sub>l) T" by (intro is_inverseI[where a=A and b=B]) ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_Rel_cs_intros cat_cs_intros )+ fix S assume prems: "is_inverse (cat_Rel \) S T" show "S = T\\<^sub>R\<^sub>e\<^sub>l" by (rule category.cat_is_inverse_eq[OF Rel.category_axioms prems inv_T_T]) qed lemmas [cat_Rel_cs_simps] = \.cat_Rel_the_inverse text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Set.thy @@ -1,2095 +1,2211 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\Set\\ theory CZH_ECAT_Set imports CZH_Foundations.CZH_SMC_Set CZH_ECAT_Par CZH_ECAT_Subcategory CZH_ECAT_PCategory begin subsection\Background\ text\ The methodology chosen for the exposition of \Set\ as a category is analogous to the one used in \cite{milehins_category_2021} for the exposition of \Set\ as a semicategory. \ named_theorems cat_Set_cs_simps named_theorems cat_Set_cs_intros lemmas (in arr_Set) [cat_Set_cs_simps] = dg_Rel_shared_cs_simps lemmas [cat_Set_cs_simps] = dg_Rel_shared_cs_simps arr_Set.arr_Set_ArrVal_vdomain arr_Set_comp_Set_id_Set_left arr_Set_comp_Set_id_Set_right lemmas [cat_Set_cs_intros] = dg_Rel_shared_cs_intros arr_Set_comp_Set (* Certain lemmas are applicable to any of the categories among Rel, Par, Set. If these lemmas are included in general-purpose collections like cat_cs_simps/cat_cs_intros, then backtracking can become slow. The following collections were created to resolve such issues. *) named_theorems cat_rel_par_Set_cs_intros named_theorems cat_rel_par_Set_cs_simps named_theorems cat_rel_Par_set_cs_intros named_theorems cat_rel_Par_set_cs_simps named_theorems cat_Rel_par_set_cs_intros named_theorems cat_Rel_par_set_cs_simps subsection\\Set\ as a category\ subsubsection\Definition and elementary properties\ definition cat_Set :: "V \ V" where "cat_Set \ = [ Vset \, set {T. arr_Set \ T}, (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrDom\), (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrCod\), (\ST\\<^sub>\composable_arrs (dg_Set \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\), VLambda (Vset \) id_Set ]\<^sub>\" text\Components.\ lemma cat_Set_components: shows "cat_Set \\Obj\ = Vset \" and "cat_Set \\Arr\ = set {T. arr_Set \ T}" and "cat_Set \\Dom\ = (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrDom\)" and "cat_Set \\Cod\ = (\T\\<^sub>\set {T. arr_Set \ T}. T\ArrCod\)" and "cat_Set \\Comp\ = (\ST\\<^sub>\composable_arrs (dg_Set \). ST\0\ \\<^sub>P\<^sub>a\<^sub>r ST\1\<^sub>\\)" and "cat_Set \\CId\ = VLambda (Vset \) id_Set" unfolding cat_Set_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma cat_smc_cat_Set: "cat_smc (cat_Set \) = smc_Set \" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ (cat_smc (cat_Set \)) = 5\<^sub>\" unfolding cat_smc_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (smc_Set \) = 5\<^sub>\" unfolding smc_Set_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cat_smc (cat_Set \)) = \\<^sub>\ (smc_Set \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (cat_smc (cat_Set \)) \ cat_smc (cat_Set \)\a\ = smc_Set \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold cat_smc_def dg_field_simps cat_Set_def smc_Set_def ) (auto simp: nat_omega_simps) qed (auto simp: cat_smc_def smc_Set_def) lemmas_with [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_Obj_iff = smc_Set_Obj_iff and cat_Set_Arr_iff[cat_Set_cs_simps] = smc_Set_Arr_iff and cat_Set_Dom_vsv[intro] = smc_Set_Dom_vsv and cat_Set_Dom_vdomain[simp] = smc_Set_Dom_vdomain and cat_Set_Dom_vrange = smc_Set_Dom_vrange and cat_Set_Dom_app = smc_Set_Dom_app and cat_Set_Cod_vsv[intro] = smc_Set_Cod_vsv and cat_Set_Cod_vdomain[simp] = smc_Set_Cod_vdomain and cat_Set_Cod_vrange = smc_Set_Cod_vrange and cat_Set_Cod_app[cat_Set_cs_simps] = smc_Set_Cod_app and cat_Set_is_arrI = smc_Set_is_arrI and cat_Set_is_arrD = smc_Set_is_arrD and cat_Set_is_arrE = smc_Set_is_arrE and cat_Set_ArrVal_vdomain[cat_cs_simps] = smc_Set_ArrVal_vdomain and cat_Set_ArrVal_app_vrange[cat_Set_cs_intros] = smc_Set_ArrVal_app_vrange lemmas [cat_cs_simps] = cat_Set_is_arrD(2,3) lemmas [cat_Set_cs_intros] = cat_Set_is_arrI lemmas_with [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_composable_arrs_dg_Set = smc_Set_composable_arrs_dg_Set and cat_Set_Comp = smc_Set_Comp and cat_Set_Comp_app[cat_Set_cs_simps] = smc_Set_Comp_app and cat_Set_Comp_vdomain[cat_Set_cs_simps] = smc_Set_Comp_vdomain lemmas_with (in \) [folded cat_smc_cat_Set, unfolded slicing_simps]: cat_Set_Hom_vifunion_in_Vset = smc_Set_Hom_vifunion_in_Vset and cat_Set_incl_Set_is_arr = smc_Set_incl_Set_is_arr and cat_Set_incl_Set_is_arr'[cat_Set_cs_intros] = smc_Set_incl_Set_is_arr' and cat_Set_Comp_ArrVal = smc_Set_Comp_ArrVal and cat_Set_Comp_vrange = smc_Set_Comp_vrange and cat_Set_is_monic_arrI = smc_Set_is_monic_arrI and cat_Set_is_monic_arr = smc_Set_is_monic_arr and cat_Set_is_epic_arrI = smc_Set_is_epic_arrI and cat_Set_is_epic_arrD = smc_Set_is_epic_arrD and cat_Set_is_epic_arr = smc_Set_is_epic_arr and cat_Set_obj_terminal = smc_Set_obj_terminal and cat_Set_obj_initial = smc_Set_obj_initial and cat_Set_obj_null = smc_Set_obj_null and cat_Set_is_zero_arr = smc_Set_is_zero_arr lemmas [cat_Set_cs_intros] = \.cat_Set_incl_Set_is_arr' lemmas [cat_cs_simps] = \.cat_Set_Comp_ArrVal subsubsection\Identity\ lemma cat_Set_CId_app[cat_Set_cs_simps]: assumes "A \\<^sub>\ Vset \" shows "cat_Set \\CId\\A\ = id_Set A" using assms unfolding cat_Set_components by simp lemma id_Par_CId_app_app[cat_cs_simps]: assumes "A \\<^sub>\ Vset \" and "a \\<^sub>\ A" shows "cat_Set \\CId\\A\\ArrVal\\a\ = a" unfolding cat_Set_CId_app[OF assms(1)] id_Rel_ArrVal_app[OF assms(2)] by simp subsubsection\\Set\ is a category\ lemma (in \) category_cat_Set: "category \ (cat_Set \)" proof(rule categoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) interpret Set: semicategory \ \cat_smc (cat_Set \)\ unfolding cat_smc_cat_Set by (simp add: semicategory_smc_Set) show "vfsequence (cat_Set \)" unfolding cat_Set_def by simp show "vcard (cat_Set \) = 6\<^sub>\" unfolding cat_Set_def by (simp add: nat_omega_simps) show "semicategory \ (smc_Set \)" by (simp add: semicategory_smc_Set) show "cat_Set \\CId\\A\ : A \\<^bsub>cat_Set \\<^esub> A" if "A \\<^sub>\ cat_Set \\Obj\" for A using that unfolding cat_Set_Obj_iff by ( cs_concl cs_simp: cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) show "cat_Set \\CId\\B\ \\<^sub>A\<^bsub>cat_Set \\<^esub> F = F" if "F : A \\<^bsub>cat_Set \\<^esub> B" for F A B proof- from that have "arr_Set \ F" "B \\<^sub>\ Vset \" by (auto elim: cat_Set_is_arrE) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) qed show "F \\<^sub>A\<^bsub>cat_Set \\<^esub> cat_Set \\CId\\B\ = F" if "F : B \\<^bsub>cat_Set \\<^esub> C" for F B C proof- from that have "arr_Set \ F" "B \\<^sub>\ Vset \" by (auto elim: cat_Set_is_arrE) with that show ?thesis by ( cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI ) qed qed (auto simp: cat_Set_components) lemma (in \) category_cat_Set': assumes "\ = \" shows "category \ (cat_Set \)" unfolding assms by (rule category_cat_Set) lemmas [cat_cs_intros] = \.category_cat_Set' subsubsection\\Set\ is a wide replete subcategory of \Par\\ lemma (in \) wide_replete_subcategory_cat_Set_cat_Par: "cat_Set \ \\<^sub>C\<^sub>.\<^sub>w\<^sub>r\<^bsub>\\<^esub> cat_Par \" proof(intro wide_replete_subcategoryI) show wide_subcategory_cat_Set_cat_Par: "cat_Set \ \\<^sub>C\<^sub>.\<^sub>w\<^sub>i\<^sub>d\<^sub>e\<^bsub>\\<^esub> cat_Par \" proof(intro wide_subcategoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) interpret Par: category \ \cat_Par \\ by (rule category_cat_Par) interpret Set: category \ \cat_Set \\ by (rule category_cat_Set) interpret wide_subsemicategory \ \smc_Set \\ \smc_Par \\ by (simp add: wide_subsemicategory_smc_Set_smc_Par) show "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Par \" proof(intro subcategoryI, unfold cat_smc_cat_Par cat_smc_cat_Set) show "smc_Set \ \\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Par \" by (simp add: subsemicategory_axioms) fix A assume "A \\<^sub>\ cat_Set \\Obj\" then show "cat_Set \\CId\\A\ = cat_Par \\CId\\A\" unfolding cat_Set_components cat_Par_components by simp qed ( auto simp: subsemicategory_axioms Par.category_axioms Set.category_axioms ) qed (rule wide_subsemicategory_smc_Set_smc_Par) show "cat_Set \ \\<^sub>C\<^sub>.\<^sub>r\<^sub>e\<^sub>p\<^bsub>\\<^esub> cat_Par \" proof(intro replete_subcategoryI) interpret wide_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_subcategory_cat_Set_cat_Par) show "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Par \" by (rule subcategory_axioms) fix A B F assume "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" note arr_Par = cat_Par_is_arr_isomorphismD[OF this] from arr_Par show "F : A \\<^bsub>cat_Set \\<^esub> B" by (intro cat_Set_is_arrI arr_Set_arr_ParI cat_Par_is_arrD[OF arr_Par(1)]) (auto simp: cat_Par_is_arrD(2)) qed qed subsubsection\\Set\ is a subcategory of \Set\\ lemma (in \) subcategory_cat_Set_cat_Set:(*TODO: generalize*) assumes "\ \" and "\ \\<^sub>\ \" shows "cat_Set \ \\<^sub>C\<^bsub>\\<^esub> cat_Set \" proof- interpret \: \ \ by (rule assms(1)) show ?thesis proof(intro subcategoryI') show "category \ (cat_Set \)" by (rule category.cat_category_if_ge_Limit, insert assms(2)) (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros)+ show "A \\<^sub>\ cat_Set \\Obj\" if "A \\<^sub>\ cat_Set \\Obj\" for A using that unfolding cat_Set_components(1) by (meson assms(2) Vset_in_mono \.Axiom_of_Extensionality(3)) show is_arr_if_is_arr: "F : A \\<^bsub>cat_Set \\<^esub> B" if "F : A \\<^bsub>cat_Set \\<^esub> B" for A B F proof- note f = cat_Set_is_arrD[OF that] interpret f: arr_Set \ F by (rule f(1)) show ?thesis proof(intro cat_Set_is_arrI arr_SetI) show "\\<^sub>\ (F\ArrVal\) \\<^sub>\ F\ArrCod\" by (auto simp: f.arr_Set_ArrVal_vrange) show "F\ArrDom\ \\<^sub>\ Vset \" by (auto intro!: f.arr_Set_ArrDom_in_Vset Vset_in_mono assms(2)) show "F\ArrCod\ \\<^sub>\ Vset \" by (auto intro!: f.arr_Set_ArrCod_in_Vset Vset_in_mono assms(2)) qed ( auto simp: f f.arr_Set_ArrVal_vdomain f.vfsequence_axioms f.arr_Set_length ) qed show "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F = G \\<^sub>A\<^bsub>cat_Set \\<^esub> F" if "G : B \\<^bsub>cat_Set \\<^esub> C" and "F : A \\<^bsub>cat_Set \\<^esub> B" for B C G A F proof- note g = cat_Set_is_arrD[OF that(1)] and f = cat_Set_is_arrD[OF that(2)] from that have \_gf_is_arr: "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F : A \\<^bsub>cat_Set \\<^esub> C" by (cs_concl cs_intro: cat_cs_intros is_arr_if_is_arr) from that have \_gf_is_arr: "G \\<^sub>A\<^bsub>cat_Set \\<^esub> F : A \\<^bsub>cat_Set \\<^esub> C" by (cs_concl cs_intro: cat_cs_intros is_arr_if_is_arr) note \_gf = cat_Set_is_arrD[OF \_gf_is_arr] and \_gf = cat_Set_is_arrD[OF \_gf_is_arr] show ?thesis proof(rule arr_Set_eqI) show "arr_Set \ (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)" by (rule \_gf(1)) then interpret arr_Set_\_gf: arr_Set \ \(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ by simp from \_gf_is_arr have dom_lhs: "\\<^sub>\ ((G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) show "arr_Set \ (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)" by (rule \_gf(1)) then interpret arr_Set_\_gf: arr_Set \ \(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ by simp from \_gf_is_arr have dom_rhs: "\\<^sub>\ ((G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) show "(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\ = (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ A" from that this show "(G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\ = (G \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros is_arr_if_is_arr ) qed auto qed (use \_gf_is_arr \_gf_is_arr in \cs_concl cs_simp: cat_cs_simps\)+ qed qed ( auto simp: assms(2) cat_Set_components Vset_trans Vset_in_mono cat_cs_intros ) qed subsection\Isomorphism\ lemma cat_Set_is_arr_isomorphismI[intro]: \\ See \cite{noauthor_nlab_nodate}\footnote{\url{ https://ncatlab.org/nlab/show/isomorphism }}). \ assumes "T : A \\<^bsub>cat_Set \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" shows "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" proof- interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF assms(1)]) note [cat_cs_intros] = cat_Par_is_arr_isomorphismI from wide_replete_subcategory_cat_Set_cat_Par assms have "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_cs_intros cat_sub_cs_intros cat_sub_fw_cs_intros) with wide_replete_subcategory_cat_Set_cat_Par assms show "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" by (cs_concl cs_simp: cat_sub_bw_cs_simps) qed lemma cat_Set_is_arr_isomorphismD[dest]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" shows "T : A \\<^bsub>cat_Set \\<^esub> B" and "v11 (T\ArrVal\)" and "\\<^sub>\ (T\ArrVal\) = A" and "\\<^sub>\ (T\ArrVal\) = B" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) from wide_replete_subcategory_cat_Set_cat_Par assms have T: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros) show "v11 (T\ArrVal\)" "\\<^sub>\ (T\ArrVal\) = A" "\\<^sub>\ (T\ArrVal\) = B" by (intro cat_Par_is_arr_isomorphismD[OF T])+ qed (rule is_arr_isomorphismD(1)[OF assms]) lemma cat_Set_is_arr_isomorphism: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B \ T : A \\<^bsub>cat_Set \\<^esub> B \ v11 (T\ArrVal\) \ \\<^sub>\ (T\ArrVal\) = A \ \\<^sub>\ (T\ArrVal\) = B" by auto subsection\The inverse arrow\ lemma cat_Set_ArrVal_app_is_arr[cat_cs_intros]: assumes "f : a \\<^bsub>\\<^esub> b" and "category \ \" (*the order of premises is important*) and "F : Hom \ a b \\<^bsub>cat_Set \\<^esub> Hom \ c d" shows "F\ArrVal\\f\ : c \\<^bsub>\\<^esub> d" proof- interpret \: category \ \ by (rule assms(2)) interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF assms(3)]) from assms have "F\ArrVal\\f\ \\<^sub>\ Hom \ c d" by (cs_concl cs_intro: cat_cs_intros cat_Set_cs_intros) then show ?thesis unfolding in_Hom_iff by simp qed abbreviation (input) converse_Set :: "V \ V" ("(_\\<^sub>S\<^sub>e\<^sub>t)" [1000] 999) where "a\\<^sub>S\<^sub>e\<^sub>t \ a\\<^sub>R\<^sub>e\<^sub>l" lemma cat_Set_the_inverse[cat_Set_cs_simps]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" shows "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>S\<^sub>e\<^sub>t" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) from wide_replete_subcategory_cat_Set_cat_Par assms have T: "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by (cs_concl cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros) from wide_replete_subcategory_cat_Set_cat_Par assms have [cat_cs_simps]: "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>C\<^bsub>cat_Par \\<^esub>" by ( cs_concl cs_full cs_simp: cat_sub_bw_cs_simps cs_intro: cat_sub_cs_intros ) from T show "T\\<^sub>C\<^bsub>cat_Set \\<^esub> = T\\<^sub>R\<^sub>e\<^sub>l" by (cs_concl cs_simp: cat_Par_cs_simps cat_cs_simps cs_intro: \_\) qed lemma cat_Set_the_inverse_app[cat_cs_intros]: assumes "T : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "a \\<^sub>\ A" and [cat_cs_simps]: "T\ArrVal\\a\ = b" shows "(T\\<^sub>C\<^bsub>cat_Set \\<^esub>)\ArrVal\\b\ = a" proof- from assms have T: "T : A \\<^bsub>cat_Set \\<^esub> B" by auto interpret arr_Set \ T by (rule cat_Set_is_arrD(1)[OF T]) note T = cat_Set_is_arr_isomorphismD[OF assms(1)] interpret T: v11 \T\ArrVal\\ by (rule T(2)) from T.v11_axioms assms(1,2) show "T\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\ = a" by ( cs_concl cs_simp: converse_Rel_components V_cs_simps cat_Set_cs_simps cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) qed lemma cat_Set_ArrVal_app_the_inverse_is_arr[cat_cs_intros]: assumes "f : c \\<^bsub>\\<^esub> d" and "category \ \" (*the order of premises is important*) and "F : Hom \ a b \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> Hom \ c d" shows "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\f\ : a \\<^bsub>\\<^esub> b" proof- interpret \: category \ \ by (rule assms(2)) from cat_Set_is_arr_isomorphismD[OF assms(3)] interpret F: arr_Set \ F by (simp add: cat_Set_is_arrD) from assms have "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\f\ \\<^sub>\ Hom \ a b" by (cs_concl cs_intro: cat_cs_intros cat_arrow_cs_intros) then show ?thesis unfolding in_Hom_iff by simp qed lemma cat_Set_app_the_inverse_app[cat_cs_simps]: assumes "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "b \\<^sub>\ B" shows "F\ArrVal\\F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\\ = b" proof- note F = cat_Set_is_arr_isomorphismD[OF assms(1)] note F = F cat_Set_is_arrD[OF F(1)] interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF F(1)]) from assms have [cat_cs_simps]: "F \\<^sub>A\<^bsub>cat_Set \\<^esub> F\\<^sub>C\<^bsub>cat_Set \\<^esub> = cat_Set \\CId\\B\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have [cat_cs_simps]: "F\ArrVal\\F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\b\\ = (F \\<^sub>A\<^bsub>cat_Set \\<^esub> F\\<^sub>C\<^bsub>cat_Set \\<^esub>)\ArrVal\\b\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) from assms F.arr_Par_ArrCod_in_Vset[unfolded F] show ?thesis by (cs_concl cs_simp: cat_cs_simps) qed lemma cat_Set_the_inverse_app_app[cat_cs_simps]: assumes "F : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "a \\<^sub>\ A" shows "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\F\ArrVal\\a\\ = a" proof- note F = cat_Set_is_arr_isomorphismD[OF assms(1)] note F = F cat_Set_is_arrD[OF F(1)] interpret F: arr_Set \ F by (rule cat_Set_is_arrD[OF F(1)]) from assms have [cat_cs_simps]: "F\\<^sub>C\<^bsub>cat_Set \\<^esub> \\<^sub>A\<^bsub>cat_Set \\<^esub> F = cat_Set \\CId\\A\" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms have [cat_cs_simps]: "F\\<^sub>C\<^bsub>cat_Set \\<^esub>\ArrVal\\F\ArrVal\\a\\ = (F\\<^sub>C\<^bsub>cat_Set \\<^esub> \\<^sub>A\<^bsub>cat_Set \\<^esub> F)\ArrVal\\a\" by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros ) from assms F.arr_Par_ArrDom_in_Vset[unfolded F] show ?thesis by (cs_concl cs_simp: cat_cs_simps) qed subsection\Projection arrows for \vtimes\\ subsubsection\Definition and elementary properties\ definition vfst_arrow :: "V \ V \ V" where "vfst_arrow A B = [(\ab\\<^sub>\A \\<^sub>\ B. vfst ab), A \\<^sub>\ B, A]\<^sub>\" definition vsnd_arrow :: "V \ V \ V" where "vsnd_arrow A B = [(\ab\\<^sub>\A \\<^sub>\ B. vsnd ab), A \\<^sub>\ B, B]\<^sub>\" text\Components.\ lemma vfst_arrow_components: shows "vfst_arrow A B\ArrVal\ = (\ab\\<^sub>\A \\<^sub>\ B. vfst ab)" and [cat_cs_simps]: "vfst_arrow A B\ArrDom\ = A \\<^sub>\ B" and [cat_cs_simps]: "vfst_arrow A B\ArrCod\ = A" unfolding vfst_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) lemma vsnd_arrow_components: shows "vsnd_arrow A B\ArrVal\ = (\ab\\<^sub>\A \\<^sub>\ B. vsnd ab)" and [cat_cs_simps]: "vsnd_arrow A B\ArrDom\ = A \\<^sub>\ B" and [cat_cs_simps]: "vsnd_arrow A B\ArrCod\ = B" unfolding vsnd_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow value\ mk_VLambda vfst_arrow_components(1) |vsv vfst_arrow_ArrVal_vsv[cat_cs_intros]| |vdomain vfst_arrow_ArrVal_vdomain[cat_cs_simps]| |app vfst_arrow_ArrVal_app'| mk_VLambda vsnd_arrow_components(1) |vsv vsnd_arrow_ArrVal_vsv[cat_cs_intros]| |vdomain vsnd_arrow_ArrVal_vdomain[cat_cs_simps]| |app vsnd_arrow_ArrVal_app'| lemma vfst_arrow_ArrVal_app[cat_cs_simps]: assumes "ab = \a, b\" and "ab \\<^sub>\ A \\<^sub>\ B" shows "vfst_arrow A B\ArrVal\\ab\ = a" using assms(2) unfolding assms(1) by (simp add: vfst_arrow_ArrVal_app') lemma vfst_arrow_vrange: "\\<^sub>\ (vfst_arrow A B\ArrVal\) \\<^sub>\ A" unfolding vfst_arrow_components proof(intro vrange_VLambda_vsubset) fix ab assume "ab \\<^sub>\ A \\<^sub>\ B" then obtain a b where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ A" by clarsimp from a show "vfst ab \\<^sub>\ A" unfolding ab_def by simp qed lemma vsnd_arrow_ArrVal_app[cat_cs_simps]: assumes "ab = \a, b\" and "ab \\<^sub>\ A \\<^sub>\ B" shows "vsnd_arrow A B\ArrVal\\ab\ = b" using assms(2) unfolding assms(1) by (simp add: vsnd_arrow_ArrVal_app') lemma vsnd_arrow_vrange: "\\<^sub>\ (vsnd_arrow A B\ArrVal\) \\<^sub>\ B" unfolding vsnd_arrow_components proof(intro vrange_VLambda_vsubset) fix ab assume "ab \\<^sub>\ A \\<^sub>\ B" then obtain a b where ab_def: "ab = \a, b\" and b: "b \\<^sub>\ B" by clarsimp from b show "vsnd ab \\<^sub>\ B" unfolding ab_def by simp qed subsubsection\Projection arrows are arrows in the category \Set\\ lemma (in \) vfst_arrow_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A" proof(intro cat_Set_is_arrI arr_SetI, unfold cat_cs_simps) show "vfsequence (vfst_arrow A B)" unfolding vfst_arrow_def by simp show "vcard (vfst_arrow A B) = 3\<^sub>\" unfolding vfst_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (vfst_arrow A B\ArrVal\) \\<^sub>\ A" by (rule vfst_arrow_vrange) qed (use assms in \cs_concl cs_intro: V_cs_intros cat_cs_intros\)+ lemma (in \) vfst_arrow_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A" using assms unfolding cat_Set_components by (rule vfst_arrow_is_cat_Set_arr_Vset) lemma (in \) vfst_arrow_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Set \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.vfst_arrow_is_cat_Set_arr' lemma (in \) vsnd_arrow_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> B" proof(intro cat_Set_is_arrI arr_SetI , unfold cat_cs_simps) show "vfsequence (vsnd_arrow A B)" unfolding vsnd_arrow_def by simp show "vcard (vsnd_arrow A B) = 3\<^sub>\" unfolding vsnd_arrow_def by (simp add: nat_omega_simps) show "\\<^sub>\ (vsnd_arrow A B\ArrVal\) \\<^sub>\ B" by (rule vsnd_arrow_vrange) qed (use assms in \cs_concl cs_intro: V_cs_intros cat_cs_intros\)+ lemma (in \) vsnd_arrow_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> B" using assms unfolding cat_Set_components by (rule vsnd_arrow_is_cat_Set_arr_Vset) lemma (in \) vsnd_arrow_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Set \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.vsnd_arrow_is_cat_Set_arr' subsubsection\Projection arrows are arrows in the category \Par\\ lemma (in \) vfst_arrow_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Par \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD vfst_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vfst_arrow_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Par \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.vfst_arrow_is_cat_Par_arr' lemma (in \) vsnd_arrow_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Par \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD vsnd_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vsnd_arrow_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Par \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.vsnd_arrow_is_cat_Par_arr' subsubsection\Projection arrows are arrows in the category \Rel\\ lemma (in \) vfst_arrow_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vfst_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD vfst_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vfst_arrow_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ B" and "A' = A" and "\' = cat_Rel \" shows "vfst_arrow A B : AB \\<^bsub>\'\<^esub> A'" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.vfst_arrow_is_cat_Rel_arr' lemma (in \) vsnd_arrow_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vsnd_arrow A B : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD vsnd_arrow_is_cat_Set_arr_Vset) auto qed lemma (in \) vsnd_arrow_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ B" and "B' = B" and "\' = cat_Rel \" shows "vsnd_arrow A B : AB \\<^bsub>\'\<^esub> B'" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.vsnd_arrow_is_cat_Rel_arr' subsubsection\Projection arrows are isomorphisms in the category \Set\\ lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism_Vset: assumes "A \\<^sub>\ Vset \" and "b \\<^sub>\ Vset \" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A" proof ( intro cat_Set_is_arr_isomorphismI arr_SetI vfst_arrow_is_cat_Set_arr_Vset assms, unfold cat_cs_simps ) show "v11 (vfst_arrow A (set {b})\ArrVal\)" proof(rule vsv.vsv_valeq_v11I, unfold cat_cs_simps) fix ab ab' assume prems: "ab \\<^sub>\ A \\<^sub>\ set {b}" "ab' \\<^sub>\ A \\<^sub>\ set {b}" "vfst_arrow A (set {b})\ArrVal\\ab\ = vfst_arrow A (set {b})\ArrVal\\ab'\" from prems obtain a where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ A" by clarsimp from prems obtain a' where ab'_def: "ab' = \a', b\" and a': "a' \\<^sub>\ A" by clarsimp from prems(3) a a' have "a = a'" unfolding ab_def ab'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: V_cs_intros) then show "ab = ab'" unfolding ab_def ab'_def by simp qed (cs_concl cs_intro: cat_cs_intros) show "\\<^sub>\ (vfst_arrow A (set {b})\ArrVal\) = A" proof(intro vsubset_antisym) show "A \\<^sub>\ \\<^sub>\ (vfst_arrow A (set {b})\ArrVal\)" proof(intro vsubsetI) fix a assume a: "a \\<^sub>\ A" then have a_def: "a = vfst_arrow A (set {b})\ArrVal\\\a, b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from a assms show "a \\<^sub>\ \\<^sub>\ (vfst_arrow A (set {b})\ArrVal\)" by (subst a_def, use nothing in \intro vsv.vsv_vimageI2\) (auto simp: cat_cs_simps cat_cs_intros) qed qed (rule vfst_arrow_vrange) qed (use assms in auto) lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism: assumes "A \\<^sub>\ cat_Set \\Obj\" and "b \\<^sub>\ cat_Set \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A" using assms unfolding cat_Set_components by (rule vfst_arrow_is_cat_Set_arr_isomorphism_Vset) lemma (in \) vfst_arrow_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "b \\<^sub>\ cat_Set \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Set \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.vfst_arrow_is_cat_Set_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism_Vset: assumes "a \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" proof ( intro cat_Set_is_arr_isomorphismI arr_SetI vsnd_arrow_is_cat_Set_arr_Vset assms, unfold cat_cs_simps ) show "v11 (vsnd_arrow (set {a}) B\ArrVal\)" proof(rule vsv.vsv_valeq_v11I, unfold cat_cs_simps) fix ab ab' assume prems: "ab \\<^sub>\ set {a} \\<^sub>\ B" "ab' \\<^sub>\ set {a} \\<^sub>\ B" "vsnd_arrow (set {a}) B\ArrVal\\ab\ = vsnd_arrow (set {a}) B\ArrVal\\ab'\" from prems obtain b where ab_def: "ab = \a, b\" and b: "b \\<^sub>\ B" by clarsimp from prems obtain b' where ab'_def: "ab' = \a, b'\" and b': "b' \\<^sub>\ B" by clarsimp from prems(3) b b' have "b = b'" unfolding ab_def ab'_def by (cs_prems cs_simp: cat_cs_simps cs_intro: V_cs_intros) then show "ab = ab'" unfolding ab_def ab'_def by simp qed (cs_concl cs_intro: cat_cs_intros) show "\\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\) = B" proof(intro vsubset_antisym) show "B \\<^sub>\ \\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\)" proof(intro vsubsetI) fix b assume b: "b \\<^sub>\ B" then have b_def: "b = vsnd_arrow (set {a}) B\ArrVal\\\a, b\\" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from b assms show "b \\<^sub>\ \\<^sub>\ (vsnd_arrow (set {a}) B\ArrVal\)" by (subst b_def, use nothing in \intro vsv.vsv_vimageI2\) (auto simp: cat_cs_simps cat_cs_intros) qed qed (rule vsnd_arrow_vrange) qed (use assms in auto) lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism: assumes "a \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" using assms unfolding cat_Set_components by (rule vsnd_arrow_is_cat_Set_arr_isomorphism_Vset) lemma (in \) vsnd_arrow_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "a \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Set \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.vsnd_arrow_is_cat_Set_arr_isomorphism' subsubsection\Projection arrows are isomorphisms in the category \Par\\ lemma (in \) vfst_arrow_is_cat_Par_arr_isomorphism: assumes "A \\<^sub>\ cat_Par \\Obj\" and "b \\<^sub>\ cat_Par \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) show "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vfst_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) qed lemma (in \) vfst_arrow_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "b \\<^sub>\ cat_Par \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Par \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.vfst_arrow_is_cat_Par_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Par_arr_isomorphism: assumes "a \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) show "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> B" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vsnd_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) qed lemma (in \) vsnd_arrow_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "a \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Par \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.vsnd_arrow_is_cat_Par_arr_isomorphism' subsubsection\Projection arrows are isomorphisms in the category \Rel\\ lemma (in \) vfst_arrow_is_cat_Rel_arr_isomorphism: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "b \\<^sub>\ cat_Rel \\Obj\" shows "vfst_arrow A (set {b}) : A \\<^sub>\ set {b} \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> A" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show ?thesis by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vfst_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) qed lemma (in \) vfst_arrow_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "b \\<^sub>\ cat_Rel \\Obj\" and "AB = A \\<^sub>\ set {b}" and "A' = A" and "\' = cat_Rel \" shows "vfst_arrow A (set {b}) : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> A" using assms(1-2) unfolding assms(3-5) by (rule vfst_arrow_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.vfst_arrow_is_cat_Rel_arr_isomorphism' lemma (in \) vsnd_arrow_is_cat_Rel_arr_isomorphism: assumes "a \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "vsnd_arrow (set {a}) B : set {a} \\<^sub>\ B \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> B" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show ?thesis by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF vsnd_arrow_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) qed lemma (in \) vsnd_arrow_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "a \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "AB = set {a} \\<^sub>\ B" and "A' = A" and "\' = cat_Rel \" shows "vsnd_arrow (set {a}) B : AB \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B" using assms(1-2) unfolding assms(3-5) by (rule vsnd_arrow_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.vsnd_arrow_is_cat_Rel_arr_isomorphism' subsection\Projection arrow for \vproduct\\ definition vprojection_arrow :: "V \ (V \ V) \ V \ V" where "vprojection_arrow I A i = [vprojection I A i, (\\<^sub>\i\\<^sub>\I. A i), A i]\<^sub>\" text\Components.\ lemma vprojection_arrow_components: shows "vprojection_arrow I A i\ArrVal\ = vprojection I A i" and "vprojection_arrow I A i\ArrDom\ = (\\<^sub>\i\\<^sub>\I. A i)" and "vprojection_arrow I A i\ArrCod\ = A i" unfolding vprojection_arrow_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Projection arrow value\ mk_VLambda vprojection_arrow_components(1)[unfolded vprojection_def] |vsv vprojection_arrow_vsv[cat_Set_cs_intros]| |vdomain vprojection_arrow_vdomain[cat_Set_cs_simps]| |app vprojection_arrow_app[cat_Set_cs_simps]| subsubsection\Projection arrow is an arrow in the category \Set\\ lemma (in \) arr_Set_vprojection_arrow: assumes "i \\<^sub>\ I" and "VLambda I A \\<^sub>\ Vset \" shows "arr_Set \ (vprojection_arrow I A i)" proof(intro arr_SetI) show "vfsequence (vprojection_arrow I A i)" unfolding vprojection_arrow_def by auto show "vcard (vprojection_arrow I A i) = 3\<^sub>\" unfolding vprojection_arrow_def by (simp add: nat_omega_simps) show "vprojection_arrow I A i\ArrCod\ \\<^sub>\ Vset \" unfolding vprojection_arrow_components proof- from assms(1) have "i \\<^sub>\ I" by simp then have "A i \\<^sub>\ \\<^sub>\ (VLambda I A)" by auto moreover from assms(2) have "\\<^sub>\ (VLambda I A) \\<^sub>\ Vset \" by (meson vrange_in_VsetI) ultimately show "A i \\<^sub>\ Vset \" by auto qed qed ( auto simp: vprojection_arrow_components intro!: assms vprojection_vrange_vsubset Limit_vproduct_in_Vset_if_VLambda_in_VsetI ) lemma (in \) vprojection_arrow_is_arr: assumes "i \\<^sub>\ I" and "VLambda I A \\<^sub>\ Vset \" shows "vprojection_arrow I A i : (\\<^sub>\i\\<^sub>\I. A i) \\<^bsub>cat_Set \\<^esub> A i" proof(intro cat_Set_is_arrI) from assms show "arr_Set \ (vprojection_arrow I A i)" by (rule arr_Set_vprojection_arrow) qed (simp_all add: vprojection_arrow_components) subsection\Product arrow value for \Rel\\ subsubsection\Definition and elementary properties\ definition prod_2_Rel_ArrVal :: "V \ V \ V" where "prod_2_Rel_ArrVal S T = set {\\a, b\, \c, d\\ | a b c d. \a, c\ \\<^sub>\ S \ \b, d\ \\<^sub>\ T}" lemma small_prod_2_Rel_ArrVal[simp]: "small {\\a, b\, \c, d\\ | a b c d. \a, c\ \\<^sub>\ S \ \b, d\ \\<^sub>\ T}" (is \small ?S\) proof(rule down) show "?S \ elts ((\\<^sub>\ S \\<^sub>\ \\<^sub>\ T) \\<^sub>\ (\\<^sub>\ S \\<^sub>\ \\<^sub>\ T))" by auto qed text\Rules.\ lemma prod_2_Rel_ArrValI: assumes "ab_cd = \\a, b\, \c, d\\" and "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" shows "ab_cd \\<^sub>\ prod_2_Rel_ArrVal S T" using assms unfolding prod_2_Rel_ArrVal_def by simp lemma prod_2_Rel_ArrValD[dest]: assumes "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" shows "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" using assms unfolding prod_2_Rel_ArrVal_def by auto lemma prod_2_Rel_ArrValE[elim]: assumes "ab_cd \\<^sub>\ prod_2_Rel_ArrVal S T" obtains a b c d where "ab_cd = \\a, b\, \c, d\\" and "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" using assms unfolding prod_2_Rel_ArrVal_def by auto text\Elementary properties\ lemma prod_2_Rel_ArrVal_vsubset_vprod: "prod_2_Rel_ArrVal S T \\<^sub>\ ((\\<^sub>\ S \\<^sub>\ \\<^sub>\ T) \\<^sub>\ (\\<^sub>\ S \\<^sub>\ \\<^sub>\ T))" by auto lemma prod_2_Rel_ArrVal_vbrelation: "vbrelation (prod_2_Rel_ArrVal S T)" using prod_2_Rel_ArrVal_vsubset_vprod by auto lemma prod_2_Rel_ArrVal_vdomain: "\\<^sub>\ (prod_2_Rel_ArrVal S T) = \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" proof(intro vsubset_antisym) show "\\<^sub>\ S \\<^sub>\ \\<^sub>\ T \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" proof(intro vsubsetI) fix ab assume "ab \\<^sub>\ \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" then obtain a b where ab_def: "ab = \a, b\" and "a \\<^sub>\ \\<^sub>\ S" and "b \\<^sub>\ \\<^sub>\ T" by auto then obtain c d where "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" by force then have "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" by (intro prod_2_Rel_ArrValI) auto then show "ab \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" unfolding ab_def by auto qed qed (use prod_2_Rel_ArrVal_vsubset_vprod in blast) lemma prod_2_Rel_ArrVal_vrange: "\\<^sub>\ (prod_2_Rel_ArrVal S T) = \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" proof(intro vsubset_antisym) show "\\<^sub>\ S \\<^sub>\ \\<^sub>\ T \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" proof(intro vsubsetI) fix cd assume "cd \\<^sub>\ \\<^sub>\ S \\<^sub>\ \\<^sub>\ T" then obtain c d where cd_def: "cd = \c, d\" and "c \\<^sub>\ \\<^sub>\ S" and "d \\<^sub>\ \\<^sub>\ T" by auto then obtain a b where "\a, c\ \\<^sub>\ S" and "\b, d\ \\<^sub>\ T" by force then have "\\a, b\, \c, d\\ \\<^sub>\ prod_2_Rel_ArrVal S T" by (intro prod_2_Rel_ArrValI) auto then show "cd \\<^sub>\ \\<^sub>\ (prod_2_Rel_ArrVal S T)" unfolding cd_def by auto qed qed (use prod_2_Rel_ArrVal_vsubset_vprod in blast) subsubsection\Further properties\ lemma assumes "vsv g" and "vsv f" shows prod_2_Rel_ArrVal_vsv: "vsv (prod_2_Rel_ArrVal g f)" and prod_2_Rel_ArrVal_app: "\a b. \ a \\<^sub>\ \\<^sub>\ g; b \\<^sub>\ \\<^sub>\ f \ \ prod_2_Rel_ArrVal g f\\a,b\\ = \g\a\, f\b\\" proof- interpret g: vsv g by (rule assms(1)) interpret f: vsv f by (rule assms(2)) show vsv_gf: "vsv (prod_2_Rel_ArrVal g f)" by (intro vsvI; (elim prod_2_Rel_ArrValE)?; (unfold prod_2_Rel_ArrVal_def)?) (auto simp: g.vsv f.vsv) fix a b assume "a \\<^sub>\ \\<^sub>\ g" "b \\<^sub>\ \\<^sub>\ f" then have a_ga: "\a, g\a\\ \\<^sub>\ g" and b_fb: "\b, f\b\\ \\<^sub>\ f" by auto from a_ga b_fb show "prod_2_Rel_ArrVal g f\\a, b\\ = \g\a\, f\b\\" by (cs_concl cs_simp: vsv.vsv_appI[OF vsv_gf] cs_intro: prod_2_Rel_ArrValI) qed lemma prod_2_Rel_ArrVal_v11: assumes "v11 g" and "v11 f" shows "v11 (prod_2_Rel_ArrVal g f)" proof- interpret g: v11 g by (rule assms(1)) interpret f: v11 f by (rule assms(2)) show ?thesis proof ( intro vsv.vsv_valeq_v11I prod_2_Rel_ArrVal_vsv g.vsv_axioms f.vsv_axioms, unfold prod_2_Rel_ArrVal_vdomain ) fix ab cd assume prems: "ab \\<^sub>\ \\<^sub>\ g \\<^sub>\ \\<^sub>\ f" "cd \\<^sub>\ \\<^sub>\ g \\<^sub>\ \\<^sub>\ f" "prod_2_Rel_ArrVal g f\ab\ = prod_2_Rel_ArrVal g f\cd\" from prems(1) obtain a b where ab_def: "ab = \a, b\" and a: "a \\<^sub>\ \\<^sub>\ g" and b: "b \\<^sub>\ \\<^sub>\ f" by auto from prems(2) obtain c d where cd_def: "cd = \c, d\" and c: "c \\<^sub>\ \\<^sub>\ g" and d: "d \\<^sub>\ \\<^sub>\ f" by auto from prems(3) a b c d have "\g\a\, f\b\\ = \g\c\, f\d\\" unfolding ab_def cd_def by (cs_prems cs_simp: prod_2_Rel_ArrVal_app cs_intro: V_cs_intros) then have "g\a\ = g\c\" and "f\b\ = f\d\" by simp_all then show "ab = cd" by (auto simp: ab_def cd_def a b c d f.v11_injective g.v11_injective) qed qed lemma prod_2_Rel_ArrVal_vcomp: "prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T = prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" proof- interpret ST': vbrelation \prod_2_Rel_ArrVal S' T'\ by (rule prod_2_Rel_ArrVal_vbrelation) interpret ST: vbrelation \prod_2_Rel_ArrVal S T\ by (rule prod_2_Rel_ArrVal_vbrelation) show ?thesis (*TODO: simplify proof*) proof(intro vsubset_antisym vsubsetI) fix aa'_cc' assume "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T" then obtain aa' bb' cc' where ac_def: "aa'_cc' = \aa', cc'\" and bc: "\bb', cc'\ \\<^sub>\ prod_2_Rel_ArrVal S' T'" and ab: "\aa', bb'\ \\<^sub>\ prod_2_Rel_ArrVal S T" by auto from bc obtain b b' c c' where bb'_cc'_def: "\bb', cc'\ = \\b, b'\, \c, c'\\" and bc: "\b, c\ \\<^sub>\ S'" and bc': "\b', c'\ \\<^sub>\ T'" by auto with ab obtain a a' where aa'_bb'_def: "\aa', bb'\ = \\a, a'\, \b, b'\\" and ab: "\a, b\ \\<^sub>\ S" and ab': "\a', b'\ \\<^sub>\ T" by auto from bb'_cc'_def have bb'_def: "bb' = \b, b'\" and cc'_def: "cc' = \c, c'\" by simp_all from aa'_bb'_def have aa'_def: "aa' = \a, a'\" and bb'_def: "bb' = \b, b'\" by simp_all from bc bc' ab ab' show "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" unfolding ac_def aa'_def cc'_def by (intro prod_2_Rel_ArrValI) (cs_concl cs_intro: prod_2_Rel_ArrValI vcompI)+ next fix aa'_cc' assume "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal (S' \\<^sub>\ S) (T' \\<^sub>\ T)" then obtain a a' c c' where aa'_cc'_def: "aa'_cc' = \\a, a'\, \c, c'\\" and ac: "\a, c\ \\<^sub>\ S' \\<^sub>\ S" and ac': "\a', c'\ \\<^sub>\ T' \\<^sub>\ T" by blast from ac obtain b where ab: "\a, b\ \\<^sub>\ S" and bc: "\b, c\ \\<^sub>\ S'" by auto from ac' obtain b' where ab': "\a', b'\ \\<^sub>\ T" and bc': "\b', c'\ \\<^sub>\ T'" by auto from ab bc ab' bc' show "aa'_cc' \\<^sub>\ prod_2_Rel_ArrVal S' T' \\<^sub>\ prod_2_Rel_ArrVal S T" unfolding aa'_cc'_def by (cs_concl cs_intro: vcompI prod_2_Rel_ArrValI) qed qed lemma prod_2_Rel_ArrVal_vid_on[cat_cs_simps]: "prod_2_Rel_ArrVal (vid_on A) (vid_on B) = vid_on (A \\<^sub>\ B)" unfolding prod_2_Rel_ArrVal_def by auto subsection\Product arrow for \Rel\\ subsubsection\Definition and elementary properties\ definition prod_2_Rel :: "V \ V \ V" where "prod_2_Rel S T = [ prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\), S\ArrDom\ \\<^sub>\ T\ArrDom\, S\ArrCod\ \\<^sub>\ T\ArrCod\ ]\<^sub>\" text\Components.\ lemma prod_2_Rel_components: shows "prod_2_Rel S T\ArrVal\ = prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)" and [cat_cs_simps]: "prod_2_Rel S T\ArrDom\ = S\ArrDom\ \\<^sub>\ T\ArrDom\" and [cat_cs_simps]: "prod_2_Rel S T\ArrCod\ = S\ArrCod\ \\<^sub>\ T\ArrCod\" unfolding prod_2_Rel_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Product arrow for \Rel\ is an arrow in \Rel\\ lemma prod_2_Rel_is_cat_Rel_arr: assumes "S : A \\<^bsub>cat_Rel \\<^esub> B" and "T : C \\<^bsub>cat_Rel \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^bsub>cat_Rel \\<^esub> B \\<^sub>\ D" proof- note S = cat_Rel_is_arrD[OF assms(1)] note T = cat_Rel_is_arrD[OF assms(2)] interpret S: arr_Rel \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Rel \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) show ?thesis proof(intro cat_Rel_is_arrI arr_RelI) show "vfsequence (prod_2_Rel S T)" unfolding prod_2_Rel_def by simp show "vcard (prod_2_Rel S T) = 3\<^sub>\" unfolding prod_2_Rel_def by (simp add: nat_omega_simps) from S have "\\<^sub>\ (S\ArrVal\) \\<^sub>\ A" and "\\<^sub>\ (S\ArrVal\) \\<^sub>\ B" by auto moreover from T have "\\<^sub>\ (T\ArrVal\) \\<^sub>\ C" and "\\<^sub>\ (T\ArrVal\) \\<^sub>\ D" by auto ultimately have "\\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\) \\<^sub>\ A \\<^sub>\ C" "\\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\) \\<^sub>\ B \\<^sub>\ D" by auto then show "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrDom\" "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrCod\" unfolding prod_2_Rel_components prod_2_Rel_ArrVal_vdomain prod_2_Rel_ArrVal_vrange by (force simp: prod_2_Rel_components)+ from S.arr_Rel_ArrDom_in_Vset T.arr_Rel_ArrDom_in_Vset S.arr_Rel_ArrCod_in_Vset T.arr_Rel_ArrCod_in_Vset show "prod_2_Rel S T\ArrDom\ \\<^sub>\ Vset \" "prod_2_Rel S T\ArrCod\ \\<^sub>\ Vset \" unfolding prod_2_Rel_components by (all\intro Limit_vtimes_in_VsetI\) auto qed (auto simp: prod_2_Rel_components intro: prod_2_Rel_ArrVal_vbrelation) qed lemma prod_2_Rel_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "S : A \\<^bsub>cat_Rel \\<^esub> B" and "T : C \\<^bsub>cat_Rel \\<^esub> D" and "A' = A \\<^sub>\ C" and "B' = B \\<^sub>\ D" and "\' = cat_Rel \" shows "prod_2_Rel S T : A' \\<^bsub>\'\<^esub> B'" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Rel_arr) subsubsection\Product arrow for \Rel\ is an arrow in \Set\\ lemma prod_2_Rel_app[cat_rel_par_Set_cs_simps]: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" and "a \\<^sub>\ A" and "c \\<^sub>\ C" and "ac = \a, c\" shows "prod_2_Rel S T\ArrVal\\ac\ = \S\ArrVal\\a\, T\ArrVal\\c\\" proof- note S = cat_Set_is_arrD[OF assms(1)] note T = cat_Set_is_arrD[OF assms(2)] interpret S: arr_Set \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Set \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) from assms(3,4) show ?thesis unfolding prod_2_Rel_components(1) assms(5) by ( cs_concl cs_simp: S.arr_Set_ArrVal_vdomain T.arr_Set_ArrVal_vdomain prod_2_Rel_ArrVal_app cs_intro: V_cs_intros ) qed lemma prod_2_Rel_is_cat_Set_arr: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> B \\<^sub>\ D" proof- note S = cat_Set_is_arrD[OF assms(1)] note T = cat_Set_is_arrD[OF assms(2)] interpret S: arr_Set \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = B" by (simp_all add: S) interpret T: arr_Set \ T rewrites [simp]: "T\ArrDom\ = C" and [simp]: "T\ArrCod\ = D" by (simp_all add: T) show ?thesis proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (prod_2_Rel S T)" unfolding prod_2_Rel_def by simp show "vcard (prod_2_Rel S T) = 3\<^sub>\" unfolding prod_2_Rel_def by (simp add: nat_omega_simps) from S.arr_Set_ArrVal_vrange T.arr_Set_ArrVal_vrange show "\\<^sub>\ (prod_2_Rel S T\ArrVal\) \\<^sub>\ prod_2_Rel S T\ArrCod\" unfolding prod_2_Rel_components prod_2_Rel_ArrVal_vdomain prod_2_Rel_ArrVal_vrange by auto from assms S.arr_Par_ArrDom_in_Vset T.arr_Par_ArrDom_in_Vset show "prod_2_Rel S T\ArrDom\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from assms S.arr_Par_ArrCod_in_Vset T.arr_Par_ArrCod_in_Vset show "prod_2_Rel S T\ArrCod\ \\<^sub>\ Vset \" by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) from assms show "prod_2_Rel S T\ArrDom\ = A \\<^sub>\ C" by (cs_concl cs_simp: cat_cs_simps) from assms show "prod_2_Rel S T\ArrCod\ = B \\<^sub>\ D" by (cs_concl cs_simp: cat_cs_simps) show "vsv (prod_2_Rel S T\ArrVal\)" unfolding prod_2_Rel_components by (intro prod_2_Rel_ArrVal_vsv S.ArrVal.vsv_axioms T.ArrVal.vsv_axioms) qed ( auto simp: cat_cs_simps cat_Set_cs_simps prod_2_Rel_ArrVal_vdomain prod_2_Rel_components(1) ) qed lemma prod_2_Rel_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "S : A \\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^bsub>cat_Set \\<^esub> D" and "AC = A \\<^sub>\ C" and "BD = B \\<^sub>\ D" and "\' = cat_Set \" shows "prod_2_Rel S T : AC \\<^bsub>\'\<^esub> BD" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Set_arr) subsubsection\Product arrow for \Rel\ is an isomorphism in \Set\\ lemma prod_2_Rel_is_cat_Set_arr_isomorphism: assumes "S : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> D" shows "prod_2_Rel S T : A \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B \\<^sub>\ D" proof- note S = cat_Set_is_arr_isomorphismD[OF assms(1)] note T = cat_Set_is_arr_isomorphismD[OF assms(2)] show ?thesis proof ( intro cat_Set_is_arr_isomorphismI prod_2_Rel_is_cat_Set_arr[OF S(1) T(1)], unfold prod_2_Rel_components ) show "\\<^sub>\ (prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)) = A \\<^sub>\ C" unfolding prod_2_Rel_ArrVal_vdomain by (cs_concl cs_simp: S(3) T(3) cs_intro: cat_cs_intros) show "\\<^sub>\ (prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)) = B \\<^sub>\ D" unfolding prod_2_Rel_ArrVal_vrange by (cs_concl cs_simp: S(4) T(4) cs_intro: cat_cs_intros) qed (use S(2) T(2) in \cs_concl cs_intro: prod_2_Rel_ArrVal_v11\) qed lemma prod_2_Rel_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "S : A \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> B" and "T : C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> D" and "AC = A \\<^sub>\ C" and "BD = B \\<^sub>\ D" and "\' = cat_Set \" shows "prod_2_Rel S T : AC \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> BD" using assms(1,2) unfolding assms(3-5) by (rule prod_2_Rel_is_cat_Set_arr_isomorphism) subsubsection\Further elementary properties\ lemma prod_2_Rel_Comp: assumes "G' : B' \\<^bsub>cat_Rel \\<^esub> B''" and "F' : A' \\<^bsub>cat_Rel \\<^esub> A''" and "G : B \\<^bsub>cat_Rel \\<^esub> B'" and "F : A \\<^bsub>cat_Rel \\<^esub> A'" shows "prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F = prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)" proof- from cat_Rel_is_arrD(1)[OF assms(1)] interpret \ \ by auto interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) note (*prefer cat_Rel*)[cat_cs_simps] = cat_Rel_is_arrD(2,3) from assms have GF'_GF: "prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F : B \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B'' \\<^sub>\ A''" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) from assms Rel.category_axioms have GG'_FF': "prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F) : B \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B'' \\<^sub>\ A''" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) show ?thesis proof(rule arr_Rel_eqI[of \]) from GF'_GF show arr_Rel_GF'_GF: "arr_Rel \ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)" by (auto dest: cat_Rel_is_arrD(1)) from GG'_FF' show arr_Rel_GG'_FF': "arr_Rel \ (prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F))" by (auto dest: cat_Rel_is_arrD(1)) show "(prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\ = prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" proof(intro vsubset_antisym vsubsetI) fix R assume "R \\<^sub>\ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\" from this assms have "R \\<^sub>\ prod_2_Rel_ArrVal (G'\ArrVal\) (F'\ArrVal\) \\<^sub>\ prod_2_Rel_ArrVal (G\ArrVal\) (F\ArrVal\)" by ( cs_prems cs_simp: prod_2_Rel_components(1) comp_Rel_components(1) cat_Rel_cs_simps cs_intro: cat_Rel_par_set_cs_intros ) from this[unfolded prod_2_Rel_ArrVal_vcomp] assms show "R \\<^sub>\ prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" by ( cs_concl cs_simp: prod_2_Rel_components comp_Rel_components(1) cat_Rel_cs_simps ) next fix R assume "R \\<^sub>\ prod_2_Rel (G' \\<^sub>A\<^bsub>cat_Rel \\<^esub> G) (F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> F)\ArrVal\" from this assms have "R \\<^sub>\ prod_2_Rel_ArrVal (G'\ArrVal\ \\<^sub>\ G\ArrVal\) (F'\ArrVal\ \\<^sub>\ F\ArrVal\)" by ( cs_prems cs_simp: comp_Rel_components prod_2_Rel_components cat_Rel_cs_simps ) from this[folded prod_2_Rel_ArrVal_vcomp] assms show "R \\<^sub>\ (prod_2_Rel G' F' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel G F)\ArrVal\" by ( cs_concl cs_simp: prod_2_Rel_components comp_Rel_components(1) cat_Rel_cs_simps cs_intro: cat_Rel_par_set_cs_intros ) qed qed ( use GF'_GF assms in (*slow*) \ cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_Rel_cs_intros \ )+ qed lemma (in \) prod_2_Rel_CId[cat_cs_simps]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" shows "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\) = cat_Rel \\CId\\A \\<^sub>\ B\" proof- interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) from assms have A_B: "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\) : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ B" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) from assms Rel.category_axioms have AB: "cat_Rel \\CId\\A \\<^sub>\ B\ : A \\<^sub>\ B \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ B" by ( cs_concl cs_simp: cat_Rel_components(1) cs_intro: V_cs_intros cat_cs_intros ) show ?thesis proof(rule arr_Rel_eqI) from A_B show arr_Rel_GF'_GF: "arr_Rel \ (prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\))" by (auto dest: cat_Rel_is_arrD(1)) from AB show arr_Rel_GG'_FF': "arr_Rel \ (cat_Rel \\CId\\A \\<^sub>\ B\)" by (auto dest: cat_Rel_is_arrD(1)) from assms show "prod_2_Rel (cat_Rel \\CId\\A\) (cat_Rel \\CId\\B\)\ArrVal\ = cat_Rel \\CId\\A \\<^sub>\ B\\ArrVal\" by ( cs_concl cs_simp: id_Rel_components prod_2_Rel_components cat_cs_simps cat_Rel_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed ( use A_B assms in \ cs_concl cs_simp: prod_2_Rel_components cat_Rel_cs_simps cs_intro: cat_cs_intros \ )+ qed +lemma (in \) cf_dag_Rel_ArrMap_app_prod_2_Rel: + assumes "S : A \\<^bsub>cat_Rel \\<^esub> B" and "T : C \\<^bsub>cat_Rel \\<^esub> D" + shows + "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\ = + prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)" +proof- + + interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) + interpret dag_Rel: is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ + by (rule cf_dag_Rel_is_iso_functor) + + note ST = prod_2_Rel_is_cat_Rel_arr[OF assms] + + from assms have dag_S: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\ : B \\<^bsub>cat_Rel \\<^esub> A" + and dag_T: "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : D \\<^bsub>cat_Rel \\<^esub> C" + by + ( + cs_concl + cs_simp: cat_Rel_cs_simps cat_op_simps cs_intro: cat_cs_intros + )+ + from assms have dag_prod: + "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\ : B \\<^sub>\ D \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ C" + by + ( + cs_concl + cs_simp: cat_Rel_cs_simps cat_op_simps + cs_intro: V_cs_intros cat_cs_intros cat_Rel_par_set_cs_intros + ) + from dag_S dag_T have prod_dag: + "prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\) : + B \\<^sub>\ D \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ C" + by (cs_concl cs_intro: cat_Rel_par_set_cs_intros) + + note [cat_cs_simps] = + prod_2_Rel_ArrVal_vdomain prod_2_Rel_ArrVal_vrange prod_2_Rel_components + from dag_prod ST have [cat_cs_simps]: + "\\<^sub>\ (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ArrVal\) = + \\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + "\\<^sub>\ (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ArrVal\) = + \\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + by (cs_concl cs_simp: cat_cs_simps)+ + + show + "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\ = + prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)" + proof(rule arr_Rel_eqI) + from dag_prod show arr_Rel_dag_prod: + "arr_Rel \ (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\)" + by (auto dest: cat_Rel_is_arrD) + then interpret dag_prod: arr_Rel \ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ by simp + from prod_dag show arr_Rel_prod_dag: + "arr_Rel \ (prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\))" + by (auto dest: cat_Rel_is_arrD) + then interpret prod_dag: + arr_Rel \ \prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ + by simp + from ST have arr_Rel_ST: "arr_Rel \ (prod_2_Rel S T)" + by (auto dest: cat_Rel_is_arrD) + show + "\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ArrVal\ = + prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ArrVal\" + proof(intro vsubset_antisym vsubsetI) + fix bd_ac assume prems: "bd_ac \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ArrVal\" + then obtain bd ac + where bd_ac_def: "bd_ac = \bd, ac\" + and bd: "bd \\<^sub>\ \\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + and ac: "ac \\<^sub>\ \\<^sub>\ (S\ArrVal\) \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + by (elim cat_Rel_is_arr_ArrValE[OF dag_prod prems, unfolded cat_cs_simps]) + have "\ac, bd\ \\<^sub>\ prod_2_Rel_ArrVal (S\ArrVal\) (T\ArrVal\)" + by + ( + rule prems[ + unfolded + bd_ac_def + cf_dag_Rel_ArrMap_app_iff[OF ST] + prod_2_Rel_components + ] + ) + then obtain a b c d + where ab: "\a, b\ \\<^sub>\ S\ArrVal\" + and cd: "\c, d\ \\<^sub>\ T\ArrVal\" + and bd_def: "bd = \b, d\" + and ac_def: "ac = \a, c\" + by auto + show "bd_ac \\<^sub>\ prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ArrVal\" + unfolding prod_2_Rel_components + proof(intro prod_2_Rel_ArrValI) + show "bd_ac = \\b, d\, \a, c\\" unfolding bd_ac_def bd_def ac_def by simp + from assms ab cd show + "\b, a\ \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\\ArrVal\" + "\d, c\ \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\" + by (cs_concl cs_simp: cat_cs_simps)+ + qed + next + fix bd_ac assume prems: + "bd_ac \\<^sub>\ prod_2_Rel (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\) (\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\)\ArrVal\" + then obtain a b c d + where bd_ac_def: "bd_ac = \\b, d\, a, c\" + and ba: "\b, a\ \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\\ArrVal\" + and dc: "\d, c\ \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\" + by (elim prod_2_Rel_ArrValE[OF prems[unfolded prod_2_Rel_components]]) + then have ab: "\a, b\ \\<^sub>\ S\ArrVal\" and cd: "\c, d\ \\<^sub>\ T\ArrVal\" + unfolding assms[THEN cf_dag_Rel_ArrMap_app_iff] by simp_all + from ST ab cd show "bd_ac \\<^sub>\ \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\prod_2_Rel S T\\ArrVal\" + unfolding bd_ac_def + by + ( + cs_concl + cs_simp: prod_2_Rel_components cat_cs_simps + cs_intro: prod_2_Rel_ArrValI cat_cs_intros + ) + qed + qed (use dag_prod prod_dag in \cs_concl cs_simp: cat_cs_simps\)+ + +qed + subsection\Product functor for \Rel\\ definition cf_prod_2_Rel :: "V \ V" where "cf_prod_2_Rel \ = [ (\AB\\<^sub>\(\ \\<^sub>C \)\Obj\. AB\0\ \\<^sub>\ AB\1\<^sub>\\), (\ST\\<^sub>\(\ \\<^sub>C \)\Arr\. prod_2_Rel (ST\0\) (ST\1\<^sub>\\)), \ \\<^sub>C \, \ ]\<^sub>\" text\Components.\ lemma cf_prod_2_Rel_components: shows "cf_prod_2_Rel \\ObjMap\ = (\AB\\<^sub>\(\ \\<^sub>C \)\Obj\. AB\0\ \\<^sub>\ AB\1\<^sub>\\)" and "cf_prod_2_Rel \\ArrMap\ = (\ST\\<^sub>\(\ \\<^sub>C \)\Arr\. prod_2_Rel (ST\0\) (ST\1\<^sub>\\))" and [cat_cs_simps]: "cf_prod_2_Rel \\HomDom\ = \ \\<^sub>C \" and [cat_cs_simps]: "cf_prod_2_Rel \\HomCod\ = \" unfolding cf_prod_2_Rel_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda cf_prod_2_Rel_components(1) |vsv cf_prod_2_Rel_ObjMap_vsv[cat_cs_intros]| |vdomain cf_prod_2_Rel_ObjMap_vdomain[cat_cs_simps]| lemma cf_prod_2_Rel_ObjMap_app[cat_cs_simps]: assumes "AB = [A, B]\<^sub>\" and "AB \\<^sub>\ (\ \\<^sub>C \)\Obj\" shows "A \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>cf_prod_2_Rel \\<^esub> B = A \\<^sub>\ B" using assms(2) unfolding assms(1) cf_prod_2_Rel_components by (simp add: nat_omega_simps) lemma (in \) cf_prod_2_Rel_ObjMap_vrange: "\\<^sub>\ (cf_prod_2_Rel (cat_Rel \)\ObjMap\) \\<^sub>\ cat_Rel \\Obj\" proof- interpret Rel: category \ \cat_Rel \\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(rule vsv.vsv_vrange_vsubset, unfold cat_cs_simps) fix AB assume prems: "AB \\<^sub>\ (cat_Rel \ \\<^sub>C cat_Rel \)\Obj\" with Rel.category_axioms obtain A B where AB_def: "AB = [A, B]\<^sub>\" and A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" by (elim cat_prod_2_ObjE[rotated 2]) from prems A B show "cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\ \\<^sub>\ cat_Rel \\Obj\" unfolding AB_def cat_Rel_components(1) by (cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: V_cs_intros) qed (cs_concl cs_intro: cat_cs_intros) qed subsubsection\Arrow map\ mk_VLambda cf_prod_2_Rel_components(2) |vsv cf_prod_2_Rel_ArrMap_vsv[cat_cs_intros]| |vdomain cf_prod_2_Rel_ArrMap_vdomain[cat_cs_simps]| lemma cf_prod_2_Rel_ArrMap_app[cat_cs_simps]: assumes "GF = [G, F]\<^sub>\" and "GF \\<^sub>\ (\ \\<^sub>C \)\Arr\" shows "G \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>cf_prod_2_Rel \\<^esub> F = prod_2_Rel G F" using assms(2) unfolding assms(1) cf_prod_2_Rel_components by (simp add: nat_omega_simps) subsubsection\Product functor for \Rel\ is a functor\ lemma (in \) cf_prod_2_Rel_is_functor: "cf_prod_2_Rel (cat_Rel \) : cat_Rel \ \\<^sub>C cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof- interpret Rel: category \ \cat_Rel \\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(rule is_functorI') show "vfsequence (cf_prod_2_Rel (cat_Rel \))" unfolding cf_prod_2_Rel_def by auto show "vcard (cf_prod_2_Rel (cat_Rel \)) = 4\<^sub>\" unfolding cf_prod_2_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (cf_prod_2_Rel (cat_Rel \)\ObjMap\) \\<^sub>\ cat_Rel \\Obj\" by (rule cf_prod_2_Rel_ObjMap_vrange) show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF\ : cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\ \\<^bsub>cat_Rel \\<^esub> cf_prod_2_Rel (cat_Rel \)\ObjMap\\CD\" if "GF : AB \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> CD" for AB CD GF proof- from that obtain G F A B C D where GF_def: "GF = [G, F]\<^sub>\" and AB_def: "AB = [A, B]\<^sub>\" and CD_def: "CD = [C, D]\<^sub>\" and G: "G : A \\<^bsub>cat_Rel \\<^esub> C" and F: "F : B \\<^bsub>cat_Rel \\<^esub> D" by (elim cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms]) from that G F show ?thesis unfolding GF_def AB_def CD_def by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros cat_prod_cs_intros ) qed show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF' \\<^sub>A\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> GF\ = cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF'\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> cf_prod_2_Rel (cat_Rel \)\ArrMap\\GF\" if "GF' : AB' \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> AB''" and "GF : AB \\<^bsub>cat_Rel \ \\<^sub>C cat_Rel \\<^esub> AB'" for AB' AB'' GF' AB GF proof- from that(2) obtain G F A A' B B' where GF_def: "GF = [G, F]\<^sub>\" and AB_def: "AB = [A, B]\<^sub>\" and AB'_def: "AB' = [A', B']\<^sub>\" and G: "G : A \\<^bsub>cat_Rel \\<^esub> A'" and F: "F : B \\<^bsub>cat_Rel \\<^esub> B'" by (elim cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms]) with that(1) obtain G' F' A'' B'' where GF'_def: "GF' = [G', F']\<^sub>\" and AB''_def: "AB'' = [A'', B'']\<^sub>\" and G': "G' : A' \\<^bsub>cat_Rel \\<^esub> A''" and F': "F' : B' \\<^bsub>cat_Rel \\<^esub> B''" by ( auto elim: cat_prod_2_is_arrE[OF Rel.category_axioms Rel.category_axioms] ) from that G F G' F' show ?thesis unfolding GF_def AB_def AB'_def GF'_def AB''_def by ( cs_concl cs_simp: cat_cs_simps cat_prod_cs_simps prod_2_Rel_Comp cs_intro: cat_cs_intros cat_prod_cs_intros ) qed show "cf_prod_2_Rel (cat_Rel \)\ArrMap\\(cat_Rel \ \\<^sub>C cat_Rel \)\CId\\AB\\ = cat_Rel \\CId\\cf_prod_2_Rel (cat_Rel \)\ObjMap\\AB\\" if "AB \\<^sub>\ (cat_Rel \ \\<^sub>C cat_Rel \)\Obj\" for AB proof- from that obtain A B where AB_def: "AB = [A, B]\<^sub>\" and A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" by (elim cat_prod_2_ObjE[OF Rel.category_axioms Rel.category_axioms]) from A B show ?thesis unfolding AB_def by ( cs_concl cs_simp: cf_prod_2_Rel_ObjMap_app cf_prod_2_Rel_ArrMap_app cat_cs_simps cat_prod_cs_simps cs_intro: V_cs_intros cat_cs_intros cat_Rel_cs_intros cat_prod_cs_intros ) qed qed ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_cs_intros cat_Rel_cs_intros )+ qed lemma (in \) cf_prod_2_Rel_is_functor'[cat_cs_intros]: assumes "\' = cat_Rel \ \\<^sub>C cat_Rel \" and "\' = cat_Rel \" and "\' = \" shows "cf_prod_2_Rel (cat_Rel \) : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule cf_prod_2_Rel_is_functor) lemmas [cat_cs_intros] = \.cf_prod_2_Rel_is_functor' subsection\Product universal property arrow for \Set\\ subsubsection\Definition and elementary properties\ definition cat_Set_obj_prod_up :: "V \ (V \ V) \ V \ (V \ V) \ V" where "cat_Set_obj_prod_up I F A \ = [(\a\\<^sub>\A. (\i\\<^sub>\I. \ i\ArrVal\\a\)), A, (\\<^sub>\i\\<^sub>\I. F i)]\<^sub>\" text\Components.\ lemma cat_Set_obj_prod_up_components: shows "cat_Set_obj_prod_up I F A \\ArrVal\ = (\a\\<^sub>\A. (\i\\<^sub>\I. \ i\ArrVal\\a\))" and [cat_Set_cs_simps]: "cat_Set_obj_prod_up I F A \\ArrDom\ = A" and [cat_Set_cs_simps]: "cat_Set_obj_prod_up I F A \\ArrCod\ = (\\<^sub>\i\\<^sub>\I. F i)" unfolding cat_Set_obj_prod_up_def arr_field_simps by (simp_all add: nat_omega_simps) text\Arrow value.\ mk_VLambda cat_Set_obj_prod_up_components(1) |vsv cat_Set_obj_prod_up_ArrVal_vsv[cat_Set_cs_intros]| |vdomain cat_Set_obj_prod_up_ArrVal_vdomain[cat_Set_cs_simps]| |app cat_Set_obj_prod_up_ArrVal_app| lemma cat_Set_obj_prod_up_ArrVal_vrange: assumes "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" unfolding cat_Set_obj_prod_up_components proof(intro vrange_VLambda_vsubset vproductI) fix a assume prems: "a \\<^sub>\ A" show "\i\\<^sub>\I. (\i\\<^sub>\I. \ i\ArrVal\\a\)\i\ \\<^sub>\ F i" proof(intro ballI) fix i assume "i \\<^sub>\ I" with assms prems show "(\i\\<^sub>\I. \ i\ArrVal\\a\)\i\ \\<^sub>\ F i" by (cs_concl cs_simp: V_cs_simps cs_intro: cat_Set_cs_intros) qed qed auto lemma cat_Set_obj_prod_up_ArrVal_app_vdomain[cat_Set_cs_simps]: assumes "a \\<^sub>\ A" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\) = I" unfolding cat_Set_obj_prod_up_ArrVal_app[OF assms] by simp lemma cat_Set_obj_prod_up_ArrVal_app_component[cat_Set_cs_simps]: assumes "a \\<^sub>\ A" and "i \\<^sub>\ I" shows "cat_Set_obj_prod_up I F A \\ArrVal\\a\\i\ = \ i\ArrVal\\a\" using assms by (cs_concl cs_simp: cat_Set_obj_prod_up_ArrVal_app V_cs_simps) lemma cat_Set_obj_prod_up_ArrVal_app_vrange: assumes "a \\<^sub>\ A" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\) \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" proof(intro vsubsetI) fix b assume prems: "b \\<^sub>\ \\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\\a\)" from assms(1) have "vsv (cat_Set_obj_prod_up I F A \\ArrVal\\a\)" by (auto simp: cat_Set_obj_prod_up_components) with prems obtain i where b_def: "b = cat_Set_obj_prod_up I F A \\ArrVal\\a\\i\" and i: "i \\<^sub>\ I" by ( auto elim: vsv.vrange_atE simp: cat_Set_obj_prod_up_ArrVal_app[OF assms(1)] ) from cat_Set_obj_prod_up_ArrVal_app_component[OF assms(1) i] b_def have b_def': "b = \ i\ArrVal\\a\" by simp from assms(1) assms(2)[OF i] have "b \\<^sub>\ F i" unfolding b_def' by (cs_concl cs_intro: cat_Set_cs_intros) with i show "b \\<^sub>\ (\\<^sub>\i\\<^sub>\I. F i)" by force qed subsubsection\Product universal property arrow for \Set\ is an arrow in \Set\\ lemma (in \) cat_Set_obj_prod_up_cat_Set_is_arr: assumes "A \\<^sub>\ Vset \" and "VLambda I F \\<^sub>\ Vset \" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" shows "cat_Set_obj_prod_up I F A \ : A \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (cat_Set_obj_prod_up I F A \)" unfolding cat_Set_obj_prod_up_def by auto show "vcard (cat_Set_obj_prod_up I F A \) = 3\<^sub>\" unfolding cat_Set_obj_prod_up_def by (auto simp: nat_omega_simps) show "\\<^sub>\ (cat_Set_obj_prod_up I F A \\ArrVal\) \\<^sub>\ cat_Set_obj_prod_up I F A \\ArrCod\" unfolding cat_Set_obj_prod_up_components(3) by (rule cat_Set_obj_prod_up_ArrVal_vrange[OF assms(3)]) show "cat_Set_obj_prod_up I F A \\ArrCod\ \\<^sub>\ Vset \" unfolding cat_Set_cs_simps by (rule Limit_vproduct_in_Vset_if_VLambda_in_VsetI) (simp_all add: cat_Set_cs_simps assms) qed (auto simp: assms cat_Set_cs_simps intro: cat_Set_cs_intros) lemma (in \) cat_Set_cf_comp_proj_obj_prod_up: assumes "A \\<^sub>\ Vset \" and "VLambda I F \\<^sub>\ Vset \" and "\i. i \\<^sub>\ I \ \ i : A \\<^bsub>cat_Set \\<^esub> F i" and "i \\<^sub>\ I" shows "\ i = vprojection_arrow I F i \\<^sub>A\<^bsub>cat_Set \\<^esub> cat_Set_obj_prod_up I F A \" (is \\ i = ?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\) proof(rule arr_Set_eqI[of \]) note \i = assms(3)[OF assms(4)] note \i = cat_Set_is_arrD[OF \i] \i have Fi: "?Fi : (\\<^sub>\i\\<^sub>\I. F i) \\<^bsub>cat_Set \\<^esub> F i" by (rule vprojection_arrow_is_arr[OF assms(4,2)]) from cat_Set_obj_prod_up_cat_Set_is_arr[OF assms(1,2,3)] have \: "cat_Set_obj_prod_up I F A \ : A \\<^bsub>cat_Set \\<^esub> (\\<^sub>\i\\<^sub>\I. F i)" by simp show "arr_Set \ (\ i)" by (rule \i(1)) interpret \i: arr_Set \ \\ i\ by (rule \i(1)) from Fi \ have Fi_\: "?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\ : A \\<^bsub>cat_Set \\<^esub> F i" by (cs_concl cs_intro: cat_cs_intros) then show arr_Set_Fi_\: "arr_Set \ (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)" by (auto simp: cat_Set_is_arrD(1)) interpret arr_Set \ \?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\\ by (rule arr_Set_Fi_\) from \i have dom_lhs: "\\<^sub>\ (\ i\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps) from Fi_\ have dom_rhs: "\\<^sub>\ ((?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\) = A" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show "\ i\ArrVal\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume prems: "a \\<^sub>\ A" from assms(4) prems \i(4) \ Fi show "\ i\ArrVal\\a\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrVal\\a\" by ( cs_concl cs_simp: cat_Set_cs_simps cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros ) qed auto from Fi \ show "\ i\ArrDom\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrDom\" by (cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps \i(2)) from Fi \ show "\ i\ArrCod\ = (?Fi \\<^sub>A\<^bsub>cat_Set \\<^esub> ?\)\ArrCod\" by (cs_concl cs_simp: cat_cs_simps cat_Set_cs_simps \i(3)) qed subsection\Equalizer object for the category \Set\\ text\ The definition of the (non-categorical concept of an) equalizer can be found in \cite{noauthor_wikipedia_2001}\footnote{ \url{https://en.wikipedia.org/wiki/Equaliser_(mathematics)} }\ definition vequalizer :: "V \ V \ V \ V" where "vequalizer X f g = set {x. x \\<^sub>\ X \ f\ArrVal\\x\ = g\ArrVal\\x\}" lemma small_vequalizer[simp]: "small {x. x \\<^sub>\ X \ f\ArrVal\\x\ = g\ArrVal\\x\}" by auto text\Rules.\ lemma vequalizerI: assumes "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" shows "x \\<^sub>\ vequalizer X f g" using assms unfolding vequalizer_def by auto lemma vequalizerD[dest]: assumes "x \\<^sub>\ vequalizer X f g" shows "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" using assms unfolding vequalizer_def by auto lemma vequalizerE[elim]: assumes "x \\<^sub>\ vequalizer X f g" obtains "x \\<^sub>\ X" and "f\ArrVal\\x\ = g\ArrVal\\x\" using assms unfolding vequalizer_def by auto text\Elementary results.\ lemma vequalizer_vsubset_vdomain[cat_Set_cs_intros]: "vequalizer a g f \\<^sub>\ a" by auto lemma Limit_vequalizer_in_Vset[cat_Set_cs_intros]: assumes "Limit \" and "a \\<^sub>\ Vset \" shows "vequalizer a g f \\<^sub>\ Vset \" using assms by auto lemma vequalizer_flip: "vequalizer a f g = vequalizer a g f" unfolding vequalizer_def by auto lemma (in \) cat_Set_incl_Set_commute: assumes "\ : \ \\<^bsub>cat_Set \\<^esub> \" and "\ : \ \\<^bsub>cat_Set \\<^esub> \" shows "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer \ \ \) \ = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> incl_Set (vequalizer \ \ \) \" (is \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl = \ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\) proof- note \ = cat_Set_is_arrD[OF assms(1)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note \ = cat_Set_is_arrD[OF assms(2)] interpret \: arr_Set \ \ rewrites "\\ArrDom\ = \" and "\\ArrCod\ = \" by (rule \(1)) (simp_all add: \) note [cat_Set_cs_intros] = \.arr_Set_ArrDom_in_Vset \.arr_Set_ArrCod_in_Vset from assms have \_incl: "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl : vequalizer \ \ \ \\<^bsub>cat_Set \\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros) then have dom_lhs: "\\<^sub>\ ((\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\) = vequalizer \ \ \" by (cs_concl cs_simp: cat_cs_simps)+ from assms have \_incl: "\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl : vequalizer \ \ \ \\<^bsub>cat_Set \\<^esub> \" by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros) then have dom_rhs: "\\<^sub>\ ((\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\) = vequalizer \ \ \" by (cs_concl cs_simp: cat_cs_simps)+ show ?thesis proof(rule arr_Set_eqI) from \_incl show arr_Set_\_incl: "arr_Set \ (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by (auto dest: cat_Set_is_arrD(1)) interpret arr_Set_\_incl: arr_Set \ \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\ by (rule arr_Set_\_incl) from \_incl show arr_Set_\_incl: "arr_Set \ (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)" by (auto dest: cat_Set_is_arrD(1)) interpret arr_Set_\_incl: arr_Set \ \\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl\ by (rule arr_Set_\_incl) show "(\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\ = (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a assume "a \\<^sub>\ vequalizer \ \ \" with assms show "(\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\\a\ = (\ \\<^sub>A\<^bsub>cat_Set \\<^esub> ?incl)\ArrVal\\a\" by ( cs_concl cs_simp: vequalizerD(2) cat_Set_cs_simps cat_cs_simps cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros ) qed auto qed (use \_incl \_incl in \cs_concl cs_simp: cat_cs_simps\)+ qed subsection\Auxiliary\ text\ This subsection is reserved for insignificant helper lemmas and rules that are used in applied formalization elsewhere. \ lemma (in \) cat_Rel_CId_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" shows "cat_Rel \\CId\\A\ : A \\<^bsub>cat_Set \\<^esub> A" proof- from assms show ?thesis unfolding cat_Rel_components cat_Set_components(6)[symmetric] by (cs_concl cs_simp: cat_Set_components(1) cs_intro: cat_cs_intros) qed lemma (in \) cat_Rel_CId_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B' = A" and "C' = A" and "\' = cat_Set \" shows "cat_Rel \\CId\\A\ : B' \\<^bsub>\'\<^esub> C'" using assms(1) unfolding assms(2-4) by (rule cat_Rel_CId_is_cat_Set_arr) text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Structure_Example.thy b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Structure_Example.thy --- a/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Structure_Example.thy +++ b/thys/CZH_Elementary_Categories/czh_ecategories/CZH_ECAT_Structure_Example.thy @@ -1,1990 +1,2347 @@ (* Copyright 2021 (C) Mihails Milehins *) section\Example: categories with additional structure\ theory CZH_ECAT_Structure_Example imports CZH_ECAT_Introduction CZH_ECAT_PCategory CZH_ECAT_Set begin subsection\Background\ text\ The examples that are presented in this section showcase how the framework developed in this article can be used for the formalization of the theory of categories with additional structure. The content of this section also indicates some of the potential future directions for this body of work. \ subsection\Dagger category\ named_theorems dag_field_simps -named_theorems catdag_cs_simps -named_theorems catdag_cs_intros +named_theorems dagcat_cs_simps +named_theorems dagcat_cs_intros definition DagCat :: V where [dag_field_simps]: "DagCat = 0" definition DagDag :: V where [dag_field_simps]: "DagDag = 1\<^sub>\" abbreviation DagDag_app :: "V \ V" (\\\<^sub>C\) where "\\<^sub>C \ \ \\DagDag\" subsubsection\Definition and elementary properties\ text\ For further information see \cite{noauthor_nlab_nodate}\footnote{\url{ https://ncatlab.org/nlab/show/dagger+category }}. \ locale dagger_category = \ \ + vfsequence \ + DagCat: category \ \\\DagCat\\ + DagDag: is_functor \ \op_cat (\\DagCat\)\ \\\DagCat\\ \\\<^sub>C \\ for \ \ + - assumes catdag_length: "vcard \ = 2\<^sub>\" - and catdag_ObjMap_identity[catdag_cs_simps]: + assumes dagcat_length: "vcard \ = 2\<^sub>\" + and dagcat_ObjMap_identity[dagcat_cs_simps]: "a \\<^sub>\ \\DagCat\\Obj\ \ (\\<^sub>C \)\ObjMap\\a\ = a" - and catdag_DagCat_idem[catdag_cs_simps]: + and dagcat_DagCat_idem[dagcat_cs_simps]: "\\<^sub>C \ \<^sub>C\<^sub>F\ \\<^sub>C \ = cf_id (\\DagCat\)" -lemmas [catdag_cs_simps] = - dagger_category.catdag_ObjMap_identity - dagger_category.catdag_DagCat_idem +lemmas [dagcat_cs_simps] = + dagger_category.dagcat_ObjMap_identity + dagger_category.dagcat_DagCat_idem text\Rules.\ -lemma (in dagger_category) dagger_category_axioms'[cat_cs_intros]: +lemma (in dagger_category) dagger_category_axioms'[dagcat_cs_intros]: assumes "\' = \" shows "dagger_category \' \" unfolding assms by (rule dagger_category_axioms) mk_ide rf dagger_category_def[unfolded dagger_category_axioms_def] |intro dagger_categoryI| |dest dagger_categoryD[dest]| |elim dagger_categoryE[elim]| -lemma category_if_dagger_category[catdag_cs_intros]: +lemma category_if_dagger_category[dagcat_cs_intros]: assumes "\' = (\\DagCat\)" and "dagger_category \ \" shows "category \ \'" unfolding assms(1) using assms(2) by (rule dagger_categoryD(3)) -lemma (in dagger_category) catdag_is_functor'[catdag_cs_intros]: +lemma (in dagger_category) dagcat_is_functor'[dagcat_cs_intros]: assumes "\' = op_cat (\\DagCat\)" and "\' = \\DagCat\" shows "\\<^sub>C \ : \' \\\<^sub>C\<^bsub>\\<^esub> \'" unfolding assms by (rule DagDag.is_functor_axioms) -lemmas [catdag_cs_intros] = dagger_category.catdag_is_functor' +lemmas [dagcat_cs_intros] = dagger_category.dagcat_is_functor' subsection\\Rel\ as a dagger category\ subsubsection\Definition and elementary properties\ text\ For further information see \cite{noauthor_nlab_nodate}\footnote{\url{ https://ncatlab.org/nlab/show/Rel }}. \ definition dagcat_Rel :: "V \ V" where "dagcat_Rel \ = [cat_Rel \, \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \]\<^sub>\" text\Components.\ lemma dagcat_Rel_components: shows "dagcat_Rel \\DagCat\ = cat_Rel \" and "dagcat_Rel \\DagDag\ = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \" unfolding dagcat_Rel_def dag_field_simps by (simp_all add: nat_omega_simps) subsubsection\\Rel\ is a dagger category\ -lemma (in \) "dagger_category \ (dagcat_Rel \)" +lemma (in \) dagger_category_dagcat_Rel: "dagger_category \ (dagcat_Rel \)" proof(intro dagger_categoryI) show "category \ (dagcat_Rel \\DagCat\)" by (cs_concl cs_simp: dagcat_Rel_components cs_intro: cat_Rel_cs_intros) show "\\<^sub>C (dagcat_Rel \) : op_cat (dagcat_Rel \\DagCat\) \\\<^sub>C\<^bsub>\\<^esub> dagcat_Rel \\DagCat\" unfolding dagcat_Rel_components by (cs_concl cs_intro: cf_cs_intros cat_cs_intros) show "vcard (dagcat_Rel \) = 2\<^sub>\" unfolding dagcat_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>C (dagcat_Rel \)\ObjMap\\a\ = a" if "a \\<^sub>\ dagcat_Rel \\DagCat\\Obj\" for a using that unfolding dagcat_Rel_components cat_Rel_components(1) by (cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps) show "\\<^sub>C (dagcat_Rel \) \<^sub>C\<^sub>F\ \\<^sub>C (dagcat_Rel \) = dghm_id (dagcat_Rel \\DagCat\)" unfolding dagcat_Rel_components by (cs_concl cs_simp: cf_cn_comp_cf_dag_Rel_cf_dag_Rel) qed (auto simp: dagcat_Rel_def) subsection\Monoidal category\ text\ For background information see Chapter 2 in \cite{etingof_tensor_2015}. \ subsubsection\Background\ named_theorems mcat_field_simps named_theorems mcat_cs_simps named_theorems mcat_cs_intros definition Mcat :: V where [mcat_field_simps]: "Mcat = 0" definition Mcf :: V where [mcat_field_simps]: "Mcf = 1\<^sub>\" definition Me :: V where [mcat_field_simps]: "Me = 2\<^sub>\" definition M\ :: V where [mcat_field_simps]: "M\ = 3\<^sub>\" definition Ml :: V where [mcat_field_simps]: "Ml = 4\<^sub>\" definition Mr :: V where [mcat_field_simps]: "Mr = 5\<^sub>\" subsubsection\Definition and elementary properties\ locale monoidal_category = \\See Definition 2.2.8 in \cite{etingof_tensor_2015}.\ \ \ + vfsequence \ + Mcat: category \ \\\Mcat\\ + Mcf: is_functor \ \(\\Mcat\) \\<^sub>C (\\Mcat\)\ \\\Mcat\\ \\\Mcf\\ + M\: is_iso_ntcf \ \\\Mcat\^\<^sub>C\<^sub>3\ \\\Mcat\\ \cf_blcomp (\\Mcf\)\ \cf_brcomp (\\Mcf\)\ \\\M\\\ + Ml: is_iso_ntcf \ \\\Mcat\\ \\\Mcat\\ \\\Mcf\\<^bsub>\\Mcat\,\\Mcat\\<^esub>(\\Me\,-)\<^sub>C\<^sub>F\ \cf_id (\\Mcat\)\ \\\Ml\\ + Mr: is_iso_ntcf \ \\\Mcat\\ \\\Mcat\\ \\\Mcf\\<^bsub>\\Mcat\,\\Mcat\\<^esub>(-,\\Me\)\<^sub>C\<^sub>F\ \cf_id (\\Mcat\)\ \\\Mr\\ for \ \ + assumes mcat_length[mcat_cs_simps]: "vcard \ = 6\<^sub>\" and mcat_Me_is_obj[mcat_cs_intros]: "\\Me\ \\<^sub>\ \\Mcat\\Obj\" and mcat_pentagon: "\ a \\<^sub>\ \\Mcat\\Obj\; b \\<^sub>\ \\Mcat\\Obj\; c \\<^sub>\ \\Mcat\\Obj\; d \\<^sub>\ \\Mcat\\Obj\ \ \ (\\Mcat\\CId\\a\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\Mcf\\<^esub> \\M\\\NTMap\\b, c, d\\<^sub>\) \\<^sub>A\<^bsub>\\Mcat\\<^esub> \\M\\\NTMap\\a, b \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\Mcf\\<^esub> c, d\\<^sub>\ \\<^sub>A\<^bsub>\\Mcat\\<^esub> (\\M\\\NTMap\\a, b, c\\<^sub>\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\Mcf\\<^esub> \\Mcat\\CId\\d\) = \\M\\\NTMap\\a, b, c \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\Mcf\\<^esub> d\\<^sub>\ \\<^sub>A\<^bsub>\\Mcat\\<^esub> \\M\\\NTMap\\a \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>\\Mcf\\<^esub> b, c, d\\<^sub>\" and mcat_triangle[mcat_cs_simps]: "\ a \\<^sub>\ \\Mcat\\Obj\; b \\<^sub>\ \\Mcat\\Obj\ \ \ (\\Mcat\\CId\\a\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\Mcf\\<^esub> \\Ml\\NTMap\\b\) \\<^sub>A\<^bsub>\\Mcat\\<^esub> \\M\\\NTMap\\a, \\Me\, b\\<^sub>\ = (\\Mr\\NTMap\\a\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\Mcf\\<^esub> \\Mcat\\CId\\b\)" lemmas [mcat_cs_intros] = monoidal_category.mcat_Me_is_obj lemmas [mcat_cs_simps] = monoidal_category.mcat_triangle text\Rules.\ -lemma (in monoidal_category) monoidal_category_axioms'[cat_cs_intros]: +lemma (in monoidal_category) monoidal_category_axioms'[mcat_cs_intros]: assumes "\' = \" shows "monoidal_category \' \" unfolding assms by (rule monoidal_category_axioms) mk_ide rf monoidal_category_def[unfolded monoidal_category_axioms_def] |intro monoidal_categoryI| |dest monoidal_categoryD[dest]| |elim monoidal_categoryE[elim]| text\Elementary properties.\ lemma mcat_eqI: assumes "monoidal_category \ \" and "monoidal_category \ \" and "\\Mcat\ = \\Mcat\" and "\\Mcf\ = \\Mcf\" and "\\Me\ = \\Me\" and "\\M\\ = \\M\\" and "\\Ml\ = \\Ml\" and "\\Mr\ = \\Mr\" shows "\ = \" proof- interpret \: monoidal_category \ \ by (rule assms(1)) interpret \: monoidal_category \ \ by (rule assms(2)) show ?thesis proof(rule vsv_eqI) have dom: "\\<^sub>\ \ = 6\<^sub>\" by (cs_concl cs_simp: mcat_cs_simps V_cs_simps) show "\\<^sub>\ \ = \\<^sub>\ \" by (cs_concl cs_simp: mcat_cs_simps V_cs_simps) show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = \\a\" for a by (unfold dom, elim_in_numeral, insert assms) (auto simp: mcat_field_simps) qed auto qed subsection\Components for \M\\ for \Rel\\ subsubsection\Definition and elementary properties\ definition M\_Rel_arrow_lr :: "V \ V \ V \ V" where "M\_Rel_arrow_lr A B C = [ (\ab_c\\<^sub>\(A \\<^sub>\ B) \\<^sub>\ C. \vfst (vfst ab_c), \vsnd (vfst ab_c), vsnd ab_c\\), (A \\<^sub>\ B) \\<^sub>\ C, A \\<^sub>\ (B \\<^sub>\ C) ]\<^sub>\" definition M\_Rel_arrow_rl :: "V \ V \ V \ V" where "M\_Rel_arrow_rl A B C = [ (\a_bc\\<^sub>\A \\<^sub>\ (B \\<^sub>\ C). \\vfst a_bc, vfst (vsnd a_bc)\, vsnd (vsnd a_bc)\), A \\<^sub>\ (B \\<^sub>\ C), (A \\<^sub>\ B) \\<^sub>\ C ]\<^sub>\" text\Components.\ lemma M\_Rel_arrow_lr_components: shows "M\_Rel_arrow_lr A B C\ArrVal\ = (\ab_c\\<^sub>\(A \\<^sub>\ B) \\<^sub>\ C. \vfst (vfst ab_c), \vsnd (vfst ab_c), vsnd ab_c\\)" and [cat_cs_simps]: "M\_Rel_arrow_lr A B C\ArrDom\ = (A \\<^sub>\ B) \\<^sub>\ C" and [cat_cs_simps]: "M\_Rel_arrow_lr A B C\ArrCod\ = A \\<^sub>\ (B \\<^sub>\ C)" unfolding M\_Rel_arrow_lr_def arr_field_simps by (simp_all add: nat_omega_simps) lemma M\_Rel_arrow_rl_components: shows "M\_Rel_arrow_rl A B C\ArrVal\ = (\a_bc\\<^sub>\A \\<^sub>\ (B \\<^sub>\ C). \\vfst a_bc, vfst (vsnd a_bc)\, vsnd (vsnd a_bc)\)" and [cat_cs_simps]: "M\_Rel_arrow_rl A B C\ArrDom\ = A \\<^sub>\ (B \\<^sub>\ C)" and [cat_cs_simps]: "M\_Rel_arrow_rl A B C\ArrCod\ = (A \\<^sub>\ B) \\<^sub>\ C" unfolding M\_Rel_arrow_rl_def arr_field_simps by (simp_all add: nat_omega_simps) subsubsection\Arrow value\ mk_VLambda M\_Rel_arrow_lr_components(1) |vsv M\_Rel_arrow_lr_ArrVal_vsv[cat_cs_intros]| |vdomain M\_Rel_arrow_lr_ArrVal_vdomain[cat_cs_simps]| |app M\_Rel_arrow_lr_ArrVal_app'| lemma M\_Rel_arrow_lr_ArrVal_app[cat_cs_simps]: assumes "ab_c = \\a, b\, c\" and "ab_c \\<^sub>\ (A \\<^sub>\ B) \\<^sub>\ C" shows "M\_Rel_arrow_lr A B C\ArrVal\\ab_c\ = \a, \b, c\\" using assms(2) unfolding assms(1) by (simp_all add: M\_Rel_arrow_lr_ArrVal_app' nat_omega_simps) mk_VLambda M\_Rel_arrow_rl_components(1) |vsv M\_Rel_arrow_rl_ArrVal_vsv[cat_cs_intros]| |vdomain M\_Rel_arrow_rl_ArrVal_vdomain[cat_cs_simps]| |app M\_Rel_arrow_rl_ArrVal_app'| lemma M\_Rel_arrow_rl_ArrVal_app[cat_cs_simps]: assumes "a_bc = \a, \b, c\\" and "a_bc \\<^sub>\ A \\<^sub>\ (B \\<^sub>\ C)" shows "M\_Rel_arrow_rl A B C\ArrVal\\a_bc\ = \\a, b\, c\" using assms(2) unfolding assms(1) by (simp_all add: M\_Rel_arrow_rl_ArrVal_app' nat_omega_simps) subsubsection\Components for \M\\ for \Rel\ are arrows\ lemma (in \) M\_Rel_arrow_lr_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "C \\<^sub>\ Vset \" shows "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (M\_Rel_arrow_lr A B C)" unfolding M\_Rel_arrow_lr_def by auto show "vcard (M\_Rel_arrow_lr A B C) = 3\<^sub>\" unfolding M\_Rel_arrow_lr_def by (simp add: nat_omega_simps) show "\\<^sub>\ (M\_Rel_arrow_lr A B C\ArrVal\) \\<^sub>\ M\_Rel_arrow_lr A B C\ArrCod\" unfolding M\_Rel_arrow_lr_components by auto qed ( use assms in \cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros\ )+ lemma (in \) M\_Rel_arrow_rl_is_cat_Set_arr_Vset: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "C \\<^sub>\ Vset \" shows "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof(intro cat_Set_is_arrI arr_SetI) show "vfsequence (M\_Rel_arrow_rl A B C)" unfolding M\_Rel_arrow_rl_def by auto show "vcard (M\_Rel_arrow_rl A B C) = 3\<^sub>\" unfolding M\_Rel_arrow_rl_def by (simp add: nat_omega_simps) show "\\<^sub>\ (M\_Rel_arrow_rl A B C\ArrVal\) \\<^sub>\ M\_Rel_arrow_rl A B C\ArrCod\" unfolding M\_Rel_arrow_rl_components by auto qed ( use assms in \cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros\ )+ lemma (in \) M\_Rel_arrow_lr_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" shows "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" using assms unfolding cat_Set_components by (rule M\_Rel_arrow_lr_is_cat_Set_arr_Vset) lemma (in \) M\_Rel_arrow_lr_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Set \" shows "M\_Rel_arrow_lr A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Set_arr' lemma (in \) M\_Rel_arrow_rl_is_cat_Set_arr: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" shows "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" using assms unfolding cat_Set_components by (rule M\_Rel_arrow_rl_is_cat_Set_arr_Vset) lemma (in \) M\_Rel_arrow_rl_is_cat_Set_arr'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Set \" shows "M\_Rel_arrow_rl A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Set_arr) lemmas [cat_rel_par_Set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Set_arr' lemma (in \) M\_Rel_arrow_lr_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" shows "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Par \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD M\_Rel_arrow_lr_is_cat_Set_arr_Vset) auto qed lemma (in \) M\_Rel_arrow_lr_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Par \" shows "M\_Rel_arrow_lr A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Par_arr' lemma (in \) M\_Rel_arrow_rl_is_cat_Par_arr: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" shows "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Par \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) from assms show ?thesis unfolding cat_Par_components(1) by (intro Set_Par.subcat_is_arrD M\_Rel_arrow_rl_is_cat_Set_arr_Vset) auto qed lemma (in \) M\_Rel_arrow_rl_is_cat_Par_arr'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Par \" shows "M\_Rel_arrow_rl A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Par_arr) lemmas [cat_rel_Par_set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Par_arr' lemma (in \) M\_Rel_arrow_lr_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" shows "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD M\_Rel_arrow_lr_is_cat_Set_arr_Vset) auto qed lemma (in \) M\_Rel_arrow_lr_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Rel \" shows "M\_Rel_arrow_lr A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Rel_arr' lemma (in \) M\_Rel_arrow_rl_is_cat_Rel_arr: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" shows "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Rel \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule subcat_trans[ OF Set_Par.subcategory_axioms Par_Rel.subcategory_axioms ] ) from assms show ?thesis unfolding cat_Rel_components(1) by (intro Set_Rel.subcat_is_arrD M\_Rel_arrow_rl_is_cat_Set_arr_Vset) auto qed lemma (in \) M\_Rel_arrow_rl_is_cat_Rel_arr'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Rel \" shows "M\_Rel_arrow_rl A B C : A' \\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Rel_arr) lemmas [cat_Rel_par_set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Rel_arr' subsubsection\Further properties\ lemma (in \) M\_Rel_arrow_rl_M\_Rel_arrow_lr[cat_cs_simps]: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "C \\<^sub>\ Vset \" shows "M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C = cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\" proof- interpret Set: category \ \cat_Set \\ by (cs_concl cs_intro: cat_cs_intros) from assms have lhs: "M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by ( cs_concl cs_simp: cat_Set_components(1) cs_intro: cat_rel_par_Set_cs_intros cat_cs_intros ) then have dom_lhs: "\\<^sub>\ ((M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C)\ArrVal\) = (A \\<^sub>\ B) \\<^sub>\ C" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms Set.category_axioms have rhs: "cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\ : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by ( cs_concl cs_simp: cat_Set_components(1) cs_intro: V_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ ((cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\)\ArrVal\) = (A \\<^sub>\ B) \\<^sub>\ C" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show ?thesis proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C)" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ (cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\)" by (auto dest: cat_Set_is_arrD(1)) show "(M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C)\ArrVal\ = cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix ab_c assume prems: "ab_c \\<^sub>\ (A \\<^sub>\ B) \\<^sub>\ C" then obtain a b c where ab_c_def: "ab_c = \\a, b\, c\" and a: "a \\<^sub>\ A" and b: "b \\<^sub>\ B" and c: "c \\<^sub>\ C" by clarsimp from assms prems a b c lhs rhs show "(M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C)\ArrVal\\ab_c\ = cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\\ArrVal\\ab_c\" unfolding ab_c_def by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cs_intro: cat_rel_par_Set_cs_intros V_cs_intros cat_cs_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed lemma (in \) M\_Rel_arrow_rl_M\_Rel_arrow_lr'[cat_cs_simps]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" shows "M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C = cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\" - using assms unfolding cat_Set_components(1) by (rule M\_Rel_arrow_rl_M\_Rel_arrow_lr) + using assms + unfolding cat_Set_components(1) + by (rule M\_Rel_arrow_rl_M\_Rel_arrow_lr) lemmas [cat_cs_simps] = \.M\_Rel_arrow_rl_M\_Rel_arrow_lr' lemma (in \) M\_Rel_arrow_lr_M\_Rel_arrow_rl[cat_cs_simps]: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "C \\<^sub>\ Vset \" shows "M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C = cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\" proof- interpret Set: category \ \cat_Set \\ by (cs_concl cs_intro: cat_cs_intros) from assms have lhs: "M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" by ( cs_concl cs_simp: cat_Set_components(1) cs_intro: cat_rel_par_Set_cs_intros cat_cs_intros ) then have dom_lhs: "\\<^sub>\ ((M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C)\ArrVal\) = A \\<^sub>\ (B \\<^sub>\ C)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) from assms Set.category_axioms have rhs: "cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\ : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" by ( cs_concl cs_simp: cat_Set_components(1) cs_intro: V_cs_intros cat_cs_intros ) then have dom_rhs: "\\<^sub>\ ((cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\)\ArrVal\) = A \\<^sub>\ (B \\<^sub>\ C)" by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros) show ?thesis proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C)" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ (cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\)" by (auto dest: cat_Set_is_arrD(1)) show "(M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C)\ArrVal\ = cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix a_bc assume prems: "a_bc \\<^sub>\ A \\<^sub>\ (B \\<^sub>\ C)" then obtain a b c where a_bc_def: "a_bc = \a, \b, c\\" and a: "a \\<^sub>\ A" and b: "b \\<^sub>\ B" and c: "c \\<^sub>\ C" by clarsimp from assms prems a b c lhs rhs show "(M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C)\ArrVal\\a_bc\ = cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\\ArrVal\\a_bc\" unfolding a_bc_def by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cs_intro: V_cs_intros cat_rel_par_Set_cs_intros cat_cs_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed lemma (in \) M\_Rel_arrow_lr_M\_Rel_arrow_rl'[cat_cs_simps]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" shows "M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C = cat_Set \\CId\\A \\<^sub>\ (B \\<^sub>\ C)\" using assms unfolding cat_Set_components(1) by (rule M\_Rel_arrow_lr_M\_Rel_arrow_rl) lemmas [cat_cs_simps] = \.M\_Rel_arrow_lr_M\_Rel_arrow_rl' subsubsection\Components for \M\\ for \Rel\ are isomorphisms\ lemma (in \) assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "C \\<^sub>\ Vset \" shows M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism_Vset: "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" and M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism_Vset: "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof- interpret Set: category \ \cat_Set \\ by (cs_concl cs_intro: cat_cs_intros) have lhs: "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by (intro M\_Rel_arrow_rl_is_cat_Set_arr_Vset assms) from assms have [cat_cs_simps]: "M\_Rel_arrow_rl A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A B C = cat_Set \\CId\\(A \\<^sub>\ B) \\<^sub>\ C\" by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cs_intro: cat_cs_intros ) from assms have [cat_cs_simps]: "M\_Rel_arrow_lr A B C \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_rl A B C = cat_Set \\CId\\A \\<^sub>\ B \\<^sub>\ C\" by ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cs_intro: cat_cs_intros ) from Set.is_arr_isomorphismI' [ OF lhs M\_Rel_arrow_lr_is_cat_Set_arr_Vset[OF assms], unfolded cat_cs_simps, simplified ] show "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" and "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by auto qed lemma (in \) assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" shows M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism: "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" and M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism: "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Set \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" using assms unfolding cat_Set_components(1) by ( all \ intro M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism_Vset M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism_Vset \ ) lemma (in \) M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Set \" shows "M\_Rel_arrow_lr A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism' lemma (in \) M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism'[cat_rel_par_Set_cs_intros]: assumes "A \\<^sub>\ cat_Set \\Obj\" and "B \\<^sub>\ cat_Set \\Obj\" and "C \\<^sub>\ cat_Set \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Set \" shows "M\_Rel_arrow_rl A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism) lemmas [cat_rel_par_Set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism' lemma (in \) assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" shows M\_Rel_arrow_lr_is_cat_Par_arr_isomorphism: "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" and M\_Rel_arrow_rl_is_cat_Par_arr_isomorphism: "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) show "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) show "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Par \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by ( rule Set_Par.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Par_components] ] ] ) qed lemma (in \) M\_Rel_arrow_lr_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Par \" shows "M\_Rel_arrow_lr A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Par_arr_isomorphism' lemma (in \) M\_Rel_arrow_rl_is_cat_Par_arr_isomorphism'[cat_rel_Par_set_cs_intros]: assumes "A \\<^sub>\ cat_Par \\Obj\" and "B \\<^sub>\ cat_Par \\Obj\" and "C \\<^sub>\ cat_Par \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Par \" shows "M\_Rel_arrow_rl A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Par_arr_isomorphism) lemmas [cat_rel_Par_set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Par_arr_isomorphism' lemma (in \) assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" shows M\_Rel_arrow_lr_is_cat_Rel_arr_isomorphism: "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" and M\_Rel_arrow_rl_is_cat_Rel_arr_isomorphism: "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show "M\_Rel_arrow_lr A B C : (A \\<^sub>\ B) \\<^sub>\ C \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> A \\<^sub>\ (B \\<^sub>\ C)" by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF M\_Rel_arrow_lr_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) show "M\_Rel_arrow_rl A B C : A \\<^sub>\ (B \\<^sub>\ C) \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> (A \\<^sub>\ B) \\<^sub>\ C" by ( rule Set_Rel.wr_subcat_is_arr_isomorphism_is_arr_isomorphism [ THEN iffD1, OF M\_Rel_arrow_rl_is_cat_Set_arr_isomorphism_Vset[ OF assms[unfolded cat_Rel_components] ] ] ) qed lemma (in \) M\_Rel_arrow_lr_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" and "A' = (A \\<^sub>\ B) \\<^sub>\ C" and "B' = A \\<^sub>\ (B \\<^sub>\ C)" and "\' = cat_Rel \" shows "M\_Rel_arrow_lr A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_lr_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.M\_Rel_arrow_lr_is_cat_Rel_arr_isomorphism' lemma (in \) M\_Rel_arrow_rl_is_cat_Rel_arr_isomorphism'[cat_Rel_par_set_cs_intros]: assumes "A \\<^sub>\ cat_Rel \\Obj\" and "B \\<^sub>\ cat_Rel \\Obj\" and "C \\<^sub>\ cat_Rel \\Obj\" and "A' = A \\<^sub>\ (B \\<^sub>\ C)" and "B' = (A \\<^sub>\ B) \\<^sub>\ C" and "\' = cat_Rel \" shows "M\_Rel_arrow_rl A B C : A' \\<^sub>i\<^sub>s\<^sub>o\<^bsub>\'\<^esub> B'" using assms(1-3) unfolding assms(4-6) by (rule M\_Rel_arrow_rl_is_cat_Rel_arr_isomorphism) lemmas [cat_Rel_par_set_cs_intros] = \.M\_Rel_arrow_rl_is_cat_Rel_arr_isomorphism' subsection\\M\\ for \Rel\\ subsubsection\Definition and elementary properties\ definition M\_Rel :: "V \ V" where "M\_Rel \ = [ (\abc\\<^sub>\(\^\<^sub>C\<^sub>3)\Obj\. M\_Rel_arrow_lr (abc\0\) (abc\1\<^sub>\\) (abc\2\<^sub>\\)), cf_blcomp (cf_prod_2_Rel \), cf_brcomp (cf_prod_2_Rel \), \^\<^sub>C\<^sub>3, \ ]\<^sub>\" text\Components.\ lemma M\_Rel_components: shows "M\_Rel \\NTMap\ = (\abc\\<^sub>\(\^\<^sub>C\<^sub>3)\Obj\. M\_Rel_arrow_lr (abc\0\) (abc\1\<^sub>\\) (abc\2\<^sub>\\))" and [cat_cs_simps]: "M\_Rel \\NTDom\ = cf_blcomp (cf_prod_2_Rel \)" and [cat_cs_simps]: "M\_Rel \\NTCod\ = cf_brcomp (cf_prod_2_Rel \)" and [cat_cs_simps]: "M\_Rel \\NTDGDom\ = \^\<^sub>C\<^sub>3" and [cat_cs_simps]: "M\_Rel \\NTDGCod\ = \" unfolding M\_Rel_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda M\_Rel_components(1) |vsv M\_Rel_NTMap_vsv[cat_cs_intros]| |vdomain M\_Rel_NTMap_vdomain[cat_cs_simps]| |app M\_Rel_NTMap_app'| lemma M\_Rel_NTMap_app[cat_cs_simps]: assumes "ABC = [A, B, C]\<^sub>\" and "ABC \\<^sub>\ (\^\<^sub>C\<^sub>3)\Obj\" shows "M\_Rel \\NTMap\\ABC\ = M\_Rel_arrow_lr A B C" using assms(2) unfolding assms(1) by (simp add: M\_Rel_NTMap_app' nat_omega_simps) subsubsection\\M\\ for \Rel\ is a natural isomorphism\ lemma (in \) M\_Rel_is_iso_ntcf: "M\_Rel (cat_Rel \) : cf_blcomp (cf_prod_2_Rel (cat_Rel \)) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_brcomp (cf_prod_2_Rel (cat_Rel \)) : cat_Rel \^\<^sub>C\<^sub>3 \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof- interpret cf_prod: is_functor \ \cat_Rel \ \\<^sub>C cat_Rel \\ \cat_Rel \\ \cf_prod_2_Rel (cat_Rel \)\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence (M\_Rel (cat_Rel \))" unfolding M\_Rel_def by auto show "vcard (M\_Rel (cat_Rel \)) = 5\<^sub>\" unfolding M\_Rel_def by (simp add: nat_omega_simps) show "M\_Rel (cat_Rel \)\NTMap\\ABC\ : cf_blcomp (cf_prod_2_Rel (cat_Rel \))\ObjMap\\ABC\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> cf_brcomp (cf_prod_2_Rel (cat_Rel \))\ObjMap\\ABC\" if "ABC \\<^sub>\ (cat_Rel \^\<^sub>C\<^sub>3)\Obj\" for ABC proof- from that category_cat_Rel obtain A B C where ABC_def: "ABC = [A, B, C]\<^sub>\" and A: "A \\<^sub>\ cat_Rel \\Obj\" and B: "B \\<^sub>\ cat_Rel \\Obj\" and C: "C \\<^sub>\ cat_Rel \\Obj\" by (elim cat_prod_3_ObjE[rotated 3]) from that A B C show ?thesis unfolding ABC_def by ( cs_concl cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros cat_prod_cs_intros cs_simp: cat_cs_simps cat_Rel_cs_simps ) qed then show "M\_Rel (cat_Rel \)\NTMap\\ABC\ : cf_blcomp (cf_prod_2_Rel (cat_Rel \))\ObjMap\\ABC\ \\<^bsub>cat_Rel \\<^esub> cf_brcomp (cf_prod_2_Rel (cat_Rel \))\ObjMap\\ABC\" if "ABC \\<^sub>\ (cat_Rel \^\<^sub>C\<^sub>3)\Obj\" for ABC using that by (simp add: cat_Rel_is_arr_isomorphismD(1)) show "M\_Rel (cat_Rel \)\NTMap\\ABC'\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> cf_blcomp (cf_prod_2_Rel (cat_Rel \))\ArrMap\\HGF\ = cf_brcomp (cf_prod_2_Rel (cat_Rel \))\ArrMap\\HGF\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> M\_Rel (cat_Rel \)\NTMap\\ABC\" if "HGF : ABC \\<^bsub>cat_Rel \^\<^sub>C\<^sub>3\<^esub> ABC'" for ABC ABC' HGF proof- from that obtain H G F A B C A' B' C' where HGF_def: "HGF = [H, G, F]\<^sub>\" and ABC_def: "ABC = [A, B, C]\<^sub>\" and ABC'_def: "ABC' = [A', B', C']\<^sub>\" and H_is_arr: "H : A \\<^bsub>cat_Rel \\<^esub> A'" and G_is_arr: "G : B \\<^bsub>cat_Rel \\<^esub> B'" and F_is_arr: "F : C \\<^bsub>cat_Rel \\<^esub> C'" by ( elim cat_prod_3_is_arrE[ OF category_cat_Rel category_cat_Rel category_cat_Rel ] ) note H = cat_Rel_is_arrD[OF H_is_arr] note G = cat_Rel_is_arrD[OF G_is_arr] note F = cat_Rel_is_arrD[OF F_is_arr] interpret H: arr_Rel \ H rewrites "H\ArrDom\ = A" and "H\ArrCod\ = A'" by (intro H)+ interpret G: arr_Rel \ G rewrites "G\ArrDom\ = B" and "G\ArrCod\ = B'" by (intro G)+ interpret F: arr_Rel \ F rewrites "F\ArrDom\ = C" and "F\ArrCod\ = C'" by (intro F)+ let ?ABC' = \M\_Rel_arrow_lr A' B' C'\ and ?ABC = \M\_Rel_arrow_lr A B C\ and ?HG_F = \ prod_2_Rel_ArrVal (prod_2_Rel_ArrVal (H\ArrVal\) (G\ArrVal\)) (F\ArrVal\) \ and ?H_GF = \ prod_2_Rel_ArrVal (H\ArrVal\) (prod_2_Rel_ArrVal (G\ArrVal\) (F\ArrVal\)) \ have [cat_cs_simps]: "?ABC' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (prod_2_Rel H G) F = prod_2_Rel H (prod_2_Rel G F) \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?ABC" proof- from H_is_arr G_is_arr F_is_arr have lhs: "?ABC' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (prod_2_Rel H G) F : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Rel \\<^esub> A' \\<^sub>\ (B' \\<^sub>\ C')" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) from H_is_arr G_is_arr F_is_arr have rhs: "prod_2_Rel H (prod_2_Rel G F) \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?ABC : (A \\<^sub>\ B) \\<^sub>\ C \\<^bsub>cat_Rel \\<^esub> A' \\<^sub>\ (B' \\<^sub>\ C')" by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros) show ?thesis proof(rule arr_Rel_eqI) from lhs show arr_Rel_lhs: "arr_Rel \ (?ABC' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (prod_2_Rel H G) F)" by (auto dest: cat_Rel_is_arrD) from rhs show arr_Rel_rhs: "arr_Rel \ (prod_2_Rel H (prod_2_Rel G F) \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?ABC)" by (auto dest: cat_Rel_is_arrD) have [cat_cs_simps]: "?ABC'\ArrVal\ \\<^sub>\ ?HG_F = ?H_GF \\<^sub>\ ?ABC\ArrVal\" proof(intro vsubset_antisym vsubsetI) fix abc_abc'' assume prems: "abc_abc'' \\<^sub>\ ?ABC'\ArrVal\ \\<^sub>\ ?HG_F" then obtain abc abc' abc'' where abc_abc''_def: "abc_abc'' = \abc, abc''\" and abc_abc': "\abc, abc'\ \\<^sub>\ ?HG_F" and abc'_abc'': "\abc', abc''\ \\<^sub>\ ?ABC'\ArrVal\" by clarsimp from abc_abc' obtain ab c ab' c' where abc_abc'_def: "\abc, abc'\ = \\ab, c\, \ab', c'\\" and ab_ab': "\ab, ab'\ \\<^sub>\ prod_2_Rel_ArrVal (H\ArrVal\) (G\ArrVal\)" and cc': "\c, c'\ \\<^sub>\ F\ArrVal\" by auto then have abc_def: "abc = \ab, c\" and abc'_def: "abc' = \ab', c'\" by auto from ab_ab' obtain a b a' b' where ab_ab'_def: "\ab, ab'\ = \\a, b\, \a', b'\\" and aa': "\a, a'\ \\<^sub>\ H\ArrVal\" and bb': "\b, b'\ \\<^sub>\ G\ArrVal\" by auto then have ab_def: "ab = \a, b\" and ab'_def: "ab' = \a', b'\" by auto from cc' F.arr_Rel_ArrVal_vdomain F.arr_Rel_ArrVal_vrange have c: "c \\<^sub>\ C" and c': "c' \\<^sub>\ C'" by auto from bb' G.arr_Rel_ArrVal_vdomain G.arr_Rel_ArrVal_vrange have b: "b \\<^sub>\ B" and b': "b' \\<^sub>\ B'" by auto from aa' H.arr_Rel_ArrVal_vdomain H.arr_Rel_ArrVal_vrange have a: "a \\<^sub>\ A" and a': "a' \\<^sub>\ A'" by auto from abc'_abc'' have "abc'' = ?ABC'\ArrVal\\abc'\" by (simp add: vsv.vsv_appI[OF M\_Rel_arrow_lr_ArrVal_vsv]) also from a' b' c' have "\ = \a', \b', c'\\" unfolding abc'_def ab'_def by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) finally have abc''_def: "abc'' = \a', \b', c'\\" by auto from aa' bb' cc' a a' b b' c c' show "abc_abc'' \\<^sub>\ ?H_GF \\<^sub>\ ?ABC\ArrVal\" unfolding abc_abc''_def abc_def abc'_def abc''_def ab'_def ab_def by (intro vcompI prod_2_Rel_ArrValI) ( cs_concl cs_simp: cat_cs_simps cs_intro: vsv.vsv_ex1_app2[THEN iffD1] V_cs_intros cat_cs_intros cat_Rel_cs_intros )+ next fix abc_abc'' assume prems: "abc_abc'' \\<^sub>\ ?H_GF \\<^sub>\ ?ABC\ArrVal\" then obtain abc abc' abc'' where abc_abc''_def: "abc_abc'' = \abc, abc''\" and abc_abc': "\abc, abc'\ \\<^sub>\ ?ABC\ArrVal\" and abc'_abc'': "\abc', abc''\ \\<^sub>\ ?H_GF" by clarsimp from abc'_abc'' obtain a' bc' a'' bc'' where abc'_abc''_def: "\abc', abc''\ = \\a', bc'\, \a'', bc''\\" and aa'': "\a', a''\ \\<^sub>\ H\ArrVal\" and bc'_bc'': "\bc', bc''\ \\<^sub>\ prod_2_Rel_ArrVal (G\ArrVal\) (F\ArrVal\)" by auto then have abc'_def: "abc' = \a', bc'\" and abc''_def: "abc'' = \a'', bc''\" by auto from bc'_bc'' obtain b' c' b'' c'' where bc'_bc''_def: "\bc', bc''\ = \\b', c'\, \b'', c''\\" and bb'': "\b', b''\ \\<^sub>\ G\ArrVal\" and cc'': "\c', c''\ \\<^sub>\ F\ArrVal\" by auto then have bc'_def: "bc' = \b', c'\" and bc''_def: "bc'' = \b'', c''\" by auto from cc'' F.arr_Rel_ArrVal_vdomain F.arr_Rel_ArrVal_vrange have c': "c' \\<^sub>\ C" and c'': "c'' \\<^sub>\ C'" by auto from bb'' G.arr_Rel_ArrVal_vdomain G.arr_Rel_ArrVal_vrange have b': "b' \\<^sub>\ B" and b'': "b'' \\<^sub>\ B'" by auto from aa'' H.arr_Rel_ArrVal_vdomain H.arr_Rel_ArrVal_vrange have a': "a' \\<^sub>\ A" and a'': "a'' \\<^sub>\ A'" by auto from abc_abc' have "abc \\<^sub>\ \\<^sub>\ (?ABC\ArrVal\)" by auto then have "abc \\<^sub>\ (A \\<^sub>\ B) \\<^sub>\ C" by (simp add: cat_cs_simps) then obtain a b c where abc_def: "abc = \\a, b\, c\" and a: "a \\<^sub>\ A" and b: "b \\<^sub>\ B" and c: "c \\<^sub>\ C" by auto from abc_abc' have "abc' = ?ABC\ArrVal\\abc\" by (simp add: vsv.vsv_appI[OF M\_Rel_arrow_lr_ArrVal_vsv]) also from a b c have "\ = \a, \b, c\\" unfolding abc_def bc'_def by (cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros) finally have abc'_def': "abc' = \a, \b, c\\" by auto with abc'_def[unfolded bc'_def] have [cat_cs_simps]: "a = a'" "b = b'" "c = c'" by auto from a'' b'' c'' have "\\a'', b''\, c''\ \\<^sub>\ (A' \\<^sub>\ B') \\<^sub>\ C'" by (cs_concl cs_intro: V_cs_intros) with aa'' bb'' cc'' a a' b b' c c' show "abc_abc'' \\<^sub>\ ?ABC'\ArrVal\ \\<^sub>\ ?HG_F" unfolding abc_abc''_def abc_def abc'_def abc''_def bc''_def by (intro vcompI prod_2_Rel_ArrValI) ( cs_concl cs_simp: cat_cs_simps cs_intro: vsv.vsv_ex1_app2[THEN iffD1] V_cs_intros cat_cs_intros cat_Rel_cs_intros )+ qed from that H_is_arr G_is_arr F_is_arr show "(?ABC' \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (prod_2_Rel H G) F)\ArrVal\ = (prod_2_Rel H (prod_2_Rel G F) \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?ABC)\ArrVal\" by ( cs_concl cs_simp: prod_2_Rel_components comp_Rel_components cat_Rel_cs_simps cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros cat_prod_cs_intros ) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed from that H_is_arr G_is_arr F_is_arr show ?thesis unfolding HGF_def ABC_def ABC'_def by ( cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros cat_prod_cs_intros cs_simp: cat_Rel_cs_simps cat_cs_simps ) qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ qed lemma (in \) M\_Rel_is_iso_ntcf'[cat_cs_intros]: assumes "\' = cf_blcomp (cf_prod_2_Rel (cat_Rel \))" and "\' = cf_brcomp (cf_prod_2_Rel (cat_Rel \))" and "\' = cat_Rel \^\<^sub>C\<^sub>3" and "\' = cat_Rel \" and "\' = \" shows "M\_Rel (cat_Rel \) : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" unfolding assms by (rule M\_Rel_is_iso_ntcf) lemmas [cat_cs_intros] = \.M\_Rel_is_iso_ntcf' subsection\\Ml\ and \Mr\ for \Rel\\ subsubsection\Definition and elementary properties\ definition Ml_Rel :: "V \ V \ V" where "Ml_Rel \ a = [ (\B\\<^sub>\\\Obj\. vsnd_arrow (set {a}) B), cf_prod_2_Rel \\<^bsub>\,\\<^esub>(set {a},-)\<^sub>C\<^sub>F, cf_id \, \, \ ]\<^sub>\" definition Mr_Rel :: "V \ V \ V" where "Mr_Rel \ b = [ (\A\\<^sub>\\\Obj\. vfst_arrow A (set {b})), cf_prod_2_Rel \\<^bsub>\,\\<^esub>(-,set {b})\<^sub>C\<^sub>F, cf_id \, \, \ ]\<^sub>\" text\Components.\ lemma Ml_Rel_components: shows "Ml_Rel \ a\NTMap\ = (\B\\<^sub>\\\Obj\. vsnd_arrow (set {a}) B)" and [cat_cs_simps]: "Ml_Rel \ a\NTDom\ = cf_prod_2_Rel \\<^bsub>\,\\<^esub>(set {a},-)\<^sub>C\<^sub>F" and [cat_cs_simps]: "Ml_Rel \ a\NTCod\ = cf_id \" and [cat_cs_simps]: "Ml_Rel \ a\NTDGDom\ = \" and [cat_cs_simps]: "Ml_Rel \ a\NTDGCod\ = \" unfolding Ml_Rel_def nt_field_simps by (simp_all add: nat_omega_simps) lemma Mr_Rel_components: shows "Mr_Rel \ b\NTMap\ = (\A\\<^sub>\\\Obj\. vfst_arrow A (set {b}))" and [cat_cs_simps]: "Mr_Rel \ b\NTDom\ = cf_prod_2_Rel \\<^bsub>\,\\<^esub>(-,set {b})\<^sub>C\<^sub>F" and [cat_cs_simps]: "Mr_Rel \ b\NTCod\ = cf_id \" and [cat_cs_simps]: "Mr_Rel \ b\NTDGDom\ = \" and [cat_cs_simps]: "Mr_Rel \ b\NTDGCod\ = \" unfolding Mr_Rel_def nt_field_simps by (simp_all add: nat_omega_simps) subsubsection\Natural transformation map\ mk_VLambda Ml_Rel_components(1) |vsv Ml_Rel_components_NTMap_vsv[cat_cs_intros]| |vdomain Ml_Rel_components_NTMap_vdomain[cat_cs_simps]| |app Ml_Rel_components_NTMap_app[cat_cs_simps]| mk_VLambda Mr_Rel_components(1) |vsv Mr_Rel_components_NTMap_vsv[cat_cs_intros]| |vdomain Mr_Rel_components_NTMap_vdomain[cat_cs_simps]| |app Mr_Rel_components_NTMap_app[cat_cs_simps]| subsubsection\\Ml\ and \Mr\ for \Rel\ are natural isomorphisms\ lemma (in \) Ml_Rel_is_iso_ntcf: assumes "a \\<^sub>\ cat_Rel \\Obj\" shows "Ml_Rel (cat_Rel \) a: cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub>(set {a},-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_id (cat_Rel \) : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof- let ?cf_prod = \cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub> (set {a},-)\<^sub>C\<^sub>F\ note [cat_cs_simps] = set_empty interpret cf_prod: is_functor \ \cat_Rel \ \\<^sub>C cat_Rel \\ \cat_Rel \\ \cf_prod_2_Rel (cat_Rel \)\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence (Ml_Rel (cat_Rel \) a)" unfolding Ml_Rel_def by auto show "vcard (Ml_Rel (cat_Rel \) a) = 5\<^sub>\" unfolding Ml_Rel_def by (simp add: nat_omega_simps) from assms show "?cf_prod : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_cs_intros V_cs_intros ) show "Ml_Rel (cat_Rel \) a\NTMap\\B\ : ?cf_prod\ObjMap\\B\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> cf_id (cat_Rel \)\ObjMap\\B\" if "B \\<^sub>\ cat_Rel \\Obj\" for B using assms that by ( cs_concl cs_simp: cat_Rel_components(1) V_cs_simps cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros V_cs_intros cat_prod_cs_intros ) with cat_Rel_is_arr_isomorphismD[OF this] show "Ml_Rel (cat_Rel \) a\NTMap\\B\ : ?cf_prod\ObjMap\\B\ \\<^bsub>cat_Rel \\<^esub> cf_id (cat_Rel \)\ObjMap\\B\" if "B \\<^sub>\ cat_Rel \\Obj\" for B using that by simp show "Ml_Rel (cat_Rel \) a\NTMap\\B\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?cf_prod\ArrMap\\F\ = cf_id (cat_Rel \)\ArrMap\\F\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> Ml_Rel (cat_Rel \) a\NTMap\\A\" if "F : A \\<^bsub>cat_Rel \\<^esub> B" for A B F proof- note F = cat_Rel_is_arrD[OF that] interpret F: arr_Rel \ F rewrites "F\ArrDom\ = A" and "F\ArrCod\ = B" by (intro F)+ have [cat_cs_simps]: "vsnd_arrow (set {a}) B \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (cat_Rel \\CId\\set {a}\) F = F \\<^sub>A\<^bsub>cat_Rel \\<^esub> vsnd_arrow (set {a}) A" (is \?B2 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?aF = F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A2\) proof- from assms that have lhs: "?B2 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?aF : set {a} \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros V_cs_intros ) from assms that have rhs: "F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A2 : set {a} \\<^sub>\ A \\<^bsub>cat_Rel \\<^esub> B" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros V_cs_intros ) have [cat_cs_simps]: "?B2\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (vid_on (set {a})) (F\ArrVal\) = F\ArrVal\ \\<^sub>\ ?A2\ArrVal\" proof(intro vsubset_antisym vsubsetI) fix xx'_z assume "xx'_z \\<^sub>\ ?B2\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (vid_on (set {a})) (F\ArrVal\)" then obtain xx' yy' z where xx'_z_def: "xx'_z = \xx', z\" and xx'_yy': "\xx', yy'\ \\<^sub>\ prod_2_Rel_ArrVal (vid_on (set {a})) (F\ArrVal\)" and yy'_z: "\yy', z\ \\<^sub>\ ?B2\ArrVal\" by auto from xx'_yy' obtain x x' y y' where "\xx', yy'\ = \\x, x'\, \y, y'\\" and "\x, y\ \\<^sub>\ vid_on (set {a})" and xy': "\x', y'\ \\<^sub>\ F\ArrVal\" by auto then have xx'_def: "xx' = \a, x'\" and yy'_def: "yy' = \a, y'\" by simp_all with yy'_z have y': "y' \\<^sub>\ B" and z_def: "z = y'" unfolding vsnd_arrow_components by auto from xy' vsubsetD have x': "x' \\<^sub>\ A" by (auto intro: F.arr_Rel_ArrVal_vdomain) show "xx'_z \\<^sub>\ F\ArrVal\ \\<^sub>\ ?A2\ArrVal\" unfolding xx'_z_def z_def xx'_def by (intro vcompI, rule xy') (auto simp: vsnd_arrow_components x' VLambda_iff2) next fix ay_z assume "ay_z \\<^sub>\ F\ArrVal\ \\<^sub>\ ?A2\ArrVal\" then obtain ay y z where xx'_z_def: "ay_z = \ay, z\" and ay_y: "\ay, y\ \\<^sub>\ ?A2\ArrVal\" and yz[cat_cs_intros]: "\y, z\ \\<^sub>\ F\ArrVal\" by auto then have ay_z_def: "ay_z = \\a, y\, z\" and y: "y \\<^sub>\ A" and ay_def: "ay = \a, y\" unfolding vsnd_arrow_components by auto from yz vsubsetD have z: "z \\<^sub>\ B" by (auto intro: F.arr_Rel_ArrVal_vrange) have [cat_cs_intros]: "\a, a\ \\<^sub>\ vid_on (set {a})" by auto show "ay_z \\<^sub>\ ?B2\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (vid_on (set {a})) (F\ArrVal\)" unfolding ay_z_def by ( intro vcompI prod_2_Rel_ArrValI, rule vsv.vsv_ex1_app1[THEN iffD1], unfold cat_cs_simps, insert z ) ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed show ?thesis proof(rule arr_Rel_eqI) from lhs show arr_Rel_lhs: "arr_Rel \ (?B2 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?aF)" by (auto dest: cat_Rel_is_arrD) from rhs show "arr_Rel \ (F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A2)" by (auto dest: cat_Rel_is_arrD) note cat_Rel_CId_app[cat_Rel_cs_simps del] note \.cat_Rel_CId_app[cat_Rel_cs_simps del] from that assms show "(?B2 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?aF)\ArrVal\ = (F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A2)\ArrVal\" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros cs_simp: id_Rel_components cat_Rel_CId_app comp_Rel_components(1) prod_2_Rel_components cat_Rel_components(1) ) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed from that assms show ?thesis by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros cat_prod_cs_intros cs_simp: cat_Rel_components(1) V_cs_simps ) qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ qed lemma (in \) Ml_Rel_is_iso_ntcf'[cat_cs_intros]: assumes "a \\<^sub>\ cat_Rel \\Obj\" and "\' = cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub>(set {a},-)\<^sub>C\<^sub>F" and "\' = cf_id (cat_Rel \)" and "\' = cat_Rel \" and "\' = cat_Rel \" and "\' = \" - shows "Ml_Rel (cat_Rel \) a : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + shows "Ml_Rel (cat_Rel \) a : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1) unfolding assms(2-6) by (rule Ml_Rel_is_iso_ntcf) lemmas [cat_cs_intros] = \.Ml_Rel_is_iso_ntcf' lemma (in \) Mr_Rel_is_iso_ntcf: assumes "b \\<^sub>\ cat_Rel \\Obj\" shows "Mr_Rel (cat_Rel \) b : cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub>(-,set {b})\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_id (cat_Rel \) : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" proof- let ?cf_prod = \cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub> (-,set {b})\<^sub>C\<^sub>F\ note [cat_cs_simps] = set_empty interpret cf_prod: is_functor \ \cat_Rel \ \\<^sub>C cat_Rel \\ \cat_Rel \\ \cf_prod_2_Rel (cat_Rel \)\ by (cs_concl cs_intro: cat_cs_intros cat_Rel_cs_intros) show ?thesis proof(intro is_iso_ntcfI is_ntcfI') show "vfsequence (Mr_Rel (cat_Rel \) b)" unfolding Mr_Rel_def by auto show "vcard (Mr_Rel (cat_Rel \) b) = 5\<^sub>\" unfolding Mr_Rel_def by (simp add: nat_omega_simps) from assms show "?cf_prod : cat_Rel \ \\\<^sub>C\<^bsub>\\<^esub> cat_Rel \" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_cs_intros V_cs_intros ) show "Mr_Rel (cat_Rel \) b\NTMap\\B\ : ?cf_prod\ObjMap\\B\ \\<^sub>i\<^sub>s\<^sub>o\<^bsub>cat_Rel \\<^esub> cf_id (cat_Rel \)\ObjMap\\B\" if "B \\<^sub>\ cat_Rel \\Obj\" for B using assms that by ( cs_concl cs_simp: cat_Rel_components(1) V_cs_simps cat_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros cat_prod_cs_intros ) with cat_Rel_is_arr_isomorphismD[OF this] show "Mr_Rel (cat_Rel \) b\NTMap\\B\ : ?cf_prod\ObjMap\\B\ \\<^bsub>cat_Rel \\<^esub> cf_id (cat_Rel \)\ObjMap\\B\" if "B \\<^sub>\ cat_Rel \\Obj\" for B using that by simp show "Mr_Rel (cat_Rel \) b\NTMap\\B\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?cf_prod\ArrMap\\F\ = cf_id (cat_Rel \)\ArrMap\\F\ \\<^sub>A\<^bsub>cat_Rel \\<^esub> Mr_Rel (cat_Rel \) b\NTMap\\A\" if "F : A \\<^bsub>cat_Rel \\<^esub> B" for A B F proof- note F = cat_Rel_is_arrD[OF that] interpret F: arr_Rel \ F rewrites "F\ArrDom\ = A" and "F\ArrCod\ = B" by (intro F)+ have [cat_cs_simps]: "vfst_arrow B (set {b}) \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel F (cat_Rel \\CId\\set {b}\) = F \\<^sub>A\<^bsub>cat_Rel \\<^esub> vfst_arrow A (set {b})" (is \?B1 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?bF = F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A1\) proof- from assms that have lhs: "?B1 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?bF : A \\<^sub>\ set {b} \\<^bsub>cat_Rel \\<^esub> B" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros ) from assms that have rhs: "F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A1 : A \\<^sub>\ set {b} \\<^bsub>cat_Rel \\<^esub> B" by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros ) have [cat_cs_simps]: "?B1\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (F\ArrVal\) (vid_on (set {b})) = F\ArrVal\ \\<^sub>\ ?A1\ArrVal\" proof(intro vsubset_antisym vsubsetI) fix xx'_z assume "xx'_z \\<^sub>\ ?B1\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (F\ArrVal\) (vid_on (set {b}))" then obtain xx' yy' z where xx'_z_def: "xx'_z = \xx', z\" and xx'_yy': "\xx', yy'\ \\<^sub>\ prod_2_Rel_ArrVal (F\ArrVal\) (vid_on (set {b}))" and yy'_z: "\yy', z\ \\<^sub>\ ?B1\ArrVal\" by auto from xx'_yy' obtain x x' y y' where "\xx', yy'\ = \\x, x'\, \y, y'\\" and "\x', y'\ \\<^sub>\ vid_on (set {b})" and xy: "\x, y\ \\<^sub>\ F\ArrVal\" by auto then have xx'_def: "xx' = \x, b\" and yy'_def: "yy' = \y, b\" by simp_all with yy'_z have y': "y \\<^sub>\ B" and z_def: "z = y" unfolding vfst_arrow_components by auto from xy vsubsetD have x: "x \\<^sub>\ A" by (auto intro: F.arr_Rel_ArrVal_vdomain) show "xx'_z \\<^sub>\ F\ArrVal\ \\<^sub>\ ?A1\ArrVal\" unfolding xx'_z_def z_def xx'_def by (intro vcompI, rule xy) (auto simp: vfst_arrow_components x VLambda_iff2) next fix xy_z assume "xy_z \\<^sub>\ F\ArrVal\ \\<^sub>\ ?A1\ArrVal\" then obtain xy y z where xx'_z_def: "xy_z = \xy, z\" and xy_y: "\xy, y\ \\<^sub>\ ?A1\ArrVal\" and yz[cat_cs_intros]: "\y, z\ \\<^sub>\ F\ArrVal\" by auto then have xy_z_def: "xy_z = \\y, b\, z\" and y: "y \\<^sub>\ A" and xy_def: "xy = \y, b\" unfolding vfst_arrow_components by auto from yz vsubsetD have z: "z \\<^sub>\ B" by (auto intro: F.arr_Rel_ArrVal_vrange) have [cat_cs_intros]: "\b, b\ \\<^sub>\ vid_on (set {b})" by auto show "xy_z \\<^sub>\ ?B1\ArrVal\ \\<^sub>\ prod_2_Rel_ArrVal (F\ArrVal\) (vid_on (set {b}))" unfolding xy_z_def by ( intro vcompI prod_2_Rel_ArrValI, rule vsv.vsv_ex1_app1[THEN iffD1], unfold cat_cs_simps, insert z ) ( cs_concl cs_simp: cat_cs_simps cs_intro: V_cs_intros cat_cs_intros ) qed show ?thesis proof(rule arr_Rel_eqI) from lhs show arr_Rel_lhs: "arr_Rel \ (?B1 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?bF)" by (auto dest: cat_Rel_is_arrD) from rhs show "arr_Rel \ (F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A1)" by (auto dest: cat_Rel_is_arrD) note cat_Rel_CId_app[cat_Rel_cs_simps del] note \.cat_Rel_CId_app[cat_Rel_cs_simps del] from that assms show "(?B1 \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?bF)\ArrVal\ = (F \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?A1)\ArrVal\" by (*slow*) ( cs_concl cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros cs_simp: id_Rel_components cat_Rel_CId_app comp_Rel_components(1) prod_2_Rel_components cat_Rel_components(1) ) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed from that assms show ?thesis by ( cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros V_cs_intros cat_prod_cs_intros cs_simp: cat_Rel_components(1) V_cs_simps ) qed qed (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+ qed lemma (in \) Mr_Rel_is_iso_ntcf'[cat_cs_intros]: assumes "b \\<^sub>\ cat_Rel \\Obj\" and "\' = cf_prod_2_Rel (cat_Rel \)\<^bsub>cat_Rel \,cat_Rel \\<^esub>(-,set {b})\<^sub>C\<^sub>F" and "\' = cf_id (cat_Rel \)" and "\' = cat_Rel \" and "\' = cat_Rel \" and "\' = \" - shows "Mr_Rel (cat_Rel \) b : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\\<^esub> \'" + shows "Mr_Rel (cat_Rel \) b : \' \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o \' : \' \\\<^sub>C\<^bsub>\'\<^esub> \'" using assms(1) unfolding assms(2-6) by (rule Mr_Rel_is_iso_ntcf) lemmas [cat_cs_intros] = \.Mr_Rel_is_iso_ntcf' subsection\\Rel\ as a monoidal category\ subsubsection\Definition and elementary properties\ text\ For further information see \cite{noauthor_wikipedia_2001}\footnote{\url{ https://en.wikipedia.org/wiki/Category_of_relations }}. \ definition mcat_Rel :: "V \ V \ V" where "mcat_Rel \ a = [ cat_Rel \, cf_prod_2_Rel (cat_Rel \), set {a}, M\_Rel (cat_Rel \), Ml_Rel (cat_Rel \) a, Mr_Rel (cat_Rel \) a ]\<^sub>\" text\Components.\ lemma mcat_Rel_components: shows "mcat_Rel \ a\Mcat\ = cat_Rel \" and "mcat_Rel \ a\Mcf\ = cf_prod_2_Rel (cat_Rel \)" and "mcat_Rel \ a\Me\ = set {a}" and "mcat_Rel \ a\M\\ = M\_Rel (cat_Rel \)" and "mcat_Rel \ a\Ml\ = Ml_Rel (cat_Rel \) a" and "mcat_Rel \ a\Mr\ = Mr_Rel (cat_Rel \) a" unfolding mcat_Rel_def mcat_field_simps by (simp_all add: nat_omega_simps) subsubsection\\Rel\ is a monoidal category\ -lemma (in \) +lemma (in \) monoidal_category_mcat_Rel: assumes "a \\<^sub>\ cat_Rel \\Obj\" shows "monoidal_category \ (mcat_Rel \ a)" proof- interpret Set_Par: wide_replete_subcategory \ \cat_Set \\ \cat_Par \\ by (rule wide_replete_subcategory_cat_Set_cat_Par) interpret Par_Rel: wide_replete_subcategory \ \cat_Par \\ \cat_Rel \\ by (rule wide_replete_subcategory_cat_Par_cat_Rel) interpret Set_Rel: wide_replete_subcategory \ \cat_Set \\ \cat_Rel \\ by ( rule wr_subcat_trans [ OF Set_Par.wide_replete_subcategory_axioms Par_Rel.wide_replete_subcategory_axioms ] ) show ?thesis proof(rule monoidal_categoryI) show "vfsequence (mcat_Rel \ a)" unfolding mcat_Rel_def by auto show "category \ (mcat_Rel \ a\Mcat\)" unfolding mcat_Rel_components by (cs_concl cs_intro: cat_cs_intros) show "mcat_Rel \ a\Mcf\ : mcat_Rel \ a\Mcat\ \\<^sub>C mcat_Rel \ a\Mcat\ \\\<^sub>C\<^bsub>\\<^esub> mcat_Rel \ a\Mcat\" unfolding mcat_Rel_components by (cs_concl cs_intro: cat_cs_intros) show "mcat_Rel \ a\M\\ : cf_blcomp (mcat_Rel \ a\Mcf\) \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_brcomp (mcat_Rel \ a\Mcf\) : mcat_Rel \ a\Mcat\^\<^sub>C\<^sub>3 \\\<^sub>C\<^bsub>\\<^esub> mcat_Rel \ a\Mcat\" unfolding mcat_Rel_components by (cs_concl cs_intro: cat_cs_intros) from assms show "mcat_Rel \ a\Ml\ : mcat_Rel \ a\Mcf\\<^bsub>mcat_Rel \ a\Mcat\,mcat_Rel \ a\Mcat\\<^esub> (mcat_Rel \ a\Me\,-)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_id (mcat_Rel \ a\Mcat\) : mcat_Rel \ a\Mcat\ \\\<^sub>C\<^bsub>\\<^esub> mcat_Rel \ a\Mcat\" unfolding mcat_Rel_components by (cs_concl cs_intro: cat_cs_intros) from assms show "mcat_Rel \ a\Mr\ : mcat_Rel \ a\Mcf\\<^bsub>mcat_Rel \ a\Mcat\,mcat_Rel \ a\Mcat\\<^esub> (-,mcat_Rel \ a\Me\)\<^sub>C\<^sub>F \\<^sub>C\<^sub>F\<^sub>.\<^sub>i\<^sub>s\<^sub>o cf_id (mcat_Rel \ a\Mcat\) : mcat_Rel \ a\Mcat\ \\\<^sub>C\<^bsub>\\<^esub> mcat_Rel \ a\Mcat\" unfolding mcat_Rel_components by (cs_concl cs_intro: cat_cs_intros) show "vcard (mcat_Rel \ a) = 6\<^sub>\" unfolding mcat_Rel_def by (simp add: nat_omega_simps) from assms show "mcat_Rel \ a\Me\ \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" unfolding mcat_Rel_components cat_Rel_components by force show "mcat_Rel \ a\Mcat\\CId\\A\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>mcat_Rel \ a\Mcf\\<^esub> mcat_Rel \ a\M\\\NTMap\\B, C, D\\<^sub>\ \\<^sub>A\<^bsub>mcat_Rel \ a\Mcat\\<^esub> mcat_Rel \ a\M\\\NTMap\\ A, B \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>mcat_Rel \ a\Mcf\\<^esub> C, D \\<^sub>\ \\<^sub>A\<^bsub>mcat_Rel \ a\Mcat\\<^esub> (mcat_Rel \ a\M\\\NTMap\\A, B, C\\<^sub>\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>mcat_Rel \ a\Mcf\\<^esub> mcat_Rel \ a\Mcat\\CId\\D\) = mcat_Rel \ a\M\\\NTMap\\ A, B, C \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>mcat_Rel \ a\Mcf\\<^esub> D \\<^sub>\ \\<^sub>A\<^bsub>mcat_Rel \ a\Mcat\\<^esub> mcat_Rel \ a\M\\\NTMap\\A \\<^sub>H\<^sub>M\<^sub>.\<^sub>O\<^bsub>mcat_Rel \ a\Mcf\\<^esub> B, C, D\\<^sub>\" if "A \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" and "B \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" and "C \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" and "D \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" for A B C D proof- have [cat_cs_simps]: "prod_2_Rel (cat_Rel \\CId\\A\) (M\_Rel_arrow_lr B C D) \\<^sub>A\<^bsub>cat_Rel \\<^esub> ( M\_Rel_arrow_lr A (B \\<^sub>\ C) D \\<^sub>A\<^bsub>cat_Rel \\<^esub> prod_2_Rel (M\_Rel_arrow_lr A B C) (cat_Rel \\CId\\D\) ) = M\_Rel_arrow_lr A B (C \\<^sub>\ D) \\<^sub>A\<^bsub>cat_Rel \\<^esub> M\_Rel_arrow_lr (A \\<^sub>\ B) C D" ( is \ ?A_BCD \\<^sub>A\<^bsub>cat_Rel \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?ABC_D) = ?A_B_CD \\<^sub>A\<^bsub>cat_Rel \\<^esub> ?AB_C_D \ ) proof- have [cat_cs_simps]: "prod_2_Rel (cat_Set \\CId\\A\) (M\_Rel_arrow_lr B C D) \\<^sub>A\<^bsub>cat_Set \\<^esub> ( ?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> prod_2_Rel (M\_Rel_arrow_lr A B C) (cat_Set \\CId\\D\) ) = ?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D" ( is \ ?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D) = ?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D \ ) proof- from that have lhs: "?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D) : ((A \\<^sub>\ B) \\<^sub>\ C) \\<^sub>\ D \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ B \\<^sub>\ C \\<^sub>\ D" unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_Set_components(1) cs_intro: cat_rel_par_Set_cs_intros cat_cs_intros V_cs_intros ) then have dom_lhs: "\\<^sub>\ ((?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D))\ArrVal\) = ((A \\<^sub>\ B) \\<^sub>\ C) \\<^sub>\ D" by (cs_concl cs_simp: cat_cs_simps) from that have rhs: "?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D : ((A \\<^sub>\ B) \\<^sub>\ C) \\<^sub>\ D \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ B \\<^sub>\ C \\<^sub>\ D" unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_Rel_components(1) cat_Set_components(1) cs_intro: cat_cs_intros V_cs_intros M\_Rel_arrow_lr_is_cat_Set_arr' ) then have dom_rhs: "\\<^sub>\ ((?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D)\ArrVal\) = ((A \\<^sub>\ B) \\<^sub>\ C) \\<^sub>\ D" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D))" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ (?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D)" by (auto dest: cat_Set_is_arrD(1)) show "(?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D))\ArrVal\ = (?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D)\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix abcd assume prems: "abcd \\<^sub>\ ((A \\<^sub>\ B) \\<^sub>\ C) \\<^sub>\ D" then obtain a b c d where abcd_def: "abcd = \\\a, b\, c\, d\" and a: "a \\<^sub>\ A" and b: "b \\<^sub>\ B" and c: "c \\<^sub>\ C" and d: "d \\<^sub>\ D" by clarsimp from that prems a b c d show "( ?A_BCD \\<^sub>A\<^bsub>cat_Set \\<^esub> (?A_BC_D \\<^sub>A\<^bsub>cat_Set \\<^esub> ?ABC_D) )\ArrVal\\abcd\ = (?A_B_CD \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AB_C_D)\ArrVal\\abcd\" unfolding abcd_def mcat_Rel_components(1) cat_Rel_components(1) by (*slow*) ( cs_concl cs_simp: cat_Set_components(1) cat_cs_simps cat_rel_par_Set_cs_simps cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed from assms that show ?thesis unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_cs_simps cat_Rel_components(1) cat_Set_components(1) Set_Rel.subcat_CId[symmetric] Set_Rel.subcat_Comp_simp[symmetric] cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros )+ qed from that show ?thesis unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_Rel_components(1) cat_cs_simps cs_intro: cat_cs_intros cat_Rel_par_set_cs_intros V_cs_intros cat_prod_cs_intros ) qed show "mcat_Rel \ a\Mcat\\CId\\A\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>mcat_Rel \ a\Mcf\\<^esub> mcat_Rel \ a\Ml\\NTMap\\B\ \\<^sub>A\<^bsub>mcat_Rel \ a\Mcat\\<^esub> mcat_Rel \ a\M\\\NTMap\\A, mcat_Rel \ a\Me\, B\\<^sub>\ = mcat_Rel \ a\Mr\\NTMap\\A\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>mcat_Rel \ a\Mcf\\<^esub> mcat_Rel \ a\Mcat\\CId\\B\" if "A \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" and "B \\<^sub>\ mcat_Rel \ a\Mcat\\Obj\" for A B proof- note [cat_cs_simps] = set_empty have [cat_cs_simps]: "prod_2_Rel (cat_Set \\CId\\A\) (vsnd_arrow (set {a}) B) \\<^sub>A\<^bsub>cat_Set \\<^esub> M\_Rel_arrow_lr A (set {a}) B = prod_2_Rel (vfst_arrow A (set {a})) (cat_Set \\CId\\B\)" (is \?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB = ?Aa_B\) proof- from assms that have lhs: "?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB : (A \\<^sub>\ set {a}) \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ B" unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_cs_simps cat_Rel_components(1) cat_Set_components(1) cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros ) then have dom_lhs: "\\<^sub>\ ((?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB)\ArrVal\) = (A \\<^sub>\ set {a}) \\<^sub>\ B" by (cs_concl cs_simp: cat_cs_simps) from assms that have rhs: "?Aa_B : (A \\<^sub>\ set {a}) \\<^sub>\ B \\<^bsub>cat_Set \\<^esub> A \\<^sub>\ B" unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_cs_simps cat_Set_components(1) cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros ) then have dom_rhs: "\\<^sub>\ (?Aa_B\ArrVal\) = (A \\<^sub>\ set {a}) \\<^sub>\ B" by (cs_concl cs_simp: cat_cs_simps) show ?thesis proof(rule arr_Set_eqI) from lhs show arr_Set_lhs: "arr_Set \ (?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB)" by (auto dest: cat_Set_is_arrD(1)) from rhs show arr_Set_rhs: "arr_Set \ ?Aa_B" by (auto dest: cat_Set_is_arrD(1)) show "(?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB)\ArrVal\ = ?Aa_B\ArrVal\" proof(rule vsv_eqI, unfold dom_lhs dom_rhs) fix xay assume "xay \\<^sub>\ (A \\<^sub>\ set {a}) \\<^sub>\ B" then obtain x y where xay_def: "xay = \\x, a\, y\" and x: "x \\<^sub>\ A" and y: "y \\<^sub>\ B" by auto from assms that x y show "(?A_aB \\<^sub>A\<^bsub>cat_Set \\<^esub> ?AaB)\ArrVal\\xay\ = ?Aa_B\ArrVal\\xay\" unfolding xay_def mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_Rel_components(1) cat_Set_components(1) cat_cs_simps cat_rel_par_Set_cs_simps cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros ) qed (use arr_Set_lhs arr_Set_rhs in auto) qed (use lhs rhs in \cs_concl cs_simp: cat_cs_simps\)+ qed from assms that show ?thesis unfolding mcat_Rel_components cat_Rel_components(1) by ( cs_concl cs_simp: cat_cs_simps cat_Rel_components(1) cat_Set_components(1) Set_Rel.subcat_CId[symmetric] Set_Rel.subcat_Comp_simp[symmetric] cs_intro: cat_cs_intros cat_rel_par_Set_cs_intros V_cs_intros cat_prod_cs_intros Set_Rel.subcat_is_arrD ) qed qed auto qed + + +subsection\Dagger monoidal categories\ + + +subsubsection\Background\ + + +text\See \cite{coecke_survey_2010} for further information.\ + + +named_theorems dmcat_field_simps + +named_theorems dmcat_cs_simps +named_theorems dmcat_cs_intros + +definition DMcat :: V where [dmcat_field_simps]: "DMcat = 0" +definition DMdag :: V where [dmcat_field_simps]: "DMdag = 1\<^sub>\" +definition DMcf :: V where [dmcat_field_simps]: "DMcf = 2\<^sub>\" +definition DMe :: V where [dmcat_field_simps]: "DMe = 3\<^sub>\" +definition DM\ :: V where [dmcat_field_simps]: "DM\ = 4\<^sub>\" +definition DMl :: V where [dmcat_field_simps]: "DMl = 5\<^sub>\" +definition DMr :: V where [dmcat_field_simps]: "DMr = 6\<^sub>\" + +abbreviation DMDag_app :: "V \ V" (\\\<^sub>M\<^sub>C\) + where "\\<^sub>M\<^sub>C \ \ \\DMdag\" + + +subsubsection\Slicing\ + + +text\Dagger category.\ + +definition dmcat_dagcat :: "V \ V" + where "dmcat_dagcat \ = [\\DMcat\, \\DMdag\]\<^sub>\" + +lemma dmcat_dagcat_components[slicing_simps]: + shows "dmcat_dagcat \\DagCat\ = \\DMcat\" + and "dmcat_dagcat \\DagDag\ = \\DMdag\" + unfolding dmcat_dagcat_def dmcat_field_simps dag_field_simps + by (auto simp: nat_omega_simps) + + +text\Monoidal category.\ + +definition dmcat_mcat :: "V \ V" + where "dmcat_mcat \ = [\\DMcat\, \\DMcf\, \\DMe\, \\DM\\, \\DMl\, \\DMr\]\<^sub>\" + +lemma dmcat_mcat_components[slicing_simps]: + shows "dmcat_mcat \\Mcat\ = \\DMcat\" + and "dmcat_mcat \\Mcf\ = \\DMcf\" + and "dmcat_mcat \\Me\ = \\DMe\" + and "dmcat_mcat \\M\\ = \\DM\\" + and "dmcat_mcat \\Ml\ = \\DMl\" + and "dmcat_mcat \\Mr\ = \\DMr\" + unfolding dmcat_mcat_def dmcat_field_simps mcat_field_simps + by (auto simp: nat_omega_simps) + + +subsubsection\Definition and elementary properties\ + +locale dagger_monoidal_category = \ \ + vfsequence \ for \ \ + + assumes dmcat_length[dmcat_cs_simps]: "vcard \ = 7\<^sub>\" + and dmcat_dagger_category: "dagger_category \ (dmcat_dagcat \)" + and dmcat_monoidal_category: "monoidal_category \ (dmcat_mcat \)" + and dmcat_compatibility: + "\ g : c \\<^bsub>\\DMcat\\<^esub> d; f : a \\<^bsub>\\DMcat\\<^esub> b \ \ + \\<^sub>M\<^sub>C \\ArrMap\\g \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\DMcf\\<^esub> f\ = + \\<^sub>M\<^sub>C \\ArrMap\\g\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>\\DMcf\\<^esub> \\<^sub>M\<^sub>C \\ArrMap\\f\" + and dmcat_M\_unital: "A \\<^sub>\ (\\DMcat\^\<^sub>C\<^sub>3)\Obj\ \ + \\<^sub>M\<^sub>C \\ArrMap\\\\DM\\\NTMap\\A\\ = (\\DM\\\NTMap\\A\)\\<^sub>C\<^bsub>\\DMcat\\<^esub>" + and dmcat_Ml_unital: "a \\<^sub>\ \\DMcat\\Obj\ \ + \\<^sub>M\<^sub>C \\ArrMap\\\\DMl\\NTMap\\a\\ = (\\DMl\\NTMap\\a\)\\<^sub>C\<^bsub>\\DMcat\\<^esub>" + and dmcat_Mr_unital: "a \\<^sub>\ \\DMcat\\Obj\ \ + \\<^sub>M\<^sub>C \\ArrMap\\\\DMr\\NTMap\\a\\ = (\\DMr\\NTMap\\a\)\\<^sub>C\<^bsub>\\DMcat\\<^esub>" + + +text\Rules.\ + +lemma (in dagger_monoidal_category) + dagger_monoidal_category_axioms'[dmcat_cs_intros]: + assumes "\' = \" + shows "dagger_monoidal_category \' \" + unfolding assms by (rule dagger_monoidal_category_axioms) + +mk_ide rf + dagger_monoidal_category_def[unfolded dagger_monoidal_category_axioms_def] + |intro dagger_monoidal_categoryI[intro]| + |dest dagger_monoidal_categoryD[dest]| + |elim dagger_monoidal_categoryE[elim]| + + +text\Elementary properties.\ + +lemma dmcat_eqI: + assumes "dagger_monoidal_category \ \" + and "dagger_monoidal_category \ \" + and "\\DMcat\ = \\DMcat\" + and "\\DMdag\ = \\DMdag\" + and "\\DMcf\ = \\DMcf\" + and "\\DMe\ = \\DMe\" + and "\\DM\\ = \\DM\\" + and "\\DMl\ = \\DMl\" + and "\\DMr\ = \\DMr\" + shows "\ = \" +proof- + interpret \: dagger_monoidal_category \ \ by (rule assms(1)) + interpret \: dagger_monoidal_category \ \ by (rule assms(2)) + show ?thesis + proof(rule vsv_eqI) + have dom: "\\<^sub>\ \ = 7\<^sub>\" by (cs_concl cs_simp: dmcat_cs_simps V_cs_simps) + show "\\<^sub>\ \ = \\<^sub>\ \" by (cs_concl cs_simp: dmcat_cs_simps V_cs_simps) + show "a \\<^sub>\ \\<^sub>\ \ \ \\a\ = \\a\" for a + by (unfold dom, elim_in_numeral, insert assms) + (auto simp: dmcat_field_simps) + qed auto +qed + + +text\Slicing.\ + +context dagger_monoidal_category +begin + +interpretation dagcat: dagger_category \ \dmcat_dagcat \\ + by (rule dmcat_dagger_category) + +sublocale DMCat: category \ \\\DMcat\\ + by (rule dagcat.DagCat.category_axioms[unfolded slicing_simps]) + +sublocale DMDag: is_functor \ \op_cat (\\DMcat\)\ \\\DMcat\\ \\\<^sub>M\<^sub>C \\ + by (rule dagcat.DagDag.is_functor_axioms[unfolded slicing_simps]) + +lemmas_with [unfolded slicing_simps]: + dmcat_Dom_vdomain[dmcat_cs_simps] = dagcat.dagcat_ObjMap_identity + and dmcat_DagCat_idem[dmcat_cs_simps] = dagcat.dagcat_DagCat_idem + and dmcat_is_functor'[dmcat_cs_intros] = dagcat.dagcat_is_functor' + +end + +lemmas [dmcat_cs_simps] = + dagger_monoidal_category.dmcat_Dom_vdomain + dagger_monoidal_category.dmcat_DagCat_idem + +lemmas [dmcat_cs_intros] = dagger_monoidal_category.dmcat_is_functor' + +context dagger_monoidal_category +begin + +interpretation mcat: monoidal_category \ \dmcat_mcat \\ + by (rule dmcat_monoidal_category) + +sublocale DMcf: is_functor \ \\\DMcat\ \\<^sub>C \\DMcat\\ \\\DMcat\\ \\\DMcf\\ + by (rule mcat.Mcf.is_functor_axioms[unfolded slicing_simps]) + +sublocale DM\: is_iso_ntcf + \ \\\DMcat\^\<^sub>C\<^sub>3\ \\\DMcat\\ \cf_blcomp (\\DMcf\)\ \cf_brcomp (\\DMcf\)\ \\\DM\\\ + by (rule mcat.M\.is_iso_ntcf_axioms[unfolded slicing_simps]) + +sublocale DMl: is_iso_ntcf + \ + \\\DMcat\\ + \\\DMcat\\ + \\\DMcf\\<^bsub>\\DMcat\,\\DMcat\\<^esub>(\\DMe\,-)\<^sub>C\<^sub>F\ + \cf_id (\\DMcat\)\ + \\\DMl\\ + by (rule mcat.Ml.is_iso_ntcf_axioms[unfolded slicing_simps]) + +sublocale DMr: is_iso_ntcf + \ + \\\DMcat\\ + \\\DMcat\\ + \\\DMcf\\<^bsub>\\DMcat\,\\DMcat\\<^esub>(-,\\DMe\)\<^sub>C\<^sub>F\ + \cf_id (\\DMcat\)\ + \\\DMr\\ + by (rule mcat.Mr.is_iso_ntcf_axioms[unfolded slicing_simps]) + +lemmas_with [unfolded slicing_simps]: + dmcat_Me_is_obj[dmcat_cs_intros] = mcat.mcat_Me_is_obj + and dmcat_pentagon = mcat.mcat_pentagon + and dmcat_triangle[dmcat_cs_simps] = mcat.mcat_triangle + +end + +lemmas [dmcat_cs_intros] = dagger_monoidal_category.dmcat_Me_is_obj +lemmas [dmcat_cs_simps] = dagger_monoidal_category.dmcat_triangle + + + +subsection\\Rel\ as a dagger monoidal category\ + + +subsubsection\Definition and elementary properties\ + +definition dmcat_Rel :: "V \ V \ V" + where "dmcat_Rel \ a = + [ + cat_Rel \, + \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \, + cf_prod_2_Rel (cat_Rel \), + set {a}, + M\_Rel (cat_Rel \), + Ml_Rel (cat_Rel \) a, + Mr_Rel (cat_Rel \) a + ]\<^sub>\" + + +text\Components.\ + +lemma dmcat_Rel_components: + shows "dmcat_Rel \ a\DMcat\ = cat_Rel \" + and "dmcat_Rel \ a\DMdag\ = \\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \" + and "dmcat_Rel \ a\DMcf\ = cf_prod_2_Rel (cat_Rel \)" + and "dmcat_Rel \ a\DMe\ = set {a}" + and "dmcat_Rel \ a\DM\\ = M\_Rel (cat_Rel \)" + and "dmcat_Rel \ a\DMl\ = Ml_Rel (cat_Rel \) a" + and "dmcat_Rel \ a\DMr\ = Mr_Rel (cat_Rel \) a" + unfolding dmcat_Rel_def dmcat_field_simps by (simp_all add: nat_omega_simps) + + +text\Slicing.\ + +lemma dmcat_dagcat_dmcat_Rel: "dmcat_dagcat (dmcat_Rel \ a) = dagcat_Rel \" +proof(rule vsv_eqI) + have dom_lhs: "\\<^sub>\ (dmcat_dagcat (dmcat_Rel \ a)) = 2\<^sub>\" + unfolding dmcat_dagcat_def by (simp add: nat_omega_simps) + have dom_rhs: "\\<^sub>\ (dagcat_Rel \) = 2\<^sub>\" + unfolding dagcat_Rel_def by (simp add: nat_omega_simps) + show "\\<^sub>\ (dmcat_dagcat (dmcat_Rel \ a)) = \\<^sub>\ (dagcat_Rel \)" + unfolding dom_lhs dom_rhs by simp + show "A \\<^sub>\ \\<^sub>\ (dmcat_dagcat (dmcat_Rel \ a)) \ + dmcat_dagcat (dmcat_Rel \ a)\A\ = dagcat_Rel \\A\" + for A + by + ( + unfold dom_lhs, + elim_in_numeral, + unfold dmcat_dagcat_def dmcat_field_simps dmcat_Rel_def dagcat_Rel_def + ) + (auto simp: nat_omega_simps) +qed (auto simp: dmcat_dagcat_def dagcat_Rel_def) + +lemma dmcat_mcat_dmcat_Rel: "dmcat_mcat (dmcat_Rel \ a) = mcat_Rel \ a" +proof(rule vsv_eqI) + have dom_lhs: "\\<^sub>\ (dmcat_mcat (dmcat_Rel \ a)) = 6\<^sub>\" + unfolding dmcat_mcat_def by (simp add: nat_omega_simps) + have dom_rhs: "\\<^sub>\ (mcat_Rel \ a) = 6\<^sub>\" + unfolding mcat_Rel_def by (simp add: nat_omega_simps) + show "\\<^sub>\ (dmcat_mcat (dmcat_Rel \ a)) = \\<^sub>\ (mcat_Rel \ a)" + unfolding dom_lhs dom_rhs by simp + show "A \\<^sub>\ \\<^sub>\ (dmcat_mcat (dmcat_Rel \ a)) \ + dmcat_mcat (dmcat_Rel \ a)\A\ = mcat_Rel \ a\A\" + for A + by + ( + unfold dom_lhs, + elim_in_numeral, + unfold dmcat_mcat_def dmcat_field_simps dmcat_Rel_def mcat_Rel_def + ) + (auto simp: nat_omega_simps) +qed (auto simp: dmcat_mcat_def mcat_Rel_def) + + +subsubsection\\Rel\ is a dagger monoidal category\ + +lemma (in \) dagger_monoidal_category_dmcat_Rel: + assumes "A \\<^sub>\ cat_Rel \\Obj\" + shows "dagger_monoidal_category \ (dmcat_Rel \ A)" +proof- + interpret Rel: category \ \cat_Rel \\ by (rule category_cat_Rel) + interpret dag_Rel: is_iso_functor \ \op_cat (cat_Rel \)\ \cat_Rel \\ \\\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ + by (rule cf_dag_Rel_is_iso_functor) + show ?thesis + proof(rule dagger_monoidal_categoryI) + show "\ \" by auto + show "vfsequence (dmcat_Rel \ A)" unfolding dmcat_Rel_def by simp + show "vcard (dmcat_Rel \ A) = 7\<^sub>\" + unfolding dmcat_Rel_def by (simp add: nat_omega_simps) + show "dagger_category \ (dmcat_dagcat (dmcat_Rel \ A))" + unfolding dmcat_dagcat_dmcat_Rel by (rule dagger_category_dagcat_Rel) + show "monoidal_category \ (dmcat_mcat (dmcat_Rel \ A))" + unfolding dmcat_mcat_dmcat_Rel by (intro monoidal_category_mcat_Rel assms) + show + "\\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\g \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>dmcat_Rel \ A\DMcf\\<^esub> f\ = + \\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\g\ \\<^sub>H\<^sub>M\<^sub>.\<^sub>A\<^bsub>dmcat_Rel \ A\DMcf\\<^esub> + \\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\f\" + if "g : c \\<^bsub>dmcat_Rel \ A\DMcat\\<^esub> d" and "f : a \\<^bsub>dmcat_Rel \ A\DMcat\\<^esub> b" + for c d g a b f + using that + unfolding dmcat_Rel_components + by + ( + cs_concl + cs_simp: cf_dag_Rel_ArrMap_app_prod_2_Rel cat_cs_simps cat_op_simps + cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros + ) + show + "\\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\dmcat_Rel \ A\DM\\\NTMap\\BCD\\ = + (dmcat_Rel \ A\DM\\\NTMap\\BCD\)\\<^sub>C\<^bsub>dmcat_Rel \ A\DMcat\\<^esub>" + if "BCD \\<^sub>\ (dmcat_Rel \ A\DMcat\^\<^sub>C\<^sub>3)\Obj\" for BCD + proof- + from that obtain B C D + where BCD_def: "BCD = [B, C, D]\<^sub>\" + and B: "B \\<^sub>\ cat_Rel \\Obj\" + and C: "C \\<^sub>\ cat_Rel \\Obj\" + and D: "D \\<^sub>\ cat_Rel \\Obj\" + unfolding dmcat_Rel_components + by + ( + elim cat_prod_3_ObjE + [ + unfolded dmcat_Rel_components, + OF Rel.category_axioms Rel.category_axioms Rel.category_axioms + ] + ) + from B C D show ?thesis + unfolding dmcat_Rel_components BCD_def + by + ( + cs_concl + cs_simp: cat_Rel_cs_simps cat_cs_simps + cs_intro: + cat_Rel_is_arrD + cat_cs_intros + cat_Rel_par_set_cs_intros + cat_prod_cs_intros + ) + qed + show + "\\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\dmcat_Rel \ A\DMl\\NTMap\\B\\ = + (dmcat_Rel \ A\DMl\\NTMap\\B\)\\<^sub>C\<^bsub>dmcat_Rel \ A\DMcat\\<^esub>" + if "B \\<^sub>\ dmcat_Rel \ A\DMcat\\Obj\" for B + using assms that + unfolding dmcat_Rel_components + by + ( + cs_concl + cs_simp: cat_Rel_cs_simps + cs_intro: cat_Rel_is_arrD cat_cs_intros cat_arrow_cs_intros + )+ + show + "\\<^sub>M\<^sub>C (dmcat_Rel \ A)\ArrMap\\dmcat_Rel \ A\DMr\\NTMap\\B\\ = + (dmcat_Rel \ A\DMr\\NTMap\\B\)\\<^sub>C\<^bsub>dmcat_Rel \ A\DMcat\\<^esub>" + if "B \\<^sub>\ dmcat_Rel \ A\DMcat\\Obj\" for B + using assms that + unfolding dmcat_Rel_components + by + ( + cs_concl + cs_simp: cat_Rel_cs_simps + cs_intro: cat_Rel_is_arrD cat_cs_intros cat_arrow_cs_intros + )+ + qed +qed + text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Elementary_Categories/document/root.bib b/thys/CZH_Elementary_Categories/document/root.bib --- a/thys/CZH_Elementary_Categories/document/root.bib +++ b/thys/CZH_Elementary_Categories/document/root.bib @@ -1,105 +1,116 @@ @book{takeuti_introduction_1971, address = {Heidelberg}, title = {Introduction to {Axiomatic} {Set} {Theory}}, isbn = {0-387-05302-6}, publisher = {Springer-Verlag}, author = {Takeuti, Gaisi and Zaring, Wilson M.}, year = {1971}, } @book{etingof_tensor_2015, address = {Providence}, series = {Mathematical {Surveys} and {Monographs}}, title = {Tensor {Categories}}, isbn = {978-1-4704-2024-6}, number = {205}, publisher = {American Mathematical Society}, author = {Etingof, Pavel I. and Gelaki, Shlomo and Nikshych, Dmitri and Ostrik, Victor}, year = {2015}, keywords = {\$K\$-theory -- Higher algebraic \$K\$-theory -- Symmetric monoidal categories, Algebraic topology, Associative rings and algebras -- Hopf algebras, quantum groups and related topics -- Hopf algebras and their applications, Category theory; homological algebra -- Categories with structure -- Monoidal categories (= multiplicative categories), symmetric monoidal categories, braided categories, Group theory and generalizations -- Linear algebraic groups and related topics -- Quantum groups (quantized function algebras) and their representations, Hopf algebras, Nonassociative rings and algebras -- Lie algebras and Lie superalgebras -- Quantum groups (quantized enveloping algebras) and related deformations, Tensor fields}, } @misc{noauthor_wikipedia_2001, title = {Wikipedia}, url = {https://www.wikipedia.org/}, year = {2001}, } @book{mac_lane_categories_2010, address = {New York}, edition = {2}, series = {Graduate {Texts} in {Mathematics}}, title = {Categories for the {Working} {Mathematician}}, isbn = {978-1-4419-3123-8}, number = {5}, publisher = {Springer}, author = {Mac Lane, Saunders}, year = {2010}, } @article{paulson_zermelo_2019, title = {Zermelo {Fraenkel} {Set} {Theory} in {Higher}-{Order} {Logic}}, journal = {Archive of Formal Proofs}, author = {Paulson, Lawrence C.}, year = {2019}, } @misc{noauthor_nlab_nodate, title = {{nLab}}, url = {https://ncatlab.org/nlab/show/HomePage}, - journal = {nLab}, } @book{bodo_categories_1970, address = {New York}, title = {Categories and {Functors}}, publisher = {Academic Press}, author = {Bodo, Pareigis}, year = {1970}, } @inproceedings{feferman_set-theoretical_1969, address = {Heidelberg}, series = {Lecture {Notes} in {Mathematics}}, title = {Set-{Theoretical} {Foundations} of {Category} {Theory}}, booktitle = {Reports of the {Midwest} {Category} {Seminar} {III}}, publisher = {Springer}, author = {Feferman, Solomon and Kreisel, Georg}, editor = {Barr, M. and Berthiaume, P. and Day, B. J. and Duskin, J. and Feferman, S. and Kelly, G. M. and Mac Lane, S. and Tierney, M. and Walters, R. F. C.}, year = {1969}, keywords = {Abelian Group, Category Theory, Closure Condition, Free Variable, Mathematical Practice}, pages = {201--247}, } @incollection{barkaoui_partizan_2006, address = {Berlin}, title = {Partizan {Games} in {Isabelle}/{HOLZF}}, volume = {4281}, isbn = {978-3-540-48815-6}, booktitle = {{ICTAC} 2006}, publisher = {Springer}, author = {Obua, Steven}, editor = {Barkaoui, Kamel and Cavalcanti, Ana and Cerone, Antonio}, year = {2006}, pages = {272--286}, } @article{paulson_natural_1986, title = {Natural {Deduction} as {Higher}-{Order} {Resolution}}, volume = {3}, number = {3}, journal = {The Journal of Logic Programming}, author = {Paulson, Lawrence C.}, year = {1986}, pages = {237--258}, } @book{riehl_category_2016, title = {Category {Theory} in {Context}}, publisher = {Emily Riehl}, author = {Riehl, Emily}, year = {2016}, keywords = {Mathematics / Logic}, } @misc{haftmann_sketch-and-explore_2021, title = {Sketch-and-{Explore}}, url = {https://isabelle.in.tum.de/library/HOL/HOL-ex/Sketch_and_Explore.html}, author = {Haftmann, Florian}, year = {2021}, } +@incollection{coecke_survey_2010, + address = {Heidelberg}, + title = {A {Survey} of {Graphical} {Languages} for {Monoidal} {Categories}}, + volume = {813}, + isbn = {978-3-642-12820-2}, + booktitle = {New {Structures} for {Physics}}, + publisher = {Springer}, + author = {Selinger, Peter}, + editor = {Coecke, Bob}, + year = {2010}, + pages = {289--355}, +} @article{milehins_category_2021, title = {Category {Theory} for {ZFC} in {HOL} {I}: {Foundations}: {Design} {Patterns}, {Set} {Theory}, {Digraphs}, {Semicategories}}, journal = {Archive of Formal Proofs}, author = {Milehins, Mihails}, year = {2021}, } diff --git a/thys/CZH_Foundations/czh_digraphs/CZH_DG_Rel.thy b/thys/CZH_Foundations/czh_digraphs/CZH_DG_Rel.thy --- a/thys/CZH_Foundations/czh_digraphs/CZH_DG_Rel.thy +++ b/thys/CZH_Foundations/czh_digraphs/CZH_DG_Rel.thy @@ -1,875 +1,938 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\Rel\ as a digraph\ theory CZH_DG_Rel imports CZH_DG_Small_DGHM begin subsection\Background\ text\ \Rel\ is usually defined as a category of sets and binary relations (e.g., see Chapter I-7 in \cite{mac_lane_categories_2010}). However, there is little that can prevent one from exposing \Rel\ as a digraph and provide additional structure gradually in subsequent installments of this work. Thus, in this section, \\\-\Rel\ is defined as a digraph of sets and binary relations in \V\<^sub>\\. \ named_theorems dg_Rel_shared_cs_simps named_theorems dg_Rel_shared_cs_intros named_theorems dg_Rel_cs_simps named_theorems dg_Rel_cs_intros subsection\Canonical arrow for \<^typ>\V\\ named_theorems arr_field_simps definition ArrVal :: V where [arr_field_simps]: "ArrVal = 0" definition ArrDom :: V where [arr_field_simps]: "ArrDom = 1\<^sub>\" definition ArrCod :: V where [arr_field_simps]: "ArrCod = 2\<^sub>\" lemma ArrVal_eq_helper: assumes "f = g" shows "f\ArrVal\\a\ = g\ArrVal\\a\" using assms by simp subsection\Arrow for \Rel\\ subsubsection\Definition and elementary properties\ locale arr_Rel = \ \ + vfsequence T + ArrVal: vbrelation \T\ArrVal\\ for \ T + assumes arr_Rel_length[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "vcard T = 3\<^sub>\" and arr_Rel_ArrVal_vdomain: "\\<^sub>\ (T\ArrVal\) \\<^sub>\ T\ArrDom\" and arr_Rel_ArrVal_vrange: "\\<^sub>\ (T\ArrVal\) \\<^sub>\ T\ArrCod\" and arr_Rel_ArrDom_in_Vset: "T\ArrDom\ \\<^sub>\ Vset \" and arr_Rel_ArrCod_in_Vset: "T\ArrCod\ \\<^sub>\ Vset \" lemmas [dg_Rel_cs_simps] = arr_Rel.arr_Rel_length text\Components.\ lemma arr_Rel_components[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: shows "[f, A, B]\<^sub>\\ArrVal\ = f" and "[f, A, B]\<^sub>\\ArrDom\ = A" and "[f, A, B]\<^sub>\\ArrCod\ = B" unfolding arr_field_simps by (simp_all add: nat_omega_simps) text\Rules.\ mk_ide rf arr_Rel_def[unfolded arr_Rel_axioms_def] |intro arr_RelI| |dest arr_RelD[dest]| |elim arr_RelE[elim!]| lemma (in \) arr_Rel_vfsequenceI: assumes "vbrelation r" and "\\<^sub>\ r \\<^sub>\ a" and "\\<^sub>\ r \\<^sub>\ b" and "a \\<^sub>\ Vset \" and "b \\<^sub>\ Vset \" shows "arr_Rel \ [r, a, b]\<^sub>\" by (intro arr_RelI) (insert assms, auto simp: nat_omega_simps arr_Rel_components) text\Elementary properties.\ lemma arr_Rel_eqI: assumes "arr_Rel \ S" and "arr_Rel \ T" and "S\ArrVal\ = T\ArrVal\" and "S\ArrDom\ = T\ArrDom\" and "S\ArrCod\ = T\ArrCod\" shows "S = T" proof- interpret S: arr_Rel \ S by (rule assms(1)) interpret T: arr_Rel \ T by (rule assms(2)) show ?thesis proof(rule vsv_eqI) show "\\<^sub>\ S = \\<^sub>\ T" by (simp add: S.vfsequence_vdomain T.vfsequence_vdomain dg_Rel_cs_simps) have dom_lhs: "\\<^sub>\ S = 3\<^sub>\" by (simp add: S.vfsequence_vdomain dg_Rel_cs_simps) show "a \\<^sub>\ \\<^sub>\ S \ S\a\ = T\a\" for a by (unfold dom_lhs, elim_in_numeral, insert assms) (auto simp: arr_field_simps) qed auto qed lemma (in arr_Rel) arr_Rel_def: "T = [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\" proof(rule vsv_eqI) have dom_lhs: "\\<^sub>\ T = 3\<^sub>\" by (simp add: vfsequence_vdomain dg_Rel_cs_simps) have dom_rhs: "\\<^sub>\ [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\ = 3\<^sub>\" by (simp add: nat_omega_simps) then show "\\<^sub>\ T = \\<^sub>\ [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ T \ T\a\ = [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\\a\" for a unfolding dom_lhs by elim_in_numeral (simp_all add: arr_field_simps nat_omega_simps) qed (auto simp: vsv_axioms) text\Size.\ lemma (in arr_Rel) arr_Rel_ArrVal_in_Vset: "T\ArrVal\ \\<^sub>\ Vset \" proof- from arr_Rel_ArrVal_vdomain arr_Rel_ArrDom_in_Vset have "\\<^sub>\ (T\ArrVal\) \\<^sub>\ Vset \" by auto moreover from arr_Rel_ArrVal_vrange arr_Rel_ArrCod_in_Vset have "\\<^sub>\ (T\ArrVal\) \\<^sub>\ Vset \" by auto ultimately show "T\ArrVal\ \\<^sub>\ Vset \" by (simp add: ArrVal.vbrelation_Limit_in_VsetI) qed lemma (in arr_Rel) arr_Rel_in_Vset: "T \\<^sub>\ Vset \" proof- note [dg_Rel_cs_intros] = arr_Rel_ArrVal_in_Vset arr_Rel_ArrDom_in_Vset arr_Rel_ArrCod_in_Vset show ?thesis by (subst arr_Rel_def) (cs_concl cs_intro: dg_Rel_cs_intros V_cs_intros) qed lemma small_arr_Rel[simp]: "small {T. arr_Rel \ T}" by (rule down[of _ \Vset \\]) (auto intro!: arr_Rel.arr_Rel_in_Vset) text\Other elementary properties.\ lemma set_Collect_arr_Rel[simp]: "x \\<^sub>\ set (Collect (arr_Rel \)) \ arr_Rel \ x" by auto lemma (in arr_Rel) arr_Rel_ArrVal_vsubset_ArrDom_ArrCod: "T\ArrVal\ \\<^sub>\ T\ArrDom\ \\<^sub>\ T\ArrCod\" proof fix ab assume "ab \\<^sub>\ T\ArrVal\" then obtain a b where "ab = \a, b\" and "a \\<^sub>\ \\<^sub>\ (T\ArrVal\)" and "b \\<^sub>\ \\<^sub>\ (T\ArrVal\)" by (blast elim: ArrVal.vbrelation_vinE) with arr_Rel_ArrVal_vdomain arr_Rel_ArrVal_vrange show "ab \\<^sub>\ T\ArrDom\ \\<^sub>\ T\ArrCod\" by auto qed subsubsection\Composition\ text\See Chapter I-7 in \cite{mac_lane_categories_2010}.\ definition comp_Rel :: "V \ V \ V" (infixl \\\<^sub>R\<^sub>e\<^sub>l\ 55) where "comp_Rel S T = [S\ArrVal\ \\<^sub>\ T\ArrVal\, T\ArrDom\, S\ArrCod\]\<^sub>\" text\Components.\ lemma comp_Rel_components: shows "(S \\<^sub>R\<^sub>e\<^sub>l T)\ArrVal\ = S\ArrVal\ \\<^sub>\ T\ArrVal\" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "(S \\<^sub>R\<^sub>e\<^sub>l T)\ArrDom\ = T\ArrDom\" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "(S \\<^sub>R\<^sub>e\<^sub>l T)\ArrCod\ = S\ArrCod\" unfolding comp_Rel_def arr_field_simps by (simp_all add: nat_omega_simps) text\Elementary properties.\ lemma comp_Rel_vsv[dg_Rel_shared_cs_intros, dg_Rel_cs_intros]: "vsv (S \\<^sub>R\<^sub>e\<^sub>l T)" unfolding comp_Rel_def by auto lemma arr_Rel_comp_Rel[dg_Rel_cs_intros]: assumes "arr_Rel \ S" and "arr_Rel \ T" shows "arr_Rel \ (S \\<^sub>R\<^sub>e\<^sub>l T)" proof- interpret S: arr_Rel \ S by (rule assms(1)) interpret T: arr_Rel \ T by (rule assms(2)) show ?thesis proof(intro arr_RelI) show "vfsequence (S \\<^sub>R\<^sub>e\<^sub>l T)" unfolding comp_Rel_def by simp show "vcard (S \\<^sub>R\<^sub>e\<^sub>l T) = 3\<^sub>\" unfolding comp_Rel_def by (simp add: nat_omega_simps) from T.arr_Rel_ArrVal_vdomain show "\\<^sub>\ ((S \\<^sub>R\<^sub>e\<^sub>l T)\ArrVal\) \\<^sub>\ (S \\<^sub>R\<^sub>e\<^sub>l T)\ArrDom\" unfolding comp_Rel_components by auto show "\\<^sub>\ ((S \\<^sub>R\<^sub>e\<^sub>l T)\ArrVal\) \\<^sub>\ (S \\<^sub>R\<^sub>e\<^sub>l T)\ArrCod\" unfolding comp_Rel_components proof(intro vsubsetI) fix z assume "z \\<^sub>\ \\<^sub>\ (S\ArrVal\ \\<^sub>\ T\ArrVal\)" then obtain x y where "\y, z\ \\<^sub>\ S\ArrVal\" and "\x, y\ \\<^sub>\ T\ArrVal\" by (meson vcomp_obtain_middle vrange_iff_vdomain) with S.arr_Rel_ArrVal_vrange show "z \\<^sub>\ S\ArrCod\" by auto qed qed ( auto simp: comp_Rel_components T.arr_Rel_ArrDom_in_Vset S.arr_Rel_ArrCod_in_Vset ) qed lemma arr_Rel_comp_Rel_assoc[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "(H \\<^sub>R\<^sub>e\<^sub>l G) \\<^sub>R\<^sub>e\<^sub>l F = H \\<^sub>R\<^sub>e\<^sub>l (G \\<^sub>R\<^sub>e\<^sub>l F)" by (simp add: comp_Rel_def vcomp_assoc arr_field_simps nat_omega_simps) subsubsection\Inclusion arrow\ text\ The definition of the inclusion arrow is based on the concept of the inclusion map, e.g., see \cite{noauthor_wikipedia_2001}\footnote{ \url{https://en.wikipedia.org/wiki/Inclusion_map} }\ definition "incl_Rel A B = [vid_on A, A, B]\<^sub>\" text\Components.\ lemma incl_Rel_components: shows "incl_Rel A B\ArrVal\ = vid_on A" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "incl_Rel A B\ArrDom\ = A" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "incl_Rel A B\ArrCod\ = B" unfolding incl_Rel_def arr_field_simps by (simp_all add: nat_omega_simps) text\Arrow value.\ lemma incl_Rel_ArrVal_vsv[dg_Rel_shared_cs_intros, dg_Rel_cs_intros]: "vsv (incl_Rel A B\ArrVal\)" unfolding incl_Rel_components by simp lemma incl_Rel_ArrVal_vdomain[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "\\<^sub>\ (incl_Rel A B\ArrVal\) = A" unfolding incl_Rel_components by simp lemma incl_Rel_ArrVal_app[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: assumes "a \\<^sub>\ A" shows "incl_Rel A B\ArrVal\\a\ = a" using assms unfolding incl_Rel_components by simp text\Elementary properties.\ lemma incl_Rel_vfsequence[dg_Rel_shared_cs_intros, dg_Rel_cs_intros]: "vfsequence (incl_Rel A B)" unfolding incl_Rel_def by simp lemma incl_Rel_vcard[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "vcard (incl_Rel A B) = 3\<^sub>\" unfolding incl_Rel_def incl_Rel_def by (simp add: nat_omega_simps) lemma (in \) arr_Rel_incl_RelI: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "A \\<^sub>\ B" shows "arr_Rel \ (incl_Rel A B)" proof(intro arr_RelI) show "vfsequence (incl_Rel A B)" unfolding incl_Rel_def by simp show "vcard (incl_Rel A B) = 3\<^sub>\" unfolding incl_Rel_def by (simp add: nat_omega_simps) qed (auto simp: incl_Rel_components assms) subsubsection\Identity\ text\See Chapter I-7 in \cite{mac_lane_categories_2010}.\ definition id_Rel :: "V \ V" where "id_Rel A = incl_Rel A A" text\Components.\ lemma id_Rel_components: shows "id_Rel A\ArrVal\ = vid_on A" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "id_Rel A\ArrDom\ = A" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "id_Rel A\ArrCod\ = A" unfolding id_Rel_def incl_Rel_components by simp_all text\Elementary properties.\ lemma id_Rel_vfsequence[dg_Rel_shared_cs_intros, dg_Rel_cs_intros]: "vfsequence (id_Rel A)" unfolding id_Rel_def by (simp add: dg_Rel_cs_intros) lemma id_Rel_vcard[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "vcard (id_Rel A) = 3\<^sub>\" unfolding id_Rel_def by (simp add: dg_Rel_cs_simps) lemma (in \) arr_Rel_id_RelI: assumes "A \\<^sub>\ Vset \" shows "arr_Rel \ (id_Rel A)" by (intro arr_RelI) (auto simp: id_Rel_components(1) assms dg_Rel_cs_intros dg_Rel_cs_simps) lemma id_Rel_ArrVal_app[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: assumes "a \\<^sub>\ A" shows "id_Rel A\ArrVal\\a\ = a" using assms unfolding id_Rel_components by simp lemma arr_Rel_comp_Rel_id_Rel_left[dg_Rel_cs_simps]: assumes "arr_Rel \ F" and "F\ArrCod\ = A" shows "id_Rel A \\<^sub>R\<^sub>e\<^sub>l F = F" proof(rule arr_Rel_eqI [of \]) interpret F: arr_Rel \ F by (rule assms(1)) from assms(2) have "A \\<^sub>\ Vset \" by (auto intro: F.arr_Rel_ArrCod_in_Vset) with assms(1) show "arr_Rel \ (id_Rel A \\<^sub>R\<^sub>e\<^sub>l F)" by (blast intro: F.arr_Rel_id_RelI intro!: arr_Rel_comp_Rel) from assms(2) F.arr_Rel_ArrVal_vrange show "(id_Rel A \\<^sub>R\<^sub>e\<^sub>l F)\ArrVal\ = F\ArrVal\" unfolding comp_Rel_components id_Rel_components by auto qed ( use assms(2) in \auto simp: assms(1) comp_Rel_components id_Rel_components\ ) lemma arr_Rel_comp_Rel_id_Rel_right[dg_Rel_cs_simps]: assumes "arr_Rel \ F" and "F\ArrDom\ = A" shows "F \\<^sub>R\<^sub>e\<^sub>l id_Rel A = F" proof(rule arr_Rel_eqI[of \]) interpret F: arr_Rel \ F by (rule assms(1)) from assms(2) have "A \\<^sub>\ Vset \" by (auto intro: F.arr_Rel_ArrDom_in_Vset) with assms(1) show "arr_Rel \ (F \\<^sub>R\<^sub>e\<^sub>l id_Rel A)" by (blast intro: F.arr_Rel_id_RelI intro!: arr_Rel_comp_Rel) show "arr_Rel \ F" by (simp add: assms(1)) from assms(2) F.arr_Rel_ArrVal_vdomain show "(F \\<^sub>R\<^sub>e\<^sub>l id_Rel A)\ArrVal\ = F\ArrVal\" unfolding comp_Rel_components id_Rel_components by auto qed (use assms(2) in \auto simp: comp_Rel_components id_Rel_components\) subsubsection\Converse\ text\ As mentioned in Chapter I-7 in \cite{mac_lane_categories_2010}, the category \Rel\ is usually equipped with an additional structure that is the operation of taking a converse of a relation. The operation is meant to be used almost exclusively as part of the dagger functor for \Rel\. \ definition converse_Rel :: "V \ V" ("(_\\<^sub>R\<^sub>e\<^sub>l)" [1000] 999) where "converse_Rel T = [(T\ArrVal\)\\<^sub>\, T\ArrCod\, T\ArrDom\]\<^sub>\" lemma converse_Rel_components: shows "T\\<^sub>R\<^sub>e\<^sub>l\ArrVal\ = (T\ArrVal\)\\<^sub>\" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "T\\<^sub>R\<^sub>e\<^sub>l\ArrDom\ = T\ArrCod\" and [dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "T\\<^sub>R\<^sub>e\<^sub>l\ArrCod\ = T\ArrDom\" unfolding converse_Rel_def arr_field_simps by (simp_all add: nat_omega_simps) text\Elementary properties.\ lemma (in arr_Rel) arr_Rel_converse_Rel: "arr_Rel \ (T\\<^sub>R\<^sub>e\<^sub>l)" proof(rule arr_RelI, unfold converse_Rel_components) show "vfsequence (T\\<^sub>R\<^sub>e\<^sub>l)" unfolding converse_Rel_def by simp show "vcard (T\\<^sub>R\<^sub>e\<^sub>l) = 3\<^sub>\" unfolding converse_Rel_def by (simp add: nat_omega_simps) qed ( auto simp: converse_Rel_components(1) arr_Rel_ArrDom_in_Vset arr_Rel_ArrCod_in_Vset arr_Rel_ArrVal_vdomain arr_Rel_ArrVal_vrange ) lemmas [dg_Rel_cs_intros] = arr_Rel.arr_Rel_converse_Rel lemma (in arr_Rel) arr_Rel_converse_Rel_converse_Rel[dg_Rel_shared_cs_simps, dg_Rel_cs_simps]: "(T\\<^sub>R\<^sub>e\<^sub>l)\\<^sub>R\<^sub>e\<^sub>l = T" proof(rule arr_Rel_eqI) from arr_Rel_axioms show "arr_Rel \ ((T\\<^sub>R\<^sub>e\<^sub>l)\\<^sub>R\<^sub>e\<^sub>l)" by (cs_intro_step dg_Rel_cs_intros)+ qed (simp_all add: arr_Rel_axioms converse_Rel_components) lemmas [dg_Rel_cs_simps] = arr_Rel.arr_Rel_converse_Rel_converse_Rel lemma arr_Rel_converse_Rel_eq_iff[dg_Rel_cs_simps]: assumes "arr_Rel \ F" and "arr_Rel \ G" shows "F\\<^sub>R\<^sub>e\<^sub>l = G\\<^sub>R\<^sub>e\<^sub>l \ F = G" proof(rule iffI) show "F\\<^sub>R\<^sub>e\<^sub>l = G\\<^sub>R\<^sub>e\<^sub>l \ F = G" by (metis arr_Rel.arr_Rel_converse_Rel_converse_Rel assms) qed simp lemma arr_Rel_converse_Rel_comp_Rel[dg_Rel_cs_simps]: assumes "arr_Rel \ G" and "arr_Rel \ F" shows "(F \\<^sub>R\<^sub>e\<^sub>l G)\\<^sub>R\<^sub>e\<^sub>l = G\\<^sub>R\<^sub>e\<^sub>l \\<^sub>R\<^sub>e\<^sub>l F\\<^sub>R\<^sub>e\<^sub>l" proof(rule arr_Rel_eqI, unfold converse_Rel_components comp_Rel_components) from assms show "arr_Rel \ (G\\<^sub>R\<^sub>e\<^sub>l \\<^sub>R\<^sub>e\<^sub>l F\\<^sub>R\<^sub>e\<^sub>l)" by (cs_concl cs_intro: dg_Rel_cs_intros) from assms show "arr_Rel \ ((F \\<^sub>R\<^sub>e\<^sub>l G)\\<^sub>R\<^sub>e\<^sub>l)" by (cs_intro_step dg_Rel_cs_intros)+ qed (simp_all add: vconverse_vcomp) lemma (in \) arr_Rel_converse_Rel_id_Rel: assumes "c \\<^sub>\ Vset \" shows "arr_Rel \ ((id_Rel c)\\<^sub>R\<^sub>e\<^sub>l)" using assms \_axioms by (cs_concl cs_intro: dg_Rel_cs_intros arr_Rel_id_RelI)+ lemma (in \) arr_Rel_converse_Rel_id_Rel_eq_id_Rel[ dg_Rel_shared_cs_simps, dg_Rel_cs_simps ]: assumes "c \\<^sub>\ Vset \" shows "(id_Rel c)\\<^sub>R\<^sub>e\<^sub>l = id_Rel c" by (rule arr_Rel_eqI[of \], unfold converse_Rel_components id_Rel_components) (simp_all add: assms arr_Rel_id_RelI arr_Rel_converse_Rel_id_Rel) lemmas [dg_Rel_shared_cs_simps, dg_Rel_cs_simps] = \.arr_Rel_converse_Rel_id_Rel_eq_id_Rel lemma arr_Rel_comp_Rel_converse_Rel_left_if_v11[dg_Rel_cs_simps]: assumes "arr_Rel \ T" and "\\<^sub>\ (T\ArrVal\) = A" and "T\ArrDom\ = A" and "v11 (T\ArrVal\)" and "A \\<^sub>\ Vset \" shows "T\\<^sub>R\<^sub>e\<^sub>l \\<^sub>R\<^sub>e\<^sub>l T = id_Rel A" proof- interpret T: arr_Rel \ T by (rule assms(1)) interpret v11: v11 \T\ArrVal\\ by (rule assms(4)) show ?thesis by (rule arr_Rel_eqI[of \]) ( auto simp: converse_Rel_components comp_Rel_components id_Rel_components assms(1,3,5) arr_Rel.arr_Rel_converse_Rel arr_Rel_comp_Rel T.arr_Rel_id_RelI v11.v11_vcomp_vconverse[unfolded assms(2)] ) qed lemma arr_Rel_comp_Rel_converse_Rel_right_if_v11[dg_Rel_cs_simps]: assumes "arr_Rel \ T" and "\\<^sub>\ (T\ArrVal\) = A" and "T\ArrCod\ = A" and "v11 (T\ArrVal\)" and "A \\<^sub>\ Vset \" shows "T \\<^sub>R\<^sub>e\<^sub>l T\\<^sub>R\<^sub>e\<^sub>l = id_Rel A" proof- interpret T: arr_Rel \ T by (rule assms(1)) interpret v11: v11 \T\ArrVal\\ by (rule assms(4)) show ?thesis by (rule arr_Rel_eqI[of \]) ( auto simp: assms(1,3,5) comp_Rel_components converse_Rel_components id_Rel_components v11.v11_vcomp_vconverse'[unfolded assms(2)] T.arr_Rel_id_RelI arr_Rel.arr_Rel_converse_Rel arr_Rel_comp_Rel ) qed subsection\\Rel\ as a digraph\ subsubsection\Definition and elementary properties\ definition dg_Rel :: "V \ V" where "dg_Rel \ = [ Vset \, set {T. arr_Rel \ T}, (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\), (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\) ]\<^sub>\" text\Components.\ lemma dg_Rel_components: shows "dg_Rel \\Obj\ = Vset \" and "dg_Rel \\Arr\ = set {T. arr_Rel \ T}" and "dg_Rel \\Dom\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\)" and "dg_Rel \\Cod\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\)" unfolding dg_Rel_def dg_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object\ lemma dg_Rel_Obj_iff: "x \\<^sub>\ dg_Rel \\Obj\ \ x \\<^sub>\ Vset \" unfolding dg_Rel_components by auto subsubsection\Arrow\ lemma dg_Rel_Arr_iff[dg_Rel_cs_simps]: "x \\<^sub>\ dg_Rel \\Arr\ \ arr_Rel \ x" unfolding dg_Rel_components by auto subsubsection\Domain\ mk_VLambda dg_Rel_components(3) |vsv dg_Rel_Dom_vsv[dg_Rel_cs_intros]| |vdomain dg_Rel_Dom_vdomain[dg_Rel_cs_simps]| |app dg_Rel_Dom_app[unfolded set_Collect_arr_Rel, dg_Rel_cs_simps]| lemma dg_Rel_Dom_vrange: "\\<^sub>\ (dg_Rel \\Dom\) \\<^sub>\ dg_Rel \\Obj\" unfolding dg_Rel_components by (rule vrange_VLambda_vsubset, unfold set_Collect_arr_Rel) auto subsubsection\Codomain\ mk_VLambda dg_Rel_components(4) |vsv dg_Rel_Cod_vsv[dg_Rel_cs_intros]| |vdomain dg_Rel_Cod_vdomain[dg_Rel_cs_simps]| |app dg_Rel_Cod_app[unfolded set_Collect_arr_Rel, dg_Rel_cs_simps]| lemma dg_Rel_Cod_vrange: "\\<^sub>\ (dg_Rel \\Cod\) \\<^sub>\ dg_Rel \\Obj\" unfolding dg_Rel_components by (rule vrange_VLambda_vsubset, unfold set_Collect_arr_Rel) auto subsubsection\Arrow with a domain and a codomain\ text\Rules.\ lemma dg_Rel_is_arrI[dg_Rel_cs_intros]: assumes "arr_Rel \ S" and "S\ArrDom\ = A" and "S\ArrCod\ = B" shows "S : A \\<^bsub>dg_Rel \\<^esub> B" using assms by (intro is_arrI, unfold dg_Rel_components) simp_all lemma dg_Rel_is_arrD: assumes "S : A \\<^bsub>dg_Rel \\<^esub> B" shows "arr_Rel \ S" and [dg_cs_simps]: "S\ArrDom\ = A" and [dg_cs_simps]: "S\ArrCod\ = B" using is_arrD[OF assms] unfolding dg_Rel_components by simp_all lemma dg_Rel_is_arrE: assumes "S : A \\<^bsub>dg_Rel \\<^esub> B" obtains "arr_Rel \ S" and "S\ArrDom\ = A" and "S\ArrCod\ = B" using is_arrD[OF assms] unfolding dg_Rel_components by simp_all text\Elementary properties.\ lemma (in \) dg_Rel_incl_Rel_is_arr: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "A \\<^sub>\ B" shows "incl_Rel A B : A \\<^bsub>dg_Rel \\<^esub> B" proof(rule dg_Rel_is_arrI) show "arr_Rel \ (incl_Rel A B)" by (intro arr_Rel_incl_RelI assms) qed (simp_all add: incl_Rel_components) lemma (in \) dg_Rel_incl_Rel_is_arr'[dg_Rel_cs_intros]: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" and "A \\<^sub>\ B" and "A' = A" and "B' = B" shows "incl_Rel A B : A' \\<^bsub>dg_Rel \\<^esub> B'" using assms(1-3) unfolding assms(4,5) by (rule dg_Rel_incl_Rel_is_arr) lemmas [dg_Rel_cs_intros] = \.dg_Rel_incl_Rel_is_arr' +lemma dg_Rel_is_arr_ArrValE: + assumes "T : A \\<^bsub>dg_Rel \\<^esub> B" and "ab \\<^sub>\ T\ArrVal\" + obtains a b + where "ab = \a, b\" and "a \\<^sub>\ \\<^sub>\ (T\ArrVal\)" and "b \\<^sub>\ \\<^sub>\ (T\ArrVal\)" +proof- + note T = dg_Rel_is_arrD[OF assms(1)] + then interpret T: arr_Rel \ T + rewrites "T\ArrDom\ = A" and "T\ArrCod\ = B" + by simp_all + from assms(2) obtain a b + where "ab = \a, b\" and "a \\<^sub>\ \\<^sub>\ (T\ArrVal\)" and "b \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + by (blast elim: T.ArrVal.vbrelation_vinE) + with that show ?thesis by simp +qed + subsubsection\\Rel\ is a digraph\ lemma (in \) dg_Rel_Hom_vifunion_in_Vset: assumes "X \\<^sub>\ Vset \" and "Y \\<^sub>\ Vset \" shows "(\\<^sub>\A\\<^sub>\X. \\<^sub>\B\\<^sub>\Y. Hom (dg_Rel \) A B) \\<^sub>\ Vset \" proof- define Q where "Q i = (if i = 0 then VPow (\\<^sub>\X \\<^sub>\ \\<^sub>\Y) else if i = 1\<^sub>\ then X else Y)" for i have "{[r, A, B]\<^sub>\ |r A B. r \\<^sub>\ \\<^sub>\X \\<^sub>\ \\<^sub>\Y \ A \\<^sub>\ X \ B \\<^sub>\ Y} \ elts (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" proof(intro subsetI, unfold mem_Collect_eq, elim exE conjE) fix F r A B assume prems: "F = [r, A, B]\<^sub>\" "r \\<^sub>\ \\<^sub>\X \\<^sub>\ \\<^sub>\Y" "A \\<^sub>\ X" "B \\<^sub>\ Y" 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: three prems(1) 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 then show "F\i\ \\<^sub>\ Q i" by cases (auto simp: Q_def prems nat_omega_simps) qed (auto simp: prems(1)) qed moreover then have small[simp]: "small {[r, A, B]\<^sub>\ | r A B. r \\<^sub>\\\<^sub>\X \\<^sub>\ \\<^sub>\Y \ A \\<^sub>\ X \ B \\<^sub>\ Y}" by (rule down) ultimately have "set {[r, A, B]\<^sub>\ |r A B. r \\<^sub>\ \\<^sub>\X \\<^sub>\ \\<^sub>\Y \ A \\<^sub>\ X \ B \\<^sub>\ Y} \\<^sub>\ (\\<^sub>\i\\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}. Q i)" by auto 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 (auto simp: three[symmetric] intro!: Axiom_of_Infinity) from assms(1,2) have "VPow (\\<^sub>\X \\<^sub>\ \\<^sub>\Y) \\<^sub>\ Vset \" by (intro Limit_VPow_in_VsetI Limit_vtimes_in_VsetI) auto then show "Q i \\<^sub>\ Vset \" if "i \\<^sub>\ set {0, 1\<^sub>\, 2\<^sub>\}" for i using that assms(1,2) unfolding Q_def by (auto simp: nat_omega_simps) qed auto moreover have "(\\<^sub>\A\\<^sub>\X. \\<^sub>\B\\<^sub>\Y. Hom (dg_Rel \) A B) \\<^sub>\ set {[r, A, B]\<^sub>\ | r A B. r \\<^sub>\\\<^sub>\X \\<^sub>\ \\<^sub>\Y \ A \\<^sub>\ X \ B \\<^sub>\ Y}" proof(rule vsubsetI) fix F assume prems: "F \\<^sub>\ (\\<^sub>\A\\<^sub>\X. \\<^sub>\B\\<^sub>\Y. Hom (dg_Rel \) A B)" then obtain A where A: "A \\<^sub>\ X" and F_b: "F \\<^sub>\ (\\<^sub>\B\\<^sub>\Y. Hom (dg_Rel \) A B)" unfolding vifunion_iff by auto then obtain B where B: "B \\<^sub>\ Y" and F_fba: "F \\<^sub>\ Hom (dg_Rel \) A B" by fastforce then have "F : A \\<^bsub>dg_Rel \\<^esub> B" by simp note F = dg_Rel_is_arrD[OF this] interpret F: arr_Rel \ F rewrites "F\ArrDom\ = A" and "F\ArrCod\ = B" by (intro F)+ show "F \\<^sub>\ set {[r, A, B]\<^sub>\ | r A B. r \\<^sub>\\\<^sub>\X \\<^sub>\ \\<^sub>\Y \ A \\<^sub>\ X \ B \\<^sub>\ Y}" proof(intro in_set_CollectI small exI conjI) from F.arr_Rel_def show "F = [F\ArrVal\, A, B]\<^sub>\" unfolding F(2,3) by simp from A B have "A \\<^sub>\ B \\<^sub>\ \\<^sub>\X \\<^sub>\ \\<^sub>\Y" by auto moreover then have "F\ArrVal\ \\<^sub>\ A \\<^sub>\ B" by (auto simp: F.arr_Rel_ArrVal_vsubset_ArrDom_ArrCod) ultimately show "F\ArrVal\ \\<^sub>\ \\<^sub>\X \\<^sub>\ \\<^sub>\Y" by auto qed (intro A B)+ qed ultimately show "(\\<^sub>\A\\<^sub>\X. \\<^sub>\B\\<^sub>\Y. Hom (dg_Rel \) A B) \\<^sub>\ Vset \" by blast qed lemma (in \) digraph_dg_Rel: "digraph \ (dg_Rel \)" proof(intro digraphI) show "vfsequence (dg_Rel \)" unfolding dg_Rel_def by clarsimp show "vcard (dg_Rel \) = 4\<^sub>\" unfolding dg_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (dg_Rel \\Dom\) \\<^sub>\ dg_Rel \\Obj\" by (simp add: dg_Rel_Dom_vrange) show "\\<^sub>\ (dg_Rel \\Cod\) \\<^sub>\ dg_Rel \\Obj\" by (simp add: dg_Rel_Cod_vrange) qed (auto simp: dg_Rel_components dg_Rel_Hom_vifunion_in_Vset dg_Rel_Dom_vrange) subsection\Canonical dagger for \Rel\\ text\ Dagger categories are exposed explicitly later. In the context of this section, the ``dagger'' is viewed merely as an explicitly defined homomorphism. A definition of a dagger functor, upon which the definition presented in this section is based, can be found in nLab \cite{noauthor_nlab_nodate}\footnote{\url{https://ncatlab.org/nlab/show/Rel})}. This reference also contains the majority of the results that are presented in this subsection. \ subsubsection\Definition and elementary properties\ definition dghm_dag_Rel :: "V \ V" (\\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l\) where "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = [ vid_on (dg_Rel \\Obj\), VLambda (dg_Rel \\Arr\) converse_Rel, op_dg (dg_Rel \), dg_Rel \ ]\<^sub>\" text\Components.\ lemma dghm_dag_Rel_components: shows "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\ = vid_on (dg_Rel \\Obj\)" and "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\ = VLambda (dg_Rel \\Arr\) converse_Rel" and "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomDom\ = op_dg (dg_Rel \)" and "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomCod\ = dg_Rel \" unfolding dghm_dag_Rel_def dghm_field_simps by (simp_all add: nat_omega_simps) subsubsection\Object map\ mk_VLambda dghm_dag_Rel_components(1)[folded VLambda_vid_on] |vsv dghm_dag_Rel_ObjMap_vsv[dg_Rel_cs_intros]| |vdomain dghm_dag_Rel_ObjMap_vdomain[unfolded dg_Rel_components, dg_Rel_cs_simps] | |app dghm_dag_Rel_ObjMap_app[unfolded dg_Rel_components, dg_Rel_cs_simps]| lemma dghm_dag_Rel_ObjMap_vrange[dg_cs_simps]: "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\) = Vset \" unfolding dghm_dag_Rel_components dg_Rel_components by simp subsubsection\Arrow map\ mk_VLambda dghm_dag_Rel_components(2) |vsv dghm_dag_Rel_ArrMap_vsv[dg_Rel_cs_intros]| |vdomain dghm_dag_Rel_ArrMap_vdomain[dg_Rel_cs_simps]| |app dghm_dag_Rel_ArrMap_app[unfolded dg_Rel_cs_simps, dg_Rel_cs_simps]| +lemma (in \) dghm_dag_Rel_ArrMap_app_vdomain[dg_cs_simps]: + assumes "T : A \\<^bsub>dg_Rel \\<^esub> B" + shows "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\) = \\<^sub>\ (T\ArrVal\)" +proof- + from assms interpret T: arr_Rel \ T by (simp add: dg_Rel_is_arrD) + from dg_Rel_is_arrD(1)[OF assms] show ?thesis + by (cs_concl cs_simp: dg_Rel_cs_simps V_cs_simps converse_Rel_components(1)) +qed + +lemmas [dg_cs_simps] = \.dghm_dag_Rel_ArrMap_app_vdomain + +lemma (in \) dghm_dag_Rel_ArrMap_app_vrange[dg_cs_simps]: + assumes "T : A \\<^bsub>dg_Rel \\<^esub> B" + shows "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\) = \\<^sub>\ (T\ArrVal\)" +proof- + from assms interpret T: arr_Rel \ T by (simp add: dg_Rel_is_arrD) + from dg_Rel_is_arrD(1)[OF assms] show ?thesis + by (cs_concl cs_simp: dg_Rel_cs_simps V_cs_simps converse_Rel_components(1)) +qed + +lemmas [dg_cs_simps] = \.dghm_dag_Rel_ArrMap_app_vrange + +lemma (in \) dghm_dag_Rel_ArrMap_app_iff[dg_cs_simps]: + assumes "T : A \\<^bsub>dg_Rel \\<^esub> B" + shows "\a, b\ \\<^sub>\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\ \ \b, a\ \\<^sub>\ T\ArrVal\" +proof- + from assms interpret T: arr_Rel \ T by (simp add: dg_Rel_is_arrD) + note T = dg_Rel_is_arrD[OF assms] + note [dg_Rel_cs_simps] = converse_Rel_components + show ?thesis + proof(intro iffI) + assume prems: "\a, b\ \\<^sub>\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\" + then have a: "a \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\)" + and b: "b \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\)" + by auto + with assms have a: "a \\<^sub>\ \\<^sub>\ (T\ArrVal\)" and b: "b \\<^sub>\ \\<^sub>\ (T\ArrVal\)" + by (simp_all add: dg_cs_simps) + from prems T(1) have "\a, b\ \\<^sub>\ (T\ArrVal\)\\<^sub>\" + by (cs_prems cs_simp: dg_Rel_cs_simps) + then show "\b, a\ \\<^sub>\ T\ArrVal\" by clarsimp + next + assume "\b, a\ \\<^sub>\ T\ArrVal\" + then have "\a, b\ \\<^sub>\ (T\ArrVal\)\\<^sub>\" by auto + with T(1) show "\a, b\ \\<^sub>\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\" + by (cs_concl cs_simp: dg_Rel_cs_simps) + qed +qed + subsubsection\Further properties\ lemma dghm_dag_Rel_ArrMap_vrange[dg_Rel_cs_simps]: "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\) = dg_Rel \\Arr\" proof(intro vsubset_antisym vsubsetI) interpret ArrMap: vsv \\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\ unfolding dghm_dag_Rel_components by simp fix T assume "T \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\)" then obtain S where T_def: "T = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\" and S: "S \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\)" by (blast dest: ArrMap.vrange_atD) from S show "T \\<^sub>\ dg_Rel \\Arr\" by ( simp add: T_def dghm_dag_Rel_components dg_Rel_components arr_Rel.arr_Rel_converse_Rel ) next interpret ArrMap: vsv \\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\ unfolding dghm_dag_Rel_components by simp fix T assume "T \\<^sub>\ dg_Rel \\Arr\" then have "arr_Rel \ T" by (simp add: dg_Rel_components) then have "(T\\<^sub>R\<^sub>e\<^sub>l)\\<^sub>R\<^sub>e\<^sub>l = T" and "arr_Rel \ (T\\<^sub>R\<^sub>e\<^sub>l)" by ( auto simp: arr_Rel.arr_Rel_converse_Rel_converse_Rel arr_Rel.arr_Rel_converse_Rel ) then have "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\<^sub>R\<^sub>e\<^sub>l\ = T" "T\\<^sub>R\<^sub>e\<^sub>l \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\)" by (simp_all add: dg_Rel_components(2) dghm_dag_Rel_components(2)) then show "T \\<^sub>\ \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\)" by blast qed lemma dghm_dag_Rel_ArrMap_app_is_arr: assumes "T : b \\<^bsub>dg_Rel \\<^esub> a" shows "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\ \\<^bsub>dg_Rel \\<^esub> \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\b\" proof(intro is_arrI) from assms have a: "a \\<^sub>\ Vset \" and b: "b \\<^sub>\ Vset \" unfolding dg_Rel_components by (fastforce simp: dg_Rel_components)+ from assms have T: "arr_Rel \ T" by (auto simp: dg_Rel_is_arrD(1)) then show dag_T: "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>\ dg_Rel \\Arr\" by (cs_concl cs_simp: dg_Rel_cs_simps cs_intro: dg_Rel_cs_intros) from a assms T show "dg_Rel \\Dom\\\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\" by (cs_concl cs_simp: dg_cs_simps dg_Rel_cs_simps cs_intro: dg_Rel_cs_intros) from b assms T show "dg_Rel \\Cod\\\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\b\" by (cs_concl cs_simp: dg_cs_simps dg_Rel_cs_simps cs_intro: dg_Rel_cs_intros) qed subsubsection\Canonical dagger for \Rel\ is a digraph isomorphism\ lemma (in \) dghm_dag_Rel_is_iso_dghm: "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_dg (dg_Rel \) \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> dg_Rel \" proof(rule is_iso_dghmI) interpret digraph \ \dg_Rel \\ by (simp add: digraph_dg_Rel) show "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_dg (dg_Rel \) \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_Rel \" proof(rule is_dghmI, unfold dg_op_simps dghm_dag_Rel_components(3,4)) show "vfsequence (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding dghm_dag_Rel_def by (simp add: nat_omega_simps) show "vcard (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = 4\<^sub>\" unfolding dghm_dag_Rel_def by (simp add: nat_omega_simps) fix T a b assume "T : b \\<^bsub>dg_Rel \\<^esub> a" then show "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\ \\<^bsub>dg_Rel \\<^esub> \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\b\" by (rule dghm_dag_Rel_ArrMap_app_is_arr) qed (auto simp: dghm_dag_Rel_components intro: dg_cs_intros dg_op_intros) show "v11 (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\)" proof ( intro vsv.vsv_valeq_v11I, unfold dghm_dag_Rel_ArrMap_vdomain dg_Rel_Arr_iff ) fix S T assume prems: "arr_Rel \ S" "arr_Rel \ T" "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\ = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\" from prems show "S = T" by ( auto simp: dg_Rel_components dg_Rel_cs_simps dghm_dag_Rel_ArrMap_app[OF prems(1)] dghm_dag_Rel_ArrMap_app[OF prems(2)] ) qed (auto intro: dg_Rel_cs_intros) show "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\) = dg_Rel \\Arr\" by (simp add: dg_Rel_cs_simps) qed (simp_all add: dghm_dag_Rel_components) subsubsection\Further properties of the canonical dagger\ lemma (in \) dghm_cn_comp_dghm_dag_Rel_dghm_dag_Rel: "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = dghm_id (dg_Rel \)" proof- interpret digraph \ \dg_Rel \\ by (simp add: digraph_dg_Rel) from dghm_dag_Rel_is_iso_dghm have dag: "\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : dg_Rel \ \<^sub>D\<^sub>G\\\<^bsub>\\<^esub> dg_Rel \" by (simp add: is_iso_dghm_def) show ?thesis proof(rule dghm_eqI) show "(\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\ = dghm_id (dg_Rel \)\ArrMap\" proof(rule vsv_eqI) show "vsv ((\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\)" by (auto simp: dghm_cn_comp_components dghm_dag_Rel_components) fix a assume "a \\<^sub>\ \\<^sub>\ ((\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\)" then have a: "arr_Rel \ a" unfolding dg_Rel_cs_simps dghm_cn_comp_ArrMap_vdomain[OF dag dag] by simp from a dghm_dag_Rel_is_iso_dghm show "(\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>D\<^sub>G\<^sub>H\<^sub>M\ \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\\a\ = dghm_id (dg_Rel \)\ArrMap\\a\" by ( cs_concl cs_simp: dg_Rel_cs_simps dg_cs_simps dg_cn_cs_simps cs_intro: dg_Rel_cs_intros dghm_cs_intros ) qed (simp_all add: dghm_cn_comp_components dghm_id_components dg_Rel_cs_simps) show "dghm_id (dg_Rel \) : dg_Rel \ \\\<^sub>D\<^sub>G\<^bsub>\\<^esub> dg_Rel \" by (simp_all add: digraph.dg_dghm_id_is_dghm digraph_axioms) qed ( auto simp: dghm_cn_comp_is_dghm[OF digraph_axioms dag dag] dghm_cn_comp_components dghm_dag_Rel_components dghm_id_components ) qed text\\newpage\ end \ No newline at end of file diff --git a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Rel.thy b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Rel.thy --- a/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Rel.thy +++ b/thys/CZH_Foundations/czh_semicategories/CZH_SMC_Rel.thy @@ -1,1063 +1,1074 @@ (* Copyright 2021 (C) Mihails Milehins *) section\\Rel\ as a semicategory\ theory CZH_SMC_Rel imports CZH_DG_Rel CZH_SMC_Semifunctor CZH_SMC_Small_Semicategory begin subsection\Background\ text\ The methodology chosen for the exposition of \Rel\ as a semicategory is analogous to the one used in the previous chapter for the exposition of \Rel\ as a digraph. The general references for this section are Chapter I-7 in \cite{mac_lane_categories_2010} and nLab \cite{noauthor_nlab_nodate}\footnote{ \url{https://ncatlab.org/nlab/show/Rel} }. \ named_theorems smc_Rel_cs_simps named_theorems smc_Rel_cs_intros lemmas (in arr_Rel) [smc_Rel_cs_simps] = dg_Rel_shared_cs_simps lemmas [smc_Rel_cs_simps] = dg_Rel_shared_cs_simps arr_Rel.arr_Rel_length arr_Rel_comp_Rel_id_Rel_left arr_Rel_comp_Rel_id_Rel_right arr_Rel.arr_Rel_converse_Rel_converse_Rel arr_Rel_converse_Rel_eq_iff arr_Rel_converse_Rel_comp_Rel arr_Rel_comp_Rel_converse_Rel_left_if_v11 arr_Rel_comp_Rel_converse_Rel_right_if_v11 lemmas [smc_Rel_cs_intros] = dg_Rel_shared_cs_intros arr_Rel_comp_Rel arr_Rel.arr_Rel_converse_Rel subsection\\Rel\ as a semicategory\ subsubsection\Definition and elementary properties\ definition smc_Rel :: "V \ V" where "smc_Rel \ = [ Vset \, set {T. arr_Rel \ T}, (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\), (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\), (\ST\\<^sub>\composable_arrs (dg_Rel \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\) ]\<^sub>\" text\Components.\ lemma smc_Rel_components: shows "smc_Rel \\Obj\ = Vset \" and "smc_Rel \\Arr\ = set {T. arr_Rel \ T}" and "smc_Rel \\Dom\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrDom\)" and "smc_Rel \\Cod\ = (\T\\<^sub>\set {T. arr_Rel \ T}. T\ArrCod\)" and "smc_Rel \\Comp\ = (\ST\\<^sub>\composable_arrs (dg_Rel \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\)" unfolding smc_Rel_def dg_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smc_dg_smc_Rel: "smc_dg (smc_Rel \) = dg_Rel \" proof(rule vsv_eqI) show "vsv (smc_dg (smc_Rel \))" unfolding smc_dg_def by auto show "vsv (dg_Rel \)" unfolding dg_Rel_def by auto have dom_lhs: "\\<^sub>\ (smc_dg (smc_Rel \)) = 4\<^sub>\" unfolding smc_dg_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (dg_Rel \) = 4\<^sub>\" unfolding dg_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (smc_dg (smc_Rel \)) = \\<^sub>\ (dg_Rel \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (smc_dg (smc_Rel \)) \ smc_dg (smc_Rel \)\a\ = dg_Rel \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold smc_dg_def dg_field_simps smc_Rel_def dg_Rel_def ) (auto simp: nat_omega_simps) qed lemmas_with [folded smc_dg_smc_Rel, unfolded slicing_simps]: smc_Rel_Obj_iff = dg_Rel_Obj_iff and smc_Rel_Arr_iff[smc_Rel_cs_simps] = dg_Rel_Arr_iff and smc_Rel_Dom_vsv[smc_Rel_cs_intros] = dg_Rel_Dom_vsv and smc_Rel_Dom_vdomain[smc_Rel_cs_simps] = dg_Rel_Dom_vdomain and smc_Rel_Dom_app[smc_Rel_cs_simps] = dg_Rel_Dom_app and smc_Rel_Dom_vrange = dg_Rel_Dom_vrange and smc_Rel_Cod_vsv[smc_Rel_cs_intros] = dg_Rel_Cod_vsv and smc_Rel_Cod_vdomain[smc_Rel_cs_simps] = dg_Rel_Cod_vdomain and smc_Rel_Cod_app[smc_Rel_cs_simps] = dg_Rel_Cod_app and smc_Rel_Cod_vrange = dg_Rel_Cod_vrange and smc_Rel_is_arrI[smc_Rel_cs_intros] = dg_Rel_is_arrI and smc_Rel_is_arrD = dg_Rel_is_arrD and smc_Rel_is_arrE = dg_Rel_is_arrE lemmas [smc_cs_simps] = smc_Rel_is_arrD(2,3) lemmas_with (in \) [folded smc_dg_smc_Rel, unfolded slicing_simps]: smc_Rel_Hom_vifunion_in_Vset = dg_Rel_Hom_vifunion_in_Vset and smc_Rel_incl_Rel_is_arr = dg_Rel_incl_Rel_is_arr and smc_Rel_incl_Rel_is_arr'[smc_Rel_cs_intros] = dg_Rel_incl_Rel_is_arr' + and smc_Rel_is_arr_ArrValE = dg_Rel_is_arr_ArrValE lemmas [smc_Rel_cs_intros] = \.smc_Rel_incl_Rel_is_arr' subsubsection\Composable arrows\ lemma smc_Rel_composable_arrs_dg_Rel: "composable_arrs (dg_Rel \) = composable_arrs (smc_Rel \)" unfolding composable_arrs_def smc_dg_smc_Rel[symmetric] slicing_simps by simp lemma smc_Rel_Comp: "smc_Rel \\Comp\ = (\ST\\<^sub>\composable_arrs (smc_Rel \). ST\0\ \\<^sub>R\<^sub>e\<^sub>l ST\1\<^sub>\\)" unfolding smc_Rel_components smc_Rel_composable_arrs_dg_Rel .. subsubsection\Composition\ lemma smc_Rel_Comp_app[smc_Rel_cs_simps]: assumes "S : b \\<^bsub>smc_Rel \\<^esub> c" and "T : a \\<^bsub>smc_Rel \\<^esub> b" shows "S \\<^sub>A\<^bsub>smc_Rel \\<^esub> T = S \\<^sub>R\<^sub>e\<^sub>l T" proof- from assms have "[S, T]\<^sub>\ \\<^sub>\ composable_arrs (smc_Rel \)" by (auto intro: smc_cs_intros) then show "S \\<^sub>A\<^bsub>smc_Rel \\<^esub> T = S \\<^sub>R\<^sub>e\<^sub>l T" unfolding smc_Rel_Comp by (simp add: nat_omega_simps) qed lemma smc_Rel_Comp_vdomain: "\\<^sub>\ (smc_Rel \\Comp\) = composable_arrs (smc_Rel \)" unfolding smc_Rel_Comp by simp lemma (in \) smc_CAT_Comp_vrange: "\\<^sub>\ (smc_Rel \\Comp\) \\<^sub>\ set {T. arr_Rel \ T}" proof(rule vsubsetI) interpret digraph \ \smc_dg (smc_Rel \)\ unfolding smc_dg_smc_Rel by (simp add: digraph_dg_Rel) fix R assume "R \\<^sub>\ \\<^sub>\ (smc_Rel \\Comp\)" then obtain ST where R_def: "R = smc_Rel \\Comp\\ST\" and "ST \\<^sub>\ \\<^sub>\ (smc_Rel \\Comp\)" unfolding smc_Rel_components by (auto intro: smc_cs_intros) then obtain S T a b c where "ST = [S, T]\<^sub>\" and S: "S : b \\<^bsub>smc_Rel \\<^esub> c" and T: "T : a \\<^bsub>smc_Rel \\<^esub> b" by (auto simp: smc_Rel_Comp_vdomain) with R_def have R_def': "R = S \\<^sub>A\<^bsub>smc_Rel \\<^esub> T" by simp note S_D = dg_is_arrD(1)[unfolded slicing_simps, OF S] note T_D = dg_is_arrD(1)[unfolded slicing_simps, OF T] from S_D T_D have "arr_Rel \ S" "arr_Rel \ T" by (simp_all add: smc_Rel_components) from this show "R \\<^sub>\ set {T. arr_Rel \ T}" unfolding R_def' smc_Rel_Comp_app[OF S T] by (auto simp: arr_Rel_comp_Rel) qed subsubsection\\Rel\ is a semicategory\ lemma (in \) semicategory_smc_Rel: "semicategory \ (smc_Rel \)" proof(rule semicategoryI, unfold smc_dg_smc_Rel) show "vfsequence (smc_Rel \)" unfolding smc_Rel_def by simp show "vcard (smc_Rel \) = 5\<^sub>\" unfolding smc_Rel_def by (simp add: nat_omega_simps) show "gf \\<^sub>\ \\<^sub>\ (smc_Rel \\Comp\) \ (\g f b c a. gf = [g, f]\<^sub>\ \ g : b \\<^bsub>smc_Rel \\<^esub> c \ f : a \\<^bsub>smc_Rel \\<^esub> b)" for gf unfolding smc_Rel_Comp_vdomain by (auto intro: composable_arrsI) show "g \\<^sub>A\<^bsub>smc_Rel \\<^esub> f : a \\<^bsub>smc_Rel \\<^esub> c" if "g : b \\<^bsub>smc_Rel \\<^esub> c" and "f : a \\<^bsub>smc_Rel \\<^esub> b" for g b c f a proof- from that have "arr_Rel \ g" and "arr_Rel \ f" by (auto simp: smc_Rel_is_arrD(1)) with that show ?thesis by ( cs_concl cs_simp: smc_cs_simps smc_Rel_cs_simps cs_intro: smc_Rel_cs_intros ) qed show "h \\<^sub>A\<^bsub>smc_Rel \\<^esub> g \\<^sub>A\<^bsub>smc_Rel \\<^esub> f = h \\<^sub>A\<^bsub>smc_Rel \\<^esub> (g \\<^sub>A\<^bsub>smc_Rel \\<^esub> f)" if "h : c \\<^bsub>smc_Rel \\<^esub> d" and "g : b \\<^bsub>smc_Rel \\<^esub> c" and "f : a \\<^bsub>smc_Rel \\<^esub> b" for h c d g b f a proof- from that have "arr_Rel \ h" and "arr_Rel \ g" and "arr_Rel \ f" by (auto simp: smc_Rel_is_arrD(1)) with that show ?thesis by ( cs_concl cs_simp: smc_cs_simps smc_Rel_cs_simps cs_intro: smc_Rel_cs_intros ) qed qed (auto simp: digraph_dg_Rel smc_Rel_components) subsection\Canonical dagger for \Rel\\ subsubsection\Definition and elementary properties\ definition smcf_dag_Rel :: "V \ V" (\\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l\) where "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = [ vid_on (smc_Rel \\Obj\), VLambda (smc_Rel \\Arr\) converse_Rel, op_smc (smc_Rel \), smc_Rel \ ]\<^sub>\" text\Components.\ lemma smcf_dag_Rel_components: shows "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\ = vid_on (smc_Rel \\Obj\)" and "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\ = VLambda (smc_Rel \\Arr\) converse_Rel" and "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomDom\ = op_smc (smc_Rel \)" and "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\HomCod\ = smc_Rel \" unfolding smcf_dag_Rel_def dghm_field_simps by (simp_all add: nat_omega_simps) text\Slicing.\ lemma smcf_dghm_smcf_dag_Rel: "smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \" proof(rule vsv_eqI) show "vsv (smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \))" unfolding smcf_dghm_def by auto show "vsv (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding dghm_dag_Rel_def by auto have dom_lhs: "\\<^sub>\ (smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) = 4\<^sub>\" unfolding smcf_dghm_def by (simp add: nat_omega_simps) have dom_rhs: "\\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = 4\<^sub>\" unfolding dghm_dag_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>\ (smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) = \\<^sub>\ (\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding dom_lhs dom_rhs by simp show "a \\<^sub>\ \\<^sub>\ (smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)) \ smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\a\ = \\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\a\" for a by ( unfold dom_lhs, elim_in_numeral, unfold dghm_field_simps[symmetric], unfold smc_dg_smc_Rel slicing_commute[symmetric] smcf_dghm_components dghm_dag_Rel_components smcf_dag_Rel_components dg_Rel_components smc_Rel_components ) simp_all qed lemmas_with [ folded smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel, unfolded slicing_simps ]: smcf_dag_Rel_ObjMap_vsv[smc_Rel_cs_intros] = dghm_dag_Rel_ObjMap_vsv and smcf_dag_Rel_ObjMap_vdomain[smc_Rel_cs_simps] = dghm_dag_Rel_ObjMap_vdomain and smcf_dag_Rel_ObjMap_app[smc_Rel_cs_simps] = dghm_dag_Rel_ObjMap_app and smcf_dag_Rel_ObjMap_vrange[smc_Rel_cs_simps] = dghm_dag_Rel_ObjMap_vrange and smcf_dag_Rel_ArrMap_vsv[smc_Rel_cs_intros] = dghm_dag_Rel_ArrMap_vsv and smcf_dag_Rel_ArrMap_vdomain[smc_Rel_cs_simps] = dghm_dag_Rel_ArrMap_vdomain and smcf_dag_Rel_ArrMap_app[smc_Rel_cs_simps] = dghm_dag_Rel_ArrMap_app and smcf_dag_Rel_ArrMap_vrange[smc_Rel_cs_simps] = dghm_dag_Rel_ArrMap_vrange lemmas_with (in \) [ folded smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel, unfolded slicing_simps ]: smcf_dag_Rel_app_is_arr = dghm_dag_Rel_ArrMap_app_is_arr + and smcf_dag_Rel_ArrMap_app_vdomain[smc_cs_simps] = + dghm_dag_Rel_ArrMap_app_vdomain + and smcf_dag_Rel_ArrMap_app_vrange[smc_cs_simps] = + dghm_dag_Rel_ArrMap_app_vrange + and smcf_dag_Rel_ArrMap_app_iff[smc_cs_simps] = dghm_dag_Rel_ArrMap_app_iff + +lemmas [smc_cs_simps] = + \.smcf_dag_Rel_ArrMap_app_vdomain + \.smcf_dag_Rel_ArrMap_app_vrange + \.smcf_dag_Rel_ArrMap_app_iff subsubsection\Canonical dagger is a contravariant isomorphism of \Rel\\ lemma (in \) smcf_dag_Rel_is_iso_semifunctor: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_smc (smc_Rel \) \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_Rel \" proof(rule is_iso_semifunctorI) interpret dag: is_iso_dghm \ \op_dg (dg_Rel \)\ \dg_Rel \\ \\\<^sub>D\<^sub>G\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ by (rule dghm_dag_Rel_is_iso_dghm) interpret Rel: semicategory \ \smc_Rel \\ by (rule semicategory_smc_Rel) show "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_smc (smc_Rel \) \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Rel \" proof ( rule is_semifunctorI, unfold smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel smc_op_simps slicing_commute[symmetric] smcf_dag_Rel_components(3,4) ) show "vfsequence (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)" unfolding smcf_dag_Rel_def by (simp add: nat_omega_simps) show "vcard (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = 4\<^sub>\" unfolding smcf_dag_Rel_def by (simp add: nat_omega_simps) show "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f \\<^sub>A\<^bsub>smc_Rel \\<^esub> g\ = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\" if "g : c \\<^bsub>smc_Rel \\<^esub> b" and "f : b \\<^bsub>smc_Rel \\<^esub> a" for g b c f a proof- from that have "arr_Rel \ g" and "arr_Rel \ f" by (auto simp: smc_Rel_is_arrD(1)) with that show ?thesis by ( cs_concl cs_simp: smc_cs_simps smc_Rel_cs_simps cs_intro: smc_Rel_cs_intros ) qed qed (auto simp: dg_cs_intros smc_op_intros semicategory_smc_Rel) show "smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) : smc_dg (op_smc (smc_Rel \)) \\\<^sub>D\<^sub>G\<^sub>.\<^sub>i\<^sub>s\<^sub>o\<^bsub>\\<^esub> smc_dg (smc_Rel \)" by ( simp add: smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel smc_op_simps slicing_simps slicing_commute[symmetric] dghm_dag_Rel_is_iso_dghm ) qed subsubsection\Further properties of the canonical dagger\ lemma (in \) smcf_cn_comp_smcf_dag_Rel_smcf_dag_Rel: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ = smcf_id (smc_Rel \)" proof(rule smcf_dghm_eqI) interpret semicategory \ \smc_Rel \\ by (simp add: semicategory_smc_Rel) from smcf_dag_Rel_is_iso_semifunctor have dag: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_smc (smc_Rel \) \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Rel \" by (simp add: is_iso_semifunctor.axioms(1)) from smcf_cn_comp_is_semifunctor[OF semicategory_axioms dag dag] show "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : smc_Rel \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Rel \" . show "smcf_id (smc_Rel \) : smc_Rel \ \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Rel \" by (auto simp: smc_smcf_id_is_semifunctor) show "smcf_dghm (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \) = smcf_dghm (smcf_id (smc_Rel \))" unfolding slicing_simps slicing_commute[symmetric] smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel by (simp add: dghm_cn_comp_dghm_dag_Rel_dghm_dag_Rel) qed simp_all lemma (in \) smcf_dag_Rel_ArrMap_smc_Rel_Comp: assumes "S : b \\<^bsub>smc_Rel \\<^esub> c" and "T : a \\<^bsub>smc_Rel \\<^esub> b" shows "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S \\<^sub>A\<^bsub>smc_Rel \\<^esub> T\ = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\S\" proof- from assms have "arr_Rel \ S" and "arr_Rel \ T" by (auto simp: smc_Rel_is_arrD(1)) with assms show ?thesis by ( cs_concl cs_simp: smc_cs_simps smc_Rel_cs_simps cs_intro: smc_Rel_cs_intros ) qed subsection\Monic arrow and epic arrow\ text\ The conditions for an arrow of \Rel\ to be either monic or epic are outlined in nLab \cite{noauthor_nlab_nodate}\footnote{ \url{https://ncatlab.org/nlab/show/Rel} }. \ context \ begin context begin private lemma smc_Rel_is_monic_arr_vsubset: assumes "T : A \\<^bsub>smc_Rel \\<^esub> B" and "R : A' \\<^bsub>smc_Rel \\<^esub> A" and "S : A' \\<^bsub>smc_Rel \\<^esub> A" and "T \\<^sub>A\<^bsub>smc_Rel \\<^esub> R = T \\<^sub>A\<^bsub>smc_Rel \\<^esub> S" and "\y z X. \ y \\<^sub>\ A; z \\<^sub>\ A; T\ArrVal\ `\<^sub>\ y = X; T\ArrVal\ `\<^sub>\ z = X \ \ y = z" shows "R\ArrVal\ \\<^sub>\ S\ArrVal\" proof- interpret Rel: semicategory \ \smc_Rel \\ by (rule semicategory_smc_Rel) interpret R: arr_Rel \ R rewrites "R\ArrDom\ = A'" and "R\ArrCod\ = A" using assms(2) by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) interpret S: arr_Rel \ S rewrites "S\ArrDom\ = A'" and "S\ArrCod\ = A" using assms(3) by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) from assms(4) have "(T \\<^sub>A\<^bsub>smc_Rel \\<^esub> R)\ArrVal\ = (T \\<^sub>A\<^bsub>smc_Rel \\<^esub> S)\ArrVal\" by simp then have eq: "T\ArrVal\ \\<^sub>\ R\ArrVal\ = T\ArrVal\ \\<^sub>\ S\ArrVal\" unfolding smc_Rel_Comp_app[OF assms(1,2)] smc_Rel_Comp_app[OF assms(1,3)] comp_Rel_components by simp show "R\ArrVal\ \\<^sub>\ S\ArrVal\" proof(rule vsubsetI) fix ab assume ab[intro]: "ab \\<^sub>\ R\ArrVal\" with R.ArrVal.vbrelation obtain a b where ab_def: "ab = \a, b\" by auto with ab R.arr_Rel_ArrVal_vrange have "a \\<^sub>\ \\<^sub>\ (R\ArrVal\)" and "b \\<^sub>\ A" by auto define B' and C' where "B' = R\ArrVal\ `\<^sub>\ set {a}" and "C' = T\ArrVal\ `\<^sub>\ B'" have ne_C': "C' \ 0" proof(rule ccontr, unfold not_not) assume prems: "C' = 0" from ab have "b \\<^sub>\ B'" unfolding ab_def B'_def by simp with C'_def[unfolded prems] have b0: "T\ArrVal\ `\<^sub>\ set {b} = 0" by auto from assms(5)[OF _ _ b0, of 0] \b \\<^sub>\ A\ show False by auto qed have cac''[intro, simp]: "c \\<^sub>\ C' \ \a, c\ \\<^sub>\ T\ArrVal\ \\<^sub>\ S\ArrVal\" for c unfolding eq[symmetric] C'_def B'_def by (metis vcomp_vimage vimage_vsingleton_iff) define A'' where "A'' = (T\ArrVal\ \\<^sub>\ S\ArrVal\) -`\<^sub>\ C'" define B'' where "B'' = S\ArrVal\ `\<^sub>\ set {a}" define C'' where "C'' = T\ArrVal\ `\<^sub>\ B''" have a'': "a \\<^sub>\ A''" proof- from ne_C' obtain c' where [intro]: "c' \\<^sub>\ C'" by (auto intro!: vsubset_antisym) then have "\a, c'\ \\<^sub>\ T\ArrVal\ \\<^sub>\ S\ArrVal\" by simp then show ?thesis unfolding A''_def by auto qed have "C' \\<^sub>\ C''" unfolding C''_def B''_def A''_def C'_def B'_def by (rule vsubsetI) (metis eq vcomp_vimage) have "C' = C''" proof(rule ccontr) assume "C' \ C''" with \C' \\<^sub>\ C''\ obtain c' where c': "c' \\<^sub>\ C'' -\<^sub>\ C'" by (auto intro!: vsubset_antisym) then obtain b'' where "b'' \\<^sub>\ B''" and "\b'', c'\ \\<^sub>\ T\ArrVal\" unfolding C''_def by auto then have "\a, c'\ \\<^sub>\ T\ArrVal\ \\<^sub>\ R\ArrVal\" unfolding eq B''_def by auto with c' show False unfolding B'_def C'_def by auto qed then have "T\ArrVal\ `\<^sub>\ B'' = T\ArrVal\ `\<^sub>\ B'" by (simp add: C''_def C'_def) moreover have "B' \\<^sub>\ A" and "B'' \\<^sub>\ A" using R.arr_Rel_ArrVal_vrange S.arr_Rel_ArrVal_vrange unfolding B'_def B''_def by auto ultimately have "B'' = B'" by (simp add: assms(5)) with ab have "b \\<^sub>\ B''" unfolding B'_def ab_def by simp then show "ab \\<^sub>\ S\ArrVal\" unfolding ab_def B''_def by simp qed qed lemma smc_Rel_is_monic_arrI: assumes "T : A \\<^bsub>smc_Rel \\<^esub> B" and "\y z X. \ y \\<^sub>\ A; z \\<^sub>\ A; T\ArrVal\ `\<^sub>\ y = X; T\ArrVal\ `\<^sub>\ z = X \ \ y = z" shows "T : A \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> B" proof(rule is_monic_arrI) interpret Rel: semicategory \ \smc_Rel \\ by (simp add: semicategory_smc_Rel) fix R S A' assume prems: "R : A' \\<^bsub>smc_Rel \\<^esub> A" "S : A' \\<^bsub>smc_Rel \\<^esub> A" "T \\<^sub>A\<^bsub>smc_Rel \\<^esub> R = T \\<^sub>A\<^bsub>smc_Rel \\<^esub> S" interpret T: arr_Rel \ T rewrites "T\ArrDom\ = A" and "T\ArrCod\ = B" using assms(1) by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) interpret R: arr_Rel \ R rewrites [simp]: "R\ArrDom\ = A'" and [simp]: "R\ArrCod\ = A" using prems(1) by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) interpret S: arr_Rel \ S rewrites [simp]: "S\ArrDom\ = A'" and [simp]: "S\ArrCod\ = A" using prems(2) by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) from assms prems have "R\ArrVal\ \\<^sub>\ S\ArrVal\" "S\ArrVal\ \\<^sub>\ R\ArrVal\" by (auto simp: smc_Rel_is_monic_arr_vsubset) then show "R = S" using R.arr_Rel_axioms S.arr_Rel_axioms by (intro arr_Rel_eqI[of \ R S]) auto qed (rule assms(1)) end end lemma (in \) smc_Rel_is_monic_arrD[dest]: assumes "T : A \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> B" and "y \\<^sub>\ A" and "z \\<^sub>\ A" and "T\ArrVal\ `\<^sub>\ y = X" and "T\ArrVal\ `\<^sub>\ z = X" shows "y = z" proof- interpret Rel: semicategory \ \smc_Rel \\ by (simp add: semicategory_smc_Rel) from assms have T: "T : A \\<^bsub>smc_Rel \\<^esub> B" by (simp add: is_monic_arr_def) interpret T: arr_Rel \ T rewrites "T\ArrDom\ = A" and [simp]: "T\ArrCod\ = B" using T by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) define R where "R = [set {0} \\<^sub>\ y, set {0}, A]\<^sub>\" define S where "S = [set {0} \\<^sub>\ z, set {0}, A]\<^sub>\" have R: "R : set {0} \\<^bsub>smc_Rel \\<^esub> A" proof(intro smc_Rel_is_arrI) show "arr_Rel \ R" unfolding R_def proof(intro arr_Rel_vfsequenceI) from assms(2) show "\\<^sub>\ (set {0} \\<^sub>\ y) \\<^sub>\ A" by auto qed (auto simp: T.arr_Rel_ArrDom_in_Vset) qed (simp_all add: R_def arr_Rel_components) from assms(3) have S: "S : set {0} \\<^bsub>smc_Rel \\<^esub> A" proof(intro smc_Rel_is_arrI) show "arr_Rel \ S" unfolding S_def proof(intro arr_Rel_vfsequenceI) from assms(3) show "\\<^sub>\ (set {0} \\<^sub>\ z) \\<^sub>\ A" by auto qed (auto simp: T.arr_Rel_ArrDom_in_Vset) qed (simp_all add: S_def arr_Rel_components) from assms(4) have "T \\<^sub>A\<^bsub>smc_Rel \\<^esub> R = [set {0} \\<^sub>\ X, set {0}, B]\<^sub>\" unfolding smc_Rel_Comp_app[OF T R] unfolding comp_Rel_components R_def comp_Rel_def arr_Rel_components by (simp add: vcomp_vimage_vtimes_right) moreover from assms have "T \\<^sub>A\<^bsub>smc_Rel \\<^esub> S = [set {0} \\<^sub>\ X, set {0}, B]\<^sub>\" unfolding smc_Rel_Comp_app[OF T S] unfolding comp_Rel_components S_def comp_Rel_def arr_Rel_components by (simp add: vcomp_vimage_vtimes_right) ultimately have "T \\<^sub>A\<^bsub>smc_Rel \\<^esub> R = T \\<^sub>A\<^bsub>smc_Rel \\<^esub> S" by simp from R S assms(1) this have "R = S" by (elim is_monic_arrE) then show "y = z" unfolding R_def S_def by auto qed lemma (in \) smc_Rel_is_monic_arr: "T : A \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> B \ T : A \\<^bsub>smc_Rel \\<^esub> B \ ( \y z X. y \\<^sub>\ A \ z \\<^sub>\ A \ (T\ArrVal\) `\<^sub>\ y = X \ (T\ArrVal\) `\<^sub>\ z = X \ y = z )" by (rule iffI allI impI) (auto simp: smc_Rel_is_monic_arrD smc_Rel_is_monic_arrI) lemma (in \) smc_Rel_is_monic_arr_is_epic_arr: assumes "T : A \\<^bsub>smc_Rel \\<^esub> B" and "(\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\\T\ : B \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> A" shows "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B" proof- interpret is_iso_semifunctor \ \op_smc (smc_Rel \)\ \smc_Rel \\ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ rewrites "(op_smc \')\Obj\ = \'\Obj\" and "(op_smc \')\Arr\ = \'\Arr\" and "f : b \\<^bsub>op_smc \'\<^esub> a \ f : a \\<^bsub>\'\<^esub> b" for \' f a b unfolding smc_op_simps by (auto simp: smcf_dag_Rel_is_iso_semifunctor) show ?thesis proof(intro HomCod.is_epic_arrI) show T: "T : A \\<^bsub>smc_Rel \\<^esub> B" by (rule assms(1)) fix f g a assume prems: "f : B \\<^bsub>smc_Rel \\<^esub> a" "g : B \\<^bsub>smc_Rel \\<^esub> a" "f \\<^sub>A\<^bsub>smc_Rel \\<^esub> T = g \\<^sub>A\<^bsub>smc_Rel \\<^esub> T" from prems(1) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\" by (auto intro: smc_cs_intros) with prems(1) HomCod.smc_is_arrD(3) T have dag_f: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ : a \\<^bsub>smc_Rel \\<^esub> B" unfolding smcf_dag_Rel_components(1) by auto from prems(2) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\" by (auto intro: smc_cs_intros) with prems(2) have dag_g: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ : a \\<^bsub>smc_Rel \\<^esub> B" unfolding smcf_dag_Rel_components(1) by (metis HomCod.smc_is_arrD(3) T vid_on_eq_atI) from prems T have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\" by (simp add: smcf_dag_Rel_ArrMap_smc_Rel_Comp[symmetric]) from is_monic_arrD(2)[OF assms(2) dag_f dag_g this] show "f = g" by (meson prems HomDom.smc_is_arrD(1) ArrMap.v11_eq_iff) qed qed lemma (in \) smc_Rel_is_epic_arr_is_monic_arr: assumes "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B" shows "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> A" proof(rule is_monic_arrI) interpret is_iso_semifunctor \ \op_smc (smc_Rel \)\ \smc_Rel \\ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a \ f : a \\<^bsub>\'\<^esub> b" for \' f a b unfolding smc_op_simps by (auto simp: smcf_dag_Rel_is_iso_semifunctor) have dag: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ : op_smc (smc_Rel \) \\\<^sub>S\<^sub>M\<^sub>C\<^bsub>\\<^esub> smc_Rel \" by (auto intro: smc_cs_intros) from HomCod.is_epic_arrD(1)[OF assms] have T: "T : A \\<^bsub>smc_Rel \\<^esub> B". from T have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\A\" by (auto intro: smc_cs_intros) with T show dag_T: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^bsub>smc_Rel \\<^esub> A" unfolding smcf_dag_Rel_components(1) by (metis HomCod.smc_is_arrD(2) HomCod.smc_is_arrD(3) vid_on_eq_atI) fix f g a :: V assume prems: "f : a \\<^bsub>smc_Rel \\<^esub> B" "g : a \\<^bsub>smc_Rel \\<^esub> B" "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> f = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> g" then have a: "a \\<^sub>\ smc_Rel \\Obj\" by auto from prems(1) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\" by (auto intro: smc_cs_intros) with prems(1) have dag_f: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ : B \\<^bsub>smc_Rel \\<^esub> a" by (cs_concl cs_intro: smc_cs_intros cs_simp: smc_Rel_cs_simps) from prems(2) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\a\" by (cs_concl cs_intro: smc_cs_intros cs_simp:) with prems(2) have dag_g: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ : B \\<^bsub>smc_Rel \\<^esub> a" by (cs_concl cs_intro: smc_cs_intros cs_simp: smc_Rel_cs_simps) from T dag have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ = (\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \ \<^sub>S\<^sub>M\<^sub>C\<^sub>F\ \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \)\ArrMap\\T\" by ( cs_concl cs_intro: smc_cs_intros cs_simp: smc_Rel_cs_simps smc_cn_cs_simps smc_cs_simps ) also from T have "\ = T" unfolding dghm_id_components smcf_cn_comp_smcf_dag_Rel_smcf_dag_Rel by auto finally have dag_dag_T: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ = T" by simp have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\f\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> T = \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\g\ \\<^sub>A\<^bsub>smc_Rel \\<^esub> T" by (metis dag_T dag_dag_T prems smcf_dag_Rel_ArrMap_smc_Rel_Comp) from HomCod.is_epic_arrD(2)[OF assms dag_f dag_g this] prems ArrMap.v11_eq_iff show "f = g" by blast qed lemma (in \) smc_Rel_is_epic_arrI: assumes "T : A \\<^bsub>smc_Rel \\<^esub> B" and "\y z X. \ y \\<^sub>\ B; z \\<^sub>\ B; T\ArrVal\ -`\<^sub>\ y = X; T\ArrVal\ -`\<^sub>\ z = X \ \ y = z" shows "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B" proof- interpret is_iso_semifunctor \ \op_smc (smc_Rel \)\ \smc_Rel \\ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a \ f : a \\<^bsub>\'\<^esub> b" for \' f a b unfolding smc_op_simps by (auto simp: smcf_dag_Rel_is_iso_semifunctor) from assms have T: "arr_Rel \ T" by (auto simp: smc_Rel_is_arrD(1)) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> A" proof(rule smc_Rel_is_monic_arrI) from assms(1) have "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\B\ \\<^bsub>smc_Rel \\<^esub> \\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ObjMap\\A\" by (cs_concl cs_intro: smc_cs_intros) with assms(1) show "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^bsub>smc_Rel \\<^esub> A" by (cs_concl cs_intro: smc_cs_intros cs_simp: smc_Rel_cs_simps) fix y z X assume "y \\<^sub>\ B" "z \\<^sub>\ B" "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\ `\<^sub>\ y = X" "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\\ArrVal\ `\<^sub>\ z = X" then show "y = z" unfolding converse_Rel_components smcf_dag_Rel_ArrMap_app[OF T] app_invimage_def[symmetric] by (rule assms(2)) qed from smc_Rel_is_monic_arr_is_epic_arr[OF assms(1) this] show ?thesis by simp qed lemma (in \) smc_Rel_is_epic_arrD[dest]: assumes "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B" and "y \\<^sub>\ B" and "z \\<^sub>\ B" and "T\ArrVal\ -`\<^sub>\ y = X" and "T\ArrVal\ -`\<^sub>\ z = X" shows "y = z" proof- interpret is_iso_semifunctor \ \op_smc (smc_Rel \)\ \smc_Rel \\ \\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ rewrites "f : b \\<^bsub>op_smc \'\<^esub> a \ f : a \\<^bsub>\'\<^esub> b" for \' f a b unfolding smc_op_simps by (auto simp: smcf_dag_Rel_is_iso_semifunctor) have dag_T: "\\<^sub>S\<^sub>M\<^sub>C\<^sub>.\<^sub>R\<^sub>e\<^sub>l \\ArrMap\\T\ : B \\<^sub>m\<^sub>o\<^sub>n\<^bsub>smc_Rel \\<^esub> A" by (rule smc_Rel_is_epic_arr_is_monic_arr[OF assms(1)]) from HomCod.is_epic_arrD(1)[OF assms(1)] have T: "T : A \\<^bsub>smc_Rel \\<^esub> B". then have T: "arr_Rel \ T" by (auto simp: smc_Rel_is_arrD(1)) from assms(4,5) smc_Rel_is_monic_arrD [ OF dag_T assms(2,3), unfolded smc_dg_smc_Rel smcf_dghm_smcf_dag_Rel converse_Rel_components smcf_dag_Rel_ArrMap_app[OF T] ] show ?thesis by (auto simp: app_invimage_def) qed lemma (in \) smc_Rel_is_epic_arr: "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B \ T : A \\<^bsub>smc_Rel \\<^esub> B \ ( \y z X. y \\<^sub>\ B \ z \\<^sub>\ B \ T\ArrVal\ -`\<^sub>\ y = X \ T\ArrVal\ -`\<^sub>\ z = X \ y = z )" proof(intro iffI allI impI conjI) show "T : A \\<^sub>e\<^sub>p\<^sub>i\<^bsub>smc_Rel \\<^esub> B \ T : A \\<^bsub>smc_Rel \\<^esub> B" by (simp add: is_epic_arr_def is_monic_arr_def op_smc_is_arr) qed (auto simp: smc_Rel_is_epic_arrI) subsection\Terminal object, initial object and null object\ text\ An object in the semicategory \Rel\ is terminal/initial/null if and only if it is the empty set (see nLab \cite{noauthor_nlab_nodate})\footnote{ \url{https://ncatlab.org/nlab/show/database+of+categories} }. \ lemma (in \) smc_Rel_obj_terminal: "obj_terminal (smc_Rel \) A \ A = 0" proof- interpret semicategory \ \smc_Rel \\ by (rule semicategory_smc_Rel) have "(\A\\<^sub>\Vset \. \!T. T : A \\<^bsub>smc_Rel \\<^esub> B) \ B = 0" for B proof(intro iffI allI ballI) assume prems[rule_format]: "\A\\<^sub>\Vset \. \!T. T : A \\<^bsub>smc_Rel \\<^esub> B" then obtain T where "T : 0 \\<^bsub>smc_Rel \\<^esub> B" by (meson vempty_is_zet) then have [simp]: "B \\<^sub>\ Vset \" by (fastforce simp: smc_Rel_components(1)) show "B = 0" proof(rule ccontr) assume "B \ 0" with trad_foundation obtain b where "b \\<^sub>\ B" by auto let ?b0B = \[set {\0, b\}, set {0}, B]\<^sub>\\ let ?z0B = \[0, set {0}, B]\<^sub>\\ have "?b0B : set {0} \\<^bsub>smc_Rel \\<^esub> B" proof(intro smc_Rel_is_arrI) show b0B: "arr_Rel \ ?b0B" by (intro arr_Rel_vfsequenceI) (force simp: \b \\<^sub>\ B\ vsubset_vsingleton_leftI)+ qed (simp_all add: arr_Rel_components) moreover have "?z0B : set {0} \\<^bsub>smc_Rel \\<^esub> B" proof(intro smc_Rel_is_arrI) show b0B: "arr_Rel \ ?z0B" by (intro arr_Rel_vfsequenceI) (force simp: \b \\<^sub>\ B\ vsubset_vsingleton_leftI)+ qed (simp_all add: arr_Rel_components) moreover have "[set {\0, b\}, set {0}, B]\<^sub>\ \ [0, set {0}, B]\<^sub>\" by simp ultimately show False by (metis prems smc_is_arrE smc_Rel_components(1)) qed next fix A assume prems[simp]: "B = 0" "A \\<^sub>\ Vset \" let ?zAz = \[0, A, 0]\<^sub>\\ have zAz: "arr_Rel \ ?zAz" by ( simp add: \.arr_Rel_vfsequenceI \_axioms smc_Rel_components(2) vbrelation_vempty ) show "\!T. T : A \\<^bsub>smc_Rel \\<^esub> B" proof(rule ex1I[of _ \?zAz\]) show "?zAz : A \\<^bsub>smc_Rel \\<^esub> B" by (intro smc_Rel_is_arrI) ( simp_all add: zAz smc_Rel_Dom_app[OF zAz] smc_Rel_Cod_app[OF zAz] arr_Rel_components ) fix T assume "T : A \\<^bsub>smc_Rel \\<^esub> B" then have T: "T : A \\<^bsub>smc_Rel \\<^esub> 0" by simp then interpret T: arr_Rel \ T by (fastforce simp: smc_Rel_components(2)) show "T = [0, A, 0]\<^sub>\" proof ( subst T.arr_Rel_def, rule arr_Rel_eqI[of \], unfold arr_Rel_components ) show "arr_Rel \ [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\" by (fold T.arr_Rel_def) (simp add: T.arr_Rel_axioms) from zAz show "arr_Rel \ ?zAz" by (simp add: arr_Rel_vfsequenceI vbrelationI) from T have "T \\<^sub>\ smc_Rel \\Arr\" by (auto intro: smc_cs_intros) with is_arrD(2,3)[OF T] show "T\ArrDom\ = A" "T\ArrCod\ = 0" using T smc_Rel_is_arrD(2,3) by auto with T.arr_Rel_ArrVal_vrange T.ArrVal.vbrelation_vintersection_vrange show "T\ArrVal\ = []\<^sub>\" by auto qed qed qed then show ?thesis apply(intro iffI obj_terminalI) subgoal by (metis smc_is_arrD(2) obj_terminalE) subgoal by blast subgoal by (metis smc_Rel_components(1)) done qed (*TODO: generalize: duality/dagger*) lemma (in \) smc_Rel_obj_initial: "obj_initial (smc_Rel \) A \ A = 0" proof- interpret semicategory \ \smc_Rel \\ by (rule semicategory_smc_Rel) have "(\B\\<^sub>\Vset \. \!T. T : A \\<^bsub>smc_Rel \\<^esub> B) \ A = 0" for A proof(intro iffI allI ballI) assume prems[rule_format]: "\B\\<^sub>\Vset \. \!T. T : A \\<^bsub>smc_Rel \\<^esub> B" then obtain T where TA0: "T : A \\<^bsub>smc_Rel \\<^esub> 0" by (meson vempty_is_zet) then have [simp]: "A \\<^sub>\ Vset \" by (fastforce simp: smc_Rel_components(1)) show "A = 0" proof(rule ccontr) assume "A \ 0" with trad_foundation obtain a where "a \\<^sub>\ A" by auto have "[set {\a, 0\}, A, set {0}]\<^sub>\ : A \\<^bsub>smc_Rel \\<^esub> set {0}" proof(intro smc_Rel_is_arrI) show "arr_Rel \ [set {\a, 0\}, A, set {0}]\<^sub>\" by (intro arr_Rel_vfsequenceI) (auto simp: \a \\<^sub>\ A\ vsubset_vsingleton_leftI) qed (simp_all add: arr_Rel_components) moreover have "[0, A, set {0}]\<^sub>\ : A \\<^bsub>smc_Rel \\<^esub> set {0}" proof(intro smc_Rel_is_arrI) show "arr_Rel \ [0, A, set {0}]\<^sub>\" by (intro arr_Rel_vfsequenceI) (auto simp: \a \\<^sub>\ A\ vsubset_vsingleton_leftI) qed (simp_all add: arr_Rel_components) moreover have "[set {\a, 0\}, A, set {0}]\<^sub>\ \ [0, A, set {0}]\<^sub>\" by simp ultimately show False by (metis prems smc_is_arrE smc_Rel_components(1)) qed next fix B assume [simp]: "A = 0" "B \\<^sub>\ Vset \" show "\!T. T : A \\<^bsub>smc_Rel \\<^esub> B" proof(rule ex1I[of _ \[0, 0, B]\<^sub>\\]) show "[0, 0, B]\<^sub>\ : A \\<^bsub>smc_Rel \\<^esub> B" by (rule is_arrI) ( simp_all add: smc_Rel_cs_simps smc_Rel_components(2) vbrelation_vempty arr_Rel_vfsequenceI ) fix T assume "T : A \\<^bsub>smc_Rel \\<^esub> B" then have T: "T : 0 \\<^bsub>smc_Rel \\<^esub> B" by simp interpret T: arr_Rel \ T using T by (fastforce simp: smc_Rel_components(2)) show "T = [0, 0, B]\<^sub>\" proof ( subst T.arr_Rel_def, rule arr_Rel_eqI[of \], unfold arr_Rel_components ) show "arr_Rel \ [T\ArrVal\, T\ArrDom\, T\ArrCod\]\<^sub>\" by (fold T.arr_Rel_def) (simp add: T.arr_Rel_axioms) show "arr_Rel \ [[]\<^sub>\, []\<^sub>\, B]\<^sub>\" by (simp add: arr_Rel_vfsequenceI vbrelationI) from T have "T \\<^sub>\ smc_Rel \\Arr\" by (auto intro: smc_cs_intros) with T is_arrD(2,3)[OF T] show "T\ArrDom\ = 0" "T\ArrCod\ = B" by (auto simp: smc_Rel_is_arrD(2,3)) with T.arr_Rel_ArrVal_vrange T.arr_Rel_ArrVal_vdomain T.ArrVal.vbrelation_vintersection_vdomain show "T\ArrVal\ = []\<^sub>\" by auto qed qed qed then show ?thesis apply(intro iffI obj_initialI, elim obj_initialE) subgoal by (metis smc_Rel_components(1)) subgoal by (simp add: smc_Rel_components(1)) subgoal by (metis smc_Rel_components(1)) done qed lemma (in \) smc_Rel_obj_terminal_obj_initial: "obj_initial (smc_Rel \) A \ obj_terminal (smc_Rel \) A" unfolding smc_Rel_obj_initial smc_Rel_obj_terminal by simp lemma (in \) smc_Rel_obj_null: "obj_null (smc_Rel \) A \ A = 0" unfolding obj_null_def smc_Rel_obj_terminal smc_Rel_obj_initial by simp subsection\Zero arrow\ text\ A zero arrow for \Rel\ is any admissible \V\-arrow, such that its value is the empty set. A reference for this result is not given, but the result is not expected to be original. \ lemma (in \) smc_Rel_is_zero_arr: assumes "A \\<^sub>\ Vset \" and "B \\<^sub>\ Vset \" shows "T : A \\<^sub>0\<^bsub>smc_Rel \\<^esub> B \ T = [0, A, B]\<^sub>\" proof(rule HOL.ext iffI) interpret Rel: semicategory \ \smc_Rel \\ by (rule semicategory_smc_Rel) fix T A B assume "T : A \\<^sub>0\<^bsub>smc_Rel \\<^esub> B" then obtain R S where T_def: "T = R \\<^sub>A\<^bsub>smc_Rel \\<^esub> S" and S: "S : A \\<^bsub>smc_Rel \\<^esub> 0" and R: "R : 0 \\<^bsub>smc_Rel \\<^esub> B" by (elim is_zero_arrE) (simp add: obj_null_def smc_Rel_obj_terminal) interpret S: arr_Rel \ S rewrites [simp]: "S\ArrDom\ = A" and [simp]: "S\ArrCod\ = 0" using S by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) interpret R: arr_Rel \ R rewrites [simp]: "R\ArrDom\ = 0" and [simp]: "R\ArrCod\ = B" using R by (all\elim Rel.smc_is_arrE\) (simp_all add: smc_Rel_components) have S_def: "S = [0, A, 0]\<^sub>\" by ( rule arr_Rel_eqI[of \], unfold arr_Rel_components, insert S.arr_Rel_ArrVal_vrange S.ArrVal.vbrelation_vintersection_vrange ) ( auto simp: S.arr_Rel_axioms S.arr_Rel_ArrDom_in_Vset arr_Rel_vfsequenceI vbrelationI ) show "T = [0, A, B]\<^sub>\" unfolding T_def smc_Rel_Comp_app[OF R S] by (rule arr_Rel_eqI[of \], unfold comp_Rel_components) ( auto simp: S_def \_axioms R.arr_Rel_axioms S.arr_Rel_axioms arr_Rel_comp_Rel arr_Rel_components R.arr_Rel_ArrCod_in_Vset S.arr_Rel_ArrDom_in_Vset \.arr_Rel_vfsequenceI vbrelation_vempty ) next assume prems: "T = [0, A, B]\<^sub>\" let ?S = \[0, A, 0]\<^sub>\\ and ?R = \[0, 0, B]\<^sub>\\ have S: "arr_Rel \ ?S" and R: "arr_Rel \ ?R" by (all\intro arr_Rel_vfsequenceI\) (auto simp: assms) have SA0: "?S : A \\<^bsub>smc_Rel \\<^esub> 0" by (intro smc_Rel_is_arrI) (simp_all add: S arr_Rel_components) moreover have R0B: "?R : 0 \\<^bsub>smc_Rel \\<^esub> B" by (intro smc_Rel_is_arrI) (simp_all add: R arr_Rel_components) moreover have "T = ?R \\<^sub>A\<^bsub>smc_Rel \\<^esub> ?S" unfolding smc_Rel_Comp_app[OF R0B SA0] proof(rule arr_Rel_eqI, unfold comp_Rel_components arr_Rel_components prems) show "arr_Rel \ [0, A, B]\<^sub>\" unfolding prems by (intro arr_Rel_vfsequenceI) (auto simp: assms) qed (use R S in \auto simp: smc_Rel_cs_intros\) ultimately show "T : A \\<^sub>0\<^bsub>smc_Rel \\<^esub> B" by (simp add: is_zero_arrI smc_Rel_obj_null) qed text\\newpage\ end \ No newline at end of file