Theory CZH_ECAT_Yoneda

(* Copyright 2021 (C) Mihails Milehins *)

section‹Yoneda Lemma›
theory CZH_ECAT_Yoneda
  imports 
    CZH_ECAT_FUNCT
    CZH_ECAT_Hom
begin



subsection‹Yoneda map›


text‹
The Yoneda map is the bijection that is used in the statement of the
Yoneda Lemma, as presented, for example, in Chapter III-2 in 
cite"mac_lane_categories_2010" or in subsection 1.15 
in cite"bodo_categories_1970".
›

definition Yoneda_map :: "V  V  V  V"
  where "Yoneda_map α 𝔎 r =
    (
      λψthese_ntcfs α (𝔎HomDom) (cat_Set α) HomO.Cα𝔎HomDom(r,-) 𝔎.
        ψNTMaprArrVal𝔎HomDomCIdr
    )"


text‹Elementary properties.›

mk_VLambda Yoneda_map_def
  |vsv Yoneda_map_vsv[cat_cs_intros]|

mk_VLambda (in is_functor) Yoneda_map_def[where α=α and 𝔎=𝔉, unfolded cf_HomDom]
  |vdomain Yoneda_map_vdomain|
  |app Yoneda_map_app[unfolded these_ntcfs_iff]|

lemmas [cat_cs_simps] = is_functor.Yoneda_map_vdomain

lemmas Yoneda_map_app[cat_cs_simps] = 
  is_functor.Yoneda_map_app[unfolded these_ntcfs_iff]



subsection‹Yoneda component›


subsubsection‹Definition and elementary properties›


text‹
The Yoneda components are the components of the natural transformations
that appear in the statement of the Yoneda Lemma (e.g., see 
Chapter III-2 in cite"mac_lane_categories_2010" or subsection 1.15 
in cite"bodo_categories_1970").
›

definition Yoneda_component :: "V  V  V  V  V"
  where "Yoneda_component 𝔎 r u d =
    [
      (λfHom (𝔎HomDom) r d. 𝔎ArrMapfArrValu),
      Hom (𝔎HomDom) r d,
      𝔎ObjMapd
    ]"


text‹Components.›

lemma (in is_functor) Yoneda_component_components: 
  shows "Yoneda_component 𝔉 r u dArrVal =
    (λfHom 𝔄 r d. 𝔉ArrMapfArrValu)"
    and "Yoneda_component 𝔉 r u dArrDom = Hom 𝔄 r d"
    and "Yoneda_component 𝔉 r u dArrCod = 𝔉ObjMapd"
  unfolding Yoneda_component_def arr_field_simps 
  by (simp_all add: nat_omega_simps cat_cs_simps)


subsubsection‹Arrow value›

mk_VLambda (in is_functor) Yoneda_component_components(1)
  |vsv Yoneda_component_ArrVal_vsv|
  |vdomain Yoneda_component_ArrVal_vdomain|
  |app Yoneda_component_ArrVal_app[unfolded in_Hom_iff]|

lemmas [cat_cs_simps] = is_functor.Yoneda_component_ArrVal_vdomain

lemmas Yoneda_component_ArrVal_app[cat_cs_simps] = 
  is_functor.Yoneda_component_ArrVal_app[unfolded in_Hom_iff]


subsubsection‹Yoneda component is an arrow in the category Set›

lemma (in category) cat_Yoneda_component_is_arr:
  assumes "𝔎 :  ↦↦Cαcat_Set α"
    and "r  Obj"
    and "u  𝔎ObjMapr"
    and "d  Obj"
  shows "Yoneda_component 𝔎 r u d : Hom  r d cat_Set α𝔎ObjMapd"   
proof-
  interpret 𝔎: is_functor α  cat_Set α 𝔎 by (rule assms(1)) 
  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI, unfold 𝔎.Yoneda_component_components)
    show "vfsequence (Yoneda_component 𝔎 r u d)" 
      unfolding Yoneda_component_def by simp
    show "vcard (Yoneda_component 𝔎 r u d) = 3"
      unfolding Yoneda_component_def by (simp add: nat_omega_simps)
    show " (λfHom  r d. 𝔎ArrMapfArrValu)  𝔎ObjMapd"
    proof(rule vrange_VLambda_vsubset)
      fix f assume "f  Hom  r d"
      then have 𝔎f: "𝔎ArrMapf : 𝔎ObjMapr cat_Set α𝔎ObjMapd" 
        by (auto simp: cat_cs_intros)
      note 𝔎f_simps = cat_Set_is_arrD[OF 𝔎f]
      interpret 𝔎f: arr_Set α 𝔎ArrMapf by (rule 𝔎f_simps(1))          
      have "u  𝒟 (𝔎ArrMapfArrVal)" 
        by (simp add: 𝔎f_simps assms cat_Set_cs_simps)
      with 𝔎f.arr_Set_ArrVal_vrange[unfolded 𝔎f_simps] show 
        "𝔎ArrMapfArrValu  𝔎ObjMapd"
        by (blast elim: 𝔎f.ArrVal.vsv_value)
    qed 
    from assms 𝔎.HomCod.cat_Obj_vsubset_Vset show "𝔎ObjMapd  Vset α"
      by (auto dest: 𝔎.cf_ObjMap_app_in_HomCod_Obj)
  qed (auto simp: assms cat_cs_intros)
qed

lemma (in category) cat_Yoneda_component_is_arr':
  assumes "𝔎 :  ↦↦Cαcat_Set α" 
    and "r  Obj"
    and "u  𝔎ObjMapr"
    and "d  Obj"
    and "s = Hom  r d"
    and "t = 𝔎ObjMapd"
    and "𝔇 = cat_Set α"
  shows "Yoneda_component 𝔎 r u d : s 𝔇t"   
  unfolding assms(5-7) using assms(1-4) by (rule cat_Yoneda_component_is_arr)

lemmas [cat_cs_intros] = category.cat_Yoneda_component_is_arr'[rotated 1]



subsection‹Yoneda arrow›


subsubsection‹Definition and elementary properties›


text‹
The Yoneda arrows are the natural transformations that 
appear in the statement of the Yoneda Lemma in Chapter III-2 in 
cite"mac_lane_categories_2010" and subsection 1.15 
in cite"bodo_categories_1970".
›

definition Yoneda_arrow :: "V  V  V  V  V" 
  where "Yoneda_arrow α 𝔎 r u =
    [
      (λd𝔎HomDomObj. Yoneda_component 𝔎 r u d),
      HomO.Cα𝔎HomDom(r,-),
      𝔎,
      𝔎HomDom,
      cat_Set α
    ]"


text‹Components.›

lemma (in is_functor) Yoneda_arrow_components:
  shows "Yoneda_arrow α 𝔉 r uNTMap = 
    (λd𝔄Obj. Yoneda_component 𝔉 r u d)"
    and "Yoneda_arrow α 𝔉 r uNTDom = HomO.Cα𝔄(r,-)"
    and "Yoneda_arrow α 𝔉 r uNTCod = 𝔉"
    and "Yoneda_arrow α 𝔉 r uNTDGDom = 𝔄"
    and "Yoneda_arrow α 𝔉 r uNTDGCod = cat_Set α"
  unfolding Yoneda_arrow_def nt_field_simps 
  by (simp_all add: nat_omega_simps cat_cs_simps)


subsubsection‹Natural transformation map›

mk_VLambda (in is_functor) Yoneda_arrow_components(1)
  |vsv Yoneda_arrow_NTMap_vsv|
  |vdomain Yoneda_arrow_NTMap_vdomain|
  |app Yoneda_arrow_NTMap_app|

lemmas [cat_cs_simps] = is_functor.Yoneda_arrow_NTMap_vdomain

lemmas Yoneda_arrow_NTMap_app[cat_cs_simps] = 
  is_functor.Yoneda_arrow_NTMap_app


subsubsection‹Yoneda arrow is a natural transformation›

lemma (in category) cat_Yoneda_arrow_is_ntcf:
  assumes "𝔎 :  ↦↦Cαcat_Set α" 
    and "r  Obj" 
    and "u  𝔎ObjMapr"
  shows "Yoneda_arrow α 𝔎 r u : HomO.Cα(r,-) CF 𝔎 :  ↦↦Cαcat_Set α"
proof-

  interpret 𝔎: is_functor α  cat_Set α 𝔎 by (rule assms(1)) 
  note 𝔎ru = cat_Yoneda_component_is_arr[OF assms]
  let ?𝔎ru = Yoneda_component 𝔎 r u

  show ?thesis
  proof(intro is_ntcfI', unfold 𝔎.Yoneda_arrow_components)

    show "vfsequence (Yoneda_arrow α 𝔎 r u)"
      unfolding Yoneda_arrow_def by simp
    show "vcard (Yoneda_arrow α 𝔎 r u) = 5" 
      unfolding Yoneda_arrow_def by (simp add: nat_omega_simps)

    show
      "(λdObj. ?𝔎ru d)a :
        HomO.Cα(r,-)ObjMapa cat_Set α𝔎ObjMapa"
      if "a  Obj" for a
      using that assms category_axioms
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps cat_op_simps V_cs_simps 
            cs_intro: cat_cs_intros
        )
    show 
      "(λdObj. ?𝔎ru d)b Acat_Set αHomO.Cα(r,-)ArrMapf =
        𝔎ArrMapf Acat_Set α(λdObj. ?𝔎ru d)a"
      if "f : a b" for a b f
    proof-

      note 𝔐a = 𝔎ru[OF cat_is_arrD(2)[OF that]]
      note 𝔐b = 𝔎ru[OF cat_is_arrD(3)[OF that]]

      from category_axioms assms that 𝔐b have b_f:
        "?𝔎ru b Acat_Set αcf_hom  [CIdr, f] :
          Hom  r a cat_Set α𝔎ObjMapb"
        by
          (
            cs_concl cs_shallow 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
      then have dom_lhs: 
        "𝒟 ((?𝔎ru b Acat_Set αcf_hom  [CIdr, f])ArrVal) =
          Hom  r a"
          by (cs_concl cs_shallow cs_simp: cat_cs_simps)
      from assms that 𝔐a have f_a: 
        "𝔎ArrMapf Acat_Set α?𝔎ru a :
          Hom  r a cat_Set α𝔎ObjMapb"
        by (cs_concl cs_shallow cs_intro: cat_cs_intros)
      then have dom_rhs: 
        "𝒟 ((𝔎ArrMapf Acat_Set α?𝔎ru a)ArrVal) = Hom  r a"
        by (cs_concl cs_shallow cs_simp: cat_cs_simps)

      have [cat_cs_simps]:
        "?𝔎ru b Acat_Set αcf_hom  [CIdr, f] =
          𝔎ArrMapf Acat_Set α?𝔎ru a"
      proof(rule arr_Set_eqI[of α])

        from b_f show arr_Set_b_f:
          "arr_Set α (?𝔎ru b Acat_Set αcf_hom  [CIdr, f])"
          by (auto simp: cat_Set_is_arrD(1))
        interpret b_f: arr_Set α ?𝔎ru b Acat_Set αcf_hom  [CIdr, f]
          by (rule arr_Set_b_f)
        from f_a show arr_Set_f_a:
          "arr_Set α (𝔎ArrMapf Acat_Set α?𝔎ru a)"
          by (auto simp: cat_Set_is_arrD(1))
        interpret f_a: arr_Set α 𝔎ArrMapf Acat_Set α?𝔎ru a
          by (rule arr_Set_f_a)

        show 
          "(?𝔎ru b Acat_Set αcf_hom  [CIdr, f])ArrVal =
            (𝔎ArrMapf Acat_Set α?𝔎ru a)ArrVal"
        proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)
          fix q assume "q : r a"
          from category_axioms assms that this 𝔐a 𝔐b show 
            "(?𝔎ru b Acat_Set αcf_hom  [CIdr, f])ArrValq =
              (𝔎ArrMapf Acat_Set α?𝔎ru a)ArrValq"
            by
              (
                cs_concl cs_shallow
                  cs_simp: cat_cs_simps cat_op_simps
                  cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
              )
        qed (use arr_Set_b_f arr_Set_f_a in auto)

      qed (use b_f f_a in cs_concl cs_shallow cs_simp: cat_cs_simps)+
  
      from that category_axioms assms 𝔐a 𝔐b show ?thesis
        by
          (
            cs_concl 
              cs_simp: V_cs_simps cat_cs_simps cat_op_simps 
              cs_intro: cat_cs_intros
          )
    
    qed

  qed (auto simp: assms(2) cat_cs_intros)

qed



subsection‹Yoneda Lemma›

text‹
The following lemma is approximately equivalent to the Yoneda Lemma 
stated in subsection 1.15 in cite"bodo_categories_1970" 
(the first two conclusions correspond to the statement of the 
Yoneda lemma in Chapter III-2 in cite"mac_lane_categories_2010").
›

lemma (in category) cat_Yoneda_Lemma: 
  assumes "𝔎 :  ↦↦Cαcat_Set α" and "r  Obj"
  shows "v11 (Yoneda_map α 𝔎 r)" 
    and " (Yoneda_map α 𝔎 r) = 𝔎ObjMapr"
    and "(Yoneda_map α 𝔎 r)¯ = (λu𝔎ObjMapr. Yoneda_arrow α 𝔎 r u)"
proof-

  interpret 𝔎: is_functor α  cat_Set α 𝔎 by (rule assms(1)) 

  from assms(2) 𝔎.HomCod.cat_Obj_vsubset_Vset 𝔎.cf_ObjMap_app_in_HomCod_Obj 
  have 𝔎r_in_Vset: "𝔎ObjMapr  Vset α"
    by auto

  show Ym: "v11 (Yoneda_map α 𝔎 r)"
  proof(intro vsv.vsv_valeq_v11I, unfold 𝔎.Yoneda_map_vdomain these_ntcfs_iff)

    fix 𝔐 𝔑
    assume prems: 
      "𝔐 : HomO.Cα(r,-) CF 𝔎 :  ↦↦Cαcat_Set α"
      "𝔑 : HomO.Cα(r,-) CF 𝔎 :  ↦↦Cαcat_Set α"
      "Yoneda_map α 𝔎 r𝔐 = Yoneda_map α 𝔎 r𝔑"
    from prems(3) have 𝔐r_𝔑r:
      "𝔐NTMaprArrValCIdr = 𝔑NTMaprArrValCIdr"
      unfolding 
        Yoneda_map_app[OF assms(1) prems(1)] 
        Yoneda_map_app[OF assms(1) prems(2)]
      by simp

    interpret 𝔐: is_ntcf α  cat_Set α HomO.Cα(r,-) 𝔎 𝔐  
      by (rule prems(1))
    interpret 𝔑: is_ntcf α  cat_Set α HomO.Cα(r,-) 𝔎 𝔑 
      by (rule prems(2))

    show "𝔐 = 𝔑"
    proof
      (
        rule ntcf_eqI[OF prems(1,2)]; 
        (rule refl)?;
        rule vsv_eqI, 
        unfold 𝔐.ntcf_NTMap_vdomain 𝔑.ntcf_NTMap_vdomain
      )

      fix d assume prems': "d  Obj"

      note 𝔐d_simps = cat_Set_is_arrD[OF 𝔐.ntcf_NTMap_is_arr[OF prems']]
      interpret 𝔐d: arr_Set α 𝔐NTMapd by (rule 𝔐d_simps(1))

      note 𝔑d_simps = cat_Set_is_arrD[OF 𝔑.ntcf_NTMap_is_arr[OF prems']]
      interpret 𝔑d: arr_Set α 𝔑NTMapd by (rule 𝔑d_simps(1))

      show "𝔐NTMapd = 𝔑NTMapd"
      proof(rule arr_Set_eqI[of α])
        show "𝔐NTMapdArrVal = 𝔑NTMapdArrVal"
        proof
          (
            rule vsv_eqI, 
            unfold 
              𝔑d.arr_Set_ArrVal_vdomain 
              𝔐d.arr_Set_ArrVal_vdomain  
              𝔐d_simps
              𝔑d_simps
          )
          fix f assume prems'': "f  HomO.Cα(r,-)ObjMapd"
          from prems'' prems' category_axioms assms(2) have f: "f : r d"
            by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_op_intros)
          from 𝔐.ntcf_Comp_commute[OF f] have 
            "(
              𝔐NTMapd Acat_Set αHomO.Cα(r,-)ArrMapf
            )ArrValCIdr =
              (𝔎ArrMapf Acat_Set α𝔐NTMapr)ArrValCIdr"
            by simp
          from this category_axioms assms(2) f prems prems' have 𝔐df:
            "𝔐NTMapdArrValf =
              𝔎ArrMapfArrVal𝔐NTMaprArrValCIdr"
            by 
              (
                cs_prems cs_shallow
                  cs_simp: cat_cs_simps cat_op_simps 
                  cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
              )
          from 𝔑.ntcf_Comp_commute[OF f] have
            "(
              𝔑NTMapd Acat_Set αHomO.Cα(r,-)ArrMapf
            )ArrValCIdr = 
              (𝔎ArrMapf Acat_Set α𝔑NTMapr)ArrValCIdr"
            by simp
          from this category_axioms assms(2) f prems prems' have 𝔑df:
            "𝔑NTMapdArrValf =
              𝔎ArrMapfArrVal𝔑NTMaprArrValCIdr"
            by
              (
                cs_prems cs_shallow
                  cs_simp: cat_cs_simps cat_op_simps 
                  cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
              )
          show "𝔐NTMapdArrValf = 𝔑NTMapdArrValf"
            unfolding 𝔐df 𝔑df 𝔐r_𝔑r by simp
        qed auto

      qed (simp_all add: 𝔐d_simps 𝔑d_simps)

    qed auto

  qed (auto simp: Yoneda_map_vsv)

  interpret Ym: v11 Yoneda_map α 𝔎 r by (rule Ym)

  have YY: "Yoneda_map α 𝔎 rYoneda_arrow α 𝔎 r a = a" 
    if "a  𝔎ObjMapr" for a
  proof-
    note cat_Yoneda_arrow_is_ntcf[OF assms that]
    moreover with assms have Ya: "Yoneda_arrow α 𝔎 r a  𝒟 (Yoneda_map α 𝔎 r)" 
      by 
        (
          cs_concl cs_shallow 
            cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros
        )
    ultimately show "Yoneda_map α 𝔎 rYoneda_arrow α 𝔎 r a = a"
      using assms that 𝔎r_in_Vset
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  qed        

  show [simp]: " (Yoneda_map α 𝔎 r) = 𝔎ObjMapr"
  proof(intro vsubset_antisym)

    show " (Yoneda_map α 𝔎 r)  𝔎ObjMapr"
      unfolding Yoneda_map_def
    proof(intro vrange_VLambda_vsubset, unfold these_ntcfs_iff 𝔎.cf_HomDom)
      fix 𝔐 assume prems: "𝔐 : HomO.Cα(r,-) CF 𝔎 :  ↦↦Cαcat_Set α"
      then interpret 𝔐: is_ntcf α  cat_Set α HomO.Cα(r,-) 𝔎 𝔐 .
      note 𝔐r_simps = cat_Set_is_arrD[OF 𝔐.ntcf_NTMap_is_arr[OF assms(2)]]
      interpret 𝔐r: arr_Set α 𝔐NTMapr by (rule 𝔐r_simps(1))
      from prems category_axioms assms(2) have 
        "CIdr  𝒟 (𝔐NTMaprArrVal)"
        unfolding 𝔐r.arr_Set_ArrVal_vdomain 𝔐r_simps
        by 
          (
            cs_concl cs_shallow 
              cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
          )
      then have "𝔐NTMaprArrValCIdr   (𝔐NTMaprArrVal)"
        by (blast elim: 𝔐r.ArrVal.vsv_value)
      then show "𝔐NTMaprArrValCIdr  𝔎ObjMapr"
        by (auto simp: 𝔐r_simps dest!: vsubsetD[OF 𝔐r.arr_Set_ArrVal_vrange])
    qed
    
    show "𝔎ObjMapr   (Yoneda_map α 𝔎 r)"
    proof(intro vsubsetI)
      fix u assume prems: "u  𝔎ObjMapr"
      from cat_Yoneda_arrow_is_ntcf[OF assms prems] have 
        "Yoneda_arrow α 𝔎 r u  𝒟 (Yoneda_map α 𝔎 r)" 
        by 
          (
            cs_concl cs_shallow 
              cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros
          )
      with YY[OF prems] show "u   (Yoneda_map α 𝔎 r)" 
        by (force dest!: vdomain_atD)
    qed

  qed

  show "(Yoneda_map α 𝔎 r)¯ = (λu𝔎ObjMapr. Yoneda_arrow α 𝔎 r u)"
  proof(rule vsv_eqI, unfold vdomain_vconverse vdomain_VLambda)
    from Ym show "vsv ((Yoneda_map α 𝔎 r)¯)" by auto
    show "(Yoneda_map α 𝔎 r)¯a = (λu𝔎ObjMapr. Yoneda_arrow α 𝔎 r u)a"
      if "a   (Yoneda_map α 𝔎 r)" for a
    proof-
      from that have a: "a  𝔎ObjMapr" by simp
      note Ya = cat_Yoneda_arrow_is_ntcf[OF assms a]
      then have "Yoneda_arrow α 𝔎 r a  𝒟 (Yoneda_map α 𝔎 r)"
        by 
          (
            cs_concl cs_shallow
              cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros
          )
      with Ya YY[OF a] a show ?thesis
        by 
          (
            intro Ym.v11_vconverse_app[
              unfolded 𝔎.Yoneda_map_vdomain these_ntcfs_iff
              ]
          )
          (simp_all add: these_ntcfs_iff cat_cs_simps)
    qed
  qed auto

qed



subsection‹Inverse of the Yoneda map›

lemma (in category) inv_Yoneda_map_v11: 
  assumes "𝔎 :  ↦↦Cαcat_Set α" and "r  Obj"
  shows "v11 ((Yoneda_map α 𝔎 r)¯)"
  using cat_Yoneda_Lemma(1)[OF assms] by (simp add: v11.v11_vconverse)

lemma (in category) inv_Yoneda_map_vdomain: 
  assumes "𝔎 :  ↦↦Cαcat_Set α" and "r  Obj"
  shows "𝒟 ((Yoneda_map α 𝔎 r)¯) = 𝔎ObjMapr"
  unfolding cat_Yoneda_Lemma(3)[OF assms] by simp

lemmas [cat_cs_simps] = category.inv_Yoneda_map_vdomain

lemma (in category) inv_Yoneda_map_app: 
  assumes "𝔎 :  ↦↦Cαcat_Set α" and "r  Obj" and "u  𝔎ObjMapr"
  shows "(Yoneda_map α 𝔎 r)¯u = Yoneda_arrow α 𝔎 r u"
  using assms(3) unfolding cat_Yoneda_Lemma(3)[OF assms(1,2)] by simp

lemmas [cat_cs_simps] = category.inv_Yoneda_map_app

lemma (in category) inv_Yoneda_map_vrange: 
  assumes "𝔎 :  ↦↦Cαcat_Set α"
  shows " ((Yoneda_map α 𝔎 r)¯) =
    these_ntcfs α  (cat_Set α) HomO.Cα(r,-) 𝔎"
proof-
  interpret 𝔎: is_functor α  cat_Set α 𝔎 by (rule assms(1)) 
  show ?thesis unfolding Yoneda_map_def by (simp add: cat_cs_simps)
qed



subsection‹
Component of a composition of a Hom›-natural transformation 
with natural transformations
›


subsubsection‹Definition and elementary properties›


text‹
The following definition is merely a technical generalization
that is used in the context of the description of the 
composition of a Hom›-natural transformation with a natural transformation
later in this section
(also see subsection 1.15 in cite"bodo_categories_1970").
›

definition ntcf_Hom_component :: "V  V  V  V  V"
  where "ntcf_Hom_component φ ψ a b =
    [
      (
        λfHom (φNTDGCod) (φNTCodObjMapa) (ψNTDomObjMapb).
          ψNTMapb AψNTDGCodf AψNTDGCodφNTMapa
      ),
      Hom (φNTDGCod) (φNTCodObjMapa) (ψNTDomObjMapb),
      Hom (φNTDGCod) (φNTDomObjMapa) (ψNTCodObjMapb)
    ]"


text‹Components.›

lemma ntcf_Hom_component_components: 
  shows "ntcf_Hom_component φ ψ a bArrVal =
    (
      λfHom (φNTDGCod) (φNTCodObjMapa) (ψNTDomObjMapb).
        ψNTMapb AψNTDGCodf AψNTDGCodφNTMapa
    )"
    and "ntcf_Hom_component φ ψ a bArrDom =
      Hom (φNTDGCod) (φNTCodObjMapa) (ψNTDomObjMapb)"
    and "ntcf_Hom_component φ ψ a bArrCod =
      Hom (φNTDGCod) (φNTDomObjMapa) (ψNTCodObjMapb)"
  unfolding ntcf_Hom_component_def arr_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow value›

mk_VLambda ntcf_Hom_component_components(1)
  |vsv ntcf_Hom_component_ArrVal_vsv[intro]|

context
  fixes α φ ψ 𝔉 𝔊 𝔉' 𝔊' 𝔄 𝔅 
  assumes φ: "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and ψ: "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
begin

interpretation φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule φ)
interpretation ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule ψ)

mk_VLambda 
  ntcf_Hom_component_components(1)
    [
      of φ ψ, 
      unfolded 
        φ.ntcf_NTDom ψ.ntcf_NTDom 
        φ.ntcf_NTCod ψ.ntcf_NTCod 
        φ.ntcf_NTDGDom ψ.ntcf_NTDGDom
        φ.ntcf_NTDGCod ψ.ntcf_NTDGCod
    ]
  |vdomain ntcf_Hom_component_ArrVal_vdomain|
  |app ntcf_Hom_component_ArrVal_app[unfolded in_Hom_iff]|

lemmas [cat_cs_simps] = 
  ntcf_Hom_component_ArrVal_vdomain
  ntcf_Hom_component_ArrVal_app

lemma ntcf_Hom_component_ArrVal_vrange:
  assumes "a  𝔄Obj" and "b  𝔅Obj"
  shows 
    " (ntcf_Hom_component φ ψ a bArrVal) 
      Hom  (𝔉ObjMapa) (𝔊'ObjMapb)"
proof
  (
    rule vsv.vsv_vrange_vsubset, 
    unfold ntcf_Hom_component_ArrVal_vdomain in_Hom_iff
  )
  fix f assume "f : 𝔊ObjMapa 𝔉'ObjMapb"
  with assms φ ψ show 
    "ntcf_Hom_component φ ψ a bArrValf : 𝔉ObjMapa 𝔊'ObjMapb"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed (rule ntcf_Hom_component_ArrVal_vsv)

end


subsubsection‹Arrow domain and codomain›

context
  fixes α φ ψ 𝔉 𝔊 𝔉' 𝔊' 𝔄 𝔅 
  assumes φ: "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and ψ: "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
begin

interpretation φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule φ)
interpretation ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule ψ)

lemma ntcf_Hom_component_ArrDom[cat_cs_simps]:
  "ntcf_Hom_component φ ψ a bArrDom = Hom  (𝔊ObjMapa) (𝔉'ObjMapb)"
  unfolding ntcf_Hom_component_components by (simp add: cat_cs_simps)

lemma ntcf_Hom_component_ArrCod[cat_cs_simps]:
  "ntcf_Hom_component φ ψ a bArrCod = Hom  (𝔉ObjMapa) (𝔊'ObjMapb)"
  unfolding ntcf_Hom_component_components by (simp add: cat_cs_simps)

end


subsubsection‹
Component of a composition of a Hom›-natural transformation 
with natural transformations is an arrow in the category Set›

lemma (in category) cat_ntcf_Hom_component_is_arr:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
    and "a  op_cat 𝔄Obj"
    and "b  𝔅Obj"
  shows 
    "ntcf_Hom_component φ ψ a b :
      Hom  (𝔊ObjMapa) (𝔉'ObjMapb) cat_Set αHom  (𝔉ObjMapa) (𝔊'ObjMapb)"
proof-
  
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  interpret ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule assms(2))

  from assms have a: "a  𝔄Obj" unfolding cat_op_simps by simp

  show ?thesis
  proof(intro cat_Set_is_arrI arr_SetI)
    show "vfsequence (ntcf_Hom_component φ ψ a b)"
      unfolding ntcf_Hom_component_def by (simp add: nat_omega_simps)
    show "vcard (ntcf_Hom_component φ ψ a b) = 3"
      unfolding ntcf_Hom_component_def by (simp add: nat_omega_simps)
    from assms ntcf_Hom_component_ArrVal_vrange[OF assms(1,2) a assms(4)] show 
      " (ntcf_Hom_component φ ψ a bArrVal)  
        ntcf_Hom_component φ ψ a bArrCod"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    from assms(1,2,4) a show "ntcf_Hom_component φ ψ a bArrDom  Vset α"
      by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    from assms(1,2,4) a show "ntcf_Hom_component φ ψ a bArrCod  Vset α"
      by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  qed (use assms in auto simp: ntcf_Hom_component_components cat_cs_simps)

qed

lemma (in category) cat_ntcf_Hom_component_is_arr':
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
    and "a  op_cat 𝔄Obj"
    and "b  𝔅Obj"
    and "𝔄' = Hom  (𝔊ObjMapa) (𝔉'ObjMapb)"
    and "𝔅' = Hom  (𝔉ObjMapa) (𝔊'ObjMapb)"
    and "ℭ' = cat_Set α"
  shows "ntcf_Hom_component φ ψ a b : 𝔄' ℭ'𝔅'"
  using assms(1-4) unfolding assms(5-7) by (rule cat_ntcf_Hom_component_is_arr)
  
lemmas [cat_cs_intros] = category.cat_ntcf_Hom_component_is_arr'


subsubsection‹
Naturality of the components of a composition of 
a Hom›-natural transformation with natural transformations
›

lemma (in category) cat_ntcf_Hom_component_nat:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
    and "g : a op_cat 𝔄a'"
    and "f : b 𝔅b'" 
  shows
    "ntcf_Hom_component φ ψ a' b' Acat_Set αcf_hom  [𝔊ArrMapg, 𝔉'ArrMapf] =
      cf_hom  [𝔉ArrMapg, 𝔊'ArrMapf] Acat_Set αntcf_Hom_component φ ψ a b"
proof-

  let ?Y_ab = ntcf_Hom_component φ ψ a b
    and ?Y_a'b' = ntcf_Hom_component φ ψ a' b'
    and ?𝔊g = 𝔊ArrMapg
    and ?𝔉'f = 𝔉'ArrMapf
    and ?𝔉g = 𝔉ArrMapg
    and ?𝔊'f = 𝔊'ArrMapf
    and ?𝔊a = 𝔊ObjMapa
    and ?𝔉'b = 𝔉'ObjMapb
    and ?𝔉a' = 𝔉ObjMapa'
    and ?𝔊'b' = 𝔊'ObjMapb'

  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  interpret ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule assms(2))
  interpret Set: category α cat_Set α by (rule category_cat_Set)

  from assms(3) have g: "g : a' 𝔄a" unfolding cat_op_simps by simp

  from Set.category_axioms category_axioms assms g have a'b_Gg𝔉'f:
    "?Y_a'b' Acat_Set αcf_hom  [?𝔊g, ?𝔉'f] :
      Hom  ?𝔊a ?𝔉'b cat_Set αHom  ?𝔉a' ?𝔊'b'"
    by 
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_op_simps 
          cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
      )
  then have dom_lhs: 
    "𝒟 ((?Y_a'b' Acat_Set αcf_hom  [?𝔊g, ?𝔉'f])ArrVal) = 
      Hom  ?𝔊a ?𝔉'b"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)
  from Set.category_axioms category_axioms assms g have 𝔉g𝔊'f_ab: 
    "cf_hom  [?𝔉g, ?𝔊'f] Acat_Set α?Y_ab : 
      Hom  ?𝔊a ?𝔉'b cat_Set αHom  ?𝔉a' ?𝔊'b'"
    by 
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_op_simps 
          cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
      )
  then have dom_rhs: 
    "𝒟 ((cf_hom  [?𝔉g, ?𝔊'f] Acat_Set α?Y_ab)ArrVal) =
      Hom  ?𝔊a ?𝔉'b"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    from a'b_Gg𝔉'f show arr_Set_a'b_Gg𝔉'f:
      "arr_Set α (?Y_a'b' Acat_Set αcf_hom  [?𝔊g, ?𝔉'f])"
      by (auto dest: cat_Set_is_arrD(1))
    from 𝔉g𝔊'f_ab show arr_Set_𝔉g𝔊'f_ab: 
      "arr_Set α (cf_hom  [?𝔉g, ?𝔊'f] Acat_Set α?Y_ab)"
      by (auto dest: cat_Set_is_arrD(1))
    show 
      "(?Y_a'b' Acat_Set αcf_hom  [?𝔊g, ?𝔉'f])ArrVal =
        (cf_hom  [?𝔉g, ?𝔊'f] Acat_Set α?Y_ab)ArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)
      fix h assume prems: "h : 𝔊ObjMapa 𝔉'ObjMapb"
      from assms(1,2) g have [cat_cs_simps]:
        "ψNTMapb' A(?𝔉'f A(h A(?𝔊g AφNTMapa'))) =
          ψNTMapb' A(?𝔉'f A(h A(φNTMapa A?𝔉g)))"
        by 
          (
            cs_concl cs_shallow 
              cs_simp: is_ntcf.ntcf_Comp_commute cs_intro: cat_cs_intros
          )
      also from assms(1,2,4) prems g have " =
        (((ψNTMapb' A?𝔉'f) Ah) AφNTMapa) A?𝔉g"
        by (cs_concl cs_shallow cs_simp: cat_Comp_assoc cs_intro: cat_cs_intros) 
      also from assms(1,2,4) have " =
        (((?𝔊'f AψNTMapb) Ah) AφNTMapa) A?𝔉g"
        by 
          (
            cs_concl cs_shallow 
              cs_simp: is_ntcf.ntcf_Comp_commute cs_intro: cat_cs_intros
          )
      also from assms(1,2,4) prems g have " =
        ?𝔊'f A(ψNTMapb A(h A(φNTMapa A?𝔉g)))"
        by (cs_concl cs_simp: cat_Comp_assoc cs_intro: cat_cs_intros) (*slow*)
      finally have nat:
        "ψNTMapb' A(?𝔉'f A(h A(?𝔊g AφNTMapa'))) =
          ?𝔊'f A(ψNTMapb A(h A(φNTMapa A?𝔉g)))".
      from prems Set.category_axioms category_axioms assms(1,2,4) g show 
        "(?Y_a'b' Acat_Set αcf_hom  [?𝔊g, ?𝔉'f])ArrValh =
          (cf_hom  [?𝔉g, ?𝔊'f] Acat_Set α?Y_ab)ArrValh"
        by (*slow*)
          (
            cs_concl 
              cs_simp: nat cat_cs_simps cat_op_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
    qed (use arr_Set_a'b_Gg𝔉'f arr_Set_𝔉g𝔊'f_ab in auto)

  qed (use a'b_Gg𝔉'f 𝔉g𝔊'f_ab in cs_concl cs_shallow cs_simp: cat_cs_simps)+

qed


subsubsection‹
Composition of the components of a composition of a Hom›-natural 
transformation with natural transformations
›

lemma (in category) cat_ntcf_Hom_component_Comp:
  assumes "φ' : 𝔊 CF  : 𝔄 ↦↦Cα" 
    and "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "ψ' : 𝔊' CF ℌ' : 𝔅 ↦↦Cα" 
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
    and "a  𝔄Obj"
    and "b  𝔅Obj"
  shows 
    "ntcf_Hom_component φ ψ' a b Acat_Set αntcf_Hom_component φ' ψ a b =
      ntcf_Hom_component (φ' NTCF φ) (ψ' NTCF ψ) a b"
    (is ?φψ' Acat_Set α?φ'ψ = ?φ'φψ'ψ)
proof-

  interpret Set: category α cat_Set α by (rule category_cat_Set)

  from assms Set.category_axioms category_axioms have φψ'_φ'ψ: 
    "?φψ' Acat_Set α?φ'ψ :
      Hom  (ObjMapa) (𝔉'ObjMapb) cat_Set αHom  (𝔉ObjMapa) (ℌ'ObjMapb)"    
    by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_op_intros)
  then have dom_lhs: 
    "𝒟 ((?φψ' Acat_Set α?φ'ψ)ArrVal) = 
      Hom  (ObjMapa) (𝔉'ObjMapb)"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)
  from assms Set.category_axioms category_axioms have φ'φψ'ψ: 
    "?φ'φψ'ψ :
      Hom  (ObjMapa) (𝔉'ObjMapb) cat_Set αHom  (𝔉ObjMapa) (ℌ'ObjMapb)"    
    by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_op_intros)
  then have dom_rhs: 
    "𝒟 (?φ'φψ'ψArrVal) = Hom  (ObjMapa) (𝔉'ObjMapb)" 
    by (cs_concl cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    from φψ'_φ'ψ show arr_Set_φψ'_φ'ψ: "arr_Set α (?φψ' Acat_Set α?φ'ψ)"
      by (auto dest: cat_Set_is_arrD(1))
    from φ'φψ'ψ show arr_Set_φ'φψ'ψ: "arr_Set α ?φ'φψ'ψ"
      by (auto dest: cat_Set_is_arrD(1))
    show "(?φψ' Acat_Set α?φ'ψ)ArrVal = ?φ'φψ'ψArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)
      fix f assume "f : ObjMapa 𝔉'ObjMapb"
      with category_axioms assms Set.category_axioms show 
        "(?φψ' Acat_Set α?φ'ψ)ArrValf = ?φ'φψ'ψArrValf"
        by 
          (
            cs_concl cs_shallow
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
    qed (use arr_Set_φ'φψ'ψ arr_Set_φψ'_φ'ψ in auto)
    
  qed (use φψ'_φ'ψ φ'φψ'ψ in cs_concl cs_simp: cat_cs_simps)+

qed

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_component_Comp


subsubsection‹
Component of a composition of Hom›-natural 
transformation with the identity natural transformations
›

lemma (in category) cat_ntcf_Hom_component_ntcf_id:
  assumes "𝔉 : 𝔄 ↦↦Cα" 
    and "𝔉': 𝔅 ↦↦Cα"
    and "a  𝔄Obj"
    and "b  𝔅Obj"
  shows 
    "ntcf_Hom_component (ntcf_id 𝔉) (ntcf_id 𝔉') a b =
      cat_Set αCIdHom  (𝔉ObjMapa) (𝔉'ObjMapb)"
    (is ?𝔉𝔉' = cat_Set αCId?𝔉a𝔉'b)
proof-

  interpret 𝔉: is_functor α 𝔄  𝔉 by (rule assms(1))
  interpret 𝔉': is_functor α 𝔅  𝔉' by (rule assms(2))
  interpret Set: category α cat_Set α by (rule category_cat_Set)

  from assms Set.category_axioms category_axioms have 𝔉𝔉': 
    "?𝔉𝔉' :
      Hom  (𝔉ObjMapa) (𝔉'ObjMapb) cat_Set αHom  (𝔉ObjMapa) (𝔉'ObjMapb)"    
    by (cs_concl cs_intro: cat_cs_intros cat_op_intros)
  then have dom_lhs: "𝒟 (?𝔉𝔉'ArrVal) = Hom  (𝔉ObjMapa) (𝔉'ObjMapb)"
    by (cs_concl cs_simp: cat_cs_simps)

  from category_axioms assms Set.category_axioms have 𝔉a𝔉'b: 
    "cat_Set αCId?𝔉a𝔉'b :
      Hom  (𝔉ObjMapa) (𝔉'ObjMapb) cat_Set αHom  (𝔉ObjMapa) (𝔉'ObjMapb)"
    by 
      (
        cs_concl 
          cs_simp: cat_Set_cs_simps cat_Set_components(1) 
          cs_intro: cat_cs_intros
      )
  then have dom_rhs: 
    "𝒟 (cat_Set αCId?𝔉a𝔉'bArrVal) = Hom  (𝔉ObjMapa) (𝔉'ObjMapb)"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    from 𝔉𝔉' show arr_Set_𝔉ψ: "arr_Set α ?𝔉𝔉'" 
      by (auto dest: cat_Set_is_arrD(1))
    from 𝔉a𝔉'b show arr_Set_𝔉a𝔉'b: "arr_Set α (cat_Set αCId?𝔉a𝔉'b)" 
      by (auto dest: cat_Set_is_arrD(1))
    show "?𝔉𝔉'ArrVal = cat_Set αCId?𝔉a𝔉'bArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)  
      fix f assume "f : 𝔉ObjMapa 𝔉'ObjMapb"
      with category_axioms Set.category_axioms assms show 
        "?𝔉𝔉'ArrValf = cat_Set αCId?𝔉a𝔉'bArrValf"
        by
          (
            cs_concl
              cs_simp: cat_cs_simps cat_Set_components(1)
              cs_intro: cat_cs_intros
          )
    qed (use arr_Set_𝔉a𝔉'b in auto)
      
  qed (use 𝔉𝔉' 𝔉a𝔉'b in cs_concl cs_simp: cat_cs_simps)+

qed

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_component_ntcf_id



subsection‹
Component of a composition of a Hom›-natural transformation 
with a natural transformation
›


subsubsection‹Definition and elementary properties›

definition ntcf_lcomp_Hom_component :: "V  V  V  V"
  where "ntcf_lcomp_Hom_component φ a b =
    ntcf_Hom_component φ (ntcf_id (cf_id (φNTDGCod))) a b"

definition ntcf_rcomp_Hom_component :: "V  V  V  V"
  where "ntcf_rcomp_Hom_component ψ a b =
    ntcf_Hom_component (ntcf_id (cf_id (ψNTDGCod))) ψ a b"


subsubsection‹Arrow value›

lemma ntcf_lcomp_Hom_component_ArrVal_vsv: 
  "vsv (ntcf_lcomp_Hom_component φ a bArrVal)"
  unfolding ntcf_lcomp_Hom_component_def by (rule ntcf_Hom_component_ArrVal_vsv)

lemma ntcf_rcomp_Hom_component_ArrVal_vsv: 
  "vsv (ntcf_rcomp_Hom_component ψ a bArrVal)"
  unfolding ntcf_rcomp_Hom_component_def by (rule ntcf_Hom_component_ArrVal_vsv)

lemma ntcf_lcomp_Hom_component_ArrVal_vdomain[cat_cs_simps]: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" and "b  Obj"
  shows "𝒟 (ntcf_lcomp_Hom_component φ a bArrVal) = Hom  (𝔊ObjMapa) b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  show ?thesis
    using assms
    unfolding ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_rcomp_Hom_component_ArrVal_vdomain[cat_cs_simps]: 
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" and "a  op_cat Obj"
  shows "𝒟 (ntcf_rcomp_Hom_component ψ a bArrVal) = Hom  a (𝔉ObjMapb)"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  show ?thesis
    using assms
    unfolding cat_op_simps ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_lcomp_Hom_component_ArrVal_app[cat_cs_simps]: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "a  op_cat 𝔄Obj"
    and "b  Obj"
    and "h : 𝔊ObjMapa b"
  shows "ntcf_lcomp_Hom_component φ a bArrValh = h AφNTMapa"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  show ?thesis
    using assms
    unfolding cat_op_simps ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_rcomp_Hom_component_ArrVal_app[cat_cs_simps]: 
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "a  op_cat Obj"
    and "b  𝔅Obj"
    and "h : a 𝔉ObjMapb"
  shows "ntcf_rcomp_Hom_component ψ a bArrValh = ψNTMapb Ah"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  show ?thesis
    using assms
    unfolding cat_op_simps ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_lcomp_Hom_component_ArrVal_vrange: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "a  op_cat 𝔄Obj" 
    and "b  Obj"
  shows " (ntcf_lcomp_Hom_component φ a bArrVal)  Hom  (𝔉ObjMapa) b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms(2) have a: "a  𝔄Obj" unfolding cat_op_simps by simp
  from assms(1,3) a have 
    " (ntcf_lcomp_Hom_component φ a bArrVal) 
      Hom  (𝔉ObjMapa) (cf_id ObjMapb)"
    by 
      (
        unfold cat_op_simps ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod, 
        intro ntcf_Hom_component_ArrVal_vrange
      ) 
      (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+
  from this assms(3) show ?thesis by (cs_prems cs_shallow cs_simp: cat_cs_simps)
qed

lemma ntcf_rcomp_Hom_component_ArrVal_vrange: 
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "a  op_cat Obj"
    and "b  𝔅Obj"
  shows " (ntcf_rcomp_Hom_component ψ a bArrVal)  Hom  a (𝔊ObjMapb)"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms(2) have a: "a  Obj" unfolding cat_op_simps by simp
  from assms(1,3) a have 
    " (ntcf_rcomp_Hom_component ψ a bArrVal) 
      Hom  (cf_id ObjMapa) (𝔊ObjMapb)"
    by 
      (
        unfold ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod, 
        intro ntcf_Hom_component_ArrVal_vrange
      ) 
      (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from this a show ?thesis by (cs_prems cs_shallow cs_simp: cat_cs_simps)
qed


subsubsection‹Arrow domain and codomain›

lemma ntcf_lcomp_Hom_component_ArrDom[cat_cs_simps]:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" and "b  Obj"
  shows "ntcf_lcomp_Hom_component φ a bArrDom = Hom  (𝔊ObjMapa) b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_rcomp_Hom_component_ArrDom[cat_cs_simps]:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" and "a  op_cat Obj"
  shows "ntcf_rcomp_Hom_component ψ a bArrDom = Hom  a (𝔉ObjMapb)"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms show ?thesis
    unfolding cat_op_simps ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_lcomp_Hom_component_ArrCod[cat_cs_simps]:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" and "b  Obj"
  shows "ntcf_lcomp_Hom_component φ a bArrCod = Hom  (𝔉ObjMapa) b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_rcomp_Hom_component_ArrCod[cat_cs_simps]:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" and "a  op_cat Obj"
  shows "ntcf_rcomp_Hom_component ψ a bArrCod = Hom  a (𝔊ObjMapb)" 
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms show ?thesis
    unfolding cat_op_simps ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed


subsubsection‹
Component of a composition of a Hom›-natural transformation 
with a natural transformation is an arrow in the category Set›

lemma (in category) cat_ntcf_lcomp_Hom_component_is_arr:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "a  op_cat 𝔄Obj" 
    and "b  Obj"
  shows "ntcf_lcomp_Hom_component φ a b :
    Hom  (𝔊ObjMapa) b cat_Set αHom  (𝔉ObjMapa) b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms have a: "a  𝔄Obj" unfolding cat_op_simps by simp
  from assms(1,3) a have 
    "ntcf_lcomp_Hom_component φ a b :
      Hom  (𝔊ObjMapa) (cf_id ObjMapb) cat_Set αHom  (𝔉ObjMapa) (cf_id ObjMapb)"
    unfolding ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (intro cat_ntcf_Hom_component_is_arr)
      (cs_concl cs_intro: cat_cs_intros cat_op_intros)+
  from this assms(1,3) a show ?thesis 
    by (cs_prems cs_shallow cs_simp: cat_cs_simps)
qed

lemma (in category) cat_ntcf_lcomp_Hom_component_is_arr':
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "a  op_cat 𝔄Obj" 
    and "b  Obj"
    and "𝔄' = Hom  (𝔊ObjMapa) b"
    and "𝔅' = Hom  (𝔉ObjMapa) b"
    and "ℭ' = cat_Set α"
  shows "ntcf_lcomp_Hom_component φ a b : 𝔄' ℭ'𝔅'"
  using assms(1-3) 
  unfolding assms(4-6) 
  by (rule cat_ntcf_lcomp_Hom_component_is_arr)

lemmas [cat_cs_intros] = category.cat_ntcf_lcomp_Hom_component_is_arr'

lemma (in category) cat_ntcf_rcomp_Hom_component_is_arr:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "a  op_cat Obj" 
    and "b  𝔅Obj"
  shows "ntcf_rcomp_Hom_component ψ a b :
    Hom  a (𝔉ObjMapb) cat_Set αHom  a (𝔊ObjMapb)"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms have a: "a  Obj" unfolding cat_op_simps by simp
  from assms(1,3) a have 
    "ntcf_rcomp_Hom_component ψ a b :
      Hom  (cf_id ObjMapa) (𝔉ObjMapb) cat_Set αHom  (cf_id ObjMapa) (𝔊ObjMapb)"
    unfolding ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (intro cat_ntcf_Hom_component_is_arr)
      (cs_concl cs_intro: cat_cs_intros cat_op_intros)
  from this assms(1,3) a show ?thesis 
    by (cs_prems cs_shallow cs_simp: cat_cs_simps)
qed

lemma (in category) cat_ntcf_rcomp_Hom_component_is_arr':
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "a  op_cat Obj" 
    and "b  𝔅Obj"
    and "𝔄' = Hom  a (𝔉ObjMapb)"
    and "𝔅' = Hom  a (𝔊ObjMapb)"
    and "ℭ' = cat_Set α"
  shows "ntcf_rcomp_Hom_component ψ a b : 𝔄' ℭ'𝔅'"
  using assms(1-3) 
  unfolding assms(4-6)
  by (rule cat_ntcf_rcomp_Hom_component_is_arr)

lemmas [cat_cs_intros] = category.cat_ntcf_rcomp_Hom_component_is_arr'



subsection‹
Composition of a Hom›-natural transformation with two natural transformations
›


subsubsection‹Definition and elementary properties›


text‹See subsection 1.15 in cite"bodo_categories_1970".›

definition ntcf_Hom :: "V  V  V  V" (HomA.Cı'(/_-,_-/'))
  where "HomA.Cα(φ-,ψ-) =
    [
      (
        λab(op_cat (φNTDGDom) ×C ψNTDGDom)Obj.
          ntcf_Hom_component φ ψ (vpfst ab) (vpsnd ab)
      ),
      HomO.CαψNTDGCod(φNTCod-,ψNTDom-),
      HomO.CαψNTDGCod(φNTDom-,ψNTCod-),
      op_cat (φNTDGDom) ×C ψNTDGDom,
      cat_Set α
    ]"


text‹Components.›

lemma ntcf_Hom_components:
  shows "HomA.Cα(φ-,ψ-)NTMap =
    (
      λab(op_cat (φNTDGDom) ×C ψNTDGDom)Obj.
        ntcf_Hom_component φ ψ (vpfst ab) (vpsnd ab)
    )"
    and "HomA.Cα(φ-,ψ-)NTDom =
      HomO.CαψNTDGCod(φNTCod-,ψNTDom-)"
    and "HomA.Cα(φ-,ψ-)NTCod =
      HomO.CαψNTDGCod(φNTDom-,ψNTCod-)"
    and "HomA.Cα(φ-,ψ-)NTDGDom = op_cat (φNTDGDom) ×C ψNTDGDom"
    and "HomA.Cα(φ-,ψ-)NTDGCod = cat_Set α"
  unfolding ntcf_Hom_def nt_field_simps by (simp_all add: nat_omega_simps)


subsubsection‹Natural transformation map›

mk_VLambda ntcf_Hom_components(1)
  |vsv ntcf_Hom_NTMap_vsv|

context
  fixes α φ ψ 𝔉 𝔊 𝔉' 𝔊' 𝔄 𝔅 
  assumes φ: "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and ψ: "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
begin

interpretation φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule φ)
interpretation ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule ψ)

mk_VLambda ntcf_Hom_components(1)[of _ φ ψ, simplified]
  |vdomain ntcf_Hom_NTMap_vdomain[unfolded in_Hom_iff]|

lemmas [cat_cs_simps] = ntcf_Hom_NTMap_vdomain

lemma ntcf_Hom_NTMap_app[cat_cs_simps]:
  assumes "[a, b]  (op_cat 𝔄 ×C 𝔅)Obj"
  shows "HomA.Cα(φ-,ψ-)NTMapa, b = ntcf_Hom_component φ ψ a b"
  using assms
  unfolding ntcf_Hom_components
  by (simp add: nat_omega_simps cat_cs_simps)

end

lemma (in category) ntcf_Hom_NTMap_vrange: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
  shows " (HomA.Cα(φ-,ψ-)NTMap)  cat_Set αArr"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  interpret ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule assms(2))
  show ?thesis
  proof
    (
      rule vsv.vsv_vrange_vsubset, 
      unfold ntcf_Hom_NTMap_vdomain[OF assms] cat_cs_simps
    )
    fix ab assume "ab  (op_cat 𝔄 ×C 𝔅)Obj"
    then obtain a b
      where ab_def: "ab = [a, b]" 
        and a: "a  op_cat 𝔄Obj" 
        and b: "b  𝔅Obj"
      by 
        (
          rule cat_prod_2_ObjE[
            OF φ.NTDom.HomDom.category_op ψ.NTDom.HomDom.category_axioms
            ]
        )
    from assms a b category_cat_Set category_axioms show 
      "HomA.Cα(φ-,ψ-)NTMapab  cat_Set αArr"
      unfolding ab_def cat_op_simps
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
        )
  qed (simp add: ntcf_Hom_NTMap_vsv)
qed


subsubsection‹
Composition of a Hom›-natural transformation with 
two natural transformations is a natural transformation
›

lemma (in category) cat_ntcf_Hom_is_ntcf: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
  shows "HomA.Cα(φ-,ψ-) :
    HomO.Cα(𝔊-,𝔉'-) CF HomO.Cα(𝔉-,𝔊'-) :
    op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
proof-

  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  interpret ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule assms(2))

  show ?thesis
  proof(intro is_ntcfI')
    show "vfsequence (HomA.Cα(φ-,ψ-))" unfolding ntcf_Hom_def by simp
    show "vcard (HomA.Cα(φ-,ψ-)) = 5"
      unfolding ntcf_Hom_def by (simp add: nat_omega_simps)
    from assms category_axioms show 
      "HomO.Cα(𝔊-,𝔉'-) : op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
      by (cs_concl cs_shallow cs_intro: cat_cs_intros)
    from assms category_axioms show 
      "HomO.Cα(𝔉-,𝔊'-) : op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
      by (cs_concl cs_shallow cs_intro: cat_cs_intros)
    from assms show "𝒟 (HomA.Cα(φ-,ψ-)NTMap) = (op_cat 𝔄 ×C 𝔅)Obj"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    show "HomA.Cα(φ-,ψ-)NTMapab :
      HomO.Cα(𝔊-,𝔉'-)ObjMapab cat_Set αHomO.Cα(𝔉-,𝔊'-)ObjMapab"
      if "ab  (op_cat 𝔄 ×C 𝔅)Obj" for ab 
    proof-
      from that obtain a b
        where ab_def: "ab = [a, b]" 
          and a: "a  op_cat 𝔄Obj" 
          and b: "b  𝔅Obj"
        by 
          (
            rule cat_prod_2_ObjE[
              OF φ.NTDom.HomDom.category_op ψ.NTDom.HomDom.category_axioms
              ]
          )
      from category_axioms assms a b show 
        "HomA.Cα(φ-,ψ-)NTMapab :
          HomO.Cα(𝔊-,𝔉'-)ObjMapab cat_Set αHomO.Cα(𝔉-,𝔊'-)ObjMapab"
        unfolding ab_def cat_op_simps
        by 
          (
            cs_concl cs_shallow 
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
    qed
    show 
      "HomA.Cα(φ-,ψ-)NTMapa'b' Acat_Set αHomO.Cα(𝔊-,𝔉'-)ArrMapgf =
        HomO.Cα(𝔉-,𝔊'-)ArrMapgf Acat_Set αHomA.Cα(φ-,ψ-)NTMapab"
      if "gf : ab op_cat 𝔄 ×C 𝔅a'b'" for ab a'b' gf
    proof-
      from that obtain g f a b a' b'
        where gf_def: "gf = [g, f]" 
          and ab_def: "ab = [a, b]" 
          and a'b'_def: "a'b' = [a', b']"
          and g: "g : a op_cat 𝔄a'"
          and f: "f : b 𝔅b'" 
        by 
          (
            elim 
              cat_prod_2_is_arrE[
                OF φ.NTDom.HomDom.category_op ψ.NTDom.HomDom.category_axioms
                ]
          )
      from assms category_axioms that g f show ?thesis
        unfolding gf_def ab_def a'b'_def cat_op_simps
        by (*slow*)
          (
            cs_concl 
              cs_simp: cat_ntcf_Hom_component_nat cat_cs_simps cat_op_simps
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
    qed
  qed (auto simp: ntcf_Hom_components cat_cs_simps)

qed

lemma (in category) cat_ntcf_Hom_is_ntcf': 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" 
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
    and "β = α"
    and "𝔄' = HomO.Cα(𝔊-,𝔉'-)"
    and "𝔅' = HomO.Cα(𝔉-,𝔊'-)"
    and "ℭ' = op_cat 𝔄 ×C 𝔅"
    and "𝔇' = cat_Set α"
  shows "HomA.Cα(φ-,ψ-) : 𝔄' CF 𝔅' : ℭ' ↦↦Cβ𝔇'"
  using assms(1-2) unfolding assms(3-7) by (rule cat_ntcf_Hom_is_ntcf)

lemmas [cat_cs_intros] = category.cat_ntcf_Hom_is_ntcf'


subsubsection‹
Composition of a Hom›-natural transformation with 
two vertical compositions of natural transformations
›

lemma (in category) cat_ntcf_Hom_vcomp:
  assumes "φ' : 𝔊 CF  : 𝔄 ↦↦Cα" 
    and "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "ψ' : 𝔊' CF ℌ' : 𝔅 ↦↦Cα" 
    and "ψ : 𝔉' CF 𝔊' : 𝔅 ↦↦Cα"
  shows 
    "HomA.Cα(φ' NTCF φ-,ψ' NTCF ψ-) =
      HomA.Cα(φ-,ψ'-) NTCF HomA.Cα(φ'-,ψ-)"
proof(rule ntcf_eqI[of α])

  interpret φ': is_ntcf α 𝔄  𝔊  φ' by (rule assms(1))
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(2))
  interpret ψ': is_ntcf α 𝔅  𝔊' ℌ' ψ' by (rule assms(3))
  interpret ψ: is_ntcf α 𝔅  𝔉' 𝔊' ψ by (rule assms(4))

  from category_axioms assms show H_vcomp:
    "HomA.Cα(φ' NTCF φ-,ψ' NTCF ψ-) :
      HomO.Cα(-,𝔉'-) CF HomO.Cα(𝔉-,ℌ'-) :
      op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms show vcomp_H:
    "HomA.Cα(φ-,ψ'-) NTCF HomA.Cα(φ'-,ψ-) :
      HomO.Cα(-,𝔉'-) CF HomO.Cα(𝔉-,ℌ'-) :
      op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms H_vcomp have dom_H_vcomp:
    "𝒟 (HomA.Cα(φ' NTCF φ-,ψ' NTCF ψ-)NTMap) = (op_cat 𝔄 ×C 𝔅)Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms H_vcomp have dom_vcomp_H:
    "𝒟 ((HomA.Cα(φ-,ψ'-) NTCF HomA.Cα(φ'-,ψ-))NTMap) =
      (op_cat 𝔄 ×C 𝔅)Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

  show "HomA.Cα(φ' NTCF φ-,ψ' NTCF ψ-)NTMap =
    (HomA.Cα(φ-,ψ'-) NTCF HomA.Cα(φ'-,ψ-))NTMap"
  proof(rule vsv_eqI, unfold dom_H_vcomp dom_vcomp_H)
    fix ab assume prems: "ab  (op_cat 𝔄 ×C 𝔅)Obj"
    then obtain a b
      where ab_def: "ab = [a, b]" 
        and a: "a  𝔄Obj" 
        and b: "b  𝔅Obj"
      by 
        ( 
          auto 
            elim: 
              cat_prod_2_ObjE[
                OF φ'.NTDom.HomDom.category_op ψ'.NTDom.HomDom.category_axioms
                ]
            simp: cat_op_simps
        )
    from 
      assms a b
      category_axioms 
      φ'.NTDom.HomDom.category_axioms
      ψ'.NTDom.HomDom.category_axioms 
    show
      "HomA.Cα(φ' NTCF φ-,ψ' NTCF ψ-)NTMapab =
        (HomA.Cα(φ-,ψ'-) NTCF HomA.Cα(φ'-,ψ-))NTMapab"
      by
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps ab_def
            cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
        )
  qed (auto simp: ntcf_Hom_NTMap_vsv cat_cs_intros)

qed simp_all

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_vcomp

lemma (in category) cat_ntcf_Hom_ntcf_id:
  assumes "𝔉 : 𝔄 ↦↦Cα" and "𝔉': 𝔅 ↦↦Cα"
  shows "HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-) = ntcf_id HomO.Cα(𝔉-,𝔉'-)"
proof(rule ntcf_eqI[of α])

  interpret 𝔉: is_functor α 𝔄  𝔉 by (rule assms(1))
  interpret 𝔉': is_functor α 𝔅  𝔉' by (rule assms(2))

  from category_axioms assms show H_id:
    "HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-) :
      HomO.Cα(𝔉-,𝔉'-) CF HomO.Cα(𝔉-,𝔉'-) :
      op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms show id_H:
    "ntcf_id HomO.Cα(𝔉-,𝔉'-) :
      HomO.Cα(𝔉-,𝔉'-) CF HomO.Cα(𝔉-,𝔉'-) :
      op_cat 𝔄 ×C 𝔅 ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

  from category_axioms assms H_id have dom_H_id:
    "𝒟 (HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-)NTMap) = (op_cat 𝔄 ×C 𝔅)Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms H_id have dom_id_H:
    "𝒟 (ntcf_id HomO.Cα(𝔉-,𝔉'-)NTMap) = (op_cat 𝔄 ×C 𝔅)Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

  show 
    "HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-)NTMap =
      ntcf_id HomO.Cα(𝔉-,𝔉'-)NTMap"
  proof(rule vsv_eqI, unfold dom_H_id dom_id_H)
    show "vsv (HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-)NTMap)" 
      by (rule ntcf_Hom_NTMap_vsv)
    from id_H show "vsv (ntcf_id HomO.Cα(𝔉-,𝔉'-)NTMap)"
      by (intro is_functor.ntcf_id_NTMap_vsv) 
        (cs_concl cs_shallow cs_intro: cat_cs_intros)
    fix ab assume "ab  (op_cat 𝔄 ×C 𝔅)Obj"
    then obtain a b
      where ab_def: "ab = [a, b]" 
        and a: "a  𝔄Obj" 
        and b: "b  𝔅Obj"
      by 
        ( 
          auto 
            elim: 
              cat_prod_2_ObjE[OF 𝔉.HomDom.category_op 𝔉'.HomDom.category_axioms]
            simp: cat_op_simps
        )
    from category_axioms assms a b H_id id_H show
      "HomA.Cα(ntcf_id 𝔉-,ntcf_id 𝔉'-)NTMapab = 
        ntcf_id HomO.Cα(𝔉-,𝔉'-)NTMapab"
      unfolding ab_def
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps cat_op_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
        )
  qed simp

qed simp_all

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_ntcf_id



subsection‹
Composition of a Hom›-natural transformation with a natural transformation
›


subsubsection‹Definition and elementary properties›


text‹See subsection 1.15 in cite"bodo_categories_1970".›

definition ntcf_lcomp_Hom :: "V  V  V" (HomA.Cı'(/_-,-/'))
  where "HomA.Cα(φ-,-) = HomA.Cα(φ-,ntcf_id (cf_id (φNTDGCod))-)"

definition ntcf_rcomp_Hom :: "V  V  V" (HomA.Cı'(/-,_-/'))
  where "HomA.Cα(-,ψ-) = HomA.Cα(ntcf_id (cf_id (ψNTDGCod))-,ψ-)"


subsubsection‹Natural transformation map›

lemma ntcf_lcomp_Hom_NTMap_vsv: "vsv (HomA.Cα(φ-,-)NTMap)"
  unfolding ntcf_lcomp_Hom_def by (rule ntcf_Hom_NTMap_vsv)

lemma ntcf_rcomp_Hom_NTMap_vsv: "vsv (HomA.Cα(-,ψ-)NTMap)"
  unfolding ntcf_rcomp_Hom_def by (rule ntcf_Hom_NTMap_vsv)

lemma ntcf_lcomp_Hom_NTMap_vdomain[cat_cs_simps]:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" 
  shows "𝒟 (HomA.Cα(φ-,-)NTMap) = (op_cat 𝔄 ×C )Obj"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_lcomp_Hom_def φ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_rcomp_Hom_NTMap_vdomain[cat_cs_simps]:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" 
  shows "𝒟 (HomA.Cα(-,ψ-)NTMap) = (op_cat  ×C 𝔅)Obj"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_rcomp_Hom_def ψ.ntcf_NTDGCod
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemma ntcf_lcomp_Hom_NTMap_app[cat_cs_simps]:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
    and "a  op_cat 𝔄Obj"
    and "b  Obj"
  shows "HomA.Cα(φ-,-)NTMapa, b = ntcf_lcomp_Hom_component φ a b"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms show ?thesis
    unfolding
      ntcf_lcomp_Hom_def ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
      cat_op_simps
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_op_simps
          cs_intro: cat_cs_intros cat_prod_cs_intros
      )
qed

lemma ntcf_rcomp_Hom_NTMap_app[cat_cs_simps]:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "a  op_cat Obj"
    and "b  𝔅Obj"
  shows "HomA.Cα(-,ψ-)NTMapa, b = ntcf_rcomp_Hom_component ψ a b"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms show ?thesis
    unfolding 
      ntcf_rcomp_Hom_def ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
      cat_op_simps
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_op_simps 
          cs_intro: cat_cs_intros cat_prod_cs_intros
      )
qed

lemma (in category) ntcf_lcomp_Hom_NTMap_vrange:
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
  shows " (HomA.Cα(φ-,-)NTMap)  cat_Set αArr"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_lcomp_Hom_def ntcf_lcomp_Hom_component_def φ.ntcf_NTDGCod
    by (intro ntcf_Hom_NTMap_vrange) (cs_concl cs_intro: cat_cs_intros)+
qed

lemma (in category) ntcf_rcomp_Hom_NTMap_vrange:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
  shows " (HomA.Cα(-,ψ-)NTMap)  cat_Set αArr"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms show ?thesis
    unfolding ntcf_rcomp_Hom_def ntcf_rcomp_Hom_component_def ψ.ntcf_NTDGCod
    by (intro ntcf_Hom_NTMap_vrange) (cs_concl cs_intro: cat_cs_intros)+
qed


subsubsection‹
Composition of a Hom›-natural transformation with 
a natural transformation is a natural transformation
›

lemma (in category) cat_ntcf_lcomp_Hom_is_ntcf: 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" 
  shows "HomA.Cα(φ-,-) :
    HomO.Cα(𝔊-,-) CF HomO.Cα(𝔉-,-) : op_cat 𝔄 ×C  ↦↦Cαcat_Set α"
proof-
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(1))
  from assms category_axioms show ?thesis
    unfolding 
      ntcf_lcomp_Hom_def cf_bcomp_Hom_cf_lcomp_Hom[symmetric] φ.ntcf_NTDGCod
    by (intro category.cat_ntcf_Hom_is_ntcf)
      (cs_concl cs_intro: cat_cs_intros)+
qed

lemma (in category) cat_ntcf_lcomp_Hom_is_ntcf': 
  assumes "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα" 
    and "β = α"
    and "𝔄' = HomO.Cα(𝔊-,-)"
    and "𝔅' = HomO.Cα(𝔉-,-)"
    and "ℭ' = op_cat 𝔄 ×C "
    and "𝔇' = cat_Set α"
  shows "HomA.Cα(φ-,-) : 𝔄' CF 𝔅' : ℭ' ↦↦Cβ𝔇'"
  using assms(1) unfolding assms(2-6) by (rule cat_ntcf_lcomp_Hom_is_ntcf)

lemmas [cat_cs_intros] = category.cat_ntcf_lcomp_Hom_is_ntcf'

lemma (in category) cat_ntcf_rcomp_Hom_is_ntcf:
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" 
  shows "HomA.Cα(-,ψ-) :
    HomO.Cα(-,𝔉-) CF HomO.Cα(-,𝔊-) : op_cat  ×C 𝔅 ↦↦Cαcat_Set α"
proof-
  interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule assms(1))
  from assms category_axioms show ?thesis
    unfolding 
      ntcf_rcomp_Hom_def cf_bcomp_Hom_cf_rcomp_Hom[symmetric] ψ.ntcf_NTDGCod
    by (intro category.cat_ntcf_Hom_is_ntcf)
      (cs_concl cs_intro: cat_cs_intros)+
qed

lemma (in category) cat_ntcf_rcomp_Hom_is_ntcf':
  assumes "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "β = α"
    and "𝔄' = HomO.Cα(-,𝔉-)"
    and "𝔅' = HomO.Cα(-,𝔊-)"
    and "ℭ' = op_cat  ×C 𝔅"  
    and "𝔇' = cat_Set α"
  shows "HomA.Cα(-,ψ-) : 𝔄' CF 𝔅' : ℭ' ↦↦Cα𝔇'"
  using assms(1) unfolding assms(2-6) by (rule cat_ntcf_rcomp_Hom_is_ntcf)

lemmas [cat_cs_intros] = category.cat_ntcf_rcomp_Hom_is_ntcf'


subsubsection‹
Component of a composition of a Hom›-natural transformation 
with a natural transformation and the Yoneda component
›

lemma (in category) cat_ntcf_lcomp_Hom_component_is_Yoneda_component:
  assumes "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "b  op_cat 𝔅Obj"
    and "c  Obj"
  shows 
    "ntcf_lcomp_Hom_component φ b c =
      Yoneda_component HomO.Cα(𝔉ObjMapb,-) (𝔊ObjMapb) (φNTMapb) c"
  (is ?lcomp = ?Yc)
proof-

  interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))

  from assms(2) have b: "b  𝔅Obj" unfolding cat_op_simps by clarsimp
  from b have 𝔉b: "𝔉ObjMapb  Obj" and 𝔊b: "𝔊ObjMapb  Obj"
    by (auto intro: cat_cs_intros)
  from assms(1,3) b category_axioms have φb:
    "φNTMapb  HomO.Cα(𝔉ObjMapb,-)ObjMap𝔊ObjMapb"
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
      )

  have lcomp:
    "?lcomp : Hom  (𝔊ObjMapb) c cat_Set αHom  (𝔉ObjMapb) c"
    by (rule cat_ntcf_lcomp_Hom_component_is_arr[OF assms])
  then have dom_lhs: "𝒟 (?lcompArrVal) = Hom  (𝔊ObjMapb) c"
    by (cs_concl cs_simp: cat_cs_simps)  

  have Yc: "?Yc :
    Hom  (𝔊ObjMapb) c cat_Set αHomO.Cα(𝔉ObjMapb,-)ObjMapc"
    by 
      (
        rule cat_Yoneda_component_is_arr[
          OF cat_cf_Hom_snd_is_functor[OF 𝔉b] 𝔊b φb assms(3)
          ]
      )
  then have dom_rhs: "𝒟 (?YcArrVal) = Hom  (𝔊ObjMapb) c"
    by (cs_concl cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
  
    from lcomp show "arr_Set α ?lcomp" by (auto dest: cat_Set_is_arrD(1))
    from Yc show "arr_Set α ?Yc" by (auto dest: cat_Set_is_arrD(1))
  
    show "?lcompArrVal = ?YcArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      from assms(1) b category_axioms show "vsv (?YcArrVal)"
        by (intro is_functor.Yoneda_component_ArrVal_vsv)
          (cs_concl cs_shallow cs_intro: cat_cs_intros)
      show "?lcompArrValf = ?YcArrValf"
        if "f  Hom  (𝔊ObjMapb) c" for f
      proof-
        from that have "f : 𝔊ObjMapb c" by simp
        with category_axioms assms(1,3) b show ?thesis
          by 
            (
              cs_concl cs_shallow
                cs_simp: cat_cs_simps cat_op_simps 
                cs_intro: cat_cs_intros cat_op_intros
            )
      qed
    qed (simp_all add: ntcf_lcomp_Hom_component_ArrVal_vsv)
    
    from Yc category_axioms assms(1,3) b have
      "?Yc : Hom  (𝔊ObjMapb) c cat_Set αHom  (𝔉ObjMapb) c"
      by 
        (
          cs_prems cs_shallow 
            cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
            ) 
    with lcomp show "?lcompArrCod = ?YcArrCod"
      by (cs_concl cs_simp: cat_cs_simps)
  
  qed (use lcomp Yc in cs_concl cs_simp: cat_cs_simps)

qed


subsubsection‹
Composition of a Hom›-natural transformation with 
a vertical composition of natural transformations
›

lemma (in category) cat_ntcf_lcomp_Hom_vcomp:
  assumes "φ' : 𝔊 CF  : 𝔄 ↦↦Cα" and "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
  shows "HomA.Cα(φ' NTCF φ-,-) = HomA.Cα(φ-,-) NTCF HomA.Cα(φ'-,-)"
proof-
  interpret φ': is_ntcf α 𝔄  𝔊  φ' by (rule assms(1))
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(2))
  from category_axioms have ntcf_id_cf_id:
    "ntcf_id (cf_id ) = ntcf_id (cf_id ) NTCF ntcf_id (cf_id )"
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms show ?thesis
    unfolding 
      ntcf_lcomp_Hom_def
      ntsmcf_vcomp_components 
      dghm_id_components 
      φ'.ntcf_NTDGCod
      φ.ntcf_NTDGCod
    by (subst ntcf_id_cf_id) 
      (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_vcomp

lemma (in category) cat_ntcf_rcomp_Hom_vcomp:
  assumes "φ' : 𝔊 CF  : 𝔄 ↦↦Cα" and "φ : 𝔉 CF 𝔊 : 𝔄 ↦↦Cα"
  shows "HomA.Cα(-,φ' NTCF φ-) = HomA.Cα(-,φ'-) NTCF HomA.Cα(-,φ-)"
proof-
  interpret φ': is_ntcf α 𝔄  𝔊  φ' by (rule assms(1))
  interpret φ: is_ntcf α 𝔄  𝔉 𝔊 φ by (rule assms(2))
  from category_axioms have ntcf_id_cf_id:
    "ntcf_id (cf_id ) = ntcf_id (cf_id ) NTCF ntcf_id (cf_id )"
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from category_axioms assms show ?thesis
    unfolding 
      ntcf_rcomp_Hom_def
      ntsmcf_vcomp_components 
      dghm_id_components 
      φ'.ntcf_NTDGCod
      φ.ntcf_NTDGCod
    by (subst ntcf_id_cf_id) 
      (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
qed

lemmas [cat_cs_simps] = category.cat_ntcf_rcomp_Hom_vcomp


subsubsection‹
Composition of a Hom›-natural transformation with an identity natural 
transformation
›

lemma (in category) cat_ntcf_lcomp_Hom_ntcf_id:
  assumes "𝔉 : 𝔄 ↦↦Cα"
  shows "HomA.Cα(ntcf_id 𝔉-,-) = ntcf_id HomO.Cα(𝔉-,-)"
proof-
  interpret 𝔉: is_functor α 𝔄  𝔉 by (rule assms(1))
  from category_axioms assms show ?thesis
    unfolding ntcf_lcomp_Hom_def ntcf_id_components 𝔉.cf_HomCod
    by
      (
        cs_concl 
          cs_simp: ntcf_lcomp_Hom_def cat_cs_simps 
          cs_intro: cat_cs_intros
      )
qed

lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_ntcf_id

lemma (in category) cat_ntcf_rcomp_Hom_ntcf_id:
  assumes "𝔉 : 𝔅 ↦↦Cα"
  shows "HomA.Cα(-,ntcf_id 𝔉-) = ntcf_id HomO.Cα(-,𝔉-)"
proof-
  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  from category_axioms assms show ?thesis
    unfolding ntcf_rcomp_Hom_def ntcf_id_components 𝔉.cf_HomCod
    by (cs_concl cs_simp: ntcf_rcomp_Hom_def cat_cs_simps cs_intro: cat_cs_intros)
qed

lemmas [cat_cs_simps] = category.cat_ntcf_rcomp_Hom_ntcf_id



subsection‹Projections of a Hom›-natural transformation›


text‹
The concept of a projection of a Hom›-natural transformation appears 
in the corollary to the Yoneda Lemma in Chapter III-2 in 
cite"mac_lane_categories_2010" (although the concept has not been given
any specific name in the aforementioned reference).
›


subsubsection‹Definition and elementary properties›

definition ntcf_Hom_snd :: "V  V  V  V" (HomA.Cı_'(/_,-/'))
  where "HomA.Cα(f,-) =
    Yoneda_arrow α (HomO.Cα(Domf,-)) (Codf) f"

definition ntcf_Hom_fst :: "V  V  V  V" (HomA.Cı_'(/-,_/'))
  where "HomA.Cα(-,f) = HomA.Cαop_cat (f,-)"


text‹Components.›

lemma (in category) cat_ntcf_Hom_snd_components:
  assumes "f : s r"
  shows "HomA.Cα(f,-)NTMap = 
    (λdObj. Yoneda_component HomO.Cα(s,-) r f d)"
    and "HomA.Cα(f,-)NTDom = HomO.Cα(r,-)"
    and "HomA.Cα(f,-)NTCod = HomO.Cα(s,-)"
    and "HomA.Cα(f,-)NTDGDom = "
    and "HomA.Cα(f,-)NTDGCod = cat_Set α"
proof-
  interpret is_functor α  cat_Set α HomO.Cα(s,-)
    using assms category_axioms by (cs_concl cs_intro: cat_cs_intros)
  show "HomA.Cα(f,-)NTMap =
    (λdObj. Yoneda_component HomO.Cα(s,-) r f d)"
    and "HomA.Cα(f,-)NTDom = HomO.Cα(r,-)"
    and "HomA.Cα(f,-)NTCod = HomO.Cα(s,-)"
    and "HomA.Cα(f,-)NTDGDom = "
    and "HomA.Cα(f,-)NTDGCod = cat_Set α"
    unfolding ntcf_Hom_snd_def cat_is_arrD[OF assms] Yoneda_arrow_components
    by simp_all
qed

lemma (in category) cat_ntcf_Hom_fst_components:
  assumes "f : r s"
  shows "HomA.Cα(-,f)NTMap =
    (λdop_cat Obj. Yoneda_component HomO.Cα(-,s) r f d)"
    and "HomA.Cα(-,f)NTDom = HomO.Cα(-,r)"
    and "HomA.Cα(-,f)NTCod = HomO.Cα(-,s)"
    and "HomA.Cα(-,f)NTDGDom = op_cat "
    and "HomA.Cα(-,f)NTDGCod = cat_Set α"
  using category_axioms assms
  unfolding 
    ntcf_Hom_fst_def
    category.cat_ntcf_Hom_snd_components[
      OF category_op, unfolded cat_op_simps, OF assms
      ]
    cat_op_simps
  by (cs_concl cs_simp: cat_op_simps cs_intro: cat_cs_intros)+


text‹Alternative definition.›

lemma (in category) ntcf_Hom_snd_def':
  assumes "f : r s"
  shows "HomA.Cα(f,-) = Yoneda_arrow α (HomO.Cα(r,-)) s f"
  using assms unfolding ntcf_Hom_snd_def by (simp add: cat_cs_simps) 

lemma (in category) ntcf_Hom_fst_def':
  assumes "f : r s"
  shows "HomA.Cα(-,f) = Yoneda_arrow α HomO.Cα(-,s) r f"
proof-
  from assms category_axioms show ?thesis
    unfolding ntcf_Hom_fst_def ntcf_Hom_snd_def cat_op_simps
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros
      )
qed


subsubsection‹Natural transformation map›

context category
begin

context
  fixes s r f
  assumes f: "f : s r"
begin

mk_VLambda cat_ntcf_Hom_snd_components(1)[OF f]
  |vsv ntcf_Hom_snd_NTMap_vsv[intro]|
  |vdomain ntcf_Hom_snd_NTMap_vdomain|
  |app ntcf_Hom_snd_NTMap_app|

end

context
  fixes s r f
  assumes f: "f : r s"
begin

mk_VLambda cat_ntcf_Hom_fst_components(1)[OF f]
  |vsv ntcf_Hom_fst_NTMap_vsv[intro]|
  |vdomain ntcf_Hom_fst_NTMap_vdomain|
  |app ntcf_Hom_fst_NTMap_app|

end

end

lemmas [cat_cs_simps] = 
  category.ntcf_Hom_snd_NTMap_vdomain
  category.ntcf_Hom_fst_NTMap_vdomain

lemmas ntcf_Hom_snd_NTMap_app[cat_cs_simps] = 
  category.ntcf_Hom_snd_NTMap_app
  category.ntcf_Hom_fst_NTMap_app


subsubsectionHom›-natural transformation projections are natural transformations
›

lemma (in category) cat_ntcf_Hom_snd_is_ntcf:
  assumes "f : s r"
  shows "HomA.Cα(f,-) :
    HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
proof-
  note f = cat_is_arrD[OF assms]
  show ?thesis
    unfolding ntcf_Hom_snd_def f
  proof(rule category.cat_Yoneda_arrow_is_ntcf)
    from assms category_axioms show "f  HomO.Cα(s,-)ObjMapr"
      by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
  qed (intro category_axioms cat_cf_Hom_snd_is_functor f)+
qed

lemma (in category) cat_ntcf_Hom_snd_is_ntcf':
  assumes "f : s r"
    and "β = α"
    and "𝔄' = HomO.Cα(r,-)"
    and "𝔅' = HomO.Cα(s,-)"
    and "ℭ' = "
    and "𝔇' = cat_Set α"
  shows "HomA.Cα(f,-) : 𝔄' CF 𝔅' : ℭ' ↦↦Cβ𝔇'"
  using assms(1) unfolding assms(2-6) by (rule cat_ntcf_Hom_snd_is_ntcf)

lemmas [cat_cs_intros] = category.cat_ntcf_Hom_snd_is_ntcf'

lemma (in category) cat_ntcf_Hom_fst_is_ntcf:
  assumes "f : r s"
  shows "HomA.Cα(-,f) :
    HomO.Cα(-,r) CF HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
proof-
  from assms have r: "r  Obj" and s: "s  Obj" by auto
  from 
    category.cat_ntcf_Hom_snd_is_ntcf[
      OF category_op, 
      unfolded cat_op_simps, 
      OF assms, 
      unfolded cat_op_cat_cf_Hom_snd[OF r] cat_op_cat_cf_Hom_snd[OF s],
      folded ntcf_Hom_fst_def
      ]
  show ?thesis .
qed

lemma (in category) cat_ntcf_Hom_fst_is_ntcf':
  assumes "f : r s"
    and "β = α"
    and "𝔄' = HomO.Cα(-,r)"
    and "𝔅' = HomO.Cα(-,s)"
    and "ℭ' = op_cat "
    and "𝔇' = cat_Set α"
  shows "HomA.Cα(-,f) : 𝔄' CF 𝔅' : ℭ' ↦↦Cβ𝔇'"
  using assms(1) unfolding assms(2-6) by (rule cat_ntcf_Hom_fst_is_ntcf)

lemmas [cat_cs_intros] = category.cat_ntcf_Hom_fst_is_ntcf'


subsubsection‹Opposite Hom›-natural transformation projections›

lemma (in category) cat_op_cat_ntcf_Hom_snd: 
  "HomA.Cαop_cat (f,-) = HomA.Cα(-,f)"
  unfolding ntcf_Hom_fst_def by simp

lemmas [cat_op_simps] = category.cat_op_cat_ntcf_Hom_snd

lemma (in category) cat_op_cat_ntcf_Hom_fst:
  "HomA.Cαop_cat (-,f) = HomA.Cα(f,-)"
  unfolding ntcf_Hom_fst_def cat_op_simps by simp

lemmas [cat_op_simps] = category.cat_op_cat_ntcf_Hom_fst


subsubsectionHom›-natural transformation projections and the Yoneda component
›

lemma (in category) cat_Yoneda_component_cf_Hom_snd_Comp:
  assumes "g : b c" and "f : a b" and "d  Obj"
  shows 
    "Yoneda_component HomO.Cα(a,-) b f d Acat_Set αYoneda_component HomO.Cα(b,-) c g d =
      Yoneda_component HomO.Cα(a,-) c (g Af) d"
    (is ?Ya b f d Acat_Set α?Yb c g d = ?Ya c (g Af) d)
proof-

  interpret Set: category α cat_Set α by (rule category_cat_Set)

  note gD = cat_is_arrD[OF assms(1)]
  note fD = cat_is_arrD[OF assms(2)]

  from assms category_axioms have Y_f:
    "?Ya b f d : Hom  b d cat_Set αHom  a d"
    by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
  moreover from assms category_axioms have Y_g: 
    "?Yb c g d : Hom  c d cat_Set αHom  b d"
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros
      )
  ultimately have Yf_Yg: 
    "?Ya b f d Acat_Set α?Yb c g d : Hom  c d cat_Set αHom  a d"
    by (auto intro: cat_cs_intros)
  from assms category_axioms have Y_gf: 
    "?Ya c (g Af) d : Hom  c d cat_Set αHom  a d"
    by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
  from Yf_Yg have dom_rhs: 
    "𝒟 ((?Ya b f d Acat_Set α?Yb c g d)ArrVal) = Hom  c d"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)
  from Y_gf have dom_lhs: "𝒟 (?Ya c (g Af) dArrVal) = Hom  c d"  
    by (cs_concl cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    from Yf_Yg show arr_Set_Yf_Yg: 
      "arr_Set α (?Ya b f d Acat_Set α?Yb c g d)"
      by (auto dest: cat_Set_is_arrD(1))
    interpret Yf_Yg: arr_Set α ?Ya b f d Acat_Set α?Yb c g d
      by (rule arr_Set_Yf_Yg)
    from Y_gf show arr_Set_Y_gf: "arr_Set α (?Ya c (g Af) d)"
      by (auto dest: cat_Set_is_arrD(1))
    interpret Yf_Yg: arr_Set α ?Ya c (g Af) d by (rule arr_Set_Y_gf)
    show
      "(?Ya b f d Acat_Set α?Yb c g d)ArrVal =
        ?Ya c (g Af) dArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)
      fix h assume "h : c d"
      with Y_gf Y_g Y_f category_axioms assms show 
        "(?Ya b f d Acat_Set α?Yb c g d)ArrValh =
          ?Ya c (g Af) dArrValh"
        (*slow*)
        by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
    qed auto
  
  qed (use Y_gf Yf_Yg in cs_concl cs_shallow cs_simp: cat_cs_simps)+

qed

lemmas [cat_cs_simps] = 
  category.cat_Yoneda_component_cf_Hom_snd_Comp[symmetric]

lemma (in category) cat_Yoneda_component_cf_Hom_snd_CId:
  assumes "c  Obj" and "d  Obj"
  shows 
    "Yoneda_component HomO.Cα(c,-) c (CIdc) d = 
      cat_Set αCIdHom  c d"
  (is ?Ycd = cat_Set αCIdHom  c d)
proof-

  interpret Set: category α cat_Set α by (rule category_cat_Set)

  from assms category_axioms have Y_CId_c: 
    "?Ycd : Hom  c d cat_Set αHom  c d"
    by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
  from Y_CId_c Set.category_axioms assms category_axioms have CId_cd:
    "cat_Set αCIdHom  c d : Hom  c d cat_Set αHom  c d"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from Y_CId_c have dom_lhs: "𝒟 (?YcdArrVal) = Hom  c d"  
    by (cs_concl cs_simp: cat_cs_simps)
  from CId_cd have dom_rhs: "𝒟 (cat_Set αCIdHom  c dArrVal) = Hom  c d"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps)

  show ?thesis
  proof(rule arr_Set_eqI[of α])
    from Y_CId_c show arr_Set_Y_CId_c: "arr_Set α ?Ycd"
      by (auto dest: cat_Set_is_arrD(1))
    interpret Yf_Yg: arr_Set α ?Ycd by (rule arr_Set_Y_CId_c)
    from CId_cd show arr_Set_CId_cd: "arr_Set α (cat_Set αCIdHom  c d)"
      by (auto dest: cat_Set_is_arrD(1))
    interpret CId_cd: arr_Set α cat_Set αCIdHom  c d
      by (rule arr_Set_CId_cd)
    show "?YcdArrVal = cat_Set αCIdHom  c dArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)
      fix h assume "h : c d"
      with CId_cd Y_CId_c category_axioms assms show 
        "?YcdArrValh = cat_Set αCIdHom  c dArrValh"
        by (cs_concl cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros)
    qed auto
  qed (use Y_CId_c CId_cd in cs_concl cs_shallow cs_simp: cat_cs_simps)+

qed

lemmas [cat_cs_simps] = category.cat_Yoneda_component_cf_Hom_snd_CId


subsubsectionHom›-natural transformation projection of a composition›

lemma (in category) cat_ntcf_Hom_snd_Comp:
  assumes "g : b c" and "f : a b"
  shows "HomA.Cα(g Af,-) = HomA.Cα(f,-) NTCF HomA.Cα(g,-)"
  (is ?H_gf = ?H_f NTCF ?H_g)
proof(rule ntcf_eqI[of α])
  from assms category_axioms show 
    "?H_gf : HomO.Cα(c,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms show "?H_f NTCF ?H_g :
    HomO.Cα(c,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms have lhs_dom: "𝒟 (?H_gfNTMap) = Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms have rhs_dom:
    "𝒟 ((?H_f NTCF ?H_g)NTMap) = Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  show "?H_gfNTMap = (?H_f NTCF ?H_g)NTMap"
  proof(rule vsv_eqI, unfold lhs_dom rhs_dom)
    fix d assume "d  Obj" 
    with assms category_axioms show 
      "?H_gfNTMapd = (?H_f NTCF ?H_g)NTMapd"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  qed (use assms in auto intro: cat_cs_intros)
qed auto

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_snd_Comp

lemma (in category) cat_ntcf_Hom_fst_Comp:
  assumes "g : b c" and "f : a b"
  shows "HomA.Cα(-,g Af) = HomA.Cα(-,g) NTCF HomA.Cα(-,f)"
proof-
  note category.cat_ntcf_Hom_snd_Comp[
      OF category_op, unfolded cat_op_simps, OF assms(2,1)
      ]
  from this category_axioms assms show ?thesis
    by (cs_prems cs_shallow cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp
qed

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_fst_Comp


subsubsectionHom›-natural transformation projection of an identity›

lemma (in category) cat_ntcf_Hom_snd_CId:
  assumes "c  Obj"
  shows "HomA.Cα(CIdc,-) = ntcf_id HomO.Cα(c,-)"
  (is ?H_c = ?id_H_c)
proof(rule ntcf_eqI[of α])
  from assms have "CIdc : c c" by (auto simp: cat_cs_intros)
  from assms category_axioms show 
    "?H_c : HomO.Cα(c,-) CF HomO.Cα(c,-) :  ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms show 
    "?id_H_c : HomO.Cα(c,-) CF HomO.Cα(c,-) :  ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms have lhs_dom: "𝒟 (?H_cNTMap) = Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms category_axioms have rhs_dom: "𝒟 (?id_H_cNTMap) = Obj"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  show "?H_cNTMap = ?id_H_cNTMap"
  proof(rule vsv_eqI, unfold lhs_dom rhs_dom)
    from assms category_axioms show "vsv (?id_H_cNTMap)"
      by (intro is_functor.ntcf_id_NTMap_vsv) 
        (cs_concl cs_shallow cs_simp: cs_intro: cat_cs_intros)
    fix d assume "d  Obj" 
    with assms category_axioms show "?H_cNTMapd = ?id_H_cNTMapd"
      by 
        (
          cs_concl cs_shallow 
            cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
        )
  qed (use assms in auto intro: cat_cs_intros)
qed auto

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_snd_CId

lemma (in category) cat_ntcf_Hom_fst_CId:
  assumes "c  Obj"
  shows "HomA.Cα(-,CIdc) = ntcf_id HomO.Cα(-,c)"
proof-
  note category.cat_ntcf_Hom_snd_CId[
      OF category_op, unfolded cat_op_simps, OF assms
      ]
  from this category_axioms assms show ?thesis
    by (cs_prems cs_shallow cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp
qed

lemmas [cat_cs_simps] = category.cat_ntcf_Hom_fst_CId


subsubsectionHom›-natural transformation and the Yoneda map›

lemma (in category) cat_Yoneda_map_of_ntcf_Hom_snd:
  assumes "f : s r"
  shows "Yoneda_map α (HomO.Cα(s,-)) rHomA.Cα(f,-) = f"
  using category_axioms assms (*slow*)
  by
    (
      cs_concl 
        cs_simp: cat_cs_simps cat_op_simps cat_Set_components(1)
        cs_intro: cat_cs_intros cat_prod_cs_intros
    ) 

lemmas [cat_cs_simps] = category.cat_Yoneda_map_of_ntcf_Hom_snd

lemma (in category) cat_Yoneda_map_of_ntcf_Hom_fst:
  assumes "f : r s"
  shows "Yoneda_map α (HomO.Cα(-,s)) rHomA.Cα(-,f) = f"
proof-
  note category.cat_Yoneda_map_of_ntcf_Hom_snd[
      OF category_op, unfolded cat_op_simps, OF assms
      ]
  from this category_axioms assms show ?thesis
    by (cs_prems cs_shallow cs_simp: cat_op_simps cs_intro: cat_cs_intros) simp
qed

lemmas [cat_cs_simps] = category.cat_Yoneda_map_of_ntcf_Hom_fst



subsection‹Evaluation arrow›


subsubsection‹Definition and elementary properties›


text‹
The evaluation arrow is a part of the definition of the evaluation functor.
The evaluation functor appears in Chapter III-2 in
cite"mac_lane_categories_2010".
›

definition cf_eval_arrow :: "V  V  V  V"
  where "cf_eval_arrow  𝔑 f =
    [
      (
        λx𝔑NTDomObjMapDomf.
          𝔑NTCodArrMapfArrVal𝔑NTMapDomfArrValx
      ),
      𝔑NTDomObjMapDomf,
      𝔑NTCodObjMapCodf
    ]"


text‹Components.›

lemma cf_eval_arrow_components:
  shows "cf_eval_arrow  𝔑 fArrVal =
    (
      λx𝔑NTDomObjMapDomf.
        𝔑NTCodArrMapfArrVal𝔑NTMapDomfArrValx
    )"
    and "cf_eval_arrow  𝔑 fArrDom = 𝔑NTDomObjMapDomf"
    and "cf_eval_arrow  𝔑 fArrCod = 𝔑NTCodObjMapCodf"
  unfolding cf_eval_arrow_def arr_field_simps by (simp_all add: nat_omega_simps)

context
  fixes α 𝔑  𝔉 𝔊 a b f  
  assumes 𝔑: "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α"
    and f: "f : a b"
begin

interpretation 𝔑: is_ntcf α  cat_Set α 𝔉 𝔊 𝔑 by (rule 𝔑)

lemmas cf_eval_arrow_components' = cf_eval_arrow_components[
    where= and 𝔑=ntcf_arrow 𝔑 and f=f, 
    unfolded 
      ntcf_arrow_components 
      cf_map_components 
      𝔑.NTDom.HomDom.cat_is_arrD[OF f]
      cat_cs_simps
    ]

lemmas [cat_cs_simps] = cf_eval_arrow_components'(2,3)

end


subsubsection‹Arrow value›

context
  fixes α 𝔑  𝔉 𝔊 a b f  
  assumes 𝔑: "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α"
    and f: "f : a b"
begin

mk_VLambda cf_eval_arrow_components'(1)[OF 𝔑 f]
  |vsv cf_eval_arrow_ArrVal_vsv[cat_cs_intros]|
  |vdomain cf_eval_arrow_ArrVal_vdomain[cat_cs_simps]|
  |app cf_eval_arrow_ArrVal_app[cat_cs_simps]|

end


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

lemma cf_eval_arrow_is_arr:
  assumes "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α" and "f : a b"
  shows "cf_eval_arrow  (ntcf_arrow 𝔑) f :
    𝔉ObjMapa cat_Set α𝔊ObjMapb"
proof-
  interpret 𝔑: is_ntcf α  cat_Set α 𝔉 𝔊 𝔑 by (rule assms)
  show ?thesis
  proof
    (
      intro cat_Set_is_arrI arr_SetI, 
      unfold cf_eval_arrow_components'(2,3)[OF assms]
    )
    show "vfsequence (cf_eval_arrow  (ntcf_arrow 𝔑) f)"
      unfolding cf_eval_arrow_def by simp
    show "vcard (cf_eval_arrow  (ntcf_arrow 𝔑) f) = 3"
      unfolding cf_eval_arrow_def by (simp add: nat_omega_simps)
    show " (cf_eval_arrow  (ntcf_arrow 𝔑) fArrVal)  𝔊ObjMapb"
      by 
      (
        unfold cf_eval_arrow_components'[OF assms], 
        intro vrange_VLambda_vsubset
      ) 
      (
        use assms in 
          cs_concl cs_intro: cat_cs_intros cat_Set_cs_intros
      )+
  qed
    (
      use assms(2) in
        cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros
    )+
qed

lemma cf_eval_arrow_is_arr'[cat_cs_intros]:
  assumes "𝔑' = ntcf_arrow 𝔑"
    and "𝔉a = 𝔉ObjMapa"
    and "𝔊b = 𝔊ObjMapb"
    and "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α" 
    and "f : a b"
  shows "cf_eval_arrow  𝔑' f : 𝔉a cat_Set α𝔊b"
  using assms(4,5) unfolding assms(1-3) by (rule cf_eval_arrow_is_arr)

lemma (in category) cat_cf_eval_arrow_ntcf_vcomp[cat_cs_simps]:
  assumes "𝔐 : 𝔊 CF  :  ↦↦Cαcat_Set α" 
    and "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α"
    and "g : b c"
    and "f : a b"
  shows 
    "cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af) =
      cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f"
proof-

  interpret 𝔐: is_ntcf α  cat_Set α 𝔊  𝔐 by (rule assms(1))
  interpret 𝔑: is_ntcf α  cat_Set α 𝔉 𝔊 𝔑 by (rule assms(2))

  have 𝔐𝔑: "𝔐 NTCF 𝔑 : 𝔉 CF  :  ↦↦Cαcat_Set α"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms(3,4) have gf: "g Af : a c"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from 𝔐𝔑 gf have cf_eval_gf:
    "cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af) :
      𝔉ObjMapa cat_Set αObjMapc"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  from assms(3,4) have cf_eval_g_cf_eval_f:
    "cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f :
      𝔉ObjMapa cat_Set αObjMapc"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  note cf_eval_gf = cf_eval_gf cat_Set_is_arrD[OF cf_eval_gf]
  note cf_eval_g_cf_eval_f = 
    cf_eval_g_cf_eval_f cat_Set_is_arrD[OF cf_eval_g_cf_eval_f]

  interpret arr_Set_cf_eval_gf:
    arr_Set α cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af)
    by (rule cf_eval_gf(2))
  interpret arr_Set_cf_eval_g_cf_eval_f:
    arr_Set 
      α 
      cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f
    by (rule cf_eval_g_cf_eval_f(2))

  show ?thesis
  proof(rule arr_Set_eqI)
    from 𝔐𝔑 gf have dom_lhs:
      "𝒟 (cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af)ArrVal) = 
        𝔉ObjMapa"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    from cf_eval_g_cf_eval_f(1) have dom_rhs: 
      "𝒟
        (
          (
            cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f
          )ArrVal
        ) = 𝔉ObjMapa"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    show
      "cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af)ArrVal =
        (
          cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f
        )ArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix 𝔉a assume prems: "𝔉a  𝔉ObjMapa"
      from 
        ArrVal_eq_helper
          [
            OF 𝔐.ntcf_Comp_commute[OF assms(4), symmetric], 
            where a=𝔑NTMapaArrVal𝔉a
          ] 
        prems 
        assms(3,4) 
      have [cat_cs_simps]:
        "ArrMapfArrVal𝔐NTMapaArrVal𝔑NTMapaArrVal𝔉a =
          𝔐NTMapbArrVal𝔊ArrMapfArrVal𝔑NTMapaArrVal𝔉a"
        by
          (
            cs_prems 
              cs_simp: cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros
          )
      from prems assms(3,4) show 
        "cf_eval_arrow  (ntcf_arrow (𝔐 NTCF 𝔑)) (g Af)ArrVal𝔉a =
          (
            cf_eval_arrow  (ntcf_arrow 𝔐) g Acat_Set αcf_eval_arrow  (ntcf_arrow 𝔑) f
          )ArrVal𝔉a"
        by
          (
            cs_concl 
              cs_simp: cat_cs_simps cs_intro: cat_Set_cs_intros cat_cs_intros
          )
    qed (cs_concl cs_shallow cs_intro: V_cs_intros)
  qed
    (
      auto
        simp: cf_eval_gf cf_eval_g_cf_eval_f 
        intro: cf_eval_gf(2) cf_eval_g_cf_eval_f(2)
    )

qed

lemmas [cat_cs_simps] = category.cat_cf_eval_arrow_ntcf_vcomp

lemma (in category) cat_cf_eval_arrow_ntcf_id[cat_cs_simps]:
  assumes "𝔉 :  ↦↦Cαcat_Set α" and "c  Obj"
  shows 
    "cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc) =
      cat_Set αCId𝔉ObjMapc"
proof-

  interpret 𝔉: is_functor α  cat_Set α 𝔉 by (rule assms)

  from assms(2) have ntcf_id_CId_c: 
    "cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc) :
      𝔉ObjMapc cat_Set α𝔉ObjMapc"
    by (cs_concl cs_intro: cat_cs_intros)
  from assms(2) have CId_𝔉c:
    "cat_Set αCId𝔉ObjMapc : 𝔉ObjMapc cat_Set α𝔉ObjMapc"
    by (cs_concl cs_intro: cat_cs_intros)

  show ?thesis
  proof(rule arr_Set_eqI[of α])

    from ntcf_id_CId_c show arr_Set_ntcf_id_CId_c:
      "arr_Set α (cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc))"
      by (auto dest: cat_Set_is_arrD(1))
    from ntcf_id_CId_c have dom_lhs:
      "𝒟 (cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc)ArrVal) =
        𝔉ObjMapc"
      by (cs_concl cs_simp: cat_cs_simps)+
    interpret ntcf_id_CId_c: 
      arr_Set α cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc)
      by (rule arr_Set_ntcf_id_CId_c)
  
    from CId_𝔉c show arr_Set_CId_𝔉c: "arr_Set α (cat_Set αCId𝔉ObjMapc)"
      by (auto dest: cat_Set_is_arrD(1))
    from CId_𝔉c assms(2) have dom_rhs: 
      "𝒟 ((cat_Set αCId𝔉ObjMapc)ArrVal) = 𝔉ObjMapc" 
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)

    show 
      "cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc)ArrVal =
        cat_Set αCId𝔉ObjMapcArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume "a  𝔉ObjMapc"
      with category_axioms assms(2) show 
        "cf_eval_arrow  (ntcf_arrow (ntcf_id 𝔉)) (CIdc)ArrVala =
          cat_Set αCId𝔉ObjMapcArrVala"
        by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    qed (use arr_Set_ntcf_id_CId_c arr_Set_CId_𝔉c in auto)

  qed (use ntcf_id_CId_c CId_𝔉c in cs_concl cs_simp: cat_cs_simps)+

qed

lemmas [cat_cs_simps] = category.cat_cf_eval_arrow_ntcf_id



subsectionHOM›-functor›


subsubsection‹Definition and elementary properties›


text‹
The following definition is a technical generalization that is used
later in this section.
›

definition cf_HOM_snd :: "V  V  V" (HOMCı'(/,_-/'))
  where "HOMCα(,𝔉-) =
    [
      (λaop_cat (𝔉HomCod)Obj. cf_map (HomO.Cα(𝔉HomCod)(a,-) CF 𝔉)),
      (
        λfop_cat (𝔉HomCod)Arr.
          ntcf_arrow (HomA.Cα(𝔉HomCod)(f,-) NTCF-CF 𝔉)
      ),
      op_cat (𝔉HomCod),
      cat_FUNCT α (𝔉HomDom) (cat_Set α)
    ]"

definition cf_HOM_fst :: "V  V  V" (HOMCı'(/_-,/'))
  where "HOMCα(𝔉-,) =
    [
      (λa(𝔉HomCod)Obj. cf_map (HomO.Cα(𝔉HomCod)(-,a) CF op_cf 𝔉)),
      (
        λf(𝔉HomCod)Arr.
          ntcf_arrow (HomA.Cα(𝔉HomCod)(-,f) NTCF-CF op_cf 𝔉)
      ),
      𝔉HomCod,
      cat_FUNCT α (op_cat (𝔉HomDom)) (cat_Set α)
    ]"


text‹Components.›

lemma cf_HOM_snd_components:
  shows "HOMCα(,𝔉-)ObjMap =
      (λaop_cat (𝔉HomCod)Obj. cf_map (HomO.Cα(𝔉HomCod)(a,-) CF 𝔉))"
    and "HOMCα(,𝔉-)ArrMap =
      (
        λfop_cat (𝔉HomCod)Arr.
          ntcf_arrow (HomA.Cα(𝔉HomCod)(f,-) NTCF-CF 𝔉)
      )"
    and [cat_cs_simps]: "HOMCα(,𝔉-)HomDom = op_cat (𝔉HomCod)"
    and [cat_cs_simps]: 
      "HOMCα(,𝔉-)HomCod = cat_FUNCT α (𝔉HomDom) (cat_Set α)"
  unfolding cf_HOM_snd_def dghm_field_simps by (simp_all add: nat_omega_simps)

lemma cf_HOM_fst_components:
  shows "HOMCα(𝔉-,)ObjMap = 
      (λa(𝔉HomCod)Obj. cf_map (HomO.Cα(𝔉HomCod)(-,a) CF op_cf 𝔉))"
    and "HOMCα(𝔉-,)ArrMap = 
      (
        λf(𝔉HomCod)Arr.
          ntcf_arrow (HomA.Cα(𝔉HomCod)(-,f) NTCF-CF op_cf 𝔉)
      )"
    and "HOMCα(𝔉-,)HomDom = 𝔉HomCod"
    and "HOMCα(𝔉-,)HomCod = cat_FUNCT α (op_cat (𝔉HomDom)) (cat_Set α)"
  unfolding cf_HOM_fst_def dghm_field_simps by (simp_all add: nat_omega_simps)

context is_functor
begin

lemmas cf_HOM_snd_components' = 
  cf_HOM_snd_components[where 𝔉=𝔉, unfolded cf_HomDom cf_HomCod]

lemmas [cat_cs_simps] = cf_HOM_snd_components'(3,4)

lemmas cf_HOM_fst_components' = 
  cf_HOM_fst_components[where 𝔉=𝔉, unfolded cf_HomDom cf_HomCod]

lemmas [cat_cs_simps] = cf_HOM_snd_components'(3,4)

end


subsubsection‹Object map›

mk_VLambda cf_HOM_snd_components(1)
  |vsv cf_HOM_snd_ObjMap_vsv[cat_cs_intros]|

mk_VLambda (in is_functor) cf_HOM_snd_components'(1)[unfolded cat_op_simps]
  |vdomain cf_HOM_snd_ObjMap_vdomain[cat_cs_simps]|
  |app cf_HOM_snd_ObjMap_app[cat_cs_simps]|

mk_VLambda cf_HOM_snd_components(1)
  |vsv cf_HOM_fst_ObjMap_vsv[cat_cs_intros]|

mk_VLambda (in is_functor) cf_HOM_fst_components'(1)[unfolded cat_op_simps]
  |vdomain cf_HOM_fst_ObjMap_vdomain[cat_cs_simps]|
  |app cf_HOM_fst_ObjMap_app[cat_cs_simps]|


subsubsection‹Arrow map›

mk_VLambda cf_HOM_snd_components(2)
  |vsv cf_HOM_snd_ArrMap_vsv[cat_cs_intros]|

mk_VLambda (in is_functor) cf_HOM_snd_components'(2)[unfolded cat_op_simps]
  |vdomain cf_HOM_snd_ArrMap_vdomain[cat_cs_simps]|
  |app cf_HOM_snd_ArrMap_app[cat_cs_simps]|

mk_VLambda cf_HOM_fst_components(2)
  |vsv cf_HOM_fst_ArrMap_vsv[cat_cs_intros]|

mk_VLambda (in is_functor) cf_HOM_fst_components'(2)[unfolded cat_op_simps]
  |vdomain cf_HOM_fst_ArrMap_vdomain[cat_cs_simps]|
  |app cf_HOM_fst_ArrMap_app[cat_cs_simps]|


subsubsection‹Opposite HOM›-functor›

lemma (in is_functor) cf_HOM_snd_op[cat_op_simps]: 
  "HOMCα(,op_cf 𝔉-) = HOMCα(𝔉-,)"
proof-
  have dom_lhs: "𝒟 HOMCα(,op_cf 𝔉-) = 4"
    unfolding cf_HOM_snd_def by (simp add: nat_omega_simps)
  have dom_rhs: "𝒟 HOMCα(𝔉-,) = 4"
    unfolding cf_HOM_fst_def by (simp add: nat_omega_simps)
  show ?thesis
  proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
    fix a assume "a  4"
    then show "HOMCα(,op_cf 𝔉-)a = HOMCα(𝔉-,)a"
    proof
      (
        elim_in_numeral, 
        use nothing in fold dghm_field_simps, unfold cat_cs_simps
      )
      show "HOMCα(,op_cf 𝔉-)ObjMap = HOMCα(𝔉-,)ObjMap"
        unfolding 
          cf_HOM_fst_components' 
          is_functor.cf_HOM_snd_components'[OF is_functor_op]
        by (rule VLambda_eqI, unfold cat_op_simps)
         (
           cs_concl cs_shallow 
            cs_simp: cat_cs_simps cat_op_simps cs_intro: cat_cs_intros
         )+
      show "HOMCα(,op_cf 𝔉-)ArrMap = HOMCα(𝔉-,)ArrMap"
        unfolding 
          cf_HOM_fst_components' 
          is_functor.cf_HOM_snd_components'[OF is_functor_op]
        by (rule VLambda_eqI, unfold cat_op_simps)
          (cs_concl cs_shallow cs_simp: cat_op_simps cs_intro: cat_cs_intros)+ 
    qed 
      (
        auto simp:
          cf_HOM_fst_components' cat_cs_simps cat_op_simps cat_op_intros
      )
  qed (auto simp: cf_HOM_snd_def cf_HOM_fst_def)
qed

lemmas [cat_op_simps] = is_functor.cf_HOM_snd_op

context is_functor
begin

lemmas cf_HOM_fst_op[cat_op_simps] = 
  is_functor.cf_HOM_snd_op[OF is_functor_op, unfolded cat_op_simps, symmetric]

end

lemmas [cat_op_simps] = is_functor.cf_HOM_fst_op


subsubsectionHOM›-functor is a functor›

lemma (in is_functor) cf_HOM_snd_is_functor: 
  assumes "𝒵 β" and "α  β"
  shows "HOMCα(,𝔉-) : op_cat 𝔅 ↦↦Cβcat_FUNCT α 𝔄 (cat_Set α)"
proof-

  interpret β: 𝒵 β by (rule assms(1))
  interpret βℭ: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+

  show ?thesis
  proof(intro is_functorI', unfold cat_op_simps)
    show "vfsequence HOMCα(,𝔉-)" unfolding cf_HOM_snd_def by auto
    show "vcard HOMCα(,𝔉-) = 4"
      unfolding cf_HOM_snd_def by (simp add: nat_omega_simps)
    show " (HOMCα(,𝔉-)ObjMap)  cat_FUNCT α 𝔄 (cat_Set α)Obj"
      unfolding cf_HOM_snd_components'
    proof(rule vrange_VLambda_vsubset, unfold cat_op_simps)
      fix b assume prems: "b  𝔅Obj"
      with assms(2) show 
        "cf_map (HomO.Cα𝔅(b,-) CF 𝔉)  cat_FUNCT α 𝔄 (cat_Set α)Obj"
        by
          (
            cs_concl 
              cs_simp: cat_FUNCT_cs_simps 
              cs_intro: cat_cs_intros cat_FUNCT_cs_intros
          )
    qed
    show 
      "HOMCα(,𝔉-)ArrMapf A𝔅g = 
        HOMCα(,𝔉-)ArrMapg Acat_FUNCT α 𝔄 (cat_Set α)HOMCα(,𝔉-)ArrMapf"
      if "g : c 𝔅b" and "f : b 𝔅a" for b c g a f
      using that 
      by 
        (
          cs_concl 
            cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps 
            cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
    show 
      "HOMCα(,𝔉-)ArrMap𝔅CIdc =
        cat_FUNCT α 𝔄 (cat_Set α)CIdHOMCα(,𝔉-)ObjMapc"
      if "c  𝔅Obj" for c 
      using that
      by
        (
          cs_concl 
            cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps
            cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
  qed 
    (
      use assms(2) in
        cs_concl 
            cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps
            cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros
    )+

qed

lemma (in is_functor) cf_HOM_snd_is_functor'[cat_cs_intros]: 
  assumes "𝒵 β" 
    and "α  β"
    and "ℭ' = op_cat 𝔅"
    and "𝔇 = cat_FUNCT α 𝔄 (cat_Set α)"
  shows "HOMCα(,𝔉-) : ℭ' ↦↦Cβ𝔇"
  using assms(1,2) unfolding assms(3,4) by (rule cf_HOM_snd_is_functor)

lemmas [cat_cs_intros] = is_functor.cf_HOM_snd_is_functor'

lemma (in is_functor) cf_HOM_fst_is_functor: 
  assumes "𝒵 β" and "α  β"
  shows "HOMCα(𝔉-,) : 𝔅 ↦↦Cβcat_FUNCT α (op_cat 𝔄) (cat_Set α)"
  by 
    (
      rule is_functor.cf_HOM_snd_is_functor[
        OF is_functor_op assms, unfolded cat_op_simps
        ]
   )

lemma (in is_functor) cf_HOM_fst_is_functor'[cat_cs_intros]: 
  assumes "𝒵 β" 
    and "α  β"
    and "ℭ' = 𝔅"
    and "𝔇 = cat_FUNCT α (op_cat 𝔄) (cat_Set α)"
  shows "HOMCα(𝔉-,) : ℭ' ↦↦Cβ𝔇"
  using assms(1,2) unfolding assms(3,4) by (rule cf_HOM_fst_is_functor)

lemmas [cat_cs_intros] = is_functor.cf_HOM_fst_is_functor'



subsection‹Evaluation functor›


subsubsection‹Definition and elementary properties›


text‹See Chapter III-2 in cite"mac_lane_categories_2010".›

definition cf_eval :: "V  V  V  V"
  where "cf_eval α β  =
    [
      (λ𝔉d(cat_FUNCT α  (cat_Set α) ×C )Obj. 𝔉d0ObjMap𝔉d1),
      (
        λ𝔑f(cat_FUNCT α  (cat_Set α) ×C )Arr.
          cf_eval_arrow  (𝔑f0) (𝔑f1)
      ),
      cat_FUNCT α  (cat_Set α) ×C ,
      cat_Set β
    ]"


text‹Components.›

lemma cf_eval_components:
  shows "cf_eval α β ObjMap =
    (λ𝔉d(cat_FUNCT α  (cat_Set α) ×C )Obj. 𝔉d0ObjMap𝔉d1)"
    and "cf_eval α β ArrMap =
      (
        λ𝔑f(cat_FUNCT α  (cat_Set α) ×C )Arr.
          cf_eval_arrow  (𝔑f0) (𝔑f1)
      )"
    and [cat_cs_simps]: 
      "cf_eval α β HomDom = cat_FUNCT α  (cat_Set α) ×C "
    and [cat_cs_simps]: "cf_eval α β HomCod = cat_Set β"
  unfolding cf_eval_def dghm_field_simps by (simp_all add: nat_omega_simps)


subsubsection‹Object map›

lemma cf_eval_ObjMap_vsv[cat_cs_intros]: "vsv (cf_eval α β ObjMap)"
  unfolding cf_eval_components by simp

lemma cf_eval_ObjMap_vdomain[cat_cs_simps]: 
  "𝒟 (cf_eval α β ObjMap) = (cat_FUNCT α  (cat_Set α) ×C )Obj"
  unfolding cf_eval_components by simp

lemma (in category) cf_eval_ObjMap_app[cat_cs_simps]: 
  assumes "𝔉c = [cf_map 𝔉, c]"
    and "𝔉 :  ↦↦Cαcat_Set α" (*the order of premises is important*)
    and "c  Obj"
  shows "cf_eval α β ObjMap𝔉c = 𝔉ObjMapc"
proof-
  interpret 𝔉: is_functor α  cat_Set α 𝔉 by (rule assms(2))
  define β where "β = α + ω"
  have "𝒵 β" and αβ: "α  β" 
    by (simp_all add: β_def 𝒵_Limit_αω 𝒵_ω_αω 𝒵_def 𝒵_α_αω)
  then interpret β: 𝒵 β by simp 
  note [cat_small_cs_intros] = cat_category_if_ge_Limit
  from assms(2,3) αβ have "𝔉c  (cat_FUNCT α  (cat_Set α) ×C )Obj"
    by 
      (
        cs_concl cs_shallow
          cs_simp: assms(1) cat_FUNCT_components(1)
          cs_intro: 
            cat_cs_intros 
            cat_small_cs_intros 
            cat_prod_cs_intros 
            cat_FUNCT_cs_intros
      )
  then show ?thesis
    by (simp add: assms(1) cf_map_components cf_eval_components nat_omega_simps)
qed

lemmas [cat_cs_simps] = category.cf_eval_ObjMap_app


subsubsection‹Arrow map›

lemma cf_eval_ArrMap_vsv[cat_cs_intros]: "vsv (cf_eval α β ArrMap)"
  unfolding cf_eval_components by simp

lemma cf_eval_ArrMap_vdomain[cat_cs_simps]: 
  "𝒟 (cf_eval α β ArrMap) = (cat_FUNCT α  (cat_Set α) ×C )Arr"
  unfolding cf_eval_components by simp

lemma (in category) cf_eval_ArrMap_app[cat_cs_simps]: 
  assumes "𝔑f = [ntcf_arrow 𝔑, f]"
    and "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α"
    and "f : a b"
  shows "cf_eval α β ArrMap𝔑f = cf_eval_arrow  (ntcf_arrow 𝔑) f"
proof-
  interpret 𝔉: is_ntcf α  cat_Set α 𝔉 𝔊 𝔑 by (rule assms(2))
  define β where "β = α + ω"
  have "𝒵 β" and αβ: "α  β" 
    by (simp_all add: β_def 𝒵_Limit_αω 𝒵_ω_αω 𝒵_def 𝒵_α_αω)
  then interpret β: 𝒵 β by simp 
  note [cat_small_cs_intros] = cat_category_if_ge_Limit
  from assms(1,3) αβ have "𝔑f  (cat_FUNCT α  (cat_Set α) ×C )Arr"
    by 
      (
        cs_concl 
          cs_simp: assms(1) cat_FUNCT_components(1)
          cs_intro: 
            cat_cs_intros 
            cat_small_cs_intros 
            cat_prod_cs_intros 
            cat_FUNCT_cs_intros
      )
  then show ?thesis
    by (simp add: assms(1) cf_map_components cf_eval_components nat_omega_simps)
qed

lemmas [cat_cs_simps] = category.cf_eval_ArrMap_app


subsubsection‹Evaluation functor is a functor›

lemma (in category) cat_cf_eval_is_functor:
  assumes "𝒵 β" and "α  β"
  shows "cf_eval α β  : cat_FUNCT α  (cat_Set α) ×C  ↦↦Cβcat_Set β"
proof-

  interpret β: 𝒵 β by (rule assms(1))
  from assms(2) cat_category_if_ge_Limit[OF assms] interpret FUNCT: 
    category β (cat_FUNCT α  (cat_Set α))
    by 
      (
        cs_concl cs_intro: 
          cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros
      )
  interpret βℭ: category β 
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret cat_Set_αβ: subcategory β cat_Set α cat_Set β
    by (rule subcategory_cat_Set_cat_Set[OF assms])
  
  show ?thesis
  proof(intro is_functorI')
    show "vfsequence (cf_eval α β )" unfolding cf_eval_def by simp
    from cat_category_if_ge_Limit[OF assms] show 
      "category β ((cat_FUNCT α  (cat_Set α)) ×C )"
      by (cs_concl cs_shallow cs_intro: cat_small_cs_intros cat_cs_intros)
    show "vcard (cf_eval α β ) = 4" 
      unfolding cf_eval_def by (simp add: nat_omega_simps)
    show " (cf_eval α β ObjMap)  cat_Set βObj"
    proof(intro vsv.vsv_vrange_vsubset, unfold cat_cs_simps)
      fix 𝔉c assume prems: "𝔉c  (cat_FUNCT α  (cat_Set α) ×C )Obj"
      then obtain 𝔉 c 
        where 𝔉c_def: "𝔉c = [𝔉, c]"
          and 𝔉: "𝔉  cf_maps α  (cat_Set α)"
          and c: "c  Obj"
        by 
          (
            auto 
              elim: cat_prod_2_ObjE[rotated 2] 
              intro: FUNCT.category_axioms βℭ.category_axioms
              simp: cat_FUNCT_components(1)
          )
      from 𝔉 obtain 𝔊 where 𝔉_def: "𝔉 = cf_map 𝔊"
        and 𝔊: "𝔊 :  ↦↦Cαcat_Set α"
        by (elim cf_mapsE)        
      interpret 𝔊: is_functor α  cat_Set α 𝔊 by (rule 𝔊)
      from 𝔊 c show "cf_eval α β ObjMap𝔉c  cat_Set βObj"
        unfolding 𝔉c_def 𝔉_def 
        by 
          (
            cs_concl cs_shallow
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_Set_αβ.subcat_Obj_vsubset
          )
    qed (cs_concl cs_shallow cs_intro: cat_cs_intros)
    show "cf_eval α β ArrMap𝔑f :
      cf_eval α β ObjMap𝔉a cat_Set βcf_eval α β ObjMap𝔊b"
      if 𝔑f: "𝔑f : 𝔉a cat_FUNCT α  (cat_Set α) ×C 𝔊b" for 𝔉a 𝔊b 𝔑f
    proof-
      obtain 𝔑 f 𝔉 a 𝔊 b
        where 𝔑f_def: "𝔑f = [𝔑, f]" 
          and 𝔉a_def: "𝔉a = [𝔉, a]"
          and 𝔊b_def: "𝔊b = [𝔊, b]" 
          and 𝔑: "𝔑 : 𝔉 cat_FUNCT α  (cat_Set α)𝔊" 
          and f: "f : a b"
        by 
          (
            auto intro: 
              cat_prod_2_is_arrE[rotated 2, OF 𝔑f] 
              FUNCT.category_axioms 
              βℭ.category_axioms
          )
      note 𝔑 = cat_FUNCT_is_arrD[OF 𝔑]
      from 𝔑(1) f assms(2) show "cf_eval α β ArrMap𝔑f :
        cf_eval α β ObjMap𝔉a cat_Set βcf_eval α β ObjMap𝔊b"
        unfolding 𝔑f_def 𝔉a_def 𝔊b_def
        by
          (
            intro cat_Set_αβ.subcat_is_arrD,
            use nothing in subst 𝔑(2), subst 𝔑(3), subst 𝔑(4)
          )
          (
            cs_concl 
              cs_simp: cat_FUNCT_cs_simps cat_cs_simps cs_intro: cat_cs_intros 
          ) (*slow*)
    qed
    show 
      "cf_eval α β ArrMap𝔐g Acat_FUNCT α  (cat_Set α) ×C 𝔑f =
        cf_eval α β ArrMap𝔐g Acat_Set βcf_eval α β ArrMap𝔑f"
      if 𝔐g: "𝔐g : 𝔊b cat_FUNCT α  (cat_Set α) ×C ℌc"
        and 𝔑f: "𝔑f : 𝔉a cat_FUNCT α  (cat_Set α) ×C 𝔊b"
      for 𝔑f 𝔐g 𝔉a 𝔊b ℌc
    proof-
      obtain 𝔑 f 𝔉 a 𝔊 b
        where 𝔑f_def: "𝔑f = [𝔑, f]" 
          and 𝔉a_def: "𝔉a = [𝔉, a]"
          and 𝔊b_def: "𝔊b = [𝔊, b]" 
          and 𝔑: "𝔑 : 𝔉 cat_FUNCT α  (cat_Set α)𝔊" 
          and f: "f : a b"
        by 
          (
            auto intro: 
              cat_prod_2_is_arrE[rotated 2, OF 𝔑f] 
              FUNCT.category_axioms 
              βℭ.category_axioms
          )
      then obtain 𝔐 g  c
        where 𝔐g_def: "𝔐g = [𝔐, g]" 
          and ℌc_def: "ℌc = [, c]" 
          and 𝔐: "𝔐 : 𝔊 cat_FUNCT α  (cat_Set α)"
          and g: "g : b c"
        by 
          (
            auto intro: 
              cat_prod_2_is_arrE[rotated 2, OF 𝔐g] 
              FUNCT.category_axioms
              βℭ.category_axioms
          )
      note 𝔑 = cat_FUNCT_is_arrD[OF 𝔑]
        and 𝔐 = cat_FUNCT_is_arrD[OF 𝔐]
      from 𝔑(1) 𝔐(1) f g show
        "cf_eval α β ArrMap𝔐g Acat_FUNCT α  (cat_Set α) ×C 𝔑f =
          cf_eval α β ArrMap𝔐g Acat_Set βcf_eval α β ArrMap𝔑f"
        unfolding 𝔐g_def 𝔑f_def 𝔉a_def 𝔊b_def ℌc_def
        by 
          (
            subst (1 2) 𝔐(2), use nothing in subst (1 2) 𝔑(2), 
            cs_concl_step cs_shallow cat_Set_αβ.subcat_Comp_simp[symmetric]
          )
          (
            cs_concl 
              cs_simp: cat_cs_simps cat_prod_cs_simps cat_FUNCT_cs_simps 
              cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros
          )
    qed
    show
      "cf_eval α β ArrMap(cat_FUNCT α  (cat_Set α) ×C )CId𝔉c =
        cat_Set βCIdcf_eval α β ObjMap𝔉c"
      if "𝔉c  (cat_FUNCT α  (cat_Set α) ×C )Obj" for 𝔉c
    proof-
      from that obtain 𝔉 c where 𝔉c_def: "𝔉c = [𝔉, c]"
        and 𝔉: "𝔉  cf_maps α  (cat_Set α)"
        and c: "c  Obj"
        by 
          (
            auto 
              elim: cat_prod_2_ObjE[rotated 2] 
              intro: FUNCT.category_axioms βℭ.category_axioms
              simp: cat_FUNCT_components(1)
          )
      from 𝔉 obtain 𝔊 where 𝔉_def: "𝔉 = cf_map 𝔊"
        and 𝔊: "𝔊 :  ↦↦Cαcat_Set α"
        by (elim cf_mapsE)
      interpret 𝔊: is_functor α  cat_Set α 𝔊 by (rule 𝔊)
      from 𝔊 c show 
        "cf_eval α β ArrMap(cat_FUNCT α  (cat_Set α) ×C )CId𝔉c =
          cat_Set βCIdcf_eval α β ObjMap𝔉c"
        unfolding 𝔉c_def 𝔉_def
        by (cs_concl_step cat_Set_αβ.subcat_CId[symmetric])
          (
            cs_concl
              cs_simp: cat_cs_simps cat_prod_cs_simps cat_FUNCT_cs_simps 
              cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros
          )
    qed

  qed (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+

qed

lemma (in category) cat_cf_eval_is_functor':
  assumes "𝒵 β" 
    and "α  β"
    and "𝔄' = cat_FUNCT α  (cat_Set α) ×C "
    and "𝔅' = cat_Set β"
    and "β' = β"
  shows "cf_eval α β  : 𝔄' ↦↦Cβ'𝔅'"
  using assms(1,2) unfolding assms(3-5) by (rule cat_cf_eval_is_functor) 

lemmas [cat_cs_intros] = category.cat_cf_eval_is_functor'



subsectionN›-functor›


subsubsection‹Definition and elementary properties›


text‹See Chapter III-2 in cite"mac_lane_categories_2010".›

definition cf_nt :: "V  V  V  V"
  where "cf_nt α β 𝔉 =
    bifunctor_flip (𝔉HomCod) (cat_FUNCT α (𝔉HomDom) (cat_Set α))
      (HomO.Cβcat_FUNCT α (𝔉HomDom) (cat_Set α)(HOMCα(,𝔉-)-,-))"


text‹Alternative definition.›

lemma (in is_functor) cf_nt_def':
  "cf_nt α β 𝔉 =
    bifunctor_flip 𝔅 (cat_FUNCT α 𝔄 (cat_Set α))
      (HomO.Cβcat_FUNCT α 𝔄 (cat_Set α)(HOMCα(,𝔉-)-,-))"
  unfolding cf_nt_def cf_HomDom cf_HomCod by simp


text‹Components.›

lemma cf_nt_components:
  shows "cf_nt α β 𝔉ObjMap =
    (
      bifunctor_flip (𝔉HomCod) (cat_FUNCT α (𝔉HomDom) (cat_Set α))
        (HomO.Cβcat_FUNCT α (𝔉HomDom) (cat_Set α)(HOMCα(,𝔉-)-,-))
    )ObjMap"
    and "cf_nt α β 𝔉ArrMap =
      (
        bifunctor_flip (𝔉HomCod) (cat_FUNCT α (𝔉HomDom) (cat_Set α))
          (HomO.Cβcat_FUNCT α (𝔉HomDom) (cat_Set α)(HOMCα(,𝔉-)-,-))
      )ArrMap"
    and "cf_nt α β 𝔉HomDom =
      (
        bifunctor_flip (𝔉HomCod) (cat_FUNCT α (𝔉HomDom) (cat_Set α))
          (HomO.Cβcat_FUNCT α (𝔉HomDom) (cat_Set α)(HOMCα(,𝔉-)-,-))
      )HomDom"
    and "cf_nt α β 𝔉HomCod =
      (
        bifunctor_flip (𝔉HomCod) (cat_FUNCT α (𝔉HomDom) (cat_Set α))
          (HomO.Cβcat_FUNCT α (𝔉HomDom) (cat_Set α)(HOMCα(,𝔉-)-,-))
      )HomCod"
  unfolding cf_nt_def by simp_all

lemma (in is_functor) cf_nt_components':
  assumes "𝒵 β" and "α  β"
  shows "cf_nt α β 𝔉ObjMap =
      (
        bifunctor_flip 𝔅 (cat_FUNCT α 𝔄 (cat_Set α))
          (HomO.Cβcat_FUNCT α 𝔄 (cat_Set α)(HOMCα(,𝔉-)-,-))
      )ObjMap"
    and "cf_nt α β 𝔉ArrMap =
      (
        bifunctor_flip 𝔅 (cat_FUNCT α 𝔄 (cat_Set α))
          (HomO.Cβcat_FUNCT α 𝔄 (cat_Set α)(HOMCα(,𝔉-)-,-))
      )ArrMap"
    and [cat_cs_simps]: 
      "cf_nt α β 𝔉HomDom = cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅"
    and [cat_cs_simps]: 
      "cf_nt α β 𝔉HomCod = cat_Set β"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  show 
    "cf_nt α β 𝔉ObjMap =
      (
        bifunctor_flip 𝔅 (cat_FUNCT α 𝔄 (cat_Set α))
          (HomO.Cβcat_FUNCT α 𝔄 (cat_Set α)(HOMCα(,𝔉-)-,-))
      )ObjMap"
    "cf_nt α β 𝔉ArrMap =
      (
        bifunctor_flip 𝔅 (cat_FUNCT α 𝔄 (cat_Set α))
          (HomO.Cβcat_FUNCT α 𝔄 (cat_Set α)(HOMCα(,𝔉-)-,-))
      )ArrMap"
    "cf_nt α β 𝔉HomDom = cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅"
    "cf_nt α β 𝔉HomCod = cat_Set β"
    unfolding cf_nt_def 
    using assms(2)
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps
          cs_intro: cat_small_cs_intros cat_cs_intros cat_FUNCT_cs_intros
      )+
qed

lemmas [cat_cs_simps] = is_functor.cf_nt_components'(3,4)


subsubsection‹Object map›

lemma cf_nt_ObjMap_vsv[cat_cs_intros]: "vsv (cf_nt α β ObjMap)"
  unfolding cf_nt_components by (cs_intro_step cat_cs_intros)

lemma (in is_functor) cf_nt_ObjMap_vdomain[cat_cs_simps]: 
  assumes "𝒵 β" and "α  β"
  shows "𝒟 (cf_nt α β 𝔉ObjMap) = (cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅)Obj"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  from assms(2) show ?thesis
    unfolding cf_nt_components
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps
          cs_intro: 
            cat_small_cs_intros
            cat_cs_intros
            cat_FUNCT_cs_intros
            cat_prod_cs_intros
      )
qed

lemmas [cat_cs_simps] = is_functor.cf_nt_ObjMap_vdomain

lemma (in is_functor) cf_nt_ObjMap_app[cat_cs_simps]: 
  assumes "𝒵 β" 
    and "α  β"
    and "𝔊b = [cf_map 𝔊, b]"
    and "𝔊 : 𝔄 ↦↦Cαcat_Set α"
    and "b  𝔅Obj"
  shows "cf_nt α β 𝔉ObjMap𝔊b = Hom
    (cat_FUNCT α 𝔄 (cat_Set α))
    (cf_map (HomO.Cα𝔅(b,-) CF 𝔉))
    (cf_map 𝔊)"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret 𝔊: is_functor α 𝔄 cat_Set α 𝔊 by (rule assms(4))
  from assms(2,5) show ?thesis
    unfolding assms(3) cf_nt_def
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps 
          cs_intro:
            cat_cs_intros
            cat_small_cs_intros
            cat_FUNCT_cs_intros
            cat_prod_cs_intros 
            cat_op_intros
      )
qed

lemmas [cat_cs_simps] = is_functor.cf_nt_ObjMap_app


subsubsection‹Arrow map›

lemma cf_nt_ArrMap_vsv[cat_cs_intros]: "vsv (cf_nt α β ArrMap)"
  unfolding cf_nt_components by (cs_intro_step cat_cs_intros)

lemma (in is_functor) cf_nt_ArrMap_vdomain[cat_cs_simps]: 
  assumes "𝒵 β" and "α  β"
  shows "𝒟 (cf_nt α β 𝔉ArrMap) = (cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅)Arr"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  from assms(2) show ?thesis
    unfolding cf_nt_components
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps
          cs_intro: 
            cat_small_cs_intros
            cat_cs_intros
            cat_FUNCT_cs_intros
            cat_prod_cs_intros
      )
qed

lemmas [cat_cs_simps] = is_functor.cf_nt_ArrMap_vdomain

lemma (in is_functor) cf_nt_ArrMap_app[cat_cs_simps]: 
  assumes "𝒵 β" 
    and "α  β"
    and "𝔑f = [ntcf_arrow 𝔑, f]"
    and "𝔑 : 𝔊 CF  : 𝔄 ↦↦Cαcat_Set α"
    and "f : a 𝔅b"
  shows "cf_nt α β 𝔉ArrMap𝔑f = cf_hom
    (cat_FUNCT α 𝔄 (cat_Set α))
    [ntcf_arrow (HomA.Cα𝔅(f,-) NTCF-CF 𝔉), ntcf_arrow 𝔑]"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret 𝔑: is_ntcf α 𝔄 cat_Set α 𝔊  𝔑 by (rule assms(4))
  from assms(2,5) show ?thesis
    unfolding assms(3) cf_nt_def
    by
      (
        cs_concl 
          cs_simp: cat_cs_simps cat_FUNCT_cs_simps cat_op_simps
          cs_intro:
            cat_cs_intros
            cat_small_cs_intros
            cat_FUNCT_cs_intros
            cat_prod_cs_intros
            cat_op_intros
      )
qed

lemmas [cat_cs_simps] = is_functor.cf_nt_ArrMap_app


subsubsectionN›-functor is a functor›

lemma (in is_functor) cf_nt_is_functor:
  assumes "𝒵 β" and "α  β"
  shows "cf_nt α β 𝔉 : cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅 ↦↦Cβcat_Set β"
proof-
  interpret β: 𝒵 β by (rule assms(1))
  interpret β𝔄: category β 𝔄
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret β𝔅: category β 𝔅
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  from assms(2) show ?thesis
    unfolding cf_nt_def'
    by 
      (
        cs_concl 
          cs_simp: cat_op_simps 
          cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros
      )
qed

lemma (in is_functor) cf_nt_is_functor':
  assumes "𝒵 β" 
    and "α  β"
    and "𝔄' = cat_FUNCT α 𝔄 (cat_Set α) ×C 𝔅"
    and "𝔅' = cat_Set β"
    and "β' = β"
  shows "cf_nt α β 𝔉 : 𝔄' ↦↦Cβ'𝔅'"
  using assms(1,2) unfolding assms(3-5) by (rule cf_nt_is_functor) 

lemmas [cat_cs_intros] = is_functor.cf_nt_is_functor'



subsection‹Yoneda natural transformation arrow›


subsubsection‹Definition and elementary properties›


text‹
The following subsection is based on the elements of the content
of Chapter III-2 in cite"mac_lane_categories_2010". 
›

definition ntcf_Yoneda_arrow :: "V  V  V  V  V"
  where "ntcf_Yoneda_arrow α  𝔉 r =
    [
      (
        λψHom (cat_FUNCT α  (cat_Set α)) (cf_map (HomO.Cα(r,-))) 𝔉.
          Yoneda_map α (cf_of_cf_map  (cat_Set α) 𝔉) r
            ntcf_of_ntcf_arrow  (cat_Set α) ψ
            
      ),
      Hom (cat_FUNCT α  (cat_Set α)) (cf_map (HomO.Cα(r,-))) 𝔉,
      𝔉ObjMapr
    ]"


text‹Components›

lemma ntcf_Yoneda_arrow_components:
  shows "ntcf_Yoneda_arrow α  𝔉 rArrVal = 
    (
      λψHom (cat_FUNCT α  (cat_Set α)) (cf_map (HomO.Cα(r,-))) 𝔉.
        Yoneda_map α (cf_of_cf_map  (cat_Set α) 𝔉) r
          ntcf_of_ntcf_arrow  (cat_Set α) ψ
          
    )"
    and [cat_cs_simps]: "ntcf_Yoneda_arrow α  𝔉 rArrDom = 
      Hom (cat_FUNCT α  (cat_Set α)) (cf_map (HomO.Cα(r,-))) 𝔉"
    and [cat_cs_simps]: "ntcf_Yoneda_arrow α  𝔉 rArrCod = 𝔉ObjMapr"
  unfolding ntcf_Yoneda_arrow_def arr_field_simps
  by (simp_all add: nat_omega_simps)


subsubsection‹Arrow map›

mk_VLambda ntcf_Yoneda_arrow_components(1)
  |vsv ntcf_Yoneda_arrow_vsv[cat_cs_intros]|
  |vdomain ntcf_Yoneda_arrow_vdomain[cat_cs_simps]|

context category
begin

context
  fixes 𝔉 :: V
begin

mk_VLambda ntcf_Yoneda_arrow_components(1)[where α=α and= and 𝔉=cf_map 𝔉]
  |app ntcf_Yoneda_arrow_app'|

lemmas ntcf_Yoneda_arrow_app =
  ntcf_Yoneda_arrow_app'[unfolded in_Hom_iff, cat_cs_simps]

end

end

lemmas [cat_cs_simps] = category.ntcf_Yoneda_arrow_app


subsubsection‹Several technical lemmas›

lemma (in vsv) vsv_vrange_VLambda_app:
  assumes "g ` elts A = elts (𝒟 r)"
  shows " (λxA. rg x) =  r"
proof(intro vsubset_antisym vsv.vsv_vrange_vsubset, unfold vdomain_VLambda)
  show "(λxA. rg x)x   r" if "x  A" for x
  proof-
    from assms that have "g x  𝒟 r" by auto
    then have "rg x   r" by force
    with that show ?thesis by simp
  qed
  show "rx   (λxA. rg x)" if "x  𝒟 r" for x
  proof-
    from that assms have "x  g ` elts A" by simp
    then obtain c where c: "c  A" and x_def: "x = g c" by clarsimp
    from c show ?thesis unfolding x_def by auto
  qed
qed auto

lemma (in vsv) vsv_vrange_VLambda_app':
  assumes "g ` elts A = elts (𝒟 r)"
    and "R =  r"
  shows " (λxA. rg x) = R"
  using assms(1) unfolding assms(2) by (rule vsv_vrange_VLambda_app)

lemma (in v11) v11_VLambda_v11_bij_betw_comp:
  assumes "bij_betw g (elts A) (elts (𝒟 r))"
  shows "v11 (λxA. rg x)"
proof(rule vsv.vsv_valeq_v11I, unfold vdomain_VLambda beta)
  fix x y assume prems: "x  A" "y  A" "rg x = rg y"
  from assms prems(1,2) have "g x  𝒟 r" and "g y  𝒟 r" by auto
  from v11_injective[OF this prems(3)] have "g x = g y".
  with assms prems(1,2) show "x = y" unfolding bij_betw_def inj_on_def by simp
qed simp


subsubsection‹
Yoneda natural transformation arrow is an arrow in the category Set›

lemma (in category) cat_ntcf_Yoneda_arrow_is_arr_isomoprhism:
  assumes "𝒵 β"
    and "α  β"
    and "𝔉 :  ↦↦Cαcat_Set α"
    and "r  Obj"
  shows "ntcf_Yoneda_arrow α  (cf_map 𝔉) r :
    Hom 
      (cat_FUNCT α  (cat_Set α))
      (cf_map (HomO.Cα(r,-)))
      (cf_map 𝔉) isocat_Set β𝔉ObjMapr"
proof-

  interpret β: 𝒵 β by (rule assms(1))
  interpret 𝔉: is_functor α  cat_Set α 𝔉 by (rule assms)

  from assms(2) interpret FUNCT: tiny_category β cat_FUNCT α  (cat_Set α)
    by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_FUNCT_cs_intros)

  let ?Hom_r = HomO.Cα(r,-)

  from assms have [cat_cs_simps]: "cf_of_cf_map  (cat_Set α) (cf_map 𝔉) = 𝔉"
    by (cs_concl cs_shallow cs_simp: cat_FUNCT_cs_simps)

  note Yoneda = cat_Yoneda_Lemma[OF assms(3,4)]

  show ?thesis
  proof
    (
      intro cat_Set_is_iso_arrI cat_Set_is_arrI arr_SetI,
      unfold cat_cs_simps cf_map_components
    )
    show "vfsequence (ntcf_Yoneda_arrow α  (cf_map 𝔉) r)"
      unfolding ntcf_Yoneda_arrow_def by simp
    show "vcard (ntcf_Yoneda_arrow α  (cf_map 𝔉) r) = 3"
      unfolding ntcf_Yoneda_arrow_def by (simp add: nat_omega_simps)
    show " (ntcf_Yoneda_arrow α  (cf_map 𝔉) rArrVal) = 𝔉ObjMapr"
      unfolding cat_cs_simps cf_map_components ntcf_Yoneda_arrow_components 
      by (intro vsv.vsv_vrange_VLambda_app', unfold Yoneda(2))
        (
          use assms(4) in
            cs_concl cs_shallow
                cs_simp:
                  cat_cs_simps bij_betwD(2)[OF bij_betw_ntcf_of_ntcf_arrow_Hom]
                cs_intro: cat_cs_intros
        )+
    then show " (ntcf_Yoneda_arrow α  (cf_map 𝔉) rArrVal)  𝔉ObjMapr"
      by auto
    from assms(4) show "v11 (ntcf_Yoneda_arrow α  (cf_map 𝔉) rArrVal)"
      unfolding ntcf_Yoneda_arrow_components
      by 
        (
          intro v11.v11_VLambda_v11_bij_betw_comp, 
          unfold cat_cs_simps 𝔉.Yoneda_map_vdomain; 
          intro Yoneda bij_betw_ntcf_of_ntcf_arrow_Hom
        )
        (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    from assms(4) show 
      "Hom (cat_FUNCT α  (cat_Set α)) (cf_map ?Hom_r) (cf_map 𝔉)  Vset β"
      by (intro FUNCT.cat_Hom_in_Vset)
        (
          cs_concl cs_shallow
            cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
    from assms(4) have "𝔉ObjMapr  Vset α" 
      by (cs_concl cs_intro: cat_cs_intros)
    then show "𝔉ObjMapr  Vset β"
      by (auto simp: assms(2) Vset_trans Vset_in_mono)
  qed (auto intro: cat_cs_intros)

qed

lemma (in category) cat_ntcf_Yoneda_arrow_is_arr_isomoprhism':
  assumes "𝒵 β"
    and "α  β"
    and "𝔉' = cf_map 𝔉"
    and "B = 𝔉ObjMapr"
    and "A = Hom 
      (cat_FUNCT α  (cat_Set α))
      (cf_map (HomO.Cα(r,-)))
      (cf_map 𝔉)"
    and "𝔉 :  ↦↦Cαcat_Set α"
    and "r  Obj"
  shows "ntcf_Yoneda_arrow α  𝔉' r : A isocat_Set βB"
  using assms(1,2,6,7) 
  unfolding assms(3-5)
  by (rule cat_ntcf_Yoneda_arrow_is_arr_isomoprhism)

lemmas [cat_arrow_cs_intros] = 
  category.cat_ntcf_Yoneda_arrow_is_arr_isomoprhism'

lemma (in category) cat_ntcf_Yoneda_arrow_is_arr:
  assumes "𝒵 β"
    and "α  β"
    and "𝔉 :  ↦↦Cαcat_Set α"
    and "r  Obj"
  shows "ntcf_Yoneda_arrow α  (cf_map 𝔉) r :
    Hom
      (cat_FUNCT α  (cat_Set α))
      (cf_map (HomO.Cα(r,-)))
      (cf_map 𝔉) cat_Set β𝔉ObjMapr"
  by 
    (
      rule cat_Set_is_iso_arrD[
        OF cat_ntcf_Yoneda_arrow_is_arr_isomoprhism[OF assms]
        ]
    )

lemma (in category) cat_ntcf_Yoneda_arrow_is_arr'[cat_cs_intros]:
  assumes "𝒵 β"
    and "α  β"
    and "𝔉' = cf_map 𝔉"
    and "B = 𝔉ObjMapr"
    and "A = Hom 
      (cat_FUNCT α  (cat_Set α))
      (cf_map (HomO.Cα(r,-)))
      (cf_map 𝔉)"
    and "𝔉 :  ↦↦Cαcat_Set α"
    and "r  Obj"
  shows "ntcf_Yoneda_arrow α  𝔉' r : A cat_Set βB"
  using assms(1,2,6,7) 
  unfolding assms(3-5)
  by (rule cat_ntcf_Yoneda_arrow_is_arr)

lemmas [cat_arrow_cs_intros] = category.cat_ntcf_Yoneda_arrow_is_arr'


subsection‹Commutativity law for the Yoneda natural transformation arrow›

lemma (in category) cat_ntcf_Yoneda_arrow_commute:
  assumes "𝒵 β"
    and "α  β"
    and "𝔑 : 𝔉 CF 𝔊 :  ↦↦Cαcat_Set α"
    and "f : a b" 
  shows 
    "ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set βcf_hom
        (cat_FUNCT α  (cat_Set α)) 
        [ntcf_arrow HomA.Cα(f,-), ntcf_arrow 𝔑] =
          cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a"
proof-

  let ?hom = 
    cf_hom
        (cat_FUNCT α  (cat_Set α)) 
        [ntcf_arrow HomA.Cα(f,-), ntcf_arrow 𝔑]

  interpret β: 𝒵 β by (rule assms(1))
  interpret 𝔑: is_ntcf α  cat_Set α 𝔉 𝔊 𝔑 by (rule assms(3))
  interpret Set: category α cat_Set α by (rule category_cat_Set)
  interpret βℭ: category β 
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  interpret cat_Set_αβ: subcategory β cat_Set α cat_Set β
    by (rule subcategory_cat_Set_cat_Set[OF assms(1,2)])

  from assms(2,4) have 𝔊b_𝔑f:
    "ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom :
      Hom
        (cat_FUNCT α  (cat_Set α))
        (cf_map (HomO.Cα(a,-)))
        (cf_map 𝔉) cat_Set β𝔊ObjMapb"
    by
      (
        cs_concl 
          cs_intro:
            cat_small_cs_intros
            cat_cs_intros
            cat_prod_cs_intros
            cat_op_intros
            cat_FUNCT_cs_intros
      )

  from assms(2,4) have 𝔑f_𝔉a:
    "cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a :
      Hom
        (cat_FUNCT α  (cat_Set α))
        (cf_map (HomO.Cα(a,-)))
        (cf_map 𝔉) cat_Set β𝔊ObjMapb"
    by (cs_concl cs_intro: cat_cs_intros cat_Set_αβ.subcat_is_arrD)

  show ?thesis
  proof(rule arr_Set_eqI[of β])

    from 𝔊b_𝔑f show arr_Set_𝔊b_𝔑f:
      "arr_Set β (ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom)"
      by (auto dest: cat_Set_is_arrD(1))
    from 𝔊b_𝔑f have dom_lhs: 
      "𝒟 ((ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom)ArrVal) = 
        Hom
          (cat_FUNCT α  (cat_Set α))
          (cf_map (HomO.Cα(a,-)))
          (cf_map 𝔉)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)+
    interpret 𝔑f_𝔉a: arr_Set 
      β ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom
      by (rule arr_Set_𝔊b_𝔑f)
  
    from 𝔑f_𝔉a show arr_Set_𝔑f_𝔉a:
      "arr_Set 
        β 
        (
          cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a
        )"
      by (auto dest: cat_Set_is_arrD(1))
    from 𝔑f_𝔉a have dom_rhs: 
      "𝒟 
        (
          (
            cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a
          )ArrVal
        ) = Hom
          (cat_FUNCT α  (cat_Set α))
          (cf_map (HomO.Cα(a,-)))
          (cf_map 𝔉)" 
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)

    show 
      "(ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom)ArrVal =
        (
          cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a
        )ArrVal"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs in_Hom_iff)

      fix 𝔐 assume prems: 
        "𝔐 : cf_map HomO.Cα(a,-) cat_FUNCT α  (cat_Set α)cf_map 𝔉"

      from assms(4) have [cat_cs_simps]:
        "cf_of_cf_map  (cat_Set α) (cf_map HomO.Cα(a,-)) = HomO.Cα(a,-)"
        "cf_of_cf_map  (cat_Set α) (cf_map 𝔉) = 𝔉"
        by (cs_concl cs_simp: cat_FUNCT_cs_simps cs_intro: cat_cs_intros)

      note 𝔐 = cat_FUNCT_is_arrD[OF prems, unfolded cat_cs_simps]

      interpret 𝔐: is_ntcf 
        α  cat_Set α HomO.Cα(a,-) 𝔉 ntcf_of_ntcf_arrow  (cat_Set α) 𝔐
        by (rule 𝔐(1))

      have 𝔊𝔑_eq_𝔑𝔉:
        "𝔊ArrMapfArrVal𝔑NTMapaArrValA =
          𝔑NTMapbArrVal𝔉ArrMapfArrValA"
        if "A  𝔉ObjMapa" for A      
        using 
          ArrVal_eq_helper[
            OF 𝔑.ntcf_Comp_commute[OF assms(4), symmetric], where a=A
            ]
          assms(4)
          that
        by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

      from 𝔐(1) assms(2,3,4) have 𝔐a_CId_a: 
        "𝔐NTMapaArrValCIda  𝔉ObjMapa"
        by (subst 𝔐)
          (
            cs_concl 
              cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps 
              cs_intro: cat_Set_cs_intros cat_cs_intros
          )

      have 𝔉f_𝔐a_eq_𝔐b:
        "𝔉ArrMapfArrVal𝔐NTMapaArrValh =
          𝔐NTMapbArrValf Ah"
        if "h : a a" for h
        using 
          ArrVal_eq_helper[
            OF 𝔐.ntcf_Comp_commute[OF assms(4), symmetric], where a=h
            ]
          that
          assms(4)
          category_axioms
        by
          (
            cs_prems cs_shallow
              cs_simp: 
                cat_FUNCT_cs_simps 
                cat_map_extra_cs_simps 
                cat_cs_simps 
                cat_op_simps
              cs_intro: cat_cs_intros cat_prod_cs_intros cat_op_intros
          )
    
      from 𝔐(1) assms(2,3,4) 𝔐a_CId_a category_axioms show 
        "(ntcf_Yoneda_arrow α  (cf_map 𝔊) b Acat_Set β?hom)ArrVal𝔐 =
          (
            cf_eval_arrow  (ntcf_arrow 𝔑) f Acat_Set βntcf_Yoneda_arrow α  (cf_map 𝔉) a
          )ArrVal𝔐" 
        by (subst (1 2) 𝔐(2)) (*very slow*)
          (
            cs_concl 
              cs_simp:
                𝔉f_𝔐a_eq_𝔐b 𝔊𝔑_eq_𝔑𝔉
                cat_map_extra_cs_simps 
                cat_FUNCT_cs_simps 
                cat_cs_simps 
                cat_op_simps
                cat_Set_components(1)
              cs_intro: 
                cat_Set_αβ.subcat_is_arrD 
                cat_small_cs_intros
                cat_cs_intros 
                cat_FUNCT_cs_intros
                cat_prod_cs_intros
                cat_op_intros
          )

    qed (use arr_Set_𝔊b_𝔑f arr_Set_𝔑f_𝔉a in auto)

  qed (use 𝔊b_𝔑f 𝔑f_𝔉a in cs_concl cs_shallow cs_simp: cat_cs_simps)+

qed



subsection‹Yoneda Lemma: naturality›


subsubsection‹
The Yoneda natural transformation: definition and elementary properties
›


text‹
The main result of this subsection corresponds to the corollary to the 
Yoneda Lemma on page 61 in cite"mac_lane_categories_2010".
›

definition ntcf_Yoneda :: "V  V  V  V"
  where "ntcf_Yoneda α β  =
    [
      (
        λ𝔉r(cat_FUNCT α  (cat_Set α) ×C )Obj. 
          ntcf_Yoneda_arrow α  (𝔉r0) (𝔉r1)
      ),
      cf_nt α β (cf_id ),
      cf_eval α β ,
      cat_FUNCT α  (cat_Set α) ×C ,
      cat_Set β
    ]"


text‹Components.›

lemma ntcf_Yoneda_components:
  shows "ntcf_Yoneda α β NTMap =
    (
      λ𝔉r(cat_FUNCT α  (cat_Set α) ×C )Obj.
        ntcf_Yoneda_arrow α  (𝔉r0) (𝔉r1)
    )"
    and [cat_cs_simps]: "ntcf_Yoneda α β NTDom = cf_nt α β (cf_id )"
    and [cat_cs_simps]: "ntcf_Yoneda α β NTCod = cf_eval α β "
    and [cat_cs_simps]: 
      "ntcf_Yoneda α β NTDGDom = cat_FUNCT α  (cat_Set α) ×C "
    and [cat_cs_simps]: "ntcf_Yoneda α β NTDGCod = cat_Set β"
  unfolding ntcf_Yoneda_def nt_field_simps by (simp_all add: nat_omega_simps)
    

subsubsection‹Natural transformation map›

mk_VLambda ntcf_Yoneda_components(1)
  |vsv ntcf_Yoneda_NTMap_vsv[cat_cs_intros]|
  |vdomain ntcf_Yoneda_NTMap_vdomain[cat_cs_intros]|

lemma (in category) ntcf_Yoneda_NTMap_app[cat_cs_simps]:
  assumes "𝒵 β"
    and "α  β" 
    and "𝔉r = [cf_map 𝔉, r]"
    and "𝔉 :  ↦↦Cαcat_Set α"
    and "r  Obj"
  shows "ntcf_Yoneda α β NTMap𝔉r = ntcf_Yoneda_arrow α  (cf_map 𝔉) r"
proof-            
  interpret β: 𝒵 β by (rule assms(1))
  interpret 𝔉: is_functor α  cat_Set α 𝔉 by (rule assms(4))
  interpret βℭ: category β 
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  from assms(2) interpret FUNCT: category β cat_FUNCT α  (cat_Set α)
    by
      (
        cs_concl cs_shallow 
          cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros
      )
  from assms(5) have "[cf_map 𝔉, r]  (cat_FUNCT α  (cat_Set α) ×C )Obj"
    by
      (
        cs_concl cs_shallow
          cs_simp: cat_FUNCT_cs_simps
          cs_intro: cat_cs_intros cat_prod_cs_intros cat_FUNCT_cs_intros
      )
  then show ?thesis
    unfolding assms(3) ntcf_Yoneda_components by (simp add: nat_omega_simps)
qed

lemmas [cat_cs_simps] = category.ntcf_Yoneda_NTMap_app


subsubsection‹The Yoneda natural transformation is a natural transformation›

lemma (in category) cat_ntcf_Yoneda_is_ntcf:
  assumes "𝒵 β" and "α  β" 
  shows "ntcf_Yoneda α β  :
    cf_nt α β (cf_id ) CF.iso cf_eval α β  :
    cat_FUNCT α  (cat_Set α) ×C  ↦↦Cβcat_Set β"
proof-

  interpret β: 𝒵 β by (rule assms(1))
  interpret βℭ: category β 
    by (rule category.cat_category_if_ge_Limit)
      (use assms(2) in cs_concl cs_shallow cs_intro: cat_cs_intros)+
  from assms(2) interpret FUNCT: category β cat_FUNCT α  (cat_Set α)
    by 
      (
        cs_concl cs_shallow 
          cs_intro: cat_cs_intros cat_small_cs_intros cat_FUNCT_cs_intros
      )

  show ?thesis
  proof(intro is_iso_ntcfI is_ntcfI')
    show "vfsequence (ntcf_Yoneda α β )" unfolding ntcf_Yoneda_def by simp
    show "vcard (ntcf_Yoneda α β ) = 5"
      unfolding ntcf_Yoneda_def by (simp add: nat_omega_simps)
    show ntcf_Yoneda_𝔉r: "ntcf_Yoneda α β NTMap𝔉r :
      cf_nt α β (cf_id )ObjMap𝔉r isocat_Set βcf_eval α β ObjMap𝔉r"
      if "𝔉r  (cat_FUNCT α  (cat_Set α) ×C )Obj" for 𝔉r
    proof-
      from that obtain 𝔉 r 
        where 𝔉r_def: "𝔉r = [𝔉, r]" 
          and 𝔉: "𝔉  cf_maps α  (cat_Set α)"
          and r: "r  Obj"
        by
          (
            auto 
              elim: cat_prod_2_ObjE[rotated 2] 
              simp: cat_FUNCT_cs_simps
              intro: cat_cs_intros
           )
      from 𝔉 obtain 𝔊
        where 𝔉_def: "𝔉 = cf_map 𝔊" and 𝔊: "𝔊 :  ↦↦Cαcat_Set α"
        by clarsimp
      from assms(2) 𝔊 r show ?thesis
        unfolding 𝔉r_def 𝔉_def 
        by
          (
            cs_concl 
              cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_arrow_cs_intros
           )      
    qed
    show "ntcf_Yoneda α β NTMap𝔉r :
      cf_nt α β (cf_id )ObjMap𝔉r cat_Set βcf_eval α β ObjMap𝔉r"
      if "𝔉r  (cat_FUNCT α  (cat_Set α) ×C )Obj" for 𝔉r
      by (rule is_iso_arrD[OF ntcf_Yoneda_𝔉r[OF that]])
    show 
      "ntcf_Yoneda α β NTMap𝔊b Acat_Set βcf_nt α β (cf_id )ArrMap𝔑f =
          cf_eval α β ArrMap𝔑f Acat_Set βntcf_Yoneda α β NTMap𝔉a"
      if 𝔑f: "𝔑f : 𝔉a cat_FUNCT α  (cat_Set α) ×C 𝔊b" for 𝔉a 𝔊b 𝔑f
    proof-
      obtain 𝔑 f 𝔉 a 𝔊 b
        where 𝔑f_def: "𝔑f = [𝔑, f]" 
          and 𝔉a_def: "𝔉a = [𝔉, a]"
          and 𝔊b_def: "𝔊b = [𝔊, b]" 
          and 𝔑: "𝔑 : 𝔉 cat_FUNCT α  (cat_Set α)𝔊" 
          and f: "f : a b"
        by 
          (
            auto intro: 
              cat_prod_2_is_arrE[rotated 2, OF 𝔑f] 
              FUNCT.category_axioms 
              βℭ.category_axioms
          )
      note 𝔑 = cat_FUNCT_is_arrD[OF 𝔑]
      note [cat_cs_simps] = 
        cat_ntcf_Yoneda_arrow_commute[OF assms 𝔑(1) f, folded 𝔑(2,3,4)]
      from 𝔑(1) assms(2) f show ?thesis
        unfolding 𝔑f_def 𝔉a_def 𝔊b_def
        by (subst (1 2) 𝔑(2), use nothing in subst 𝔑(3), subst 𝔑(4))
          (
            cs_concl 
              cs_simp: 𝔑(2,3,4)[symmetric] cat_cs_simps cs_intro: cat_cs_intros
          )+
    qed

  qed (use assms(2) in cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)+

qed



subsectionHom›-map›

text‹
This subsection presents some of the results stated as Corollary 2 
in subsection 1.15 in cite"bodo_categories_1970" and the corollary 
following the statement of the Yoneda Lemma on 
page 61 in cite"mac_lane_categories_2010" in a variety of forms.
›


subsubsection‹Definition and elementary properties›


text‹
The following function makes an explicit appearance in subsection 1.15 in 
cite"bodo_categories_1970".
›

definition ntcf_Hom_map :: "V  V  V  V  V"
  where "ntcf_Hom_map α  a b = (λfHom  a b. HomA.Cα(f,-))"


text‹Elementary properties.›

mk_VLambda ntcf_Hom_map_def
  |vsv ntcf_Hom_map_vsv|
  |vdomain ntcf_Hom_map_vdomain[cat_cs_simps]|
  |app ntcf_Hom_map_app[unfolded in_Hom_iff, cat_cs_simps]|


subsubsectionHom›-map is a bijection›

lemma (in category) cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique:
  ―‹The following lemma approximately corresponds to the corollary on 
page 61 in \cite{mac_lane_categories_2010}.›
  assumes "r  Obj" 
    and "s  Obj"
    and "𝔑 : HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
  shows "Yoneda_map α HomO.Cα(s,-) r𝔑 : s r"
    and "𝔑 = HomA.Cα(Yoneda_map α HomO.Cα(s,-) r𝔑,-)"
    and "f.  f  Arr; 𝔑 = HomA.Cα(f,-)  
      f = Yoneda_map α HomO.Cα(s,-) r𝔑"
proof-

  interpret 𝔑: is_ntcf α  cat_Set α HomO.Cα(r,-) HomO.Cα(s,-) 𝔑
    by (rule assms(3))
  let ?Y_Hom_s = Yoneda_map α HomO.Cα(s,-) r
  note Yoneda = 
    cat_Yoneda_Lemma[OF cat_cf_Hom_snd_is_functor[OF assms(2)] assms(1)]
  interpret Y: v11 ?Y_Hom_s by (rule Yoneda(1))

  from category_axioms assms have 𝔑_in_vdomain: "𝔑  𝒟 (?Y_Hom_s)" 
    by (cs_concl cs_shallow cs_simp: these_ntcfs_iff cat_cs_simps cs_intro: cat_cs_intros) 
  then have "?Y_Hom_s𝔑   (?Y_Hom_s)" by (simp add: Y.vsv_vimageI2)
  from this category_axioms assms show Ym_𝔑: "?Y_Hom_s𝔑 : s r"
    unfolding Yoneda(2) 
    by (cs_prems cs_shallow cs_simp: cat_cs_simps cat_op_simps) 
  then have "?Y_Hom_s𝔑  Arr" by (simp add: cat_cs_intros)

  have "HomA.Cα(?Y_Hom_s𝔑,-) :
    HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
    by (intro cat_ntcf_Hom_snd_is_ntcf Ym_𝔑)

  from assms Ym_𝔑 this category_axioms assms have 
    "(?Y_Hom_s)¯?Y_Hom_s𝔑 =
      Yoneda_arrow α HomO.Cα(s,-) r (?Y_Hom_s𝔑)"
    by (intro category.inv_Yoneda_map_app)
      (
        cs_concl cs_shallow 
          cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
      ) 
  then have "(?Y_Hom_s)¯?Y_Hom_s𝔑 = HomA.Cα(?Y_Hom_s𝔑,-)"
    by (simp add: ntcf_Hom_snd_def'[OF Ym_𝔑])
  with 𝔑_in_vdomain show "𝔑 = HomA.Cα(?Y_Hom_s𝔑,-)" by auto

  fix f assume prems: "f  Arr" "𝔑 = HomA.Cα(f,-)"
  then obtain a b where f: "f : a b" by auto

  have "𝔑 : HomO.Cα(b,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
    by (rule cat_ntcf_Hom_snd_is_ntcf[OF f, folded prems(2)])
  with f 𝔑.ntcf_NTDom 𝔑.ntcf_NTCod assms cat_is_arrD(2,3)[OF f] 
  have ba_simps: "b = r" "a = s"
    by 
      (
        simp_all add: 
          prems(2) cat_cf_Hom_snd_inj cat_ntcf_Hom_snd_components(2,3)
      )
  from f have "f : s r" unfolding ba_simps .

  with category_axioms show "f = ?Y_Hom_s𝔑"
    unfolding prems(2) 
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cat_op_simps)

qed

lemma (in category) cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique:
  assumes "r  Obj" 
    and "s  Obj"
    and "𝔑 : HomO.Cα(-,r) CF HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
  shows "Yoneda_map α HomO.Cα(-,s) r𝔑 : r s"
    and "𝔑 = HomA.Cα(-,Yoneda_map α HomO.Cα(-,s) r𝔑)"
    and "f.  f  Arr; 𝔑 = HomA.Cα(-,f)  
      f = Yoneda_map α HomO.Cα(-,s) r𝔑"
  by 
    (
      intro  
        category.cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[
          OF category_op, 
          unfolded cat_op_simps cat_op_cat_ntcf_Hom_snd, 
          OF assms(1,2), 
          unfolded assms(1,2)[THEN cat_op_cat_cf_Hom_snd],
          OF assms(3)
          ]
    )+

lemma (in category) cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique':
  assumes "r  Obj" 
    and "s  Obj"
    and "𝔑 : HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
  shows "∃!f. f  Arr  𝔑 = HomA.Cα(f,-)"
  using cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms] by blast

lemma (in category) cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique':
  assumes "r  Obj"
    and "s  Obj"
    and "𝔑 : HomO.Cα(-,r) CF HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
  shows "∃!f. f  Arr  𝔑 = HomA.Cα(-,f)"
  using cat_ntcf_Hom_fst_is_ntcf_Hom_fst_unique[OF assms] by blast

lemma (in category) cat_ntcf_Hom_snd_inj:
  assumes "HomA.Cα(g,-) = HomA.Cα(f,-)" 
    and "g : a b" 
    and "f : a b" 
  shows "g = f"
proof-
  from assms have 
    "Yoneda_map α (HomO.Cα(a,-)) bHomA.Cα(g,-) =
      Yoneda_map α (HomO.Cα(a,-)) bHomA.Cα(f,-)"
    by simp
  from this assms category_axioms show "g = f"
    by 
      (
        cs_prems cs_shallow 
          cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
      ) 
      simp (*slow*)
qed

lemma (in category) cat_ntcf_Hom_fst_inj:
  assumes "HomA.Cα(-,g) = HomA.Cα(-,f)" 
    and "g : a b" 
    and "f : a b" 
  shows "g = f"
proof-
  from category.cat_ntcf_Hom_snd_inj
    [
      OF category_op, 
      unfolded cat_op_simps,
      unfolded cat_op_cat_ntcf_Hom_snd,
      OF assms
    ]
  show ?thesis .
qed

lemma (in category) cat_ntcf_Hom_map: 
  assumes "a  Obj" and "b  Obj"
  shows "v11 (ntcf_Hom_map α  a b)" 
    and " (ntcf_Hom_map α  a b) =
      these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-)"
    and "(ntcf_Hom_map α  a b)¯ =
      (λ𝔑these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-).
        Yoneda_map α HomO.Cα(a,-) b𝔑)"
proof-

  show "v11 (ntcf_Hom_map α  a b)"
  proof(rule vsv.vsv_valeq_v11I, unfold ntcf_Hom_map_vdomain in_Hom_iff)
    show "vsv (ntcf_Hom_map α  a b)" unfolding ntcf_Hom_map_def by simp
    fix g f assume prems: 
      "g : a b" 
      "f : a b"
      "ntcf_Hom_map α  a bg = ntcf_Hom_map α  a bf"
    from prems(3,1,2) have "HomA.Cα(g,-) = HomA.Cα(f,-)"
      by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    with prems(1,2) show "g = f" by (intro cat_ntcf_Hom_snd_inj[of g f])
  qed
  then interpret Hm: v11 ntcf_Hom_map α  a b .

  show Hm_vrange: " (ntcf_Hom_map α  a b) =
    these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-)"
  proof(intro vsubset_antisym)
    show " (ntcf_Hom_map α  a b) 
      these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-)"
      by 
        (
          unfold ntcf_Hom_map_def,
          intro vrange_VLambda_vsubset, 
          unfold these_ntcfs_iff in_Hom_iff, 
          intro cat_ntcf_Hom_snd_is_ntcf
        )
    show "these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-) 
       (ntcf_Hom_map α  a b)"
    proof(intro vsubsetI, unfold these_ntcfs_iff)
      fix 𝔑 assume prems: 
        "𝔑 : HomO.Cα(b,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
      note unique = 
        cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) prems]
      from unique(1) have 
        "Yoneda_map α HomO.Cα(a,-) b𝔑  𝒟 (ntcf_Hom_map α  a b)"
        by (cs_concl cs_simp: cat_cs_simps)
      moreover from 
        cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(1,2)[OF assms(2,1) prems] 
      have 𝔑_def: "𝔑 = ntcf_Hom_map α  a bYoneda_map α HomO.Cα(a,-) b𝔑"
        by (cs_concl cs_shallow cs_simp: cat_cs_simps)
      ultimately show "𝔑   (ntcf_Hom_map α  a b)" by force
    qed 
  qed

  show "(ntcf_Hom_map α  a b)¯ =
    (
      λ𝔑these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-).
        Yoneda_map α HomO.Cα(a,-) b𝔑
    )"
  proof
    (
      rule vsv_eqI, 
      unfold vdomain_vconverse vdomain_VLambda Hm_vrange these_ntcfs_iff
    )

    from Hm.v11_axioms show "vsv ((ntcf_Hom_map α  a b)¯)" by auto
    show "vsv 
      (
        λ𝔑these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-).
          Yoneda_map α HomO.Cα(a,-) b𝔑
      )"
      by simp

    fix 𝔑 assume prems: 
      "𝔑 : HomO.Cα(b,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
    then have 𝔑: 
      "𝔑  these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-)"
      unfolding these_ntcfs_iff by simp
    show "(ntcf_Hom_map α  a b)¯𝔑 =
      (
        λ𝔑these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-).
          Yoneda_map α HomO.Cα(a,-) b𝔑
      )𝔑"
    proof
      (
        intro Hm.v11_vconverse_app, 
        unfold ntcf_Hom_map_vdomain in_Hom_iff beta[OF 𝔑]
      )
      note unique = 
        cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) prems]
      show "Yoneda_map α HomO.Cα(a,-) b𝔑 : a b" by (rule unique(1))
      then show 
        "ntcf_Hom_map α  a bYoneda_map α HomO.Cα(a,-) b𝔑 = 𝔑"
        by (cs_concl cs_simp: unique(2)[symmetric] cat_cs_simps)
    qed

  qed simp

qed


subsubsection‹Inverse of a Hom›-map›

lemma (in category) inv_ntcf_Hom_map_v11: 
  assumes "a  Obj" and "b  Obj"
  shows "v11 ((ntcf_Hom_map α  a b)¯)"
  using cat_ntcf_Hom_map(1)[OF assms] by (simp add: v11.v11_vconverse)

lemma (in category) inv_ntcf_Hom_map_vdomain: 
  assumes "a  Obj" and "b  Obj"
  shows "𝒟 ((ntcf_Hom_map α  a b)¯) =
    these_ntcfs α  (cat_Set α) HomO.Cα(b,-) HomO.Cα(a,-)"
  unfolding cat_ntcf_Hom_map(3)[OF assms] by simp

lemmas [cat_cs_simps] = category.inv_ntcf_Hom_map_vdomain

lemma (in category) inv_ntcf_Hom_map_app: 
  assumes "a  Obj" 
    and "b  Obj"
    and "𝔑 : HomO.Cα(b,-) CF HomO.Cα(a,-) :  ↦↦Cαcat_Set α"
  shows "(ntcf_Hom_map α  a b)¯𝔑 = Yoneda_map α HomO.Cα(a,-) b𝔑"
  using assms(3) unfolding cat_ntcf_Hom_map(3)[OF assms(1,2)] by simp

lemmas [cat_cs_simps] = category.inv_ntcf_Hom_map_app

lemma inv_ntcf_Hom_map_vrange: " ((ntcf_Hom_map α  a b)¯) = Hom  a b"
  unfolding ntcf_Hom_map_def by simp


subsubsectionHom›-natural transformation and isomorphisms›


text‹
This subsection presents further results that were stated 
as Corollary 2 in subsection 1.15 in cite"bodo_categories_1970".
›

lemma (in category) cat_is_iso_arr_ntcf_Hom_snd_is_iso_ntcf:
  assumes "f : s isor"
  shows "HomA.Cα(f,-) :
    HomO.Cα(r,-) CF.iso HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
proof-
  from assms obtain g 
    where iso_g: "g : r isos" 
      and gf: "g Af = CIds"
      and fg: "f Ag = CIdr"
    by 
      (
        auto intro:
          cat_the_inverse_Comp_CId_left 
          cat_the_inverse_Comp_CId_right 
          cat_the_inverse_is_iso_arr'
      )
  then have g: "g : r s" by auto
  show ?thesis
  proof(intro is_iso_arr_is_iso_ntcf)
    from assms have f: "f : s r" by auto
    with category_axioms show "HomA.Cα(f,-) :
      HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros) 
    from category_axioms g show "HomA.Cα(g,-) :
      HomO.Cα(s,-) CF HomO.Cα(r,-) :  ↦↦Cαcat_Set α"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)    
    from category_axioms f g have 
      "HomA.Cα(f,-) NTCF HomA.Cα(g,-) = HomA.Cα(g Af,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    also from category_axioms f g have " = ntcf_id HomO.Cα(s,-)"
      by (cs_concl cs_simp: gf cat_cs_simps cs_intro: cat_cs_intros)
    finally show 
      "HomA.Cα(f,-) NTCF HomA.Cα(g,-) = ntcf_id HomO.Cα(s,-)"
      by simp
    from category_axioms f g have 
      "HomA.Cα(g,-) NTCF HomA.Cα(f,-) = HomA.Cα(f Ag,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    also from category_axioms f g have " = ntcf_id HomO.Cα(r,-)"
      by (cs_concl cs_simp: fg cat_cs_simps cs_intro: cat_cs_intros)
    finally show 
      "HomA.Cα(g,-) NTCF HomA.Cα(f,-) = ntcf_id HomO.Cα(r,-)"
      by simp
  qed
qed

lemma (in category) cat_is_iso_arr_ntcf_Hom_fst_is_iso_ntcf:
  assumes "f : r isos"
  shows "HomA.Cα(-,f) :
    HomO.Cα(-,r) CF.iso HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
proof-
  from assms have r: "r  Obj" and s: "s  Obj" by auto
  from 
    category.cat_is_iso_arr_ntcf_Hom_snd_is_iso_ntcf
      [
        OF category_op, 
        unfolded cat_op_simps,
        OF assms,
        unfolded
          category.cat_op_cat_cf_Hom_snd[OF category_axioms r]
          category.cat_op_cat_cf_Hom_snd[OF category_axioms s]
          category.cat_op_cat_ntcf_Hom_snd[OF category_axioms]
      ]
  show ?thesis.
qed

lemma (in category) cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique:
  assumes "r  Obj" 
    and "s  Obj"
    and "𝔑 : HomO.Cα(r,-) CF.iso HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
  shows "Yoneda_map α HomO.Cα(s,-) r𝔑 : s isor"
    and "𝔑 = HomA.Cα(Yoneda_map α HomO.Cα(s,-) r𝔑,-)"
    and "f.  f  Arr; 𝔑 = HomA.Cα(f,-)  
      f = Yoneda_map α HomO.Cα(s,-) r𝔑"
proof-

  let ?Ym_𝔑 = Yoneda_map α HomO.Cα(s,-) r𝔑
    and ?Ym_inv_𝔑 = Yoneda_map α HomO.Cα(r,-) sinv_ntcf 𝔑

  from assms(3) have 𝔑:
    "𝔑 : HomO.Cα(r,-) CF HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
    by auto
  from iso_ntcf_is_iso_arr[OF assms(3)] 
  have iso_inv_𝔑: "inv_ntcf 𝔑 : 
    HomO.Cα(s,-) CF.iso HomO.Cα(r,-) :  ↦↦Cαcat_Set α"
    and [simp]: "𝔑 NTCF inv_ntcf 𝔑 = ntcf_id HomO.Cα(s,-)"
    and [simp]: "inv_ntcf 𝔑 NTCF 𝔑 = ntcf_id HomO.Cα(r,-)"
    by auto
  from iso_inv_𝔑 have inv_𝔑: 
    "inv_ntcf 𝔑 : HomO.Cα(s,-) CF HomO.Cα(r,-) :  ↦↦Cαcat_Set α"
    by auto 
  note unique = cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(1,2) 𝔑]
    and inv_unique = 
    cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique[OF assms(2,1) inv_𝔑]
  have Ym_𝔑: "?Ym_𝔑 : s r" by (rule unique(1))

  show "𝔑 = HomA.Cα(Yoneda_map α HomO.Cα(s,-) r𝔑,-)"
    and "f.  f  Arr; 𝔑 = HomA.Cα(f,-)  
      f = Yoneda_map α HomO.Cα(s,-) r𝔑"
    by (intro unique)+

  show "Yoneda_map α HomO.Cα(s,-) r𝔑 : s isor"
  proof(intro is_iso_arrI[OF Ym_𝔑, of ?Ym_inv_𝔑] is_inverseI)

    show Ym_inv_𝔑: "?Ym_inv_𝔑 : r s" by (rule inv_unique(1))
    
    have "ntcf_id HomO.Cα(s,-) = 𝔑 NTCF inv_ntcf 𝔑" by simp
    also have " = HomA.Cα(?Ym_𝔑,-) NTCF HomA.Cα(?Ym_inv_𝔑,-)"
      by (subst unique(2), subst inv_unique(2)) simp
    also from category_axioms Ym_𝔑 inv_unique(1) assms(3) have 
      " = HomA.Cα(?Ym_inv_𝔑 A?Ym_𝔑,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    finally have "HomA.Cα(?Ym_inv_𝔑 A?Ym_𝔑,-) = ntcf_id HomO.Cα(s,-)"
      by simp
    also from category_axioms assms(1,2) have " = HomA.Cα(CIds,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    finally have "HomA.Cα(?Ym_inv_𝔑 A?Ym_𝔑,-) = HomA.Cα(CIds,-)"
      by simp
    then show "?Ym_inv_𝔑 A?Ym_𝔑 = CIds"
      by (rule cat_ntcf_Hom_snd_inj)
        (
          alluse category_axioms Ym_𝔑 Ym_inv_𝔑 assms in 
              cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros
        )
    have "ntcf_id HomO.Cα(r,-) = inv_ntcf 𝔑 NTCF 𝔑" by simp
    also have " = HomA.Cα(?Ym_inv_𝔑,-) NTCF HomA.Cα(?Ym_𝔑,-)"
      by (subst unique(2), subst inv_unique(2)) simp
    also from category_axioms Ym_𝔑 inv_unique(1) have 
      " = HomA.Cα(?Ym_𝔑 A?Ym_inv_𝔑,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)
    finally have 
      "HomA.Cα(?Ym_𝔑 A?Ym_inv_𝔑,-) = ntcf_id HomO.Cα(r,-)"
      by simp
    also from category_axioms assms(1,2) have " = HomA.Cα(CIdr,-)"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    finally have 
      "HomA.Cα(?Ym_𝔑 A?Ym_inv_𝔑,-) = HomA.Cα(CIdr,-)"
      by simp
    then show "?Ym_𝔑 A?Ym_inv_𝔑 = CIdr"
      by (rule cat_ntcf_Hom_snd_inj)
        (
          alluse category_axioms Ym_𝔑 Ym_inv_𝔑 assms in 
              cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros
        )

  qed (intro Ym_𝔑)

qed

lemma (in category) cat_ntcf_Hom_fst_is_iso_ntcf_Hom_fst_unique:
  assumes "r  Obj" 
    and "s  Obj"
    and "𝔑 : 
    HomO.Cα(-,r) CF.iso HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
  shows "Yoneda_map α HomO.Cα(-,s) r𝔑 : r isos"
    and "𝔑 = HomA.Cα(-,Yoneda_map α HomO.Cα(-,s) r𝔑)"
    and "f.  f  Arr; 𝔑 = HomA.Cα(-,f)  
      f = Yoneda_map α HomO.Cα(-,s) r𝔑"
  by 
    (
      intro  
        category.cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique[
          OF category_op, 
          unfolded cat_op_simps cat_op_cat_ntcf_Hom_snd, 
          OF assms(1,2), 
          unfolded assms(1,2)[THEN cat_op_cat_cf_Hom_snd],
          OF assms(3)
          ]
    )+

lemma (in category) cat_is_iso_arr_if_ntcf_Hom_snd_is_iso_ntcf:
  assumes "f : s r"
    and "HomA.Cα(f,-) :
      HomO.Cα(r,-) CF.iso HomO.Cα(s,-) :  ↦↦Cαcat_Set α"
  shows "f : s isor"
proof-
  from assms(1) have r: "r  Obj" and s: "s  Obj" by auto
  note unique = cat_ntcf_Hom_snd_is_iso_ntcf_Hom_snd_unique[OF r s assms(2)]
  from unique(1) have Ym_Hf: 
    "Yoneda_map α HomO.Cα(s,-) rHomA.Cα(f,-) : s r"
    by auto
  from unique(1) show ?thesis
    unfolding cat_ntcf_Hom_snd_inj[OF unique(2) assms(1) Ym_Hf, symmetric]
    by simp
qed

lemma (in category) cat_is_iso_arr_if_ntcf_Hom_fst_is_iso_ntcf:
  assumes "f : r s"
    and "HomA.Cα(-,f) :
      HomO.Cα(-,r) CF.iso HomO.Cα(-,s) : op_cat  ↦↦Cαcat_Set α"
  shows "f : r isos"
proof-
  from assms(1) have r: "r  Obj" and s: "s  Obj" by auto
  note unique = cat_ntcf_Hom_fst_is_iso_ntcf_Hom_fst_unique[OF r s assms(2)]
  from unique(1) have Ym_Hf: 
    "Yoneda_map α HomO.Cα(-,s) rHomA.Cα(-,f) : r s"
    by auto
  from unique(1) show ?thesis
    unfolding cat_ntcf_Hom_fst_inj[OF unique(2) assms(1) Ym_Hf, symmetric]
    by simp
qed


subsubsection‹
The relationship between a Hom›-natural transformation and the compositions 
of a Hom›-natural transformation and a natural transformation
›

lemma (in category) cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app:
  assumes "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "b  𝔅Obj"
    and "c  Obj"
  shows "HomA.Cα(φ-,-)NTMapb, c = HomA.Cα(φNTMapb,-)NTMapc"
proof-      
  interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))
  from assms(2) have b: "b  𝔅Obj" unfolding cat_op_simps by simp
  from category_axioms assms(1,3) b show ?thesis
    by 
      (
        cs_concl cs_shallow
          cs_simp: 
            cat_ntcf_lcomp_Hom_component_is_Yoneda_component cat_cs_simps 
          cs_intro: cat_cs_intros cat_op_intros
      )
qed

lemmas [cat_cs_simps] = category.cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app

lemma (in category) cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd:
  assumes "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "b  𝔅Obj"
  shows "HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF = HomA.Cα(φNTMapb,-)"
proof-
  interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))  
  show ?thesis
  proof(rule ntcf_eqI[of α])
    from category_axioms assms show 
      "HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF : 
      HomO.Cα(𝔊ObjMapb,-) CF HomO.Cα(𝔉ObjMapb,-) : 
       ↦↦Cαcat_Set α"      
      by 
        (
          cs_concl cs_shallow 
            cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
        )
    from assms this have dom_lhs:
      "𝒟 ((HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF)NTMap) = Obj"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    from category_axioms assms show 
      "HomA.Cα(φNTMapb,-) :
        HomO.Cα(𝔊ObjMapb,-) CF HomO.Cα(𝔉ObjMapb,-) :
         ↦↦Cαcat_Set α"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    from assms this have dom_rhs: 
      "𝒟 (HomA.Cα(φNTMapb,-)NTMap) = Obj"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    show 
      "(HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF)NTMap = 
        HomA.Cα(φNTMapb,-)NTMap"
    proof(rule vsv_eqI, unfold dom_lhs dom_rhs)
      fix a assume "a  Obj"
      with category_axioms assms show
        "(HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF)NTMapa =
          HomA.Cα(φNTMapb,-)NTMapa"
        by (cs_concl cs_simp: cat_cs_simps) 
    qed (use assms(2) in auto intro: cat_cs_intros)
  qed simp_all
qed

lemmas [cat_cs_simps] = category.cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd


subsubsection‹
The relationship between the Hom›-natural isomorphisms and the compositions
of a Hom›-natural isomorphism and a natural transformation
›

lemma (in category) cat_ntcf_lcomp_Hom_if_ntcf_Hom_snd_is_iso_ntcf:
  assumes "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "b. b  𝔅Obj  HomA.Cα(φNTMapb,-) :
      HomO.Cα(𝔊ObjMapb,-) CF.iso HomO.Cα(𝔉ObjMapb,-) :
       ↦↦Cαcat_Set α"
  shows "HomA.Cα(φ-,-) :
    HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
    op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
proof-
  interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))
  have "HomA.Cα(φ-,-)op_cat 𝔅,(b,-)NTCF : 
    HomO.Cα(𝔊-,-)op_cat 𝔅,(b,-)CF CF.iso 
    HomO.Cα(𝔉-,-)op_cat 𝔅,(b,-)CF : 
     ↦↦Cαcat_Set α"
    if "b  𝔅Obj" for b
    unfolding 
      cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd[OF assms(1) that]
      cat_cf_lcomp_Hom_cf_Hom_snd[OF φ.NTDom.is_functor_axioms that]
      cat_cf_lcomp_Hom_cf_Hom_snd[OF φ.NTCod.is_functor_axioms that]
    by (intro assms(2) that)
  from 
    is_iso_ntcf_if_bnt_proj_snd_is_iso_ntcf[
      OF 
        φ.NTDom.HomDom.category_op category_axioms 
        cat_ntcf_lcomp_Hom_is_ntcf[OF assms(1)], 
      unfolded cat_op_simps, OF this
      ]
  show ?thesis .
qed

lemma (in category) cat_ntcf_Hom_snd_if_ntcf_lcomp_Hom_is_iso_ntcf:
  assumes "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    and "HomA.Cα(φ-,-) :
      HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    and "b  𝔅Obj"
  shows "HomA.Cα(φNTMapb,-) :
    HomO.Cα(𝔊ObjMapb,-) CF.iso HomO.Cα(𝔉ObjMapb,-) :
     ↦↦Cαcat_Set α"
proof-
  interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))
  from category_axioms assms show ?thesis
    by  
      (
        fold 
          cat_bnt_proj_snd_tcf_lcomp_Hom_ntcf_Hom_snd[OF assms(1,3)]
          cat_cf_lcomp_Hom_cf_Hom_snd[OF φ.NTDom.is_functor_axioms assms(3)]
          cat_cf_lcomp_Hom_cf_Hom_snd[OF φ.NTCod.is_functor_axioms assms(3)],
        intro bnt_proj_snd_is_iso_ntcf_if_is_iso_ntcf
      )    
      (cs_concl cs_shallow cs_simp: cat_op_simps cs_intro: cat_cs_intros)
qed



subsection‹Yoneda map for arbitrary functors›


text‹
The concept of the Yoneda map for arbitrary functors was developed based
on the function that was used in the statement of Lemma 3 in 
subsection 1.15 in cite"bodo_categories_1970".
›

definition af_Yoneda_map :: "V  V  V  V"
  where "af_Yoneda_map α 𝔉 𝔊 =
    (λφthese_ntcfs α (𝔉HomDom) (𝔉HomCod) 𝔉 𝔊. HomA.Cα(φ-,-))"


text‹Elementary properties.›

context
  fixes α 𝔅  𝔉 𝔊
  assumes 𝔉: "𝔉 : 𝔅 ↦↦Cα"
    and 𝔊: "𝔊 : 𝔅 ↦↦Cα"
begin

interpretation 𝔉: is_functor α 𝔅  𝔉 by (rule 𝔉)
interpretation 𝔊: is_functor α 𝔅  𝔊 by (rule 𝔊)

mk_VLambda 
  af_Yoneda_map_def[where 𝔉=𝔉 and 𝔊=𝔊, unfolded 𝔉.cf_HomDom 𝔉.cf_HomCod]
  |vsv af_Yoneda_map_vsv|
  |vdomain af_Yoneda_map_vdomain[cat_cs_simps]|
  |app af_Yoneda_map_app[unfolded these_ntcfs_iff, cat_cs_simps]|

end



subsection‹Yoneda arrow for arbitrary functors›


subsubsection‹Definition and elementary properties›


text‹
The following natural transformation is used in the proof of Lemma 3 in 
subsection 1.15 in cite"bodo_categories_1970".
›

definition af_Yoneda_arrow :: "V  V  V  V  V"
  where "af_Yoneda_arrow α 𝔉 𝔊 𝔑 =
    [
      (
        λb(𝔉HomDom)Obj.
          Yoneda_map α HomO.Cα𝔉HomCod(𝔉ObjMapb,-) (𝔊ObjMapb)
            𝔑op_cat (𝔉HomDom),𝔉HomCod(b,-)NTCF
            
      ),
      𝔉,
      𝔊,
      𝔉HomDom,
      𝔉HomCod
    ]"


text‹Components.›

lemma af_Yoneda_arrow_components:
  shows "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTMap =
      (
        λb𝔉HomDomObj.
          Yoneda_map α HomO.Cα𝔉HomCod(𝔉ObjMapb,-) (𝔊ObjMapb)
            𝔑op_cat (𝔉HomDom),𝔉HomCod(b,-)NTCF
            
      )"
    and "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTDom = 𝔉"
    and "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTCod = 𝔊"
    and "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTDGDom = 𝔉HomDom"
    and "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTDGCod = 𝔉HomCod"
  unfolding af_Yoneda_arrow_def nt_field_simps by (simp_all add: nat_omega_simps)


subsubsection‹Natural transformation map›

mk_VLambda af_Yoneda_arrow_components(1)
  |vsv af_Yoneda_arrow_NTMap_vsv|

context
  fixes α 𝔅  𝔉
  assumes 𝔉: "𝔉 : 𝔅 ↦↦Cα"
begin

interpretation 𝔉: is_functor α 𝔅  𝔉 by (rule 𝔉)

mk_VLambda 
  af_Yoneda_arrow_components(1)[where 𝔉=𝔉, unfolded 𝔉.cf_HomDom 𝔉.cf_HomCod]
  |vdomain af_Yoneda_arrow_NTMap_vdomain[cat_cs_simps]|
  |app af_Yoneda_arrow_NTMap_app[cat_cs_simps]|

end

lemma (in category) cat_af_Yoneda_arrow_is_ntcf:
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
  shows "af_Yoneda_arrow α 𝔉 𝔊 𝔑 : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
proof-

  let ?H𝔊 = HomO.Cα(𝔊-,-)
    and ?H𝔉 = HomO.Cα(𝔉-,-)
    and ?Set = cat_Set α
    and ?Ym = 
      λb. Yoneda_map
          α HomO.Cα(𝔉ObjMapb,-) (𝔊ObjMapb)𝔑op_cat 𝔅,(b,-)NTCF

  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))
  interpret 𝔑: is_ntcf 
    α op_cat 𝔅 ×C  cat_Set α HomO.Cα(𝔊-,-) HomO.Cα(𝔉-,-) 𝔑 
    by (rule assms)

  have comm[unfolded cat_op_simps]:
    "(𝔑NTMapc, d)ArrValf A(q A𝔊ArrMapg) =
      f A((𝔑NTMapa, b)ArrValq A𝔉ArrMapg)"
    if "g : a op_cat 𝔅c" and "f : b d" and "q : 𝔊ObjMapa b" 
    for q g f a b c d
  proof-
    from that(1) have g: "g : c 𝔅a" unfolding cat_op_simps by simp
    from category_axioms assms g that(2) have ab:
      "[a, b]  (op_cat 𝔅 ×C )Obj"
      by (cs_concl cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros)
    from 𝔑.ntcf_NTMap_is_arr[OF ab] category_axioms assms g that(2) have 𝔑ab: 
      "𝔑NTMapa, b :
        Hom  (𝔊ObjMapa) b cat_Set αHom  (𝔉ObjMapa) b"
      by 
        (
          cs_prems  
            cs_simp: cat_cs_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
        )
    have 𝔑_abq: "(𝔑NTMapa, b)ArrValq : 𝔉ObjMapa b"
      by 
        (
          rule cat_Set_ArrVal_app_vrange[
            OF 𝔑ab, unfolded in_Hom_iff, OF that(3)
            ]
        )
    have "[g, f] : [a, b] op_cat 𝔅 ×C [c, d]"
      by 
        (
          rule 
            cat_prod_2_is_arrI[
              OF 𝔉.HomDom.category_op category_axioms that(1,2)
              ]
        )
    then have 
      "𝔑NTMapc, d Acat_Set αHomO.Cα(𝔊-,-)ArrMapg, f =
        HomO.Cα(𝔉-,-)ArrMapg, f Acat_Set α𝔑NTMapa, b"
      by (rule is_ntcf.ntcf_Comp_commute[OF assms(3)])
    then have 
      "(𝔑NTMapc, d A?Set?H𝔊ArrMapg, f)ArrValq =
        (?H𝔉ArrMapg, f A?Set𝔑NTMapa, b)ArrValq"
      by auto

    from 
      this that(2,3) assms
      category_axioms 𝔉.HomDom.category_axioms 𝔉.HomDom.category_op category_op
      g 𝔑ab 𝔑_abq 
    show
      "(𝔑NTMapc, d)ArrValf A(q A𝔊ArrMapg) =
        f A((𝔑NTMapa, b)ArrValq A𝔉ArrMapg)" 
      by 
        (
          cs_prems 
            cs_simp: cat_cs_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
        )
  qed

  show ?thesis
  proof(rule is_ntcfI')

    show "vfsequence (af_Yoneda_arrow α 𝔉 𝔊 𝔑)"
      unfolding af_Yoneda_arrow_def by simp
    show "vcard (af_Yoneda_arrow α 𝔉 𝔊 𝔑) = 5"
      unfolding af_Yoneda_arrow_def by (simp add: nat_omega_simps)

    have 𝔑b: "𝔑op_cat 𝔅,(b,-)NTCF :
      HomO.Cα(𝔊ObjMapb,-) CF HomO.Cα(𝔉ObjMapb,-) :
       ↦↦Cαcat_Set α"
      if "b  𝔅Obj" for b
      by 
        (
          rule 
            bnt_proj_snd_is_ntcf
              [
                OF 𝔉.HomDom.category_op category_axioms assms(3),
                unfolded cat_op_simps, 
                OF that,
                unfolded 
                  cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(1) that]
                  cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(2) that]
              ]
        )

    show "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTMapb : 𝔉ObjMapb 𝔊ObjMapb"
      if "b  𝔅Obj" for b
    proof-
      let ?𝔊b = 𝔊ObjMapb
        and ?𝔉b = 𝔉ObjMapb
        and ?ℭ𝔊b = CId𝔊ObjMapb
      from that have ℭ𝔊b: "?ℭ𝔊b : ?𝔊b ?𝔊b" by (auto simp: cat_cs_intros)
      from assms that have "[b, ?𝔊b]  (op_cat 𝔅 ×C )Obj"
        by 
          (
            cs_concl cs_shallow
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
      from 𝔑.ntcf_NTMap_is_arr[OF this] category_axioms assms that have 𝔑_b𝔊b: 
        "𝔑NTMapb, ?𝔊b : Hom  ?𝔊b ?𝔊b cat_Set αHom  ?𝔉b ?𝔊b"
        by 
          (
            cs_prems cs_shallow
              cs_simp: cat_cs_simps cat_op_simps 
              cs_intro: cat_cs_intros cat_prod_cs_intros
          )
      from ℭ𝔊b have 𝔑_b𝔊b_ℭ𝔊b: 
        "(𝔑NTMapb, ?𝔊b)ArrVal?ℭ𝔊b : ?𝔉b ?𝔊b"
        by (rule cat_Set_ArrVal_app_vrange[OF 𝔑_b𝔊b, unfolded in_Hom_iff])
      with category_axioms assms that 𝔑b[OF that] show ?thesis
        by 
          (
            cs_concl cs_shallow 
              cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
          )
    qed

    show 
      "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTMapb A𝔉ArrMapf = 
        𝔊ArrMapf Aaf_Yoneda_arrow α 𝔉 𝔊 𝔑NTMapa"
      if "f : a 𝔅b" for a b f
    proof-

      from that have a: "a  𝔅Obj" and b: "b  𝔅Obj" by auto
      
      let ?𝔅a = 𝔅CIda
        and ?𝔅b = 𝔅CIdb
        and ?𝔊a = 𝔊ObjMapa
        and ?𝔊b = 𝔊ObjMapb
        and ?𝔉a = 𝔉ObjMapa
        and ?𝔉b = 𝔉ObjMapb
        and ?ℭ𝔊a = CId𝔊ObjMapa
        and ?ℭ𝔊b = CId𝔊ObjMapb

      from that have ℭ𝔊a: "?ℭ𝔊a : ?𝔊a ?𝔊a" by (auto intro: cat_cs_intros)
      from that have ℭ𝔊b: "?ℭ𝔊b : ?𝔊b ?𝔊b" by (auto intro: cat_cs_intros)
      from that have 𝔅a: "?𝔅a : a 𝔅a" by (auto intro: cat_cs_intros)

      from assms that have "[b, ?𝔊b]  (op_cat 𝔅 ×C )Obj"
        by 
          (
            cs_concl cs_shallow
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
      from 𝔑.ntcf_NTMap_is_arr[OF this] category_axioms assms that have 𝔑_b𝔊b: 
        "𝔑NTMapb, ?𝔊b : Hom  ?𝔊b ?𝔊b cat_Set αHom  ?𝔉b ?𝔊b"
        by 
          (
            cs_prems cs_shallow
              cs_simp: cat_cs_simps cat_op_simps
              cs_intro: cat_cs_intros cat_prod_cs_intros
          )
      from ℭ𝔊b have 𝔑_b𝔊b_ℭ𝔊b: 
        "(𝔑NTMapb, ?𝔊b)ArrVal?ℭ𝔊b : ?𝔉b ?𝔊b"
        by (rule cat_Set_ArrVal_app_vrange[OF 𝔑_b𝔊b, unfolded in_Hom_iff])

      from assms that have "[a, ?𝔊a]  (op_cat 𝔅 ×C )Obj"
        by 
          (
            cs_concl  
              cs_simp: cat_cs_simps 
              cs_intro: cat_cs_intros cat_op_intros cat_prod_cs_intros
          )
      from 𝔑.ntcf_NTMap_is_arr[OF this] category_axioms assms that have 𝔑_a𝔊a: 
        "𝔑NTMapa, ?𝔊a : Hom  ?𝔊a ?𝔊a cat_Set αHom  ?𝔉a ?𝔊a"
        by 
          (
            cs_prems  
              cs_simp: cat_cs_simps cat_op_simps 
              cs_intro: cat_cs_intros cat_prod_cs_intros
          )
      from ℭ𝔊a have 𝔑_a𝔊a_ℭ𝔊a: 
        "(𝔑NTMapa, ?𝔊a)ArrVal?ℭ𝔊a : ?𝔉a ?𝔊a"
        by (rule cat_Set_ArrVal_app_vrange[OF 𝔑_a𝔊a, unfolded in_Hom_iff])

      from 
        comm[OF 𝔅a 𝔊.cf_ArrMap_is_arr[OF that] ℭ𝔊a] 
        category_axioms assms that 𝔑_a𝔊a_ℭ𝔊a
      have 𝔑_a_𝔊b[symmetric, cat_cs_simps]:
        "(𝔑NTMapa, ?𝔊b)ArrVal𝔊ArrMapf =
          𝔊ArrMapf A(𝔑NTMapa, ?𝔊a)ArrVal?ℭ𝔊a"
        by (cs_prems cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

      from comm[OF that ℭ𝔊b ℭ𝔊b] category_axioms assms that 𝔑_b𝔊b_ℭ𝔊b
      have 𝔑_a_𝔊b'[cat_cs_simps]:
        "(𝔑NTMapa, ?𝔊b)ArrVal𝔊ArrMapf =
          (𝔑NTMapb, ?𝔊b)ArrVal?ℭ𝔊b A𝔉ArrMapf"
        by (cs_prems cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)

      from category_axioms assms that 𝔑b[OF a] 𝔑b[OF b] show ?thesis
        by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros)

    qed

  qed (auto simp: af_Yoneda_arrow_components cat_cs_simps intro: cat_cs_intros)

qed

lemma (in category) cat_af_Yoneda_arrow_is_ntcf':
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    and "β = α"
    and "𝔉' = 𝔉"
    and "𝔊' = 𝔊"
  shows "af_Yoneda_arrow α 𝔉 𝔊 𝔑 : 𝔉' CF 𝔊' : 𝔅 ↦↦Cβ"
  using assms(1-3) unfolding assms(4-6) by (rule cat_af_Yoneda_arrow_is_ntcf)

lemmas [cat_cs_intros] = category.cat_af_Yoneda_arrow_is_ntcf'


subsubsection‹Yoneda Lemma for arbitrary functors›


text‹
The following lemmas correspond to variants of the elements of Lemma 3 
in subsection 1.15 in cite"bodo_categories_1970".
›

lemma (in category) cat_af_Yoneda_map_af_Yoneda_arrow_app:
  assumes "𝔉 : 𝔅 ↦↦Cα" 
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
  shows "𝔑 = HomA.Cα(af_Yoneda_arrow α 𝔉 𝔊 𝔑-,-)"
proof-

  let ?H𝔊 = HomO.Cα(𝔊-,-)
    and ?H𝔉 = HomO.Cα(𝔉-,-)
    and ?aYa = λ𝔑. af_Yoneda_arrow α 𝔉 𝔊 𝔑

  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))

  interpret 𝔑: is_ntcf α op_cat 𝔅 ×C  cat_Set α ?H𝔊 ?H𝔉 𝔑
    by (rule assms(3))
  interpret aY𝔑: is_ntcf α 𝔅  𝔉 𝔊 ?aYa 𝔑
    by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms])
  interpret HY𝔑: is_ntcf 
    α op_cat 𝔅 ×C  cat_Set α ?H𝔊 ?H𝔉 HomA.Cα(?aYa 𝔑-,-)
    by (rule cat_ntcf_lcomp_Hom_is_ntcf[OF aY𝔑.is_ntcf_axioms])
  
  show [cat_cs_simps]: "𝔑 = HomA.Cα(?aYa 𝔑-,-)"
  proof 
    (
      rule sym,
      rule ntcf_eqI[OF HY𝔑.is_ntcf_axioms assms(3)], 
      rule vsv_eqI;
      (intro HY𝔑.NTMap.vsv_axioms 𝔑.NTMap.vsv_axioms)?;
      (unfold 𝔑.ntcf_NTMap_vdomain HY𝔑.ntcf_NTMap_vdomain)?
    )
    fix bc assume prems': "bc  (op_cat 𝔅 ×C )Obj"
    then obtain b c
      where bc_def: "bc = [b, c]" 
        and op_b: "b  op_cat 𝔅Obj" 
        and c: "c  Obj"
      by (auto intro: cat_prod_2_ObjE cat_cs_intros)
    from op_b have b: "b  𝔅Obj" unfolding cat_op_simps by simp

    then have 𝔊b: "𝔊ObjMapb  Obj" and 𝔉b: "𝔉ObjMapb  Obj"
      by (auto intro: cat_cs_intros)
    have Ym_𝔑:
      "Yoneda_map α HomO.Cα(𝔉ObjMapb,-) (𝔊ObjMapb)
        𝔑op_cat 𝔅,(b,-)NTCF
         = ?aYa 𝔑NTMapb"
      unfolding af_Yoneda_arrow_NTMap_app[OF assms(1) b] by simp
    
    from 
      bnt_proj_snd_is_ntcf
        [
          OF 𝔉.HomDom.category_op category_axioms assms(3) op_b,
          unfolded 
            cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(1) b]
            cat_cf_lcomp_Hom_cf_Hom_snd[OF assms(2) b]
        ]
    have 𝔑b: "𝔑op_cat 𝔅,(b,-)NTCF :
      HomO.Cα(𝔊ObjMapb,-) CF HomO.Cα(𝔉ObjMapb,-) :
       ↦↦Cαcat_Set α"
      by simp
    from c show "HomA.Cα(?aYa 𝔑-,-)NTMapbc = 𝔑NTMapbc"
      unfolding 
        bc_def 
        cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app[OF aY𝔑.is_ntcf_axioms b c]
        cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(2)[
          OF 𝔊b 𝔉b 𝔑b, unfolded Ym_𝔑, symmetric
          ]
      by (cs_concl cs_shallow cs_simp: cat_cs_simps)

  qed simp_all

qed

lemma (in category) cat_af_Yoneda_Lemma:
  assumes "𝔉 : 𝔅 ↦↦Cα" and "𝔊 : 𝔅 ↦↦Cα"
  shows "v11 (af_Yoneda_map α 𝔉 𝔊)"
    and " (af_Yoneda_map α 𝔉 𝔊) =
    these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) HomO.Cα(𝔊-,-) HomO.Cα(𝔉-,-)"
    and "(af_Yoneda_map α 𝔉 𝔊)¯ =
      (
        λ𝔑these_ntcfs
          α (op_cat 𝔅 ×C ) (cat_Set α) HomO.Cα(𝔊-,-) HomO.Cα(𝔉-,-).
          af_Yoneda_arrow α 𝔉 𝔊 𝔑
      )"
proof-

  let ?H𝔊 = HomO.Cα(𝔊-,-)
    and ?H𝔉 = HomO.Cα(𝔉-,-)
    and ?aYm = af_Yoneda_map α 𝔉 𝔊
    and ?aYa = λ𝔑. af_Yoneda_arrow α 𝔉 𝔊 𝔑

  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))

  show v11_aY: "v11 ?aYm"
  proof
    (
      intro vsv.vsv_valeq_v11I,
      unfold af_Yoneda_map_vdomain[OF assms] these_ntcfs_iff
    )
    
    show "vsv (af_Yoneda_map α 𝔉 𝔊)" by (rule af_Yoneda_map_vsv[OF assms])

    fix φ ψ assume prems:
      "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
      "ψ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα" 
      "?aYmφ = ?aYmψ"

    interpret φ: is_ntcf α 𝔅  𝔉 𝔊 φ by (rule prems(1))
    interpret ψ: is_ntcf α 𝔅  𝔉 𝔊 ψ by (rule prems(2))

    from prems(3) have Hφ_Hψ: "HomA.Cα(φ-,-) = HomA.Cα(ψ-,-)"
      unfolding 
        af_Yoneda_map_app[OF assms prems(1)]
        af_Yoneda_map_app[OF assms prems(2)]
      by simp

    show "φ = ψ"
    proof
      (
        rule ntcf_eqI[OF prems(1,2)], 
        rule vsv_eqI, 
        unfold φ.ntcf_NTMap_vdomain ψ.ntcf_NTMap_vdomain
      )
      fix b assume prems': "b  𝔅Obj"
      from prems' have φb: "φNTMapb : 𝔉ObjMapb 𝔊ObjMapb" 
        and ψb: "ψNTMapb : 𝔉ObjMapb 𝔊ObjMapb" 
        and 𝔊b: "𝔊ObjMapb  Obj" 
        and 𝔉b: "𝔉ObjMapb  Obj"
        by (auto intro: cat_cs_intros cat_prod_cs_intros)
      have "HomA.Cα(φNTMapb,-) = HomA.Cα(ψNTMapb,-)"
      proof
        (
          rule 
            ntcf_eqI
              [
                OF 
                  cat_ntcf_Hom_snd_is_ntcf[OF φb] 
                  cat_ntcf_Hom_snd_is_ntcf[OF ψb]
              ]
        )
        show "HomA.Cα(φNTMapb,-)NTMap = HomA.Cα(ψNTMapb,-)NTMap"
        proof
          (
            rule vsv_eqI, 
            unfold 
              ntcf_Hom_snd_NTMap_vdomain[OF φb]
              ntcf_Hom_snd_NTMap_vdomain[OF ψb]
          )
          fix c assume prems'': "c  Obj"
          note H = cat_ntcf_lcomp_Hom_ntcf_Hom_snd_NTMap_app
          show 
            "HomA.Cα(φNTMapb,-)NTMapc =
              HomA.Cα(ψNTMapb,-)NTMapc"
            unfolding 
              H[OF prems(1) prems' prems'', symmetric]
              H[OF prems(2) prems' prems'', symmetric]
              Hφ_Hψ
            by simp
        qed 
          (
            simp_all add: 
              ntcf_Hom_snd_NTMap_vsv[OF ψb] ntcf_Hom_snd_NTMap_vsv[OF φb]
          )
      qed simp_all
      with φb ψb show "φNTMapb = ψNTMapb"
        by (auto intro: cat_ntcf_Hom_snd_inj)
    qed auto

  qed

  interpret aYm: v11 ?aYm by (rule v11_aY)

  have [cat_cs_simps]: "?aYm?aYa 𝔑 = 𝔑"
    if "𝔑 : ?H𝔊 CF ?H𝔉 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α" for 𝔑
    using category_axioms assms that 
    by 
      (
        cs_concl cs_shallow 
          cs_simp: 
            cat_af_Yoneda_map_af_Yoneda_arrow_app[symmetric] cat_cs_simps
          cs_intro: cat_cs_intros
      )

  show aYm_vrange: 
    " ?aYm = these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉"
  proof(intro vsubset_antisym)
    
    show " ?aYm  these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉"
    proof
      (
        rule vsv.vsv_vrange_vsubset, 
        unfold these_ntcfs_iff af_Yoneda_map_vdomain[OF assms]
      )
      fix φ assume "φ : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
      with category_axioms assms show 
        "?aYmφ : ?H𝔊 CF ?H𝔉 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
        by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    qed (auto intro: af_Yoneda_map_vsv)

    show "these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉   ?aYm"
    proof(rule vsubsetI, unfold these_ntcfs_iff)
      fix 𝔑 assume prems: 
        "𝔑 : ?H𝔊 CF ?H𝔉 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
      interpret aY𝔑: is_ntcf α 𝔅  𝔉 𝔊 ?aYa 𝔑
        by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms prems])
      from prems have 𝔑_def: "𝔑 = ?aYm?aYa 𝔑" 
        by (cs_concl cs_shallow cs_simp: cat_cs_simps)
      from assms aY𝔑.is_ntcf_axioms have "?aYa 𝔑  𝒟 ?aYm"
        by (cs_concl cs_shallow cs_simp: these_ntcfs_iff cat_cs_simps)
      then show "𝔑   ?aYm" by (subst 𝔑_def, intro aYm.vsv_vimageI2) auto
    qed

  qed

  show "?aYm¯ =
    (λ𝔑these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉. ?aYa 𝔑)"
  proof
    (
      rule vsv_eqI, 
      unfold vdomain_vconverse vdomain_VLambda aYm_vrange these_ntcfs_iff
    )
    from aYm.v11_axioms show "vsv ((af_Yoneda_map α 𝔉 𝔊)¯)" by auto
    fix 𝔑 assume prems: "𝔑 : ?H𝔊 CF ?H𝔉 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    then have 𝔑: "𝔑  these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉" 
      by simp
    show "?aYm¯𝔑 =
      (λ𝔑these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) ?H𝔊 ?H𝔉. ?aYa 𝔑)𝔑"
    proof
      (
        intro aYm.v11_vconverse_app, 
        unfold beta[OF 𝔑] af_Yoneda_map_vdomain[OF assms] these_ntcfs_iff
      )
      from prems show 𝔑_def: "?aYm?aYa 𝔑 = 𝔑" 
        by (cs_concl cs_shallow cs_simp: cat_cs_simps)
      show "?aYa 𝔑 : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
        by (rule cat_af_Yoneda_arrow_is_ntcf[OF assms prems])
    qed
  qed simp_all

qed


subsubsection‹Inverse of the Yoneda map for arbitrary functors›

lemma (in category) inv_af_Yoneda_map_v11: 
  assumes "𝔉 : 𝔅 ↦↦Cα" and "𝔊 : 𝔅 ↦↦Cα"
  shows "v11 ((af_Yoneda_map α 𝔉 𝔊)¯)"
  using cat_af_Yoneda_Lemma(1)[OF assms] by (simp add: v11.v11_vconverse)

lemma (in category) inv_af_Yoneda_map_vdomain: 
  assumes "𝔉 : 𝔅 ↦↦Cα" and "𝔊 : 𝔅 ↦↦Cα"
  shows "𝒟 ((af_Yoneda_map α 𝔉 𝔊)¯) =
    these_ntcfs α (op_cat 𝔅 ×C ) (cat_Set α) HomO.Cα(𝔊-,-) HomO.Cα(𝔉-,-)"
  unfolding cat_af_Yoneda_Lemma(3)[OF assms] by simp

lemmas [cat_cs_simps] = category.inv_af_Yoneda_map_vdomain

lemma (in category) inv_af_Yoneda_map_app: 
  assumes "𝔉 : 𝔅 ↦↦Cα" and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF  HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
  shows "(af_Yoneda_map α 𝔉 𝔊)¯𝔑 = af_Yoneda_arrow α 𝔉 𝔊 𝔑"
  using assms(3) unfolding cat_af_Yoneda_Lemma(3)[OF assms(1,2)] by simp

lemmas [cat_cs_simps] = category.inv_af_Yoneda_map_app

lemma (in category) inv_af_Yoneda_map_vrange: 
  assumes "𝔉 : 𝔅 ↦↦Cα" and "𝔊 : 𝔅 ↦↦Cα"
  shows " ((af_Yoneda_map α 𝔉 𝔊)¯) = these_ntcfs α 𝔅  𝔉 𝔊"
proof-
  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))
  from assms show ?thesis 
    unfolding af_Yoneda_map_def by (simp add: cat_cs_simps)
qed


subsubsection‹Yoneda map for arbitrary functors and natural isomorphisms›


text‹
The following lemmas correspond to variants of the elements of
Lemma 3 in subsection 1.15 in cite"bodo_categories_1970".
›

lemma (in category) cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf:
  assumes "φ : 𝔉 CF.iso 𝔊 : 𝔅 ↦↦Cα"
  shows "HomA.Cα(φ-,-) :
    HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
    op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
proof-
  interpret φ: is_iso_ntcf α 𝔅  𝔉 𝔊 φ by (rule assms(1))
  show ?thesis
  proof(intro cat_ntcf_lcomp_Hom_if_ntcf_Hom_snd_is_iso_ntcf)
    fix b assume "b  𝔅Obj"
    then show "HomA.Cα(φNTMapb,-) :
      HomO.Cα(𝔊ObjMapb,-) CF.iso HomO.Cα(𝔉ObjMapb,-) :
       ↦↦Cαcat_Set α"
      by 
        (
          auto intro!: 
            cat_is_iso_arr_ntcf_Hom_snd_is_iso_ntcf cat_arrow_cs_intros
        )
  qed (auto simp: cat_cs_intros)
qed

lemma (in category) cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf':
  assumes "φ : 𝔉 CF.iso 𝔊 : 𝔅 ↦↦Cα"
    and "β = α"
    and "𝔊' = HomO.Cα(𝔊-,-)"
    and "𝔉' = HomO.Cα(𝔉-,-)"
    and "𝔅' = op_cat 𝔅 ×C "
    and "ℭ' = cat_Set α"
  shows "HomA.Cα(φ-,-) : 𝔊' CF.iso 𝔉' : 𝔅' ↦↦Cβℭ'"
  using assms(1)
  unfolding assms(2-6) 
  by (rule cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf)

lemmas [cat_cs_intros] = 
  category.cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf'

lemma (in category) cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf:
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
  shows "af_Yoneda_arrow α 𝔉 𝔊 𝔑 : 𝔉 CF.iso 𝔊 : 𝔅 ↦↦Cα"
proof-
  
  let ?aYa = af_Yoneda_arrow α 𝔉 𝔊 𝔑
  
  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))
  interpret 𝔑: is_iso_ntcf 
    α op_cat 𝔅 ×C  cat_Set α HomO.Cα(𝔊-,-) HomO.Cα(𝔉-,-) 𝔑
    by (rule assms(3))

  from assms(1,2) 𝔑.is_ntcf_axioms have 𝔑_def: "𝔑 = HomA.Cα(?aYa-,-)" 
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_af_Yoneda_map_af_Yoneda_arrow_app[symmetric]
      )

  from category_axioms assms have aYa: "?aYa : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
    by (cs_concl cs_shallow cs_intro: cat_cs_intros)
  have Hom_aYa: "HomA.Cα(?aYa-,-) :
    HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
    op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    by (auto intro: assms(3) simp add: 𝔑_def[symmetric])
  have Hb:
    "HomA.Cα(?aYaNTMapb,-) :
      HomO.Cα(𝔊ObjMapb,-) CF.iso HomO.Cα(𝔉ObjMapb,-) :
       ↦↦Cαcat_Set α"
    if "b  𝔅Obj" for b
    by 
      ( 
        rule cat_ntcf_Hom_snd_if_ntcf_lcomp_Hom_is_iso_ntcf[
          OF aYa Hom_aYa that
          ]
      )

  show ?thesis
  proof(intro is_iso_ntcfI)
    from category_axioms assms show 
      "af_Yoneda_arrow α 𝔉 𝔊 𝔑 : 𝔉 CF 𝔊 : 𝔅 ↦↦Cα"
      by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
    fix b assume prems: "b  𝔅Obj"
    then have 𝔊b: "𝔊ObjMapb  Obj" and 𝔉b: "𝔉ObjMapb  Obj"
      by (auto intro: cat_cs_intros)
    from assms(1,2) aYa prems have aYa_b: 
      "?aYaNTMapb : 𝔉ObjMapb 𝔊ObjMapb"
      by (cs_concl cs_shallow cs_intro: cat_cs_intros cs_simp: cat_cs_simps)
    show "af_Yoneda_arrow α 𝔉 𝔊 𝔑NTMapb : 𝔉ObjMapb iso𝔊ObjMapb"
      by 
        (
          rule cat_is_iso_arr_if_ntcf_Hom_snd_is_iso_ntcf[
            OF aYa_b Hb[OF prems]
            ]
        )
  qed

qed

lemma (in category) cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf':
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔑 :
      HomO.Cα(𝔊-,-) CF.iso HomO.Cα(𝔉-,-) :
      op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    and "β = α"
    and "𝔉' = 𝔉"
    and "𝔊' = 𝔊"
  shows "af_Yoneda_arrow α 𝔉 𝔊 𝔑 : 𝔉' CF.iso 𝔊' : 𝔅 ↦↦Cα"
  using assms(1-3) 
  unfolding assms(4-6) 
  by (rule cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf)

lemmas [cat_cs_intros] = 
  category.cat_aYa_is_iso_ntcf_if_ntcf_lcomp_Hom_is_iso_ntcf'

lemma (in category) cat_iso_functor_if_cf_lcomp_Hom_iso_functor:
  assumes "𝔉 : 𝔅 ↦↦Cα" 
    and "𝔊 : 𝔅 ↦↦Cα"
    and "HomO.Cα(𝔉-,-) CFαHomO.Cα(𝔊-,-)"
  shows "𝔉 CFα𝔊"
proof-
  let ?H𝔊 = HomO.Cα(𝔊-,-)
    and ?H𝔉 = HomO.Cα(𝔉-,-)
    and ?aYa = λ𝔑. af_Yoneda_arrow α 𝔉 𝔊 𝔑
  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))
  from assms(3) obtain 𝔑 𝔄 𝔇 where 𝔑: "𝔑 : ?H𝔉 CF.iso ?H𝔊 : 𝔄 ↦↦Cα𝔇"
    by auto
  interpret 𝔑: is_iso_ntcf α 𝔄 𝔇 ?H𝔉 ?H𝔊 𝔑 by (rule 𝔑)
  from category_axioms assms have "?H𝔉 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    by (cs_concl cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  then have 𝔄_def: "𝔄 = op_cat 𝔅 ×C " and 𝔇_def: "𝔇 = cat_Set α"
    by (force simp: cat_cs_simps)+
  note 𝔑 = 𝔑[unfolded 𝔄_def 𝔇_def]
  from 𝔑 have "𝔑 : ?H𝔉 CF ?H𝔊 : op_cat 𝔅 ×C  ↦↦Cαcat_Set α"
    by 
      (
        cs_concl cs_shallow 
          cs_simp: cat_cs_simps cs_intro: cat_cs_intros ntcf_cs_intros
      )
  from category_axioms assms 𝔑 have 
    "af_Yoneda_arrow α 𝔊 𝔉 𝔑 : 𝔊 CF.iso 𝔉 : 𝔅 ↦↦Cα"
    by (cs_concl cs_shallow cs_simp: cat_cs_simps cs_intro: cat_cs_intros)
  then have "𝔊 CFα𝔉" by (clarsimp intro!: iso_functorI)
  then show ?thesis by (rule iso_functor_sym)
qed

lemma (in category) cat_cf_lcomp_Hom_iso_functor_if_iso_functor:
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔉 CFα𝔊"
  shows "HomO.Cα(𝔉-,-) CFαHomO.Cα(𝔊-,-)"
proof-
  let ?H𝔊 = HomO.Cα(𝔊-,-)
    and ?H𝔉 = HomO.Cα(𝔉-,-)
    and ?aYa = λ𝔑. af_Yoneda_arrow α 𝔉 𝔊 𝔑
  interpret 𝔉: is_functor α 𝔅  𝔉 by (rule assms(1))
  interpret 𝔊: is_functor α 𝔅  𝔊 by (rule assms(2))
  from assms obtain 𝔅' ℭ' φ where φ: "φ : 𝔉 CF.iso 𝔊 : 𝔅' ↦↦Cαℭ'"
    by auto
  interpret φ: is_iso_ntcf α 𝔅' ℭ' 𝔉 𝔊 φ by (rule φ)
  from assms φ.NTDom.is_functor_axioms 
  have 𝔅'_def: "𝔅' = 𝔅" and ℭ'_def: "ℭ' = " 
    by fast+
  note φ = φ[unfolded 𝔅'_def ℭ'_def]
  show ?thesis
    by (rule iso_functor_sym) 
      (
        intro iso_functorI[
          OF cat_ntcf_lcomp_Hom_is_iso_ntcf_if_is_iso_ntcf[OF φ]
          ]
      )
qed

lemma (in category) cat_cf_lcomp_Hom_iso_functor_if_iso_functor':
  assumes "𝔉 : 𝔅 ↦↦Cα"
    and "𝔊 : 𝔅 ↦↦Cα"
    and "𝔉 CFα𝔊"
    and "α' = α"
    and "ℭ' = "
  shows "HomO.Cα(𝔉-,-) CFαHomO.Cα'ℭ'(𝔊-,-)"
  using assms(1-3) 
  unfolding assms(4,5) 
  by (rule cat_cf_lcomp_Hom_iso_functor_if_iso_functor)

lemmas [cat_cs_intros] = 
  category.cat_cf_lcomp_Hom_iso_functor_if_iso_functor'



subsection‹The Yoneda Functor›


subsubsection‹Definition and elementary properties›


text‹See Chapter III-2 in cite"mac_lane_categories_2010".›


definition Yoneda_functor :: "V  V  V"
  where "Yoneda_functor α 𝔇 =
    [
      (λrop_cat 𝔇Obj. cf_map (HomO.Cα𝔇(r,-))),
      (λfop_cat 𝔇Arr. ntcf_arrow (HomA.Cα𝔇(f,-))),
      op_cat 𝔇,
      cat_FUNCT α 𝔇 (cat_Set α)
    ]"


text‹Components.›

lemma Yoneda_functor_components: 
  shows "Yoneda_functor α 𝔇ObjMap =
      (λrop_cat 𝔇Obj. cf_map (HomO.Cα𝔇(r,-)))"
    and "Yoneda_functor α 𝔇ArrMap =
      (λfop_cat 𝔇Arr. ntcf_arrow (HomA.Cα𝔇(f,-)))"
    and "Yoneda_functor α 𝔇HomDom = op_cat 𝔇"
    and "Yoneda_functor α 𝔇HomCod = cat_FUNCT α 𝔇 (cat_Set α)"
  unfolding Yoneda_functor_def dghm_field_simps 
  by (simp_all add: nat_omega_simps)


subsubsection‹Object map›

mk_VLambda Yoneda_functor_components(1)
  |vsv Yoneda_functor_ObjMap_vsv[cat_cs_intros]|
  |vdomain Yoneda_functor_ObjMap_vdomain[cat_cs_simps]|
  |app Yoneda_functor_ObjMap_app[cat_cs_simps]|

lemma (in category) Yoneda_functor_ObjMap_vrange:
  " (Yoneda_functor α ObjMap)  cat_FUNCT α  (cat_Set α)Obj"
proof
  (
    unfold Yoneda_functor_components, 
    rule vrange_VLambda_vsubset, 
    unfold cat_op_simps
  )
  fix c assume "c  Obj"
  with category_axioms show 
    "cf_map HomO.Cα(c,-)  cat_FUNCT α  (cat_Set α)Obj" 
    unfolding cat_op_simps cat_FUNCT_components
    by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_FUNCT_cs_intros)
qed


subsubsection‹Arrow map›

mk_VLambda Yoneda_functor_components(2)
  |vsv Yoneda_functor_ArrMap_vsv[cat_cs_intros]|
  |vdomain Yoneda_functor_ArrMap_vdomain[cat_cs_simps]|
  |app Yoneda_functor_ArrMap_app[cat_cs_simps]|

lemma (in category) Yoneda_functor_ArrMap_vrange:
  " (Yoneda_functor α ArrMap)  cat_FUNCT α  (cat_Set α)Arr"
proof
  (
    unfold Yoneda_functor_components, 
    rule vrange_VLambda_vsubset, 
    unfold cat_op_simps
  )
  fix f assume "f  Arr"
  then obtain a b where f: "f : a b" by auto
  define β where "β = α + ω"
  have 𝒵β: "𝒵 β" and αβ: "α  β"
    by (simp_all add: 𝒵_α_αω 𝒵.intro 𝒵_Limit_αω 𝒵_ω_αω β_def)
  from tiny_category_cat_FUNCT category_axioms 𝒵β αβ f show
    "ntcf_arrow HomA.Cα(f,-)  cat_FUNCT α  (cat_Set α)Arr" 
    unfolding cat_op_simps
    by (cs_concl cs_shallow cs_intro: cat_cs_intros cat_FUNCT_cs_intros)
qed


subsubsection‹The Yoneda Functor is a fully faithful functor›

lemma (in category) cat_Yoneda_functor_is_functor:
  assumes "𝒵 β" and "α  β"
  shows "Yoneda_functor α  : op_cat  ↦↦C.ffβcat_FUNCT α  (cat_Set α)"
proof
  (
    intro 
      is_ff_functorI 
      is_ft_functorI' 
      is_fl_functorI' 
      vsubset_antisym 
      vsubsetI,
    unfold cat_op_simps in_Hom_iff, 
    tacticdistinct_subgoals_tac
  )

  interpret Set: category α cat_Set α by (rule category_cat_Set)

  let ?Yf = Yoneda_functor α  and ?FUNCT = cat_FUNCT α  (cat_Set α)

  show Yf: "?Yf : op_cat  ↦↦Cβ?FUNCT"
  proof(intro is_functorI')
    show "vfsequence ?Yf" unfolding Yoneda_functor_def by simp
    from assms have "category β " by (intro cat_category_if_ge_Limit)
    then show "category β (op_cat )" by (intro category.category_op)
    from assms show "category β ?FUNCT" 
      by 
        (
          cs_concl cs_shallow 
            cs_intro: cat_small_cs_intros tiny_category_cat_FUNCT
        )
    show "vcard ?Yf = 4"
      unfolding Yoneda_functor_def by (simp add: nat_omega_simps)
    show " (?YfObjMap)  ?FUNCTObj" 
      by (rule Yoneda_functor_ObjMap_vrange)
    show 
      "?YfArrMapf : ?YfObjMapa cat_FUNCT α  (cat_Set α)?YfObjMapb"
      if "f : a op_cat b" for a b f
      using that category_axioms
      unfolding cat_op_simps
      by 
        (
          cs_concl 
            cs_simp: cat_cs_simps cat_op_simps 
            cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
    show "?YfArrMapg Aop_cat f = 
      ?YfArrMapg A?FUNCT?YfArrMapf"
      if "g : b op_cat c" and "f : a op_cat b" for b c g a f
      using that category_axioms
      unfolding cat_op_simps
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps 
            cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
    show "?YfArrMapop_cat CIdc = ?FUNCTCId?YfObjMapc"
      if "c  op_cat Obj" for c 
      using that category_axioms
      unfolding cat_op_simps
      by 
        (
          cs_concl cs_shallow
            cs_simp: cat_cs_simps cat_op_simps cat_FUNCT_cs_simps 
            cs_intro: cat_cs_intros cat_FUNCT_cs_intros
        )
  qed (auto simp: assms(1) Yoneda_functor_components 𝒵.intro 𝒵_Limit_αω 𝒵_ω_αω)

  interpret Yf: is_functor β op_cat  ?FUNCT ?Yf by (rule Yf)

  show "v11 (?YfArrMap l Hom  b a)"
    if "a  Obj" and "b  Obj" for a b
  proof-
    from that have dom_Y_ba: "𝒟 (?YfArrMap l Hom  b a) = Hom  b a"
      by 
        (
          fastforce simp: 
            cat_op_simps 
            in_Hom_iff vdomain_vlrestriction Yoneda_functor_components
        )

    show "v11 (?YfArrMap l Hom  b a)"
    proof(intro vsv.vsv_valeq_v11I, unfold dom_Y_ba in_Hom_iff)
      fix g f assume prems:
        "g : b a" 
        "f : b a" 
        "(?YfArrMap l Hom  b a)g = (?YfArrMap l Hom  b a)f"
      from 
        prems(3) category_axioms prems(1,2) Yoneda_functor_ArrMap_vsv[of α ] 
      have "HomA.Cα(g,-) = HomA.Cα(f,-)"
        by 
          (
            cs_prems cs_shallow
              cs_simp: V_cs_simps cat_cs_simps cat_op_simps cat_FUNCT_cs_simps
              cs_intro: cat_cs_intros
          )
      from this prems(1,2) show "g = f" by (rule cat_ntcf_Hom_snd_inj)
    qed (auto simp: Yoneda_functor_components)
  qed

  fix a b assume prems: "a  Obj" "b  Obj"
  show "𝔑 : ?YfObjMapa cat_FUNCT α  (cat_Set α)?YfObjMapb"
    if "𝔑  ?YfArrMap ` Hom  b a" for 𝔑
  proof-
    from that obtain f where "?YfArrMapf = 𝔑" and f: "f : b a"
      by (force elim!: Yf.ArrMap.vsv_vimageE)
    then have 𝔑_def: "𝔑 = ntcf_arrow HomA.Cα(f,-)"
      unfolding 
        Yoneda_functor_ArrMap_app[
          unfolded cat_op_simps, OF cat_is_arrD(1)[OF f]
          ]
      by (simp add: cat_cs_simps cat_op_simps cat_cs_intros)
    from category_axioms f show ?thesis
      unfolding 𝔑_def
      by 
        (
          cs_concl  
            cs_simp: cat_cs_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_FUNCT_cs_intros
        )
  qed
  show "𝔑  ?YfArrMap ` Hom  b a"
    if "𝔑 : ?YfObjMapa cat_FUNCT α  (cat_Set α)?YfObjMapb" for 𝔑
  proof-
    note 𝔑 = cat_FUNCT_is_arrD[OF that]
    from 𝔑(1) category_axioms prems have ntcf_𝔑:
      "ntcf_of_ntcf_arrow  (cat_Set α) 𝔑 : 
        HomO.Cα(a,-) CF HomO.Cα(b,-) :  ↦↦Cαcat_Set α"
      by (subst (asm) 𝔑(3), use nothing in subst (asm) 𝔑(4))
        (
          cs_prems cs_shallow 
            cs_simp: cat_cs_simps cat_FUNCT_cs_simps 
            cs_intro: cat_cs_intros cat_op_intros cat_FUNCT_cs_intros
        )
    from cat_ntcf_Hom_snd_is_ntcf_Hom_snd_unique(1,2)[OF prems ntcf_𝔑] obtain f 
      where f: "f : b a" 
        and 𝔑_def: "ntcf_of_ntcf_arrow  (cat_Set α) 𝔑 = HomA.Cα(f,-)"
      by auto
    from 𝔑(2) f show "𝔑  Yoneda_functor α ArrMap ` Hom  b a"
      unfolding 𝔑_def
      by (intro Yf.ArrMap.vsv_vimage_eqI[of f]) 
        (
          cs_concl cs_shallow 
            cs_simp: cat_cs_simps cs_intro: cat_cs_intros cat_op_intros
        )+
  qed
qed

text‹\newpage›

end