Theory CZH_ECAT_Set

(* Copyright 2021 (C) Mihails Milehins *)

sectionSet›
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 (in arr_Set) [cat_cs_intros, cat_Set_cs_intros] = 
  arr_Set_axioms'

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



subsectionSet› as a category›


subsubsection‹Definition and elementary properties›

definition cat_Set :: "V  V"
  where "cat_Set α =
    [
      Vset α,
      set {T. arr_Set α T},
      (λTset {T. arr_Set α T}. TArrDom),
      (λTset {T. arr_Set α T}. TArrCod),
      (λSTcomposable_arrs (dg_Set α). ST0 Rel ST1),
      VLambda (Vset α) id_Set 
    ]"


text‹Components.›

lemma cat_Set_components:
  shows "cat_Set αObj = Vset α"
    and "cat_Set αArr = set {T. arr_Set α T}"
    and "cat_Set αDom = (λTset {T. arr_Set α T}. TArrDom)"
    and "cat_Set αCod = (λTset {T. arr_Set α T}. TArrCod)"
    and "cat_Set αComp =
      (λSTcomposable_arrs (dg_Set α). ST0 Par ST1)"
    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: "𝒟 (cat_smc (cat_Set α)) = 5" 
    unfolding cat_smc_def by (simp add: nat_omega_simps)
  have dom_rhs: "𝒟 (smc_Set α) = 5"
    unfolding smc_Set_def by (simp add: nat_omega_simps)
  show "𝒟 (cat_smc (cat_Set α)) = 𝒟 (smc_Set α)"
    unfolding dom_lhs dom_rhs by simp
  show "a  𝒟 (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
  and cat_Set_is_monic_arrI = smc_Set_is_monic_arrI
  and cat_Set_is_monic_arrD = smc_Set_is_monic_arrD
  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

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_Comp_ArrVal = smc_Set_Comp_ArrVal
  and cat_Set_Comp_vrange = smc_Set_Comp_vrange
  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_cs_simps] = 
  𝒵.cat_Set_Comp_ArrVal

lemma (in 𝒵) cat_Set_incl_Set_is_arr'[cat_cs_intros, cat_Set_cs_intros]:
  assumes "A  cat_Set αObj"
    and "B  cat_Set αObj"
    and "A  B"
    and "A' = A"
    and "B' = B"
    and "ℭ' = cat_Set α"
  shows "incl_Set A B : A' ℭ'B'"
  using assms(1-3) unfolding assms(4-6) by (rule cat_Set_incl_Set_is_arr)

lemmas [cat_Set_cs_intros] = 𝒵.cat_Set_incl_Set_is_arr'


subsubsection‹Identity›

lemma cat_Set_CId_app[cat_Set_cs_simps]:
  assumes "A  Vset α"
  shows "cat_Set αCIdA = id_Set A"
  using assms unfolding cat_Set_components by simp

lemma cat_Set_CId_app_app[cat_cs_simps]:
  assumes "A  cat_Set αObj" and "a  A"
  shows "cat_Set αCIdAArrVala = a"
  unfolding 
    cat_Set_CId_app[OF assms(1)[unfolded cat_Set_components(1)]] 
    id_Rel_ArrVal_app[OF assms(2)] 
  by simp


subsubsectionSet› 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"
    unfolding cat_Set_def by (simp add: nat_omega_simps)
  show "semicategory α (smc_Set α)" by (simp add: semicategory_smc_Set)
  show "cat_Set αCIdA : A cat_Set αA"
    if "A  cat_Set αObj" for A
    using that
    unfolding cat_Set_Obj_iff
    by 
      (
        cs_concl cs_shallow
          cs_simp: cat_Set_cs_simps cs_intro: cat_Set_cs_intros arr_Set_id_SetI
      )

  show "cat_Set αCIdB Acat_Set αF = F" 
    if "F : A cat_Set αB" for F A B
  proof-
    from that have "arr_Set α F" "B  Vset α" by (auto elim: cat_Set_is_arrE)
    with that show ?thesis
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps cat_Set_cs_simps
            cs_intro: cat_Set_cs_intros arr_Set_id_SetI
        )
  qed

  show "F Acat_Set αcat_Set αCIdB = F"
    if "F : B cat_Set αC" for F B C
  proof-
    from that have "arr_Set α F" "B  Vset α" by (auto elim: cat_Set_is_arrE)
    with that show ?thesis
      by 
        (
          cs_concl cs_shallow
            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'


subsubsectionSet› is a wide replete subcategory of Par›

lemma (in 𝒵) wide_replete_subcategory_cat_Set_cat_Par: 
  "cat_Set α C.wrαcat_Par α"
proof(intro wide_replete_subcategoryI)
  show wide_subcategory_cat_Set_cat_Par: "cat_Set α C.wideα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 α Cαcat_Par α"
    proof(intro subcategoryI, unfold cat_smc_cat_Par cat_smc_cat_Set)
      show "smc_Set α SMCαsmc_Par α" by (simp add: subsemicategory_axioms)
      fix A assume "A  cat_Set αObj"
      then show "cat_Set αCIdA = cat_Par αCIdA"
        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 α C.repαcat_Par α"
  proof(intro replete_subcategoryI)
    interpret wide_subcategory α cat_Set α cat_Par α
      by (rule wide_subcategory_cat_Set_cat_Par)
    show "cat_Set α Cαcat_Par α" by (rule subcategory_axioms)    
    fix A B F assume "F : A isocat_Par αB"
    note arr_Par = cat_Par_is_iso_arrD[OF this]
    from arr_Par show "F : A cat_Set α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


subsubsectionSet› is a subcategory of Set›

lemma (in 𝒵) subcategory_cat_Set_cat_Set:(*TODO: generalize*)
  assumes "𝒵 β" and "α  β"
  shows "cat_Set α Cβ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  cat_Set βObj" if "A  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 cat_Set βB" if "F : A cat_Set α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 " (FArrVal)  FArrCod"  
           by (auto simp: f.arr_Set_ArrVal_vrange)
        show "FArrDom  Vset β"
          by (auto intro!: f.arr_Set_ArrDom_in_Vset Vset_in_mono assms(2))
        show "FArrCod  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 Acat_Set αF = G Acat_Set βF"
      if "G : B cat_Set αC" and "F : A cat_Set α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 Acat_Set αF : A cat_Set βC"
        by (cs_concl cs_intro: cat_cs_intros is_arr_if_is_arr)
      from that have β_gf_is_arr: "G Acat_Set βF : A cat_Set β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 Acat_Set αF)" by (rule α_gf(1))
        then interpret arr_Set_α_gf: arr_Set β (G Acat_Set αF) by simp
        from α_gf_is_arr have dom_lhs: "𝒟 ((G Acat_Set αF)ArrVal) = A"
          by (cs_concl cs_shallow cs_simp: cat_cs_simps)
        show "arr_Set β (G Acat_Set βF)" by (rule β_gf(1))
        then interpret arr_Set_β_gf: arr_Set β (G Acat_Set βF) by simp
        from β_gf_is_arr have dom_rhs: "𝒟 ((G Acat_Set βF)ArrVal) = A"
          by (cs_concl cs_shallow cs_simp: cat_cs_simps)
        show "(G Acat_Set αF)ArrVal = (G Acat_Set βF)ArrVal"
        proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
          fix a assume "a  A"
          from that this show 
            "(G Acat_Set αF)ArrVala = (G Acat_Set βF)ArrVala"
            by 
              (
                cs_concl cs_shallow
                  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_shallow cs_simp: cat_cs_simps
        )+
    qed
  qed 
    (
      auto simp: 
        assms(2) cat_Set_components Vset_trans Vset_in_mono cat_cs_intros
    )
qed


subsubsection‹Further properties›

lemma cat_Set_Comp_ArrVal_vrange: (*FIXME: generalize/migrate*)
  assumes "S : B cat_Set αC" and "T : A cat_Set αB" 
  shows " ((S Acat_Set αT)ArrVal)   (SArrVal)" 
proof(intro vsubsetI)
  note SD = cat_Set_is_arrD[OF assms(1)]
  interpret S: arr_Set α S 
    rewrites "SArrDom = B" and "SArrCod = C"
    by (intro SD)+
  from assms(1,2) have "S Acat_Set αT : A cat_Set αC"
    by (cs_concl cs_intro: cat_cs_intros)
  note ST = cat_Set_is_arrD[OF this]
  interpret ST: arr_Set α S Acat_Set αT
    rewrites "(S Acat_Set αT)ArrDom = A" 
      and "(S Acat_Set αT)ArrCod = C"
    by (intro ST)+
  fix y assume prems: "y   ((S Acat_Set αT)ArrVal)"
  with ST.arr_Set_ArrVal_vdomain obtain x 
    where x: "x  A" and y_def: "y = (S Acat_Set αT)ArrValx"
    by force
  show "y   (SArrVal)"
  proof(intro S.ArrVal.vsv_vimageI2', unfold cat_Set_cs_simps)
    from assms(1,2) x show "y = SArrValTArrValx"
      unfolding y_def 
      by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    from assms(2) x show "TArrValx  B"
      by (cs_concl cs_intro: cat_Set_cs_intros)
  qed
qed



subsection‹Isomorphism›

lemma cat_Set_is_iso_arrI[intro]:
  ―‹
  See \cite{noauthor_nlab_nodate}\footnote{\url{
  https://ncatlab.org/nlab/show/isomorphism
  }}).
  ›
  assumes "T : A cat_Set αB" 
    and "v11 (TArrVal)"
    and "𝒟 (TArrVal) = A"
    and " (TArrVal) = B"
  shows "T : A isocat_Set αB"
proof-
  interpret T: arr_Set α T by (rule cat_Set_is_arrD(1)[OF assms(1)])
  note [cat_cs_intros] = cat_Par_is_iso_arrI
  from T.wide_replete_subcategory_cat_Set_cat_Par assms have 
    "T : A isocat_Par αB"
    by (cs_concl cs_intro: cat_cs_intros cat_sub_cs_intros cat_sub_fw_cs_intros)
  with T.wide_replete_subcategory_cat_Set_cat_Par assms show 
    "T : A isocat_Set αB"
    by (cs_concl cs_shallow cs_simp: cat_sub_bw_cs_simps)
qed

lemma cat_Set_is_iso_arrD[dest]:
  assumes "T : A isocat_Set αB"
  shows "T : A cat_Set αB"
    and "v11 (TArrVal)"
    and "𝒟 (TArrVal) = A"
    and " (TArrVal) = B"
proof-
  from assms have T: "T : A cat_Set αB" by auto
  interpret T: arr_Set α T by (rule cat_Set_is_arrD(1)[OF T])
  from T.wide_replete_subcategory_cat_Set_cat_Par assms have T: 
    "T : A isocat_Par αB"
    by (cs_concl cs_shallow cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros)
  show "v11 (TArrVal)" "𝒟 (TArrVal) = A" " (TArrVal) = B"
    by (intro cat_Par_is_iso_arrD[OF T])+
qed (rule is_iso_arrD(1)[OF assms])

lemma cat_Set_is_iso_arr:
  "T : A isocat_Set αB  
    T : A cat_Set αB 
    v11 (TArrVal)  
    𝒟 (TArrVal) = A  
     (TArrVal) = B"
  by auto

lemma (in 𝒵) cat_Set_is_iso_arr_if_monic_and_epic:
  assumes "F : A moncat_Set αB" and "F : A epicat_Set αB"
  shows "F : A isocat_Set αB"
proof-
  note cat_Set_is_monic_arrD[OF assms(1)] cat_Set_is_epic_arrD[OF assms(2)]
  note FD = this(1,2,3,5) cat_Set_is_arrD[OF this(1)]
  show ?thesis by (intro cat_Set_is_iso_arrI FD)
qed



subsection‹The inverse arrow›

lemma cat_Set_ArrVal_app_is_arr[cat_cs_intros]:
  assumes "f : a 𝔄b" 
    and "category α 𝔄" (*the order of premises is important*)
    and "F : Hom 𝔄 a b cat_Set αHom 𝔅 c d"
  shows "FArrValf : c 𝔅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 "FArrValf  Hom 𝔅 c d"
    by (cs_concl cs_shallow 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" ("(_¯Set)" [1000] 999)
  where "a¯Set  a¯Rel"

lemma cat_Set_the_inverse[cat_Set_cs_simps]:
  assumes "T : A isocat_Set αB"
  shows "T¯Ccat_Set α= T¯Set"
proof-
  from assms have T: "T : A cat_Set α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 isocat_Par αB"
    by (cs_concl cs_shallow cs_intro: cat_sub_cs_intros cat_sub_fw_cs_intros)
  from wide_replete_subcategory_cat_Set_cat_Par assms 
  have [symmetric, cat_cs_simps]: "T¯Ccat_Par α= T¯Ccat_Set α⇙"
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_sub_bw_cs_simps cs_intro: cat_sub_cs_intros
      )
  from T show "T¯Ccat_Set α= T¯Set"
    by (cs_concl cs_shallow cs_simp: cat_Par_cs_simps cat_cs_simps cs_intro: 𝒵_β)
qed

lemma cat_Set_the_inverse_app[cat_cs_intros]:
  assumes "T : A isocat_Set αB"
    and "a  A"
    and [cat_cs_simps]: "TArrVala = b"
  shows "(T¯Ccat_Set α)ArrValb = a"
proof-
  from assms have T: "T : A cat_Set αB" by auto
  interpret arr_Set α T by (rule cat_Set_is_arrD(1)[OF T])
  note T = cat_Set_is_iso_arrD[OF assms(1)]
  interpret T: v11 TArrVal by (rule T(2))
  from T.v11_axioms assms(1,2) show "T¯Ccat_Set αArrValb = a"
    by
      (
        cs_concl cs_shallow
          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 𝔅d" 
    and "category α 𝔅" (*the order of premises is important*)
    and "F : Hom 𝔄 a b isocat_Set αHom 𝔅 c d"
  shows "F¯Ccat_Set αArrValf : a 𝔄b"
proof-
  interpret 𝔅: category α 𝔅 by (rule assms(2))
  from cat_Set_is_iso_arrD[OF assms(3)] interpret F: arr_Set α F 
    by (simp add: cat_Set_is_arrD)  
  from assms have "F¯Ccat_Set αArrValf  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 isocat_Set αB" and "b  B"
  shows "FArrValF¯Ccat_Set αArrValb = b"
proof-
  note F = cat_Set_is_iso_arrD[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 Acat_Set αF¯Ccat_Set α= cat_Set αCIdB"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms have [cat_cs_simps]: 
    "FArrValF¯Ccat_Set αArrValb = 
      (F Acat_Set αF¯Ccat_Set α)ArrValb"
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros
      )
  from assms F(1) F.arr_Par_ArrCod_in_Vset[unfolded F] show ?thesis
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma cat_Set_the_inverse_app_app[cat_cs_simps]:
  assumes "F : A isocat_Set αB" and "a  A"
  shows "F¯Ccat_Set αArrValFArrVala = a"
proof-
  note F = cat_Set_is_iso_arrD[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¯Ccat_Set αAcat_Set αF = cat_Set αCIdA"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms have [cat_cs_simps]: 
    "F¯Ccat_Set αArrValFArrVala =
      (F¯Ccat_Set αAcat_Set αF)ArrVala"
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cs_intro: cat_arrow_cs_intros cat_cs_intros
      )
  from assms F(1) F.arr_Par_ArrDom_in_Vset[unfolded F] show ?thesis
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed



subsection‹Conversion of a single-valued relation to an arrow in Set›


subsubsection‹Definition and elementary properties›


definition cat_Set_arr_of_vsv :: "V  V  V"
  where "cat_Set_arr_of_vsv f B = [f, 𝒟 f, B]"


text‹Components.›

lemma cat_Set_arr_of_vsv_components:
  shows [cat_Set_cs_simps]: "cat_Set_arr_of_vsv f BArrVal = f"
    and [cat_Set_cs_simps]: "cat_Set_arr_of_vsv f BArrDom = 𝒟 f"
    and [cat_cs_simps, cat_Set_cs_simps]: "cat_Set_arr_of_vsv f BArrCod = B"
  unfolding cat_Set_arr_of_vsv_def arr_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹
Conversion of a single-valued relation to an arrow in Set› is an arrow in Set›

lemma (in 𝒵) cat_Set_arr_of_vsv_is_arr:
  assumes "vsv r" 
    and " r  B" 
    and "𝒟 r  cat_Set αObj" 
    and "B  cat_Set αObj"
  shows "cat_Set_arr_of_vsv r B : 𝒟 r cat_Set αB"
proof-
  interpret r: vsv r by (rule assms)
  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI, unfold cat_Set_arr_of_vsv_components)
    show "vfsequence (cat_Set_arr_of_vsv r B)"
      unfolding cat_Set_arr_of_vsv_def by auto
    show "vcard (cat_Set_arr_of_vsv r B) = 3"
      unfolding cat_Set_arr_of_vsv_def by (auto simp: nat_omega_simps)
  qed (use assms in auto simp: cat_Set_components)
qed



subsection‹Left restriction for Set›


subsubsection‹Definition and elementary properties›

definition vlrestriction_Set :: "V  V  V" (infixr lSet 80)
  where "T lSet C = [TArrVal l C, C, TArrCod]"


text‹Components.›

lemma vlrestriction_Set_components:
  shows [cat_Set_cs_simps]: "(T lSet C)ArrVal = TArrVal l C"
    and [cat_cs_simps, cat_Set_cs_simps]: "(T lSet C)ArrDom = C"
    and [cat_cs_simps, cat_Set_cs_simps]: "(T lSet C)ArrCod = TArrCod"
  unfolding vlrestriction_Set_def arr_field_simps
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

lemma vlrestriction_Set_ArrVal_vdomain[cat_cs_simps]:
  assumes "T : A cat_Set αB" and "C  A" 
  shows "𝒟 ((T lSet C)ArrVal) = C"
proof-
  note TD = cat_Set_is_arrD[OF assms(1)]
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro TD)+
  from assms show ?thesis
    unfolding vlrestriction_Set_components
    by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: V_cs_intros)
qed

lemma vlrestriction_Set_ArrVal_app[cat_cs_simps]:
  assumes "T : A cat_Set αB" and "C  A" and "x  C" 
  shows "(T lSet C)ArrValx = TArrValx"
proof-
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro cat_Set_is_arrD[OF assms(1)])+
  from assms have x: "x  A" by auto
  with assms show ?thesis 
    unfolding vlrestriction_Set_components
    by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: V_cs_intros)
qed


subsubsection‹Left restriction for Set› is an arrow in Set›

lemma vlrestriction_Set_is_arr:
  assumes "T : A cat_Set αB" and "C  A"
  shows "T lSet C : C cat_Set αB"
proof-
  note TD = cat_Set_is_arrD[OF assms(1)]
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro TD)+
  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI, unfold cat_Set_cs_simps TD(2,3))
    show "vfsequence (T lSet C)"
      unfolding vlrestriction_Set_def by auto
    show "vcard (T lSet C) = 3"
      unfolding vlrestriction_Set_def by (simp add: nat_omega_simps)
    from assms show "𝒟 (TArrVal l C) = C"
      by (cs_concl cs_simp: V_cs_simps cat_cs_simps cs_intro: cat_cs_intros)
    show " (TArrVal l C)  B"
      unfolding app_vimage_def[symmetric]
    proof(intro vsubsetI)
      fix x assume prems: "x  TArrVal ` C"
      then obtain c where "c  C" and x_def: "x = TArrValc" by auto
      with assms(2) have c: "c  A" by auto
      from c assms show "x  B"
        unfolding x_def by (cs_concl cs_intro: cat_Set_cs_intros)
    qed
    from assms(2) show "C  Vset α"
      using vsubset_in_VsetI by (auto simp: T.arr_Set_ArrDom_in_Vset)
  qed (auto simp: T.arr_Set_ArrCod_in_Vset)
qed

lemma (in 𝒵) vlrestriction_Set_is_monic_arr:
  assumes "T : A moncat_Set αB" and "C  A"
  shows "T lSet C : C moncat_Set αB"
proof-
  note cat_Set_is_monic_arrD[OF assms(1)]
  note TD = this cat_Set_is_arrD[OF this(1)]
  interpret F: arr_Set α T by (intro TD)+
  interpret ArrVal: v11 TArrVal by (rule TD(2))
  show ?thesis
  proof
    (
      intro 
        cat_Set_is_monic_arrI 
        vlrestriction_Set_is_arr[OF TD(1) assms(2)],
      unfold cat_Set_cs_simps
    )
    from TD(1) assms(2) show "𝒟 (TArrVal l C) = C"
      by (cs_concl cs_simp: V_cs_simps cat_cs_simps)
  qed auto
qed



subsection‹Right restriction for Set›


subsubsection‹Definition and elementary properties›

definition vrrestriction_Set :: "V  V  V" (infixr rSet 80)
  where "T rSet C = [TArrVal r C, TArrDom, C]"


text‹Components.›

lemma vrrestriction_Set_components:
  shows [cat_Set_cs_simps]: "(T rSet C)ArrVal = TArrVal r C"
    and [cat_cs_simps, cat_Set_cs_simps]: "(T rSet C)ArrDom = TArrDom"
    and [cat_cs_simps, cat_Set_cs_simps]: "(T rSet C)ArrCod = C"
  unfolding vrrestriction_Set_def arr_field_simps
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

lemma vrrestriction_Set_ArrVal_app[cat_cs_simps]:
  assumes "T : A cat_Set αB" and " (TArrVal)  C" 
  shows "(T rSet C)ArrVal = TArrVal"
proof-
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro cat_Set_is_arrD[OF assms(1)])+
  from assms show ?thesis unfolding cat_Set_cs_simps by simp
qed


subsubsection‹Right restriction for Set› is an arrow in Set›

lemma vrrestriction_Set_is_arr:
  assumes "T : A cat_Set αB" 
    and " (TArrVal)  C" 
    and "C  cat_Set αObj"
  shows "T rSet C : A cat_Set αC"
proof-
  note TD = cat_Set_is_arrD[OF assms(1)]
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro TD)+
  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI, unfold cat_Set_cs_simps)
    show "vfsequence (T rSet C)" unfolding vrrestriction_Set_def by auto
    show "vcard (T rSet C) = 3"
      unfolding vrrestriction_Set_def by (simp add: nat_omega_simps)
  qed
    (
      use assms(2,3) in
        auto simp:
            TD(2)
            cat_Set_components
            T.arr_Set_ArrVal_vdomain
            T.arr_Set_ArrDom_in_Vset
    )
qed

lemma vrrestriction_Set_is_arr'[cat_cs_intros]:
  assumes "T : A cat_Set αB" 
    and " (TArrVal)  C" 
    and "C  cat_Set αObj"
    and "C' = C"
    and "ℭ' = cat_Set α"
  shows "T rSet C : A ℭ'C'"
  using assms(1-3) unfolding assms(4,5) by (rule vrrestriction_Set_is_arr)


subsubsection‹Further properties›

lemma 
  assumes "T : A cat_Set αB" 
  shows vrrestriction_Set_vrange_is_arr: 
      "T rSet  (TArrVal) : A cat_Set α (TArrVal)"
    and vrrestriction_Set_vrange_ArrVal_app[cat_cs_simps, cat_Set_cs_simps]: 
      "(T rSet  (TArrVal))ArrVal = TArrVal"
proof(intro vrrestriction_Set_is_arr, rule assms)
  note TD = cat_Set_is_arrD[OF assms(1)]
  interpret T: arr_Set α T
    rewrites "TArrDom = A" and "TArrCod = B"
    by (intro TD)+
  show " (TArrVal)  cat_Set αObj"
    by (auto simp: cat_Set_components T.arr_Rel_ArrVal_in_Vset vrange_in_VsetI)
qed (auto intro: vrrestriction_Set_ArrVal_app[OF assms])

lemma (in 𝒵) vrrestriction_Set_vrange_is_iso_arr:
  assumes "T : A moncat_Set αB" 
  shows "T rSet  (TArrVal) : A isocat_Set α (TArrVal)"
proof-
  note cat_Set_is_monic_arrD[OF assms]
  note TD = this cat_Set_is_arrD[OF this(1)]
  interpret T: arr_Set α T by (intro TD)+
  show ?thesis
    by 
      (
        intro cat_Set_is_iso_arrI vrrestriction_Set_vrange_is_arr[OF TD(1)],
        unfold cat_Set_cs_simps
      )
      (simp_all add: TD(2,3))
qed


subsubsection‹Connections›

lemma cat_Set_Comp_vrrestriction_Set:
  assumes "S : B cat_Set αC" 
    and "T : A cat_Set αB" 
    and " (SArrVal)  D"
    and "D  cat_Set αObj"
  shows "S rSet D Acat_Set αT = (S Acat_Set αT) rSet D"
proof-

  note SD = cat_Set_is_arrD[OF assms(1)]
  interpret S: arr_Set α S 
    rewrites [cat_cs_simps]: "SArrDom = B" and [cat_cs_simps]: "SArrCod = C"
    by (intro SD)+
  note TD = cat_Set_is_arrD[OF assms(2)]
  interpret T: arr_Set α T 
    rewrites [cat_cs_simps]: "TArrDom = A" and [cat_cs_simps]: "TArrCod = B"
    by (intro TD)+

  from assms(3) S.arr_Par_ArrVal_vrange have RS_D: " (SArrVal)  D" by auto

  from assms(1,2) have "S Acat_Set αT : A cat_Set αC"
    by (cs_concl cs_intro: cat_cs_intros)

  from assms(1,2) have " ((S Acat_Set αT)ArrVal)   (SArrVal)" 
    by (intro cat_Set_Comp_ArrVal_vrange)
  with assms(3) have RST: " ((S Acat_Set αT)ArrVal)  D" by auto

  from assms(1,2,4) RS_D have SD_T: 
    "S rSet D Acat_Set αT : A cat_Set αD" 
    by (cs_concl cs_intro: cat_cs_intros) 
  then have dom_lhs: "𝒟 ((S rSet D Acat_Set αT)ArrVal) = A"
    by (simp add: cat_cs_simps)

  from assms(1,2,4) RST have ST_D: 
    "(S Acat_Set αT) rSet D : A cat_Set αD"
    by (cs_concl cs_intro: cat_cs_intros)
  then have dom_rhs: "𝒟 (((S Acat_Set αT) rSet D)ArrVal) = A"
    by (simp add: cat_cs_simps)

  show "S rSet D Acat_Set αT = (S Acat_Set αT) rSet D"
  proof(rule arr_Set_eqI[of α])
    show 
      "(S rSet D Acat_Set αT)ArrVal =
        ((S Acat_Set αT) rSet D)ArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume "a  A"
      with assms(1,2,4) RST RS_D show
        "(S rSet D Acat_Set αT)ArrVala = 
          ((S Acat_Set αT) rSet D)ArrVala"
        by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    qed (use SD_T ST_D in auto dest: cat_Set_is_arrD) 
  qed (use SD_T ST_D in auto simp: cat_Set_is_arrD) 

qed

lemma (in 𝒵) cat_Set_CId_vrrestriction_Set[cat_cs_simps]:
  assumes "A  B" and "B  cat_Set αObj"
  shows "cat_Set αCIdA rSet B = incl_Set A B"
proof-

  from assms have A: "A  cat_Set αObj"
    unfolding cat_Set_components by auto
  from A have CId_A: "cat_Set αCIdA : A cat_Set αA"
    by (cs_concl cs_intro: cat_cs_intros)
  with cat_Set_is_arrD[OF CId_A] assms(1) have RA_B:
    " (cat_Set αCIdAArrVal)  B"
    by (auto intro: arr_Set.arr_Set_ArrVal_vrange)

  with assms A assms(1,2) have lhs_is_arr:
    "cat_Set αCIdA rSet B : A cat_Set αB"
    by (cs_concl cs_intro: cat_cs_intros)
  then have dom_lhs: "𝒟 ((cat_Set αCIdA rSet B)ArrVal) = A"
    by (simp add: cat_cs_simps)

  from A assms(1,2) have rhs_is_arr: "incl_Set A B : A cat_Set αB"
    by (cs_concl cs_intro: cat_cs_intros)
  then have dom_rhs: "𝒟 ((incl_Set A B)ArrVal) = A"
    by (simp add: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    show "(cat_Set αCIdA rSet B)ArrVal = incl_Rel A BArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume "a  A"
      with A RA_B show 
        "(cat_Set αCIdA rSet B)ArrVala = incl_Rel A BArrVala"
        by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    qed (use lhs_is_arr rhs_is_arr in auto dest: cat_Set_is_arrD)
  qed (use lhs_is_arr rhs_is_arr in auto simp: cat_Set_is_arrD)

qed

lemma cat_Set_Comp_incl_Rel_vrrestriction_Set[cat_cs_simps]:
  assumes "F : A cat_Set αB" and "C  B" and " (FArrVal)  C"
  shows "incl_Rel C B Acat_Set αF rSet C = F"
proof-
  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret F: arr_Set α F 
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+
  from assms(2) have C: "C  cat_Set αObj"
    unfolding cat_Set_components(1) by (auto intro: F.arr_Par_ArrCod_in_Vset)
  from assms C have lhs_is_arr:
    "incl_Rel C B Acat_Set αF rSet C : A cat_Set αB"
    by (cs_concl cs_intro: cat_cs_intros)
  then have dom_lhs: "𝒟 ((incl_Rel C B Acat_Set αF rSet C)ArrVal) = A"
    by (cs_concl cs_simp: cat_cs_simps)
  from assms(1) have dom_rhs: "𝒟 (FArrVal) = A" 
    by (cs_concl cs_simp: cat_cs_simps)
  show ?thesis
  proof(rule arr_Set_eqI[of α])
    show "(incl_Rel C B Acat_Set αF rSet C)ArrVal = FArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume prems: "a  A"
      with assms F.ArrVal.vsv_vimageI2 have "FArrVala  C"
        by (auto simp: F.arr_Set_ArrVal_vdomain)
      with prems assms C show 
        "(incl_Rel C B Acat_Set αF rSet C)ArrVala = FArrVala"
        by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    qed (use assms(1) lhs_is_arr in auto dest: cat_Set_is_arrD)
  qed (use assms(1) lhs_is_arr in auto dest: cat_Set_is_arrD)
qed



subsection‹Projection arrows for vtimes›


subsubsection‹Definition and elementary properties›

definition vfst_arrow :: "V  V  V"
  where "vfst_arrow A B = [(λabA × B. vfst ab), A × B, A]"

definition vsnd_arrow :: "V  V  V"
  where "vsnd_arrow A B = [(λabA × B. vsnd ab), A × B, B]"


text‹Components.›

lemma vfst_arrow_components: 
  shows "vfst_arrow A BArrVal = (λabA × B. vfst ab)"
    and [cat_cs_simps]: "vfst_arrow A BArrDom = A × B"
    and [cat_cs_simps]: "vfst_arrow A BArrCod = A"
  unfolding vfst_arrow_def arr_field_simps by (simp_all add: nat_omega_simps)

lemma vsnd_arrow_components: 
  shows "vsnd_arrow A BArrVal = (λabA × B. vsnd ab)"
    and [cat_cs_simps]: "vsnd_arrow A BArrDom = A × B"
    and [cat_cs_simps]: "vsnd_arrow A BArrCod = 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  A × B"
  shows "vfst_arrow A BArrValab = a"
  using assms(2) unfolding assms(1) by (simp add: vfst_arrow_ArrVal_app')

lemma vfst_arrow_vrange: " (vfst_arrow A BArrVal)  A"
  unfolding vfst_arrow_components
proof(intro vrange_VLambda_vsubset)
  fix ab assume "ab  A × B"
  then obtain a b where ab_def: "ab = a, b" and a: "a  A" by clarsimp
  from a show "vfst ab  A" unfolding ab_def by simp
qed

lemma vsnd_arrow_ArrVal_app[cat_cs_simps]:
  assumes "ab = a, b" and "ab  A × B"
  shows "vsnd_arrow A BArrValab = b"
  using assms(2) unfolding assms(1) by (simp add: vsnd_arrow_ArrVal_app')

lemma vsnd_arrow_vrange: " (vsnd_arrow A BArrVal)  B"
  unfolding vsnd_arrow_components
proof(intro vrange_VLambda_vsubset)
  fix ab assume "ab  A × B"
  then obtain a b where ab_def: "ab = a, b" and b: "b  B" by clarsimp
  from b show "vsnd ab  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  Vset α" and "B  Vset α"
  shows "vfst_arrow A B : A × B cat_Set α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"
    unfolding vfst_arrow_def by (simp add: nat_omega_simps)
  show " (vfst_arrow A BArrVal)  A" by (rule vfst_arrow_vrange)
qed (use assms in cs_concl cs_shallow cs_intro: V_cs_intros cat_cs_intros)+

lemma (in 𝒵) vfst_arrow_is_cat_Set_arr:
  assumes "A  cat_Set αObj" and "B  cat_Set αObj"
  shows "vfst_arrow A B : A × B cat_Set α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  cat_Set αObj" 
    and "B  cat_Set αObj"
    and "AB = A × B"
    and "A' = A"
    and "ℭ' = cat_Set α"
  shows "vfst_arrow A B : AB ℭ'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  Vset α" and "B  Vset α"
  shows "vsnd_arrow A B : A × B cat_Set α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"
    unfolding vsnd_arrow_def by (simp add: nat_omega_simps)
  show " (vsnd_arrow A BArrVal)  B" by (rule vsnd_arrow_vrange)
qed (use assms in cs_concl cs_shallow cs_intro: V_cs_intros cat_cs_intros)+

lemma (in 𝒵) vsnd_arrow_is_cat_Set_arr:
  assumes "A  cat_Set αObj" and "B  cat_Set αObj"
  shows "vsnd_arrow A B : A × B cat_Set α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  cat_Set αObj" 
    and "B  cat_Set αObj"
    and "AB = A × B"
    and "B' = B"
    and "ℭ' = cat_Set α"
  shows "vsnd_arrow A B : AB ℭ'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  cat_Par αObj" and "B  cat_Par αObj"
  shows "vfst_arrow A B : A × B cat_Par α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  cat_Par αObj" 
    and "B  cat_Par αObj"
    and "AB = A × B"
    and "A' = A"
    and "ℭ' = cat_Par α"
  shows "vfst_arrow A B : AB ℭ'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  cat_Par αObj" and "B  cat_Par αObj"
  shows "vsnd_arrow A B : A × B cat_Par α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  cat_Par αObj" 
    and "B  cat_Par αObj"
    and "AB = A × B"
    and "B' = B"
    and "ℭ' = cat_Par α"
  shows "vsnd_arrow A B : AB ℭ'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  cat_Rel αObj" and "B  cat_Rel αObj"
  shows "vfst_arrow A B : A × B cat_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: 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  cat_Rel αObj" 
    and "B  cat_Rel αObj"
    and "AB = A × B"
    and "A' = A"
    and "ℭ' = cat_Rel α"
  shows "vfst_arrow A B : AB ℭ'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  cat_Rel αObj" and "B  cat_Rel αObj"
  shows "vsnd_arrow A B : A × B cat_Rel α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  cat_Rel αObj" 
    and "B  cat_Rel αObj"
    and "AB = A × B"
    and "B' = B"
    and "ℭ' = cat_Rel α"
  shows "vsnd_arrow A B : AB ℭ'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_iso_arr_Vset:
  assumes "A  Vset α" and "b  Vset α"
  shows "vfst_arrow A (set {b}) : A × set {b} isocat_Set αA"
proof
  (
    intro 
      cat_Set_is_iso_arrI 
      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  A × set {b}"
      "ab'  A × set {b}"
      "vfst_arrow A (set {b})ArrValab = vfst_arrow A (set {b})ArrValab'"
    from prems obtain a where ab_def: "ab = a, b" and a: "a  A" 
      by clarsimp
    from prems obtain a' where ab'_def: "ab' = a', b" and a': "a'  A" 
      by clarsimp
    from prems(3) a a' have "a = a'"
      unfolding ab_def ab'_def
      by (cs_prems cs_shallow 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_shallow cs_intro: cat_cs_intros)
  show " (vfst_arrow A (set {b})ArrVal) = A"
  proof(intro vsubset_antisym)
    show "A   (vfst_arrow A (set {b})ArrVal)"
    proof(intro vsubsetI)
      fix a assume a: "a  A"
      then have a_def: "a = vfst_arrow A (set {b})ArrVala, b"
        by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: V_cs_intros)
      from a assms show "a   (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_iso_arr:
  assumes "A  cat_Set αObj" and "b  cat_Set αObj"
  shows "vfst_arrow A (set {b}) : A × set {b} isocat_Set αA"
  using assms 
  unfolding cat_Set_components 
  by (rule vfst_arrow_is_cat_Set_iso_arr_Vset)

lemma (in 𝒵) vfst_arrow_is_cat_Set_iso_arr'[cat_rel_par_Set_cs_intros]:
  assumes "A  cat_Set αObj" 
    and "b  cat_Set αObj"
    and "AB = A × set {b}"
    and "A' = A"
    and "ℭ' = cat_Set α"
  shows "vfst_arrow A (set {b}) : AB isoℭ'A"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vfst_arrow_is_cat_Set_iso_arr)

lemmas [cat_rel_par_Set_cs_intros] = 𝒵.vfst_arrow_is_cat_Set_iso_arr'

lemma (in 𝒵) vsnd_arrow_is_cat_Set_iso_arr_Vset:
  assumes "a  Vset α" and "B  Vset α"
  shows "vsnd_arrow (set {a}) B : set {a} × B isocat_Set αB"
proof
  (
    intro 
      cat_Set_is_iso_arrI 
      arr_SetI 
      vsnd_arrow_is_cat_Set_arr_Vset 
      assms,
    unfold cat_cs_simps
  )
  show "v11 (vsnd_arrow (set {a}) BArrVal)"
  proof(rule vsv.vsv_valeq_v11I, unfold cat_cs_simps)
    fix ab ab' assume prems:
      "ab  set {a} × B"
      "ab'  set {a} × B"
      "vsnd_arrow (set {a}) BArrValab = vsnd_arrow (set {a}) BArrValab'"
    from prems obtain b where ab_def: "ab = a, b" and b: "b  B" 
      by clarsimp
    from prems obtain b' where ab'_def: "ab' = a, b'" and b': "b'  B" 
      by clarsimp
    from prems(3) b b' have "b = b'"
      unfolding ab_def ab'_def
      by (cs_prems cs_shallow 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_shallow cs_intro: cat_cs_intros)
  show " (vsnd_arrow (set {a}) BArrVal) = B"
  proof(intro vsubset_antisym)
    show "B   (vsnd_arrow (set {a}) BArrVal)"
    proof(intro vsubsetI)
      fix b assume b: "b  B"
      then have b_def: "b = vsnd_arrow (set {a}) BArrVala, b"
        by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: V_cs_intros)
      from b assms show "b   (vsnd_arrow (set {a}) BArrVal)"
        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_iso_arr:
  assumes "a  cat_Set αObj" and "B  cat_Set αObj"
  shows "vsnd_arrow (set {a}) B : set {a} × B isocat_Set αB"
  using assms 
  unfolding cat_Set_components 
  by (rule vsnd_arrow_is_cat_Set_iso_arr_Vset)

lemma (in 𝒵) vsnd_arrow_is_cat_Set_iso_arr'[cat_rel_par_Set_cs_intros]:
  assumes "a  cat_Set αObj" 
    and "B  cat_Set αObj"
    and "AB = set {a} × B"
    and "A' = A"
    and "ℭ' = cat_Set α"
  shows "vsnd_arrow (set {a}) B : AB isoℭ'B"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vsnd_arrow_is_cat_Set_iso_arr)

lemmas [cat_rel_par_Set_cs_intros] = 𝒵.vsnd_arrow_is_cat_Set_iso_arr'


subsubsection‹Projection arrows are isomorphisms in the category Par›

lemma (in 𝒵) vfst_arrow_is_cat_Par_iso_arr:
  assumes "A  cat_Par αObj" and "b  cat_Par αObj"
  shows "vfst_arrow A (set {b}) : A × set {b} isocat_Par α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 × set {b} isocat_Par αA"
    by 
      (
        rule Set_Par.wr_subcat_is_iso_arr_is_iso_arr
          [
            THEN iffD1, 
            OF vfst_arrow_is_cat_Set_iso_arr_Vset[
              OF assms[unfolded cat_Par_components]
              ]
          ]
      )
qed

lemma (in 𝒵) vfst_arrow_is_cat_Par_iso_arr'[cat_rel_Par_set_cs_intros]:
  assumes "A  cat_Par αObj" 
    and "b  cat_Par αObj"
    and "AB = A × set {b}"
    and "A' = A"
    and "ℭ' = cat_Par α"
  shows "vfst_arrow A (set {b}) : AB isoℭ'A"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vfst_arrow_is_cat_Par_iso_arr)

lemmas [cat_rel_Par_set_cs_intros] = 𝒵.vfst_arrow_is_cat_Par_iso_arr'

lemma (in 𝒵) vsnd_arrow_is_cat_Par_iso_arr:
  assumes "a  cat_Par αObj" and "B  cat_Par αObj"
  shows "vsnd_arrow (set {a}) B : set {a} × B isocat_Par α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} × B isocat_Par αB"
    by 
      (
        rule Set_Par.wr_subcat_is_iso_arr_is_iso_arr
          [
            THEN iffD1, 
            OF vsnd_arrow_is_cat_Set_iso_arr_Vset[
              OF assms[unfolded cat_Par_components]
              ]
          ]
      )
qed

lemma (in 𝒵) vsnd_arrow_is_cat_Par_iso_arr'[cat_rel_Par_set_cs_intros]:
  assumes "a  cat_Par αObj" 
    and "B  cat_Par αObj"
    and "AB = set {a} × B"
    and "A' = A"
    and "ℭ' = cat_Par α"
  shows "vsnd_arrow (set {a}) B : AB isoℭ'B"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vsnd_arrow_is_cat_Par_iso_arr)

lemmas [cat_rel_Par_set_cs_intros] = 𝒵.vsnd_arrow_is_cat_Par_iso_arr'


subsubsection‹Projection arrows are isomorphisms in the category Rel›

lemma (in 𝒵) vfst_arrow_is_cat_Rel_iso_arr:
  assumes "A  cat_Rel αObj" and "b  cat_Rel αObj"
  shows "vfst_arrow A (set {b}) : A × set {b} isocat_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
    by 
      (
        rule Set_Rel.wr_subcat_is_iso_arr_is_iso_arr
          [
            THEN iffD1, 
            OF vfst_arrow_is_cat_Set_iso_arr_Vset[
              OF assms[unfolded cat_Rel_components]
              ]
          ]
      )
qed

lemma (in 𝒵) vfst_arrow_is_cat_Rel_iso_arr'[cat_Rel_par_set_cs_intros]:
  assumes "A  cat_Rel αObj" 
    and "b  cat_Rel αObj"
    and "AB = A × set {b}"
    and "A' = A"
    and "ℭ' = cat_Rel α"
  shows "vfst_arrow A (set {b}) : AB isoℭ'A"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vfst_arrow_is_cat_Rel_iso_arr)

lemmas [cat_Rel_par_set_cs_intros] = 𝒵.vfst_arrow_is_cat_Rel_iso_arr'

lemma (in 𝒵) vsnd_arrow_is_cat_Rel_iso_arr:
  assumes "a  cat_Rel αObj" and "B  cat_Rel αObj"
  shows "vsnd_arrow (set {a}) B : set {a} × B isocat_Rel α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_iso_arr_is_iso_arr
          [
            THEN iffD1, 
            OF vsnd_arrow_is_cat_Set_iso_arr_Vset[
              OF assms[unfolded cat_Rel_components]
              ]
          ]
      )
qed

lemma (in 𝒵) vsnd_arrow_is_cat_Rel_iso_arr'[cat_Rel_par_set_cs_intros]:
  assumes "a  cat_Rel αObj" 
    and "B  cat_Rel αObj"
    and "AB = set {a} × B"
    and "A' = A"
    and "ℭ' = cat_Rel α"
  shows "vsnd_arrow (set {a}) B : AB isoℭ'B"
  using assms(1-2) 
  unfolding assms(3-5)
  by (rule vsnd_arrow_is_cat_Rel_iso_arr)

lemmas [cat_Rel_par_set_cs_intros] = 𝒵.vsnd_arrow_is_cat_Rel_iso_arr'



subsection‹Projection arrow for vproduct›

definition vprojection_arrow :: "V  (V  V)  V  V"
  where "vprojection_arrow I A i = [vprojection I A i, (iI. A i), A i]"


text‹Components.›

lemma vprojection_arrow_components:
  shows "vprojection_arrow I A iArrVal = vprojection I A i"
    and "vprojection_arrow I A iArrDom = (iI. A i)"
    and "vprojection_arrow I A iArrCod = 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_ArrVal_vsv[cat_Set_cs_intros]|
  |vdomain vprojection_arrow_ArrVal_vdomain[cat_Set_cs_simps]|
  |app vprojection_arrow_ArrVal_app[cat_Set_cs_simps]|


subsubsection‹Projection arrow is an arrow in the category Set›

lemma (in 𝒵) arr_Set_vprojection_arrow:
  assumes "i  I" and "VLambda I A  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"
    unfolding vprojection_arrow_def by (simp add: nat_omega_simps)
  show "vprojection_arrow I A iArrCod  Vset α"
    unfolding vprojection_arrow_components
  proof-
    from assms(1) have "i  I" by simp
    then have "A i   (VLambda I A)" by auto
    moreover from assms(2) have " (VLambda I A)  Vset α"
      by (meson vrange_in_VsetI)
    ultimately show "A i  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  I" and "VLambda I A  Vset α"
  shows "vprojection_arrow I A i : (iI. A i) cat_Set α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‹Canonical injection arrow for vdunion›

definition vcinjection_arrow :: "V  (V  V)  V  V"
  where "vcinjection_arrow I A i = [vcinjection A i, A i, (iI. A i)]"


text‹Components.›

lemma vcinjection_arrow_components:
  shows "vcinjection_arrow I A iArrVal = vcinjection A i"
    and "vcinjection_arrow I A iArrDom = A i"
    and "vcinjection_arrow I A iArrCod = (iI. A i)"
  unfolding vcinjection_arrow_def arr_field_simps
  by (simp_all add: nat_omega_simps)


subsubsection‹Canonical injection arrow value›

mk_VLambda vcinjection_arrow_components(1)[unfolded vcinjection_def]
  |vsv vcinjection_arrow_ArrVal_vsv[cat_Set_cs_intros]|
  |vdomain vcinjection_arrow_ArrVal_vdomain[cat_Set_cs_simps]|
  |app vcinjection_arrow_ArrVal_app[cat_Set_cs_simps]|


subsubsection‹Canonical injection arrow is an arrow in the category Set›

lemma (in 𝒵) arr_Set_vcinjection_arrow:
  assumes "i  I" and "VLambda I A  Vset α"
  shows "arr_Set α (vcinjection_arrow I A i)"
proof(intro arr_SetI)
  show "vfsequence (vcinjection_arrow I A i)"
    unfolding vcinjection_arrow_def by auto
  show "vcard (vcinjection_arrow I A i) = 3"
    unfolding vcinjection_arrow_def by (simp add: nat_omega_simps)
  show "vcinjection_arrow I A iArrDom  Vset α"
    unfolding vcinjection_arrow_components
  proof-
    from assms(1) have Ai_def: "A i = VLambda I Ai" by simp
    with assms(1) have "A i   (VLambda I A)" by auto
    with assms(2) Limit_α show "A i  Vset α"
      unfolding Ai_def by (auto intro: vrange_in_VsetI)
  qed
  show "vcinjection_arrow I A iArrCod  Vset α"
    unfolding vcinjection_arrow_components
    by (intro Limit_vdunion_in_Vset_if_VLambda_in_VsetI Limit_α assms)
qed 
  (
    auto 
      simp: vcinjection_arrow_components 
      intro!: assms vcinjection_vrange_vsubset 
  )

lemma (in 𝒵) vcinjection_arrow_is_arr:
  assumes "i  I" and "VLambda I A  Vset α"
  shows "vcinjection_arrow I A i : A i cat_Set α(iI. A i)"
proof(intro cat_Set_is_arrI)
  from assms show "arr_Set α (vcinjection_arrow I A i)"
    by (rule arr_Set_vcinjection_arrow)
qed (simp_all add: vcinjection_arrow_components)

lemma (in 𝒵) vcinjection_arrow_is_arr'[cat_cs_intros]:
  assumes "i  I" 
    and "VLambda I A  Vset α"
    and "A' = A i"
    and "ℭ' = cat_Set α"
    and "P' = (iI. A i)"
  shows "vcinjection_arrow I A i : A' ℭ'P'"
  using assms(1,2) unfolding assms(3-5) by (rule vcinjection_arrow_is_arr)



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  S  b, d  T}"

lemma small_prod_2_Rel_ArrVal[simp]:
  "small {a, b, c, d | a b c d. a, c  S  b, d  T}"
  (is small ?S)
proof(rule down)
  show "?S  elts ((𝒟 S × 𝒟 T) × ( S ×  T))" by auto
qed


text‹Rules.›

lemma prod_2_Rel_ArrValI:
  assumes "ab_cd = a, b, c, d"
    and "a, c  S"
    and "b, d  T"
  shows "ab_cd  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  prod_2_Rel_ArrVal S T"
  shows "a, c  S" and "b, d  T"
  using assms unfolding prod_2_Rel_ArrVal_def by auto

lemma prod_2_Rel_ArrValE[elim!]:
  assumes "ab_cd  prod_2_Rel_ArrVal S T"
  obtains a b c d where "ab_cd = a, b, c, d" 
    and "a, c  S"
    and "b, d  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  ((𝒟 S × 𝒟 T) × ( S ×  T))"
  by (intro vsubsetI) 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: "𝒟 (prod_2_Rel_ArrVal S T) = 𝒟 S × 𝒟 T"
proof(intro vsubset_antisym)
  show "𝒟 S × 𝒟 T  𝒟 (prod_2_Rel_ArrVal S T)"
  proof(intro vsubsetI)
    fix ab assume "ab  𝒟 S × 𝒟 T"
    then obtain a b
      where ab_def: "ab = a, b" 
        and "a  𝒟 S"
        and "b  𝒟 T"
      by auto
    then obtain c d where "a, c  S" and "b, d  T" by force
    then have "a, b, c, d  prod_2_Rel_ArrVal S T"
      by (intro prod_2_Rel_ArrValI) auto
    then show "ab  𝒟 (prod_2_Rel_ArrVal S T)"
      unfolding ab_def by (simp add: app_vdomainI)
  qed
qed (use prod_2_Rel_ArrVal_vsubset_vprod in blast)

lemma prod_2_Rel_ArrVal_vrange: " (prod_2_Rel_ArrVal S T) =  S ×  T"
proof(intro vsubset_antisym)
  show " S ×  T   (prod_2_Rel_ArrVal S T)"
  proof(intro vsubsetI)
    fix cd assume "cd   S ×  T"
    then obtain c d
      where cd_def: "cd = c, d" 
        and "c   S"
        and "d   T"
      by auto
    then obtain a b where "a, c  S" and "b, d  T" by force
    then have "a, b, c, d  prod_2_Rel_ArrVal S T"
      by (intro prod_2_Rel_ArrValI) auto
    then show "cd   (prod_2_Rel_ArrVal S T)"
      unfolding cd_def by (simp add: app_vrangeI)
  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  𝒟 g; b  𝒟 f   
        prod_2_Rel_ArrVal g fa,b = ga, fb"
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  𝒟 g" "b  𝒟 f"
  then have a_ga: "a, ga  g" and b_fb: "b, fb  f" by auto
  from a_ga b_fb show "prod_2_Rel_ArrVal g fa, b = ga, fb"
    by 
      (
        cs_concl cs_shallow 
          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  𝒟 g × 𝒟 f"  
      "cd  𝒟 g × 𝒟 f"
      "prod_2_Rel_ArrVal g fab = prod_2_Rel_ArrVal g fcd"
    from prems(1) obtain a b
      where ab_def: "ab = a, b" and a: "a  𝒟 g" and b: "b  𝒟 f"
      by auto
    from prems(2) obtain c d
      where cd_def: "cd = c, d" and c: "c  𝒟 g" and d: "d  𝒟 f"
      by auto
    from prems(3) a b c d have "ga, fb = gc, fd"
      unfolding ab_def cd_def
      by 
        (
          cs_prems cs_shallow 
            cs_simp: prod_2_Rel_ArrVal_app cs_intro: V_cs_intros
        )
    then have "ga = gc" and "fb = fd" 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'  prod_2_Rel_ArrVal S T =
    prod_2_Rel_ArrVal (S'  S) (T'  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'  prod_2_Rel_ArrVal S' T'  prod_2_Rel_ArrVal S T"
    then obtain aa' bb' cc' where ac_def: "aa'_cc' = aa', cc'" 
      and bc: "bb', cc'  prod_2_Rel_ArrVal S' T'"
      and ab: "aa', bb'  prod_2_Rel_ArrVal S T"
      by (elim vcompE)
    from bc obtain b b' c c' 
      where bb'_cc'_def: "bb', cc' = b, b', c, c'"
        and bc: "b, c  S'"
        and bc': "b', c'  T'"
      by auto
    with ab obtain a a' 
      where aa'_bb'_def: "aa', bb' = a, a', b, b'"
        and ab: "a, b  S"
        and ab': "a', b'  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'  prod_2_Rel_ArrVal (S'  S) (T'  T)"
      unfolding ac_def aa'_def cc'_def
      by (intro prod_2_Rel_ArrValI)
        (cs_concl cs_shallow cs_intro: prod_2_Rel_ArrValI vcompI)+
  next
    fix aa'_cc' assume "aa'_cc'  prod_2_Rel_ArrVal (S'  S) (T'  T)"
    then obtain a a' c c'
      where aa'_cc'_def: "aa'_cc' = a, a', c, c'"
        and ac: "a, c  S'  S"
        and ac': "a', c'  T'  T"
      by blast
    from ac obtain b where ab: "a, b  S" and bc: "b, c  S'" 
      by auto
    from ac' obtain b' where ab': "a', b'  T" and bc': "b', c'  T'" 
      by auto
    from ab bc ab' bc' show 
      "aa'_cc'  prod_2_Rel_ArrVal S' T'  prod_2_Rel_ArrVal S T"
      unfolding aa'_cc'_def 
      by (cs_concl cs_shallow 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 × 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" (infixr A×Rel 80)
  where "prod_2_Rel S T =
    [
      prod_2_Rel_ArrVal (SArrVal) (TArrVal),
      SArrDom × TArrDom,
      SArrCod × TArrCod
    ]"

abbreviation (input) prod_2_Par :: "V  V  V" (infixr A×Par 80)
  where "prod_2_Par  prod_2_Rel"
abbreviation (input) prod_2_Set :: "V  V  V" (infixr A×Set 80)
  where "prod_2_Set  prod_2_Rel"


text‹Components.›

lemma prod_2_Rel_components: 
  shows "(S A×Rel T)ArrVal = prod_2_Rel_ArrVal (SArrVal) (TArrVal)"
    and [cat_cs_simps]: "(S A×Rel T)ArrDom = SArrDom × TArrDom"
    and [cat_cs_simps]: "(S A×Rel T)ArrCod = SArrCod × TArrCod"
  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 cat_Rel αB" and "T : C cat_Rel αD"    
  shows "S A×Rel T : A × C cat_Rel αB × 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]: "SArrDom = A" and [simp]: "SArrCod = B"
    by (simp_all add: S)
  interpret T: arr_Rel α T 
    rewrites [simp]: "TArrDom = C" and [simp]: "TArrCod = D"
    by (simp_all add: T)
  show ?thesis
  proof(intro cat_Rel_is_arrI arr_RelI)
    show "vfsequence (S A×Rel T)"
      unfolding prod_2_Rel_def by simp
    show "vcard (S A×Rel T) = 3"
      unfolding prod_2_Rel_def by (simp add: nat_omega_simps)
    from S have "𝒟 (SArrVal)  A" and " (SArrVal)  B" by auto
    moreover from T have "𝒟 (TArrVal)  C" and " (TArrVal)  D" 
      by auto
    ultimately have 
      "𝒟 (SArrVal) × 𝒟 (TArrVal)  A × C"
      " (SArrVal) ×  (TArrVal)  B × D"
      by auto
    then show 
      "𝒟 ((S A×Rel T)ArrVal)  (S A×Rel T)ArrDom"
      " ((S A×Rel T)ArrVal)  (S A×Rel 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 "(S A×Rel T)ArrDom  Vset α" "(S A×Rel T)ArrCod  Vset α"
      unfolding prod_2_Rel_components 
      by (allintro 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 cat_Rel αB"
    and "T : C cat_Rel αD"
    and "A' = A × C"
    and "B' = B × D"
    and "ℭ' = cat_Rel α"
  shows "S A×Rel T : A' ℭ'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 cat_Set αB" 
    and "T : C cat_Set αD"    
    and "a  A"
    and "c  C"
    and "ac = a, c"
  shows "(S A×Set T)ArrValac = SArrVala, TArrValc"
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]: "SArrDom = A" and [simp]: "SArrCod = B"
    by (simp_all add: S)
  interpret T: arr_Set α T 
    rewrites [simp]: "TArrDom = C" and [simp]: "TArrCod = D"
    by (simp_all add: T)
  from assms(3,4) show ?thesis
    unfolding prod_2_Rel_components(1) assms(5)
    by 
      (
        cs_concl cs_shallow
          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 cat_Set αB" and "T : C cat_Set αD"    
  shows "S A×Set T : A × C cat_Set αB × 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]: "SArrDom = A" and [simp]: "SArrCod = B"
    by (simp_all add: S)
  interpret T: arr_Set α T 
    rewrites [simp]: "TArrDom = C" and [simp]: "TArrCod = D"
    by (simp_all add: T)

  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI)
    show "vfsequence (S A×Set T)"
      unfolding prod_2_Rel_def by simp
    show "vcard (S A×Set T) = 3"
      unfolding prod_2_Rel_def by (simp add: nat_omega_simps)
    from S.arr_Set_ArrVal_vrange T.arr_Set_ArrVal_vrange show 
      " ((S A×Set T)ArrVal)  (S A×Set 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 
      "(S A×Set T)ArrDom  Vset α"
      by (cs_concl cs_shallow 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 
      "(S A×Set T)ArrCod  Vset α"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: V_cs_intros)
    from assms show "(S A×Set T)ArrDom = A × C"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    from assms show "(S A×Set T)ArrCod = B × D"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    show "vsv ((S A×Set 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 cat_Set αB" 
    and "T : C cat_Set αD"
    and "AC = A × C"
    and "BD = B × D"
    and "ℭ' = cat_Set α"
  shows "S A×Set T : AC ℭ'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_iso_arr:
  assumes "S : A isocat_Set αB" and "T : C isocat_Set αD"    
  shows "S A×Set T : A × C isocat_Set αB × D"
proof-
  note S = cat_Set_is_iso_arrD[OF assms(1)]
  note T = cat_Set_is_iso_arrD[OF assms(2)]
  show ?thesis
  proof
    (
      intro cat_Set_is_iso_arrI prod_2_Rel_is_cat_Set_arr[OF S(1) T(1)], 
      unfold prod_2_Rel_components
    )
    show "𝒟 (prod_2_Rel_ArrVal (SArrVal) (TArrVal)) = A × C"
      unfolding prod_2_Rel_ArrVal_vdomain
      by (cs_concl cs_shallow cs_simp: S(3) T(3) cs_intro: cat_cs_intros)
    show " (prod_2_Rel_ArrVal (SArrVal) (TArrVal)) = B × D"
      unfolding prod_2_Rel_ArrVal_vrange
      by (cs_concl cs_shallow cs_simp: S(4) T(4) cs_intro: cat_cs_intros)
  qed (use S(2) T(2) in cs_concl cs_shallow cs_intro: prod_2_Rel_ArrVal_v11)
qed

lemma prod_2_Rel_is_cat_Set_iso_arr'[cat_rel_par_Set_cs_intros]:
  assumes "S : A isocat_Set αB" 
    and "T : C isocat_Set αD"    
    and "AC = A × C"
    and "BD = B × D"
    and "ℭ' = cat_Set α"
  shows "S A×Set T : AC isoℭ'BD"
  using assms(1,2) 
  unfolding assms(3-5) 
  by (rule prod_2_Rel_is_cat_Set_iso_arr)


subsubsection‹Further elementary properties›

lemma prod_2_Rel_Comp:
  assumes "G' : B' cat_Rel αB''" 
    and "F' : A' cat_Rel αA''" 
    and "G : B cat_Rel αB'"
    and "F : A cat_Rel αA'"
  shows
    "G' A×Rel F' Acat_Rel αG A×Rel F =
      (G' Acat_Rel αG) A×Rel (F' Acat_Rel α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: 
    "G' A×Rel F' Acat_Rel αG A×Rel F :
      B × A cat_Rel αB'' × A''"
    by (cs_concl cs_shallow cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros)
  from assms Rel.category_axioms have GG'_FF':
    "(G' Acat_Rel αG) A×Rel (F' Acat_Rel αF) : 
      B × A cat_Rel αB'' × A''"
    by (cs_concl cs_shallow 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 α (G' A×Rel F' Acat_Rel αG A×Rel F)"
      by (auto dest: cat_Rel_is_arrD(1))
    from GG'_FF' show arr_Rel_GG'_FF':
      "arr_Rel α ((G' Acat_Rel αG) A×Rel (F' Acat_Rel αF))"
      by (auto dest: cat_Rel_is_arrD(1))
    show "(G' A×Rel F' Acat_Rel αG A×Rel F)ArrVal = 
      ((G' Acat_Rel αG) A×Rel (F' Acat_Rel αF))ArrVal"
    proof(intro vsubset_antisym vsubsetI)
      fix R assume
        "R  (G' A×Rel F' Acat_Rel αG A×Rel F)ArrVal"
      from this assms have "R 
        prod_2_Rel_ArrVal (G'ArrVal) (F'ArrVal) 
        prod_2_Rel_ArrVal (GArrVal) (FArrVal)"
        by 
          (
            cs_prems cs_shallow
              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  ((G' Acat_Rel αG) A×Rel (F' Acat_Rel αF))ArrVal"
        by 
          (
            cs_concl cs_shallow cs_simp: 
              prod_2_Rel_components comp_Rel_components(1) cat_Rel_cs_simps 
          )
    next
      fix R assume
        "R  ((G' Acat_Rel αG) A×Rel (F' Acat_Rel αF))ArrVal"
      from this assms have 
        "R  prod_2_Rel_ArrVal (G'ArrVal  GArrVal) (F'ArrVal  FArrVal)"
        by 
          (
            cs_prems cs_shallow cs_simp:
              comp_Rel_components prod_2_Rel_components cat_Rel_cs_simps
          )
      from this[folded prod_2_Rel_ArrVal_vcomp] assms show
        "R  ((G' A×Rel F') Acat_Rel α(G A×Rel F))ArrVal"
        by
          (
            cs_concl cs_shallow
              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  cat_Rel αObj" and "B  cat_Rel αObj"
  shows 
    "(cat_Rel αCIdA) A×Rel (cat_Rel αCIdB) = cat_Rel αCIdA × B"
proof-
  interpret Rel: category α cat_Rel α by (rule category_cat_Rel)
  from assms have A_B: 
    "(cat_Rel αCIdA) A×Rel (cat_Rel αCIdB) :
      A × B cat_Rel αA × B"
    by (cs_concl cs_intro: cat_Rel_par_set_cs_intros cat_cs_intros)
  from assms Rel.category_axioms have AB:
    "cat_Rel αCIdA × B : A × B cat_Rel αA × 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 α ((cat_Rel αCIdA) A×Rel (cat_Rel αCIdB))"
      by (auto dest: cat_Rel_is_arrD(1))
    from AB show arr_Rel_GG'_FF': "arr_Rel α (cat_Rel αCIdA × B)"
      by (auto dest: cat_Rel_is_arrD(1))
    from assms show 
      "((cat_Rel αCIdA) A×Rel (cat_Rel αCIdB))ArrVal =
        cat_Rel αCIdA × BArrVal"
      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 cf_dag_Rel_ArrMap_app_prod_2_Rel:
  assumes "S : A cat_Rel αB" and "T : C cat_Rel αD"
  shows
    "C.Rel αArrMapS A×Rel T =
      (C.Rel αArrMapS) A×Rel (C.Rel αArrMapT)"
proof-

  interpret S: arr_Rel α S by (intro cat_Rel_is_arrD[OF assms(1)])
  interpret Rel: category α cat_Rel α by (rule S.category_cat_Rel)
  interpret dag_Rel: is_iso_functor α op_cat (cat_Rel α) cat_Rel α C.Rel α
    by (rule S.cf_dag_Rel_is_iso_functor)

  note ST = prod_2_Rel_is_cat_Rel_arr[OF assms]

  from assms have dag_S: "C.Rel αArrMapS : B cat_Rel αA"
    and dag_T: "C.Rel αArrMapT : D cat_Rel αC"
    by
      (
        cs_concl 
          cs_simp: cat_Rel_cs_simps cat_op_simps cs_intro: cat_cs_intros 
      )+
  from assms have dag_prod:
    "C.Rel αArrMapS A×Rel T : B × D cat_Rel αA × 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:
    "(C.Rel αArrMapS) A×Rel (C.Rel αArrMapT) :
      B × D cat_Rel αA × C" 
    by (cs_concl cs_shallow 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]:
    "𝒟 (C.Rel αArrMapS A×Rel TArrVal) =  (SArrVal) ×  (TArrVal)"
    " (C.Rel αArrMapS A×Rel TArrVal) = 𝒟 (SArrVal) × 𝒟 (TArrVal)"
    by (cs_concl cs_simp: cat_cs_simps)+

 show
    "C.Rel αArrMapS A×Rel T =
      (C.Rel αArrMapS) A×Rel (C.Rel αArrMapT)"
  proof(rule arr_Rel_eqI)
    from dag_prod show arr_Rel_dag_prod: 
      "arr_Rel α (C.Rel αArrMapS A×Rel T)"
      by (auto dest: cat_Rel_is_arrD)
    then interpret dag_prod: arr_Rel α C.Rel αArrMapS A×Rel T by simp
    from prod_dag show arr_Rel_prod_dag:
      "arr_Rel α ((C.Rel αArrMapS) A×Rel (C.Rel αArrMapT))"
      by (auto dest: cat_Rel_is_arrD)
    then interpret prod_dag: 
      arr_Rel α (C.Rel αArrMapS) A×Rel (C.Rel αArrMapT) 
      by simp
    from ST have arr_Rel_ST: "arr_Rel α (S A×Rel T)" 
      by (auto dest: cat_Rel_is_arrD)
    show
      "C.Rel αArrMapS A×Rel TArrVal =
        ((C.Rel αArrMapS) A×Rel (C.Rel αArrMapT))ArrVal"
    proof(intro vsubset_antisym vsubsetI)
      fix bd_ac assume prems: "bd_ac  C.Rel αArrMapS A×Rel TArrVal"
      then obtain bd ac 
        where bd_ac_def: "bd_ac = bd, ac" 
          and bd: "bd   (SArrVal) ×  (TArrVal)" 
          and ac: "ac  𝒟 (SArrVal) × 𝒟 (TArrVal)"
        by (elim cat_Rel_is_arr_ArrValE[OF dag_prod prems, unfolded cat_cs_simps])
      have "ac, bd  prod_2_Rel_ArrVal (SArrVal) (TArrVal)"
        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  SArrVal"
          and cd: "c, d  TArrVal"
          and bd_def: "bd = b, d" 
          and ac_def: "ac = a, c"
        by auto
      show "bd_ac  ((C.Rel αArrMapS) A×Rel (C.Rel αArrMapT))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  C.Rel αArrMapSArrVal"
          "d, c  C.Rel αArrMapTArrVal"
          by (cs_concl cs_shallow cs_simp: cat_cs_simps)+
      qed
    next
      fix bd_ac assume prems:
        "bd_ac  ((C.Rel αArrMapS) A×Rel (C.Rel αArrMapT))ArrVal"
      then obtain a b c d 
        where bd_ac_def: "bd_ac = b, d, a, c"
          and ba: "b, a  C.Rel αArrMapSArrVal"
          and dc: "d, c  C.Rel αArrMapTArrVal"
        by (elim prod_2_Rel_ArrValE[OF prems[unfolded prod_2_Rel_components]])
      then have ab: "a, b  SArrVal" and cd: "c, d  TArrVal"
        unfolding assms[THEN cf_dag_Rel_ArrMap_app_iff] by simp_all
      from ST ab cd show "bd_ac  C.Rel αArrMapS A×Rel TArrVal"
        unfolding bd_ac_def 
        by
          (
            cs_concl cs_shallow
              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(𝔄 ×C 𝔄)Obj. AB0 × AB1),
      (λST(𝔄 ×C 𝔄)Arr. (ST0) A×Rel (ST1)),
      𝔄 ×C 𝔄,
      𝔄
    ]"


text‹Components.›

lemma cf_prod_2_Rel_components: 
  shows "cf_prod_2_Rel 𝔄ObjMap = (λAB(𝔄 ×C 𝔄)Obj. AB0 × AB1)"
    and "cf_prod_2_Rel 𝔄ArrMap =
      (λST(𝔄 ×C 𝔄)Arr. (ST0) A×Rel (ST1))"
    and [cat_cs_simps]: "cf_prod_2_Rel 𝔄HomDom = 𝔄 ×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]" and "AB  (𝔄 ×C 𝔄)Obj"
  shows "A HM.Ocf_prod_2_Rel 𝔄B = A × 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: 
  " (cf_prod_2_Rel (cat_Rel α)ObjMap)  cat_Rel αObj"
proof-
  interpret Rel: category α cat_Rel α
    by (cs_concl cs_shallow 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  (cat_Rel α ×C cat_Rel α)Obj"
    with Rel.category_axioms obtain A B where AB_def: "AB = [A, B]"
      and A: "A  cat_Rel αObj"
      and B: "B  cat_Rel αObj"
      by (elim cat_prod_2_ObjE[rotated 2])
    from prems A B show "cf_prod_2_Rel (cat_Rel α)ObjMapAB  cat_Rel αObj"
      unfolding AB_def cat_Rel_components(1)
      by 
        (
          cs_concl cs_shallow 
            cs_simp: cat_cs_simps cat_Rel_cs_simps cs_intro: V_cs_intros
        )
  qed (cs_concl cs_shallow 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]" and "GF  (𝔄 ×C 𝔄)Arr"
  shows "G HM.Acf_prod_2_Rel 𝔄F = G A×Rel 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 α ×C cat_Rel α ↦↦Cαcat_Rel α"
proof-

  interpret Rel: category α cat_Rel α
    by (cs_concl cs_shallow 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"
      unfolding cf_prod_2_Rel_def by (simp add: nat_omega_simps)
    show " (cf_prod_2_Rel (cat_Rel α)ObjMap)  cat_Rel αObj"
      by (rule cf_prod_2_Rel_ObjMap_vrange)
    show "cf_prod_2_Rel (cat_Rel α)ArrMapGF :
      cf_prod_2_Rel (cat_Rel α)ObjMapAB cat_Rel αcf_prod_2_Rel (cat_Rel α)ObjMapCD"
      if "GF : AB cat_Rel α ×C cat_Rel αCD" for AB CD GF
    proof-
      from that obtain G F A B C D
        where GF_def: "GF = [G, F]"
          and AB_def: "AB = [A, B]"
          and CD_def: "CD = [C, D]"
          and G: "G : A cat_Rel αC"
          and F: "F : B cat_Rel α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 α)ArrMapGF' Acat_Rel α ×C cat_Rel αGF =
        cf_prod_2_Rel (cat_Rel α)ArrMapGF' Acat_Rel αcf_prod_2_Rel (cat_Rel α)ArrMapGF"
      if "GF' : AB' cat_Rel α ×C cat_Rel αAB''"
        and "GF : AB cat_Rel α ×C cat_Rel αAB'"
      for AB' AB'' GF' AB GF
    proof-
      from that(2) obtain G F A A' B B' 
        where GF_def: "GF = [G, F]"
          and AB_def: "AB = [A, B]"
          and AB'_def: "AB' = [A', B']"
          and G: "G : A cat_Rel αA'"
          and F: "F : B cat_Rel α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']"
          and AB''_def: "AB'' = [A'', B'']"
          and G': "G' : A' cat_Rel αA''"
          and F': "F' : B' cat_Rel α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_shallow
              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 α ×C cat_Rel α)CIdAB =
        cat_Rel αCIdcf_prod_2_Rel (cat_Rel α)ObjMapAB"
      if "AB  (cat_Rel α ×C cat_Rel α)Obj" for AB 
    proof-
      from that obtain A B 
        where AB_def: "AB = [A, B]"
          and A: "A  cat_Rel αObj"
          and B: "B  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_shallow
        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 α ×C cat_Rel α"
    and "𝔅' = cat_Rel α"
    and "α' = α"
  shows "cf_prod_2_Rel (cat_Rel α) : 𝔄' ↦↦Cα'𝔅'"
  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 φ =
    [(λaA. (λiI. φ iArrVala)), A, (iI. F i)]"


text‹Components.›

lemma cat_Set_obj_prod_up_components: 
  shows "cat_Set_obj_prod_up I F A φArrVal = 
    (λaA. (λiI. φ iArrVala))"
    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 = (iI. F i)"
  unfolding cat_Set_obj_prod_up_def arr_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹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  I  φ i : A cat_Set αF i"
  shows " (cat_Set_obj_prod_up I F A φArrVal)  (iI. F i)"
  unfolding cat_Set_obj_prod_up_components 
proof(intro vrange_VLambda_vsubset vproductI)
  fix a assume prems: "a  A"
  show "iI. (λiI. φ iArrVala)i  F i"
  proof(intro ballI)
    fix i assume "i  I"
    with assms prems show "(λiI. φ iArrVala)i  F i"
      by (cs_concl cs_shallow 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  A"
  shows "𝒟 (cat_Set_obj_prod_up I F A φArrVala) = 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  A" and "i  I"
  shows "cat_Set_obj_prod_up I F A φArrValai = φ iArrVala"
  using assms 
  by (cs_concl cs_shallow cs_simp: cat_Set_obj_prod_up_ArrVal_app V_cs_simps)

lemma cat_Set_obj_prod_up_ArrVal_app_vrange: 
  assumes "a  A" and "i. i  I  φ i : A cat_Set αF i"
  shows " (cat_Set_obj_prod_up I F A φArrVala)  (iI. F i)"
proof(intro vsubsetI)
  fix b assume prems: "b   (cat_Set_obj_prod_up I F A φArrVala)"
  from assms(1) have "vsv (cat_Set_obj_prod_up I F A φArrVala)"
    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 φArrValai" 
      and i: "i  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 = φ iArrVala"
    by simp
  from assms(1) assms(2)[OF i] have "b  F i" 
    unfolding b_def' by (cs_concl cs_shallow cs_intro: cat_Set_cs_intros)
  with i show "b  (iI. 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  cat_Set αObj" 
    and "VLambda I F  Vset α" 
    and "i. i  I  φ i : A cat_Set αF i"
  shows "cat_Set_obj_prod_up I F A φ : A cat_Set α(iI. 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"
    unfolding cat_Set_obj_prod_up_def by (auto simp: nat_omega_simps)
  show 
    " (cat_Set_obj_prod_up I F A φArrVal) 
      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  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[unfolded cat_Set_components(1)] cat_Set_cs_simps 
      intro: cat_Set_cs_intros
  )


subsubsection‹Further properties›

lemma (in 𝒵) cat_Set_cf_comp_proj_obj_prod_up: 
  assumes "A  cat_Set αObj" 
    and "VLambda I F  Vset α"
    and "i. i  I  φ i : A cat_Set αF i" 
    and "i  I"
  shows 
    "φ i = vprojection_arrow I F i Acat_Set αcat_Set_obj_prod_up I F A φ"
    (is φ i = ?Fi Acat_Set α)
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 : (iI. F i) cat_Set α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 cat_Set α(iI. 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 Acat_Set α : A cat_Set αF i"
    by (cs_concl cs_shallow cs_intro: cat_cs_intros)
  then show arr_Set_Fi_φ: "arr_Set α (?Fi Acat_Set α)"
    by (auto simp: cat_Set_is_arrD(1))
  interpret arr_Set α ?Fi Acat_Set α by (rule arr_Set_Fi_φ)
  from φi have dom_lhs: "𝒟 (φ iArrVal) = A"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)
  from Fi_φ have dom_rhs: "𝒟 ((?Fi Acat_Set α)ArrVal) = A"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  show "φ iArrVal = (?Fi Acat_Set α)ArrVal"
  proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
    fix a assume prems: "a  A"
    from assms(4) prems φi(4) φ Fi show 
      "φ iArrVala = (?Fi Acat_Set α)ArrVala"
      by 
        ( 
          cs_concl cs_shallow
            cs_simp: cat_Set_cs_simps cat_cs_simps 
            cs_intro: cat_Set_cs_intros cat_cs_intros
        )
  qed auto
  from Fi φ show
    "φ iArrDom = (?Fi Acat_Set α)ArrDom"
    "φ iArrCod = (?Fi Acat_Set α)ArrCod"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cat_Set_cs_simps φi(2,3))+
qed



subsection‹Coproduct universal property arrow for Set›


subsubsection‹Definition and elementary properties›

definition cat_Set_obj_coprod_up :: "V  (V  V)  V  (V  V)  V"
  where "cat_Set_obj_coprod_up I F A φ =
    [(λix(iI. F i). φ (vfst ix)ArrValvsnd ix), (iI. F i), A]"


text‹Components.›

lemma cat_Set_obj_coprod_up_components: 
  shows "cat_Set_obj_coprod_up I F A φArrVal = 
    (λix(iI. F i). φ (vfst ix)ArrValvsnd ix)"
    and [cat_Set_cs_simps]: 
      "cat_Set_obj_coprod_up I F A φArrDom = (iI. F i)"
    and [cat_Set_cs_simps]: 
      "cat_Set_obj_coprod_up I F A φArrCod = A"
  unfolding cat_Set_obj_coprod_up_def arr_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

mk_VLambda cat_Set_obj_coprod_up_components(1)
  |vsv cat_Set_obj_coprod_up_ArrVal_vsv[cat_Set_cs_intros]|
  |vdomain cat_Set_obj_coprod_up_ArrVal_vdomain[cat_Set_cs_simps]|
  |app cat_Set_obj_coprod_up_ArrVal_app'|

lemma cat_Set_obj_coprod_up_ArrVal_app[cat_cs_simps]:
  assumes "ix = i, x" and "i, x  (iI. F i)"
  shows "cat_Set_obj_coprod_up I F A φArrValix = φ iArrValx"
  using assms by (auto simp: cat_Set_obj_coprod_up_ArrVal_app')

lemma cat_Set_obj_coprod_up_ArrVal_vrange:
  assumes "i. i  I  φ i : F i cat_Set αA"
  shows " (cat_Set_obj_coprod_up I F A φArrVal)  A"
proof
  (
    intro vsv.vsv_vrange_vsubset cat_Set_obj_coprod_up_ArrVal_vsv, 
    unfold cat_Set_cs_simps
  )
  fix ix assume "ix  (iI. F i)"
  then obtain i x where ix_def: "ix = i, x" and i: "i  I" and x: "x  F i" 
    by auto
  show "cat_Set_obj_coprod_up I F A φArrValix  A"
  proof(cs_concl_step cat_Set_obj_coprod_up_ArrVal_app)
    show "ix = i, x" by (rule ix_def)
    from i x show "i, x  (iI. F i)" by auto
    from i x assms[OF i] show "φ iArrValx  A"
      by (auto intro: cat_Set_ArrVal_app_vrange)
  qed
qed


subsubsection‹Coproduct universal property arrow for Set› is an arrow in Set›

lemma (in 𝒵) cat_Set_obj_coprod_up_cat_Set_is_arr:
  assumes "A  cat_Set αObj" 
    and "VLambda I F  Vset α" 
    and "i. i  I  φ i : F i cat_Set αA"
  shows "cat_Set_obj_coprod_up I F A φ : (iI. F i) cat_Set αA"
proof(intro cat_Set_is_arrI arr_SetI)
  show "vfsequence (cat_Set_obj_coprod_up I F A φ)"
    unfolding cat_Set_obj_coprod_up_def by auto
  show "vcard (cat_Set_obj_coprod_up I F A φ) = 3"
    unfolding cat_Set_obj_coprod_up_def by (auto simp: nat_omega_simps)
  show 
    " (cat_Set_obj_coprod_up I F A φArrVal) 
      cat_Set_obj_coprod_up I F A φArrCod"
    unfolding cat_Set_obj_coprod_up_components(3)
    by (rule cat_Set_obj_coprod_up_ArrVal_vrange[OF assms(3)])
  show "cat_Set_obj_coprod_up I F A φArrCod  Vset α"
    by (simp_all add: cat_Set_cs_simps assms[unfolded cat_Set_components(1)])
qed 
  (
    auto simp: 
      assms 
      cat_Set_obj_coprod_up_components 
      Limit_vdunion_in_Vset_if_VLambda_in_VsetI
  ) 


subsubsection‹Further properties›

lemma (in 𝒵) cat_Set_cf_comp_coprod_up_vcia:
  assumes "A  cat_Set αObj"
    and "VLambda I F  Vset α"
    and "i. i  I  φ i : F i cat_Set αA" 
    and "i  I"
  shows 
    "φ i = cat_Set_obj_coprod_up I F A φ Acat_Set αvcinjection_arrow I F i"
    (is φ i =  Acat_Set α?Fi)
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 : F i cat_Set α(iI. F i)"
    by (rule vcinjection_arrow_is_arr[OF assms(4,2)])
  from cat_Set_obj_coprod_up_cat_Set_is_arr[OF assms(1,2,3)] have φ:
    "cat_Set_obj_coprod_up I F A φ : (iI. F i) cat_Set αA"
    by simp
  show "arr_Set α (φ i)" by (rule φi(1))
  then interpret φi: arr_Set α φ i .
  from Fi φ have Fi_φ: " Acat_Set α?Fi : F i cat_Set αA"
    by (cs_concl cs_shallow cs_intro: cat_cs_intros)
  then show arr_Set_Fi_φ: "arr_Set α ( Acat_Set α?Fi)"
    by (auto simp: cat_Set_is_arrD(1))
  interpret arr_Set α  Acat_Set α?Fi by (rule arr_Set_Fi_φ)
  from φi have dom_lhs: "𝒟 (φ iArrVal) = F i"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)
  from Fi_φ have dom_rhs: "𝒟 (( Acat_Set α?Fi)ArrVal) = F i"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  show "φ iArrVal = ( Acat_Set α?Fi)ArrVal"
  proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
    fix a assume "a  F i"
    from assms(4) this φi(4) φ Fi show 
      "φ iArrVala = ( Acat_Set α?Fi)ArrVala"
      by
        (
          cs_concl cs_shallow
            cs_simp: cat_Set_cs_simps cat_cs_simps 
            cs_intro: vdunionI cat_Set_cs_intros cat_cs_intros
        )
  qed auto
  from Fi φ show 
    "φ iArrDom = ( Acat_Set α?Fi)ArrDom"
    "φ iArrCod = ( Acat_Set α?Fi)ArrCod"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cat_Set_cs_simps φi(2,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  X  fArrValx = gArrValx}"

lemma small_vequalizer[simp]: 
  "small {x. x  X  fArrValx = gArrValx}"
  by auto


text‹Rules.›

lemma vequalizerI:
  assumes "x  X" and "fArrValx = gArrValx"
  shows "x  vequalizer X f g"
  using assms unfolding vequalizer_def by auto

lemma vequalizerD[dest]:
  assumes "x  vequalizer X f g"
  shows "x  X" and "fArrValx = gArrValx"
  using assms unfolding vequalizer_def by auto

lemma vequalizerE[elim]:
  assumes "x  vequalizer X f g"
  obtains "x  X" and "fArrValx = gArrValx"
  using assms unfolding vequalizer_def by auto


text‹Elementary results.›

lemma vequalizer_vsubset_vdomain[cat_Set_cs_intros]: "vequalizer a g f  a" 
  by auto
  
lemma Limit_vequalizer_in_Vset[cat_Set_cs_intros]:
  assumes "Limit α" and "a  cat_Set αObj"
  shows "vequalizer a g f  cat_Set αObj"
  using assms unfolding cat_Set_components(1) by auto

lemma vequalizer_flip: "vequalizer a f g = vequalizer a g f"
  unfolding vequalizer_def by auto

lemma cat_Set_incl_Set_commute:
  assumes "𝔤 : 𝔞 cat_Set α𝔟" and "𝔣 : 𝔞 cat_Set α𝔟" 
  shows 
    "𝔤 Acat_Set αincl_Set (vequalizer 𝔞 𝔣 𝔤) 𝔞 =
      𝔣 Acat_Set αincl_Set (vequalizer 𝔞 𝔣 𝔤) 𝔞"
  (is 𝔤 Acat_Set α?incl = 𝔣 Acat_Set α?incl)
proof-

  interpret 𝔤: arr_Set α 𝔤 
    rewrites "𝔤ArrDom = 𝔞" and "𝔤ArrCod = 𝔟"
    by (intro cat_Set_is_arrD[OF assms(1)])+
  interpret 𝔣: arr_Set α 𝔣 
    rewrites "𝔣ArrDom = 𝔞" and "𝔣ArrCod = 𝔟"
    by (intro cat_Set_is_arrD[OF assms(2)])+

  note [cat_Set_cs_intros] = 𝔤.arr_Set_ArrDom_in_Vset 𝔣.arr_Set_ArrCod_in_Vset

  from assms have 𝔤_incl: 
    "𝔤 Acat_Set α?incl : vequalizer 𝔞 𝔣 𝔤 cat_Set α𝔟"
    by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros)
  then have dom_lhs: "𝒟 ((𝔤 Acat_Set α?incl)ArrVal) = vequalizer 𝔞 𝔣 𝔤"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)+
  from assms have 𝔣_incl: 
    "𝔣 Acat_Set α?incl : vequalizer 𝔞 𝔣 𝔤 cat_Set α𝔟"
    by (cs_concl cs_intro: V_cs_intros cat_Set_cs_intros cat_cs_intros)
  then have dom_rhs: "𝒟 ((𝔣 Acat_Set α?incl)ArrVal) = vequalizer 𝔞 𝔣 𝔤"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)+

  show ?thesis
  proof(rule arr_Set_eqI)
    from 𝔤_incl show arr_Set_𝔤_incl: "arr_Set α (𝔤 Acat_Set α?incl)"
      by (auto dest: cat_Set_is_arrD(1))
    interpret arr_Set_𝔤_incl: arr_Set α 𝔤 Acat_Set α?incl
      by (rule arr_Set_𝔤_incl)
    from 𝔣_incl show arr_Set_𝔣_incl: "arr_Set α (𝔣 Acat_Set α?incl)"
      by (auto dest: cat_Set_is_arrD(1))
    interpret arr_Set_𝔣_incl: arr_Set α 𝔣 Acat_Set α?incl
      by (rule arr_Set_𝔣_incl)
    show "(𝔤 Acat_Set α?incl)ArrVal = (𝔣 Acat_Set α?incl)ArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume "a  vequalizer 𝔞 𝔣 𝔤"
      with assms show 
        "(𝔤 Acat_Set α?incl)ArrVala = (𝔣 Acat_Set α?incl)ArrVala"
        by (*very slow*)
          (
            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_shallow cs_simp: cat_cs_simps)+

qed



subsection‹Application of a function to a finite sequence as an arrow in Set›

definition vfsequence_map :: "V  V"
  where "vfsequence_map F =
    [
      (λxsvfsequences_on (FArrDom). FArrVal  xs),
      vfsequences_on (FArrDom),
      vfsequences_on (FArrCod)
    ]"


text‹Components.›

lemma vfsequence_map_components:
  shows "vfsequence_map FArrVal =
    (λxsvfsequences_on (FArrDom). FArrVal  xs)"
    and [cat_cs_simps]: "vfsequence_map FArrDom = vfsequences_on (FArrDom)"
    and [cat_cs_simps]: "vfsequence_map FArrCod = vfsequences_on (FArrCod)"
  unfolding vfsequence_map_def arr_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

mk_VLambda vfsequence_map_components(1)
  |vsv vfsequence_map_ArrVal_vsv[cat_cs_intros, cat_Set_cs_intros]|
  |vdomain vfsequence_map_ArrVal_vdomain[cat_cs_simps, cat_Set_cs_simps]|
  |app vfsequence_map_ArrVal_app|

lemma vfsequence_map_ArrVal_app_app:
  assumes "F : A cat_Set αB" 
    and "xs  vfsequences_on A"
    and "i  𝒟 xs"
  shows "vfsequence_map FArrValxsi = FArrValxsi"
proof-
  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret arr_Set α F 
    rewrites "FArrDom = A" and "FArrCod = B"
    by (intro FD)+
  note xsD = vfsequences_onD[OF assms(2)]
  interpret xs: vfsequence xs by (rule xsD(1))
  from assms xsD(2)[OF assms(3)] show ?thesis
    by
      (
        cs_concl
          cs_simp: V_cs_simps cat_cs_simps vfsequence_map_ArrVal_app
          cs_intro: V_cs_intros
      )
qed


subsubsection‹
Application of a function to a finite sequence is an arrow in Set›

lemma vfsequence_map_is_arr:
  assumes "F : A cat_Set αB"
  shows "vfsequence_map F : vfsequences_on A cat_Set αvfsequences_on B"
proof-

  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret arr_Set α F 
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+

  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI , unfold cat_cs_simps)
    show "vfsequence (vfsequence_map F)"
      unfolding vfsequence_map_def by auto
    show "vcard (vfsequence_map F) = 3"
      unfolding vfsequence_map_def by (simp_all add: nat_omega_simps)
    show " (vfsequence_map FArrVal)  vfsequences_on B"
      unfolding vfsequence_map_components
    proof
      (
        intro vrange_VLambda_vsubset vfsequences_onI; 
        elim vfsequences_onE; 
        unfold cat_cs_simps
      )
      fix xs assume prems: "vfsequence xs" "i  𝒟 xs  xsi  A" for i
      interpret xs: vfsequence xs by (rule prems(1))
      have [intro]: "x  𝒟 (FArrVal)" if "x   xs" for x
      proof-
        from that obtain i where i: "i  𝒟 xs" and x_def: "x = xsi"
          by (auto dest: xs.vrange_atD)
        from prems(2)[OF i] show "x  𝒟 (FArrVal)"
          unfolding x_def arr_Set_ArrVal_vdomain .
      qed
      show "vfsequence (FArrVal  xs)"
        by (intro vfsequence_vcomp_vsv_vfsequence vsubsetI)
          (auto intro: prems(1))
      fix i assume prems': "i  𝒟 (FArrVal  xs)"
      moreover have "𝒟 (FArrVal  xs) = 𝒟 xs"
        by (intro vdomain_vcomp_vsubset vsubsetI) (auto intro: prems(1))
      ultimately have i: "i  𝒟 xs" by simp
      with assms(1) prems(2)[OF i] show "(FArrVal  xs)i  B"
        by 
          (
            cs_concl
              cs_simp: V_cs_simps cat_cs_simps 
              cs_intro: V_cs_intros cat_Set_cs_intros
          )
    qed

  qed 
    (
      auto intro: 
        vfsequences_on_in_VsetI 
        arr_Set_ArrDom_in_Vset 
        arr_Set_ArrCod_in_Vset
        cat_cs_intros
    )

qed

lemma (in 𝒵) vfsequence_map_is_monic_arr:
  assumes "F : A moncat_Set αB"
  shows "vfsequence_map F : vfsequences_on A moncat_Set αvfsequences_on B"
proof-
  
  note cat_Set_is_monic_arrD[OF assms]
  note FD = this cat_Set_is_arrD[OF this(1)]
  interpret F: arr_Set α F 
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+

  show ?thesis
  proof
    (
      intro cat_Set_is_monic_arrI vfsequence_map_is_arr FD(1) vsv.vsv_valeq_v11I,
      unfold cat_cs_simps;
      (elim vfsequences_onE)?
    )
  
    fix xs ys assume prems:
      "vfsequence_map FArrValxs = vfsequence_map FArrValys"
      "vfsequence xs"
      "i. i  𝒟 xs  xsi  A"
      "vfsequence ys"
      "i. i  𝒟 ys  ysi  A"

    interpret xs: vfsequence xs by (rule prems(2))
    interpret ys: vfsequence ys by (rule prems(4))

    have "xs  vfsequences_on (FArrDom)"
      unfolding cat_cs_simps by (intro vfsequences_onI prems(2,3))
    from vfsequence_map_ArrVal_app[OF this] have F_xs:
      "vfsequence_map FArrValxs = FArrVal  xs" 
      by simp
    from prems(3) have rxs: " xs  A"
      by (intro vsubsetI) (auto dest: xs.vrange_atD)
    from xs.vfsequence_vdomain_in_omega have dxs: "𝒟 xs  Vset α"
      by (auto intro!: Axiom_of_Infinity)
    note xs_is_arr = cat_Set_arr_of_vsv_is_arr
      [
        OF xs.vsv_axioms rxs,
        unfolded cat_Set_components(1), 
        OF dxs F.arr_Par_ArrDom_in_Vset
      ]

    have ys: "ys  vfsequences_on (FArrDom)"
      unfolding cat_cs_simps by (intro vfsequences_onI prems(4,5))
    from vfsequence_map_ArrVal_app[OF this] have F_ys:
      "vfsequence_map FArrValys = FArrVal  ys" 
      by simp
    from prems(5) have rys: " ys  A"
      by (intro vsubsetI) (auto dest: ys.vrange_atD)
    from ys.vfsequence_vdomain_in_omega have dys: "𝒟 ys  Vset α"
      by (auto intro!: Axiom_of_Infinity)
    note ys_is_arr = cat_Set_arr_of_vsv_is_arr
      [
        OF ys.vsv_axioms rys,
        unfolded cat_Set_components(1), 
        OF dys F.arr_Par_ArrDom_in_Vset
      ]

    note Fxs_Fys = prems(1)[unfolded F_xs F_ys]

    from rxs have dom_rxs: "𝒟 (FArrVal  xs) = 𝒟 xs"
      by (intro vdomain_vcomp_vsubset vsubsetI, unfold F.arr_Set_ArrVal_vdomain)
        auto
    moreover from rys have dom_rys: "𝒟 (FArrVal  ys) = 𝒟 ys"
      by (intro vdomain_vcomp_vsubset vsubsetI, unfold F.arr_Set_ArrVal_vdomain)
        auto
    ultimately have dxs_dys: "𝒟 xs = 𝒟 ys"
      by (simp add: prems(1)[unfolded F_xs F_ys])

    from FD(1) xs_is_arr have lhs_is_arr:
      "F Acat_Set αcat_Set_arr_of_vsv xs A : 𝒟 xs cat_Set αB"
      by (cs_concl cs_intro: cat_cs_intros)
    then have dom_lhs: 
      "𝒟 ((F Acat_Set αcat_Set_arr_of_vsv xs A)ArrVal) = 𝒟 xs"
      by (simp add: cat_cs_simps)

    from FD(1) ys_is_arr have rhs_is_arr:
      "F Acat_Set αcat_Set_arr_of_vsv ys A : 𝒟 xs cat_Set αB"
      by (cs_concl cs_simp: dxs_dys cs_intro: cat_cs_intros)
    then have dom_rhs: 
      "𝒟 ((F Acat_Set αcat_Set_arr_of_vsv ys A)ArrVal) = 𝒟 xs"
      by (simp add: cat_cs_simps)

    have F_xs_F_ys:
      "F Acat_Set αcat_Set_arr_of_vsv xs A = 
        F Acat_Set αcat_Set_arr_of_vsv ys A"
    proof(rule arr_Set_eqI[of α])
      show
        "(F Acat_Set αcat_Set_arr_of_vsv xs A)ArrVal = 
          (F Acat_Set αcat_Set_arr_of_vsv ys A)ArrVal"
      proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
        fix i assume prems: "i  𝒟 xs"
        from prems rxs have xsi: "xsi  A" 
          by (auto dest: xs.vdomain_atD)
        from prems rys have ysi: "ysi  A" 
          by (auto simp: dxs_dys dest: ys.vdomain_atD)
        from arg_cong[OF Fxs_Fys, where f=λx. xi] prems FD(1) xsi ysi
        have "FArrValxsi = FArrValysi"
          by
            (
              cs_prems 
                cs_simp: V_cs_simps cat_cs_simps dxs_dys[symmetric] 
                cs_intro: V_cs_intros cat_cs_intros
            )
        with prems FD(1) xs_is_arr ys_is_arr show 
          "(F Acat_Set αcat_Set_arr_of_vsv xs A)ArrVali = 
            (F Acat_Set αcat_Set_arr_of_vsv ys A)ArrVali"
          by
            (
              cs_concl 
                cs_simp: cat_Set_cs_simps cat_cs_simps dxs_dys[symmetric] 
                cs_intro: cat_cs_intros
            )
      qed (use lhs_is_arr rhs_is_arr in auto dest: cat_Set_is_arrD)
    qed
      (
        use lhs_is_arr rhs_is_arr in 
          auto simp: cat_cs_simps dest: cat_Set_is_arrD(1)
      )+
    have "cat_Set_arr_of_vsv xs A = cat_Set_arr_of_vsv ys A"
      by 
        (
          rule is_monic_arrD(2)[
            OF assms(1) xs_is_arr, unfolded dxs_dys, OF ys_is_arr, OF F_xs_F_ys
            ]
        )
    from arg_cong [OF this, where f=λx. xArrVal, unfolded cat_Set_cs_simps]
    show "xs = ys" .

  qed (auto intro: cat_cs_intros)

qed

lemma (in 𝒵) vfsequence_map_is_epic_arr:
  assumes "F : A epicat_Set αB"
  shows "vfsequence_map F : vfsequences_on A epicat_Set αvfsequences_on B"
proof-

  note cat_Set_is_epic_arrD[OF assms]
  note FD = this cat_Set_is_arrD[OF this(1)]

  interpret F: arr_Set α F 
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+
  interpret SF: arr_Set α vfsequence_map F
    rewrites "vfsequence_map FArrDom = vfsequences_on A"
      and "vfsequence_map FArrCod = vfsequences_on B"
    by (intro cat_Set_is_arrD[OF vfsequence_map_is_arr[OF FD(1)]])+
  
  show ?thesis
  proof
    (
      intro cat_Set_is_epic_arrI, 
      rule vfsequence_map_is_arr[OF FD(1)], 
      rule vsubset_antisym, 
      rule SF.arr_Par_ArrVal_vrange,
      rule vsubsetI
    )
    fix xs assume prems: "xs  vfsequences_on B"
    note xsD = vfsequences_onD[OF prems]
    interpret vfsequence xs by (rule xsD(1))
    define ys where "ys = (λi𝒟 xs. SOME x. x  A  xsi = FArrValx)"
    have ys_vdomain: "𝒟 ys = 𝒟 xs" unfolding ys_def by simp
    interpret ys: vfsequence ys
      by (rule vfsequenceI)
        (auto intro: vfsequence_vdomain_in_omega simp: ys_def)
    have ysi: "ysi = (SOME x. x  A  xsi = FArrValx)"
      if "i  𝒟 xs" for i
      using that unfolding ys_def by simp
    have ysi: "ysi  A" 
      and xsi_def: "xsi = FArrValysi"
      if "i  𝒟 xs" for i
    proof-
      have "xsi   (FArrVal)" by (rule xsD(2)[OF that, folded FD(2)])
      then obtain x where x: "x  A" and xsi_def: "xsi = FArrValx"
        by (auto elim: F.ArrVal.vrange_atE simp: F.arr_Set_ArrVal_vdomain)
      show "ysi  A" and "xsi = FArrValysi"
        unfolding ysi[OF that]
        by 
          (
            allrule someI2_ex, intro exI conjI; (elim conjE)?, 
            tacticdistinct_subgoals_tac
          )
          (auto simp: x xsi_def)
    qed
    show "xs   (vfsequence_map FArrVal)"
    proof
      (
        intro vsv.vsv_vimageI2' cat_cs_intros, 
        cs_concl_step vfsequence_map_ArrVal_app, 
        unfold cat_cs_simps,
        tacticdistinct_subgoals_tac
      )
      show "ys  vfsequences_on A"
        by (intro vfsequences_onI ys.vfsequence_axioms)
          (auto intro: ysi simp: ys_vdomain)
      show "xs = FArrVal  ys"  
      proof(rule vsv_eqI)
        show "𝒟 xs = 𝒟 (FArrVal  ys)"
          unfolding ys_vdomain[symmetric]
        proof(intro vdomain_vcomp_vsubset[symmetric] vsubsetI)
          fix y assume "y   ys"
          then obtain i where i: "i  𝒟 ys" and y_def: "y = ysi"
            by (auto dest: ys.vrange_atD)
          from i show "y  𝒟 (FArrVal)"
            unfolding y_def F.arr_Set_ArrVal_vdomain ys_vdomain by (rule ysi)
        qed
        show "xsi = (FArrVal  ys)i"
          if "i  𝒟 xs" for i 
          using FD(1) that 
          by
            (
              cs_concl
                cs_simp: V_cs_simps cat_cs_simps xsi_def ys_vdomain
                cs_intro: V_cs_intros ysi
            )
      qed (auto intro: vsv_vcomp)
    qed
  qed

qed

lemma vfsequence_map_is_iso_arr:
  assumes "F : A isocat_Set αB"
  shows "vfsequence_map F : vfsequences_on A isocat_Set αvfsequences_on B"
proof-
  note cat_Set_is_iso_arrD[OF assms]
  note FD = this cat_Set_is_arrD[OF this(1)]
  interpret F: arr_Set α F 
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+
  interpret Set: category α cat_Set α by (cs_concl cs_intro: cat_cs_intros)
  show ?thesis
    by 
      ( 
        intro 
        F.cat_Set_is_iso_arr_if_monic_and_epic
        F.vfsequence_map_is_monic_arr[
          OF Set.cat_is_iso_arr_is_monic_arr[OF assms]
          ]
        F.vfsequence_map_is_epic_arr[
          OF Set.cat_is_iso_arr_is_epic_arr[OF assms]
          ]
      )
qed



subsection‹An injection from the range of an arrow in Set› into its domain›


subsubsection‹Definition and elementary properties›

definition vrange_iso :: "V  V"
  where "vrange_iso F =
    [
      (λy (FArrVal). (SOME x. x  FArrDom  y = FArrValx)),
       (FArrVal),
      FArrDom
    ]"


text‹Components.›

lemma vrange_iso_components:
  shows "vrange_iso FArrVal =
    (λy (FArrVal). (SOME x. x  FArrDom  y = FArrValx))"
    and [cat_cs_simps]: "vrange_iso FArrDom =  (FArrVal)"
    and [cat_cs_simps]: "vrange_iso FArrCod = FArrDom"
  unfolding vrange_iso_def arr_field_simps by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

mk_VLambda vrange_iso_components(1)
  |vsv vrange_iso_ArrVal_vsv[cat_cs_intros]|
  |vdomain vrange_iso_ArrVal_vdomain[cat_cs_simps]|
  |app vrange_iso_ArrVal_app|

lemma vrange_iso_ArrVal_rules:
  assumes "F : A cat_Set αB" and "y   (FArrVal)"
  shows "vrange_iso FArrValy  A"
    and "y = FArrValvrange_iso FArrValy"
proof-
  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret F: arr_Set α F
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+
  from assms(2) have vri_Fy_def:
    "vrange_iso FArrValy = (SOME x. x  FArrDom  y = FArrValx)"
    by (cs_concl cs_simp: vrange_iso_ArrVal_app)
  from assms(2) F.arr_Set_ArrVal_vdomain obtain x 
    where x: "x  A" and y_def: "y = FArrValx"
    by (auto elim: F.ArrVal.vrange_atE)
  show "vrange_iso FArrValy  A"
    and "y = FArrValvrange_iso FArrValy"
    unfolding vri_Fy_def cat_cs_simps 
    by (allrule someI2_ex; (intro exI conjI)?; (elim conjE)?)
      (simp_all add: x y_def)
qed


subsubsection‹
An injection from the range of a function into its domain is a monic in Set›

lemma vrange_iso_is_arr:
  assumes "F : A cat_Set αB"
  shows "vrange_iso F :  (FArrVal) cat_Set αA"
proof-

  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret F: arr_Set α F
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+

  show "vrange_iso F :  (FArrVal) cat_Set αA"
  proof(intro cat_Set_is_arrI arr_SetI, unfold cat_cs_simps)
    show "vfsequence (vrange_iso F)" 
      unfolding vrange_iso_def by (simp_all add: nat_omega_simps)
    show "vsv (vrange_iso FArrVal)"
      by (cs_concl cs_intro: cat_cs_intros)
    then interpret vsv vrange_iso FArrVal 
      rewrites "𝒟 (vrange_iso FArrVal) =  (FArrVal)"
      unfolding cat_cs_simps by simp_all
    show "vcard (vrange_iso F) = 3"
      unfolding vrange_iso_def by (simp_all add: nat_omega_simps)
    show " (vrange_iso FArrVal)  A"
    proof(intro vsubsetI)
      fix x assume "x   (vrange_iso FArrVal)"
      then obtain y where y: "y   (FArrVal)"
        and x_def: "x = vrange_iso FArrValy" 
        by (auto dest: vrange_atD)
      show "x  A"
        unfolding x_def
        by (rule vrange_iso_ArrVal_rules(1)[OF assms y, unfolded cat_cs_simps])
    qed
  qed 
    (
      auto 
        simp: F.arr_Set_ArrDom_in_Vset 
        intro: vrange_in_VsetI F.arr_Rel_ArrVal_in_Vset
    )

qed

lemma vrange_iso_is_arr':
  assumes "F : A cat_Set αB" 
    and "B' =  (FArrVal)"
    and "ℭ' = cat_Set α"
  shows "vrange_iso F : B' ℭ'A"
  using assms(1) unfolding assms(2,3) by (rule vrange_iso_is_arr)

lemma vrange_iso_is_monic_arr:
  assumes "F : A cat_Set αB"
  shows "vrange_iso F :  (FArrVal) moncat_Set αA"
proof-
  note FD = cat_Set_is_arrD[OF assms(1)]
  interpret F: arr_Set α F
    rewrites [cat_cs_simps]: "FArrDom = A" and [cat_cs_simps]: "FArrCod = B"
    by (intro FD)+
  show ?thesis
  proof
    (
      intro cat_Set_is_monic_arrI vrange_iso_is_arr, 
      rule assms, 
      rule vsv.vsv_valeq_v11I[OF vrange_iso_ArrVal_vsv], 
      unfold cat_cs_simps
    )
    fix x y assume prems:
      "x   (FArrVal)"
      "y   (FArrVal)"
      "vrange_iso FArrValx = vrange_iso FArrValy"    
    show "x = y"
      by 
        (
          rule vrange_iso_ArrVal_rules(2)
            [
              OF assms prems(1), 
              unfolded prems(3), 
              folded vrange_iso_ArrVal_rules(2)[OF assms prems(2)]
            ]
        )
  qed simp
qed

lemma vrange_iso_is_monic_arr':
  assumes "F : A cat_Set αB"
    and "B' =  (FArrVal)"
    and "ℭ' = cat_Set α"
  shows "vrange_iso F : B' monℭ'A"
  using assms(1) unfolding assms(2,3) by (rule vrange_iso_is_monic_arr)



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  cat_Rel αObj"
  shows "cat_Rel αCIdA : A cat_Set αA"
proof-
  from assms show ?thesis
    unfolding cat_Rel_components cat_Set_components(6)[symmetric]
    by 
      (
        cs_concl cs_shallow 
          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  cat_Rel αObj" 
    and "B' = A"
    and "C' = A"
    and "ℭ' = cat_Set α"
  shows "cat_Rel αCIdA : B' ℭ'C'"
  using assms(1) unfolding assms(2-4) by (rule cat_Rel_CId_is_cat_Set_arr)

text‹\newpage›

end